diff --git a/thys/Stone_Algebras/Filters.thy b/thys/Stone_Algebras/Filters.thy --- a/thys/Stone_Algebras/Filters.thy +++ b/thys/Stone_Algebras/Filters.thy @@ -1,911 +1,943 @@ (* Title: Filters Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Filters\ text \ This theory develops filters based on orders, semilattices, lattices and distributive lattices. We prove the ultrafilter lemma for orders with a least element. We show the following structure theorems: \begin{itemize} \item The set of filters over a directed semilattice forms a lattice with a greatest element. \item The set of filters over a bounded semilattice forms a bounded lattice. \item The set of filters over a distributive lattice with a greatest element forms a bounded distributive lattice. \end{itemize} Another result is that in a distributive lattice ultrafilters are prime filters. We also prove a lemma of Gr\"atzer and Schmidt about principal filters. We apply these results in proving the construction theorem for Stone algebras (described in a separate theory). See, for example, \cite{BalbesDwinger1974,Birkhoff1967,Blyth2005,DaveyPriestley2002,Graetzer1971} for further results about filters. \ theory Filters imports Lattice_Basics begin subsection \Orders\ text \ This section gives the basic definitions related to filters in terms of orders. The main result is the ultrafilter lemma. \ context ord begin abbreviation down :: "'a \ 'a set" ("\_" [81] 80) where "\x \ { y . y \ x }" abbreviation down_set :: "'a set \ 'a set" ("\_" [81] 80) where "\X \ { y . \x\X . y \ x }" abbreviation is_down_set :: "'a set \ bool" where "is_down_set X \ \x\X . \y . y \ x \ y\X" abbreviation is_principal_down :: "'a set \ bool" where "is_principal_down X \ \x . X = \x" abbreviation up :: "'a \ 'a set" ("\_" [81] 80) where "\x \ { y . x \ y }" abbreviation up_set :: "'a set \ 'a set" ("\_" [81] 80) where "\X \ { y . \x\X . x \ y }" abbreviation is_up_set :: "'a set \ bool" where "is_up_set X \ \x\X . \y . x \ y \ y\X" abbreviation is_principal_up :: "'a set \ bool" where "is_principal_up X \ \x . X = \x" text \ A filter is a non-empty, downward directed, up-closed set. \ definition filter :: "'a set \ bool" where "filter F \ (F \ {}) \ (\x\F . \y\F . \z\F . z \ x \ z \ y) \ is_up_set F" abbreviation proper_filter :: "'a set \ bool" where "proper_filter F \ filter F \ F \ UNIV" abbreviation ultra_filter :: "'a set \ bool" where "ultra_filter F \ proper_filter F \ (\G . proper_filter G \ F \ G \ F = G)" +abbreviation filters :: "'a set set" + where "filters \ { F::'a set . filter F }" + +lemma filter_map_filter: + assumes "filter F" + and "mono f" + and "\x y . f x \ y \ (\z . x \ z \ y = f z)" + shows "filter (f ` F)" +proof (unfold ord_class.filter_def, intro conjI) + show "f ` F \ {}" + using assms(1) ord_class.filter_def by auto +next + show "\x\f ` F . \y\f ` F . \z\f ` F . z \ x \ z \ y" + proof (intro ballI) + fix x y + assume "x \ f ` F" and "y \ f ` F" + then obtain u v where 1: "x = f u \ u \ F \ y = f v \ v \ F" + by auto + then obtain w where "w \ u \ w \ v \ w \ F" + by (meson assms(1) ord_class.filter_def) + thus "\z\f ` F . z \ x \ z \ y" + using 1 assms(2) mono_def image_eqI by blast + qed +next + show "is_up_set (f ` F)" + proof + fix x + assume "x \ f ` F" + then obtain u where 1: "x = f u \ u \ F" + by auto + show "\y . x \ y \ y \ f ` F" + proof (rule allI, rule impI) + fix y + assume "x \ y" + hence "f u \ y" + using 1 by simp + then obtain z where "u \ z \ y = f z" + using assms(3) by auto + thus "y \ f ` F" + by (meson 1 assms(1) image_iff ord_class.filter_def) + qed + qed +qed + end context order begin lemma self_in_downset [simp]: "x \ \x" by simp lemma self_in_upset [simp]: "x \ \x" by simp lemma up_filter [simp]: "filter (\x)" using filter_def order_lesseq_imp by auto lemma up_set_up_set [simp]: "is_up_set (\X)" using order.trans by fastforce lemma up_injective: "\x = \y \ x = y" using antisym by auto lemma up_antitone: "x \ y \ \y \ \x" by auto end context order_bot begin lemma bot_in_downset [simp]: "bot \ \x" by simp lemma down_bot [simp]: "\bot = {bot}" by (simp add: bot_unique) lemma up_bot [simp]: "\bot = UNIV" by simp text \ The following result is the ultrafilter lemma, generalised from \cite[10.17]{DaveyPriestley2002} to orders with a least element. Its proof uses Isabelle/HOL's \Zorn_Lemma\, which requires closure under union of arbitrary (possibly empty) chains. Actually, the proof does not use any of the underlying order properties except \bot_least\. \ lemma ultra_filter: assumes "proper_filter F" shows "\G . ultra_filter G \ F \ G" proof - let ?A = "{ G . (proper_filter G \ F \ G) \ G = {} }" have "\C \ chains ?A . \C \ ?A" proof fix C :: "'a set set" let ?D = "C - {{}}" assume 1: "C \ chains ?A" hence 2: "\x\\?D . \H\?D . x \ H \ proper_filter H" using chainsD2 by fastforce have 3: "\?D = \C" by blast have "\?D \ ?A" proof (cases "?D = {}") assume "?D = {}" thus ?thesis by auto next assume 4: "?D \ {}" then obtain G where "G \ ?D" by auto hence 5: "F \ \?D" using 1 chainsD2 by blast have 6: "is_up_set (\?D)" proof fix x assume "x \ \?D" then obtain H where "x \ H \ H \ ?D \ filter H" using 2 by auto thus "\y . x \ y \ y\\?D" using filter_def UnionI by fastforce qed have 7: "\?D \ UNIV" proof (rule ccontr) assume "\ \?D \ UNIV" then obtain H where "bot \ H \ proper_filter H" using 2 by blast thus False by (meson UNIV_I bot_least filter_def subsetI subset_antisym) qed { fix x y assume "x\\?D \ y\\?D" then obtain H I where 8: "x \ H \ H \ ?D \ filter H \ y \ I \ I \ ?D \ filter I" using 2 by metis have "\z\\?D . z \ x \ z \ y" proof (cases "H \ I") assume "H \ I" hence "\z\I . z \ x \ z \ y" using 8 by (metis subsetCE filter_def) thus ?thesis using 8 by (metis UnionI) next assume "\ (H \ I)" hence "I \ H" using 1 8 by (meson DiffE chainsD) hence "\z\H . z \ x \ z \ y" using 8 by (metis subsetCE filter_def) thus ?thesis using 8 by (metis UnionI) qed } thus ?thesis using 4 5 6 7 filter_def by auto qed thus "\C \ ?A" using 3 by simp qed hence "\M\?A . \X\?A . M \ X \ X = M" by (rule Zorn_Lemma) then obtain M where 9: "M \ ?A \ (\X\?A . M \ X \ X = M)" by auto hence 10: "M \ {}" using assms filter_def by auto { fix G assume 11: "proper_filter G \ M \ G" hence "F \ G" using 9 10 by blast hence "M = G" using 9 11 by auto } thus ?thesis using 9 10 by blast qed end context order_top begin lemma down_top [simp]: "\top = UNIV" by simp lemma top_in_upset [simp]: "top \ \x" by simp lemma up_top [simp]: "\top = {top}" by (simp add: top_unique) lemma filter_top [simp]: "filter {top}" using filter_def top_unique by auto lemma top_in_filter [simp]: "filter F \ top \ F" using filter_def by fastforce end text \ The existence of proper filters and ultrafilters requires that the underlying order contains at least two elements. \ context non_trivial_order begin lemma proper_filter_exists: "\F . proper_filter F" proof - from consistent obtain x y :: 'a where "x \ y" by auto hence "\x \ UNIV \ \y \ UNIV" using antisym by blast hence "proper_filter (\x) \ proper_filter (\y)" by simp thus ?thesis by blast qed end context non_trivial_order_bot begin lemma ultra_filter_exists: "\F . ultra_filter F" using ultra_filter proper_filter_exists by blast end context non_trivial_bounded_order begin lemma proper_filter_top: "proper_filter {top}" using bot_not_top filter_top by blast lemma ultra_filter_top: "\G . ultra_filter G \ top \ G" using ultra_filter proper_filter_top by fastforce end subsection \Lattices\ text \ This section develops the lattice structure of filters based on a semilattice structure of the underlying order. The main results are that filters over a directed semilattice form a lattice with a greatest element and that filters over a bounded semilattice form a bounded lattice. \ context semilattice_sup begin abbreviation prime_filter :: "'a set \ bool" where "prime_filter F \ proper_filter F \ (\x y . x \ y \ F \ x \ F \ y \ F)" end context semilattice_inf begin lemma filter_inf_closed: "filter F \ x \ F \ y \ F \ x \ y \ F" by (meson filter_def inf.boundedI) lemma filter_univ: "filter UNIV" by (meson UNIV_I UNIV_not_empty filter_def inf.cobounded1 inf.cobounded2) text \ The operation \filter_sup\ is the join operation in the lattice of filters. \ -abbreviation "filter_sup F G \ { z . \x\F . \y\G . x \ y \ z }" +definition "filter_sup F G \ { z . \x\F . \y\G . x \ y \ z }" lemma filter_sup: assumes "filter F" and "filter G" shows "filter (filter_sup F G)" proof - have "F \ {} \ G \ {}" using assms filter_def by blast hence 1: "filter_sup F G \ {}" - by blast + using filter_sup_def by blast have 2: "\x\filter_sup F G . \y\filter_sup F G . \z\filter_sup F G . z \ x \ z \ y" proof fix x assume "x\filter_sup F G" then obtain t u where 3: "t \ F \ u \ G \ t \ u \ x" - by auto + using filter_sup_def by auto show "\y\filter_sup F G . \z\filter_sup F G . z \ x \ z \ y" proof fix y assume "y\filter_sup F G" then obtain v w where 4: "v \ F \ w \ G \ v \ w \ y" - by auto + using filter_sup_def by auto let ?z = "(t \ v) \ (u \ w)" have 5: "?z \ x \ ?z \ y" using 3 4 by (meson order.trans inf.cobounded1 inf.cobounded2 inf_mono) have "?z \ filter_sup F G" - using assms 3 4 filter_inf_closed by blast + unfolding filter_sup_def using assms 3 4 filter_inf_closed by blast thus "\z\filter_sup F G . z \ x \ z \ y" using 5 by blast qed qed have "\x\filter_sup F G . \y . x \ y \ y \ filter_sup F G" - using order_trans by blast + unfolding filter_sup_def using order_trans by blast thus ?thesis using 1 2 filter_def by presburger qed lemma filter_sup_left_upper_bound: assumes "filter G" shows "F \ filter_sup F G" proof - from assms obtain y where "y\G" using all_not_in_conv filter_def by auto thus ?thesis - using inf.cobounded1 by blast + unfolding filter_sup_def using inf.cobounded1 by blast qed lemma filter_sup_symmetric: "filter_sup F G = filter_sup G F" - using inf.commute by fastforce + unfolding filter_sup_def using inf.commute by fastforce lemma filter_sup_right_upper_bound: "filter F \ G \ filter_sup F G" using filter_sup_symmetric filter_sup_left_upper_bound by simp lemma filter_sup_least_upper_bound: assumes "filter H" and "F \ H" and "G \ H" shows "filter_sup F G \ H" proof fix x assume "x \ filter_sup F G" then obtain y z where 1: "y \ F \ z \ G \ y \ z \ x" - by auto + using filter_sup_def by auto hence "y \ H \ z \ H" using assms(2-3) by auto hence "y \ z \ H" by (simp add: assms(1) filter_inf_closed) thus "x \ H" using 1 assms(1) filter_def by auto qed lemma filter_sup_left_isotone: "G \ H \ filter_sup G F \ filter_sup H F" - by blast + unfolding filter_sup_def by blast lemma filter_sup_right_isotone: "G \ H \ filter_sup F G \ filter_sup F H" - by blast + unfolding filter_sup_def by blast lemma filter_sup_right_isotone_var: "filter_sup F (G \ H) \ filter_sup F H" - by blast + unfolding filter_sup_def by blast lemma up_dist_inf: "\(x \ y) = filter_sup (\x) (\y)" proof show "\(x \ y) \ filter_sup (\x) (\y)" - by blast + unfolding filter_sup_def by blast next show "filter_sup (\x) (\y) \ \(x \ y)" proof fix z assume "z \ filter_sup (\x) (\y)" then obtain u v where "u\\x \ v\\y \ u \ v \ z" - by auto + using filter_sup_def by auto hence "x \ y \ z" using order.trans inf_mono by blast thus "z \ \(x \ y)" by blast qed qed text \ The following result is part of \cite[Exercise 2.23]{DaveyPriestley2002}. \ lemma filter_inf_filter [simp]: assumes "filter F" shows "filter (\{ y . \z\F . x \ z = y})" proof - let ?G = "\{ y . \z\F . x \ z = y}" have "F \ {}" using assms filter_def by simp hence 1: "?G \ {}" by blast have 2: "is_up_set ?G" by auto { fix y z assume "y \ ?G \ z \ ?G" then obtain v w where "v \ F \ w \ F \ x \ v \ y \ x \ w \ z" by auto hence "v \ w \ F \ x \ (v \ w) \ y \ z" by (meson assms filter_inf_closed order.trans inf.boundedI inf.cobounded1 inf.cobounded2) hence "\u\?G . u \ y \ u \ z" by auto } hence "\x\?G . \y\?G . \z\?G . z \ x \ z \ y" by auto thus ?thesis using 1 2 filter_def by presburger qed end context directed_semilattice_inf begin text \ Set intersection is the meet operation in the lattice of filters. \ lemma filter_inf: assumes "filter F" and "filter G" shows "filter (F \ G)" proof (unfold filter_def, intro conjI) from assms obtain x y where 1: "x\F \ y\G" using all_not_in_conv filter_def by auto from ub obtain z where "x \ z \ y \ z" by auto hence "z \ F \ G" using 1 by (meson assms Int_iff filter_def) thus "F \ G \ {}" by blast next show "is_up_set (F \ G)" by (meson assms Int_iff filter_def) next show "\x\F \ G . \y\F \ G . \z\F \ G . z \ x \ z \ y" by (metis assms Int_iff filter_inf_closed inf.cobounded2 inf.commute) qed end text \ We introduce the following type of filters to instantiate the lattice classes and thereby inherit the results shown about lattices. \ typedef (overloaded) 'a filter = "{ F::'a::order set . filter F }" by (meson mem_Collect_eq up_filter) lemma simp_filter [simp]: "filter (Rep_filter x)" using Rep_filter by simp setup_lifting type_definition_filter text \ The set of filters over a directed semilattice forms a lattice with a greatest element. \ instantiation filter :: (directed_semilattice_inf) bounded_lattice_top begin lift_definition top_filter :: "'a filter" is UNIV by (simp add: filter_univ) lift_definition sup_filter :: "'a filter \ 'a filter \ 'a filter" is filter_sup by (simp add: filter_sup) lift_definition inf_filter :: "'a filter \ 'a filter \ 'a filter" is inter by (simp add: filter_inf) lift_definition less_eq_filter :: "'a filter \ 'a filter \ bool" is subset_eq . lift_definition less_filter :: "'a filter \ 'a filter \ bool" is subset . instance apply intro_classes - apply (simp add: less_eq_filter.rep_eq less_filter.rep_eq inf.less_le_not_le) - apply (simp add: less_eq_filter.rep_eq) - apply (simp add: less_eq_filter.rep_eq) - apply (simp add: Rep_filter_inject less_eq_filter.rep_eq) - apply (simp add: inf_filter.rep_eq less_eq_filter.rep_eq) - apply (simp add: inf_filter.rep_eq less_eq_filter.rep_eq) - apply (simp add: inf_filter.rep_eq less_eq_filter.rep_eq) - apply (simp add: less_eq_filter.rep_eq filter_sup_left_upper_bound sup_filter.rep_eq) - apply (simp add: less_eq_filter.rep_eq filter_sup_right_upper_bound sup_filter.rep_eq) - apply (simp add: less_eq_filter.rep_eq filter_sup_least_upper_bound sup_filter.rep_eq) - by (simp add: less_eq_filter.rep_eq top_filter.rep_eq) + subgoal apply transfer by (simp add: less_le_not_le) + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by (simp add: filter_sup_left_upper_bound) + subgoal apply transfer by (simp add: filter_sup_right_upper_bound) + subgoal apply transfer by (simp add: filter_sup_least_upper_bound) + subgoal apply transfer by simp + done end context bounded_semilattice_inf_top begin abbreviation "filter_complements F G \ filter F \ filter G \ filter_sup F G = UNIV \ F \ G = {top}" end text \ The set of filters over a bounded semilattice forms a bounded lattice. \ instantiation filter :: (bounded_semilattice_inf_top) bounded_lattice begin lift_definition bot_filter :: "'a filter" is "{top}" by simp instance - by intro_classes (simp add: less_eq_filter.rep_eq bot_filter.rep_eq) + apply intro_classes + apply transfer + by simp end context lattice begin lemma up_dist_sup: "\(x \ y) = \x \ \y" by auto end text \ For convenience, the following function injects principal filters into the filter type. We cannot define it in the \order\ class since the type filter requires the sort constraint \order\ that is not available in the class. The result of the function is a filter by lemma \up_filter\. \ abbreviation up_filter :: "'a::order \ 'a filter" where "up_filter x \ Abs_filter (\x)" lemma up_filter_dist_inf: "up_filter ((x::'a::lattice) \ y) = up_filter x \ up_filter y" by (simp add: eq_onp_def sup_filter.abs_eq up_dist_inf) lemma up_filter_dist_sup: "up_filter ((x::'a::lattice) \ y) = up_filter x \ up_filter y" by (metis eq_onp_def inf_filter.abs_eq up_dist_sup up_filter) lemma up_filter_injective: "up_filter x = up_filter y \ x = y" by (metis Abs_filter_inject mem_Collect_eq up_filter up_injective) lemma up_filter_antitone: "x \ y \ up_filter y \ up_filter x" by (metis eq_onp_same_args less_eq_filter.abs_eq up_antitone up_filter) text \ The following definition applies a function to each element of a filter. The subsequent lemma gives conditions under which the result of this application is a filter. \ abbreviation filter_map :: "('a::order \ 'b::order) \ 'a filter \ 'b filter" where "filter_map f F \ Abs_filter (f ` Rep_filter F)" lemma filter_map_filter: assumes "mono f" and "\x y . f x \ y \ (\z . x \ z \ y = f z)" shows "filter (f ` Rep_filter F)" -proof (unfold filter_def, intro conjI) - show "f ` Rep_filter F \ {}" - by (metis empty_is_image filter_def simp_filter) -next - show "\x\f ` Rep_filter F . \y\f ` Rep_filter F . \z\f ` Rep_filter F . z \ x \ z \ y" - proof (intro ballI) - fix x y - assume "x \ f ` Rep_filter F" and "y \ f ` Rep_filter F" - then obtain u v where 1: "x = f u \ u \ Rep_filter F \ y = f v \ v \ Rep_filter F" - by auto - then obtain w where "w \ u \ w \ v \ w \ Rep_filter F" - by (meson filter_def simp_filter) - thus "\z\f ` Rep_filter F . z \ x \ z \ y" - using 1 assms(1) mono_def rev_image_eqI by blast - qed -next - show "is_up_set (f ` Rep_filter F)" - proof - fix x - assume "x \ f ` Rep_filter F" - then obtain u where 1: "x = f u \ u \ Rep_filter F" - by auto - show "\y . x \ y \ y \ f ` Rep_filter F" - proof (rule allI, rule impI) - fix y - assume "x \ y" - hence "f u \ y" - using 1 by simp - then obtain z where "u \ z \ y = f z" - using assms(2) by auto - thus "y \ f ` Rep_filter F" - using 1 by (meson image_iff filter_def simp_filter) - qed - qed -qed + by (simp add: assms inf.filter_map_filter) subsection \Distributive Lattices\ text \ In this section we additionally assume that the underlying order forms a distributive lattice. Then filters form a bounded distributive lattice if the underlying order has a greatest element. Moreover ultrafilters are prime filters. We also prove a lemma of Gr\"atzer and Schmidt about principal filters. \ context distrib_lattice begin lemma filter_sup_left_dist_inf: assumes "filter F" and "filter G" and "filter H" shows "filter_sup F (G \ H) = filter_sup F G \ filter_sup F H" proof show "filter_sup F (G \ H) \ filter_sup F G \ filter_sup F H" - using filter_sup_right_isotone_var by blast + unfolding filter_sup_def using filter_sup_right_isotone_var by blast next show "filter_sup F G \ filter_sup F H \ filter_sup F (G \ H)" proof fix x assume "x \ filter_sup F G \ filter_sup F H" then obtain t u v w where 1: "t \ F \ u \ G \ v \ F \ w \ H \ t \ u \ x \ v \ w \ x" - by auto + using filter_sup_def by auto let ?y = "t \ v" let ?z = "u \ w" have 2: "?y \ F" using 1 by (simp add: assms(1) filter_inf_closed) have 3: "?z \ G \ H" using 1 by (meson assms(2-3) Int_iff filter_def sup_ge1 sup_ge2) have "?y \ ?z = (t \ v \ u) \ (t \ v \ w)" by (simp add: inf_sup_distrib1) also have "... \ (t \ u) \ (v \ w)" by (metis inf.cobounded1 inf.cobounded2 inf.left_idem inf_mono sup.mono) also have "... \ x" using 1 by (simp add: le_supI) finally show "x \ filter_sup F (G \ H)" - using 2 3 by blast + unfolding filter_sup_def using 2 3 by blast qed qed lemma filter_inf_principal_rep: "F \ G = \z \ (\x\F . \y\G . z = x \ y)" by force lemma filter_sup_principal_rep: assumes "filter F" and "filter G" and "filter_sup F G = \z" shows "\x\F . \y\G . z = x \ y" proof - from assms(3) obtain x y where 1: "x\F \ y\G \ x \ y \ z" - using order_refl by blast + unfolding filter_sup_def using order_refl by blast hence 2: "x \ z \ F \ y \ z \ G" by (meson assms(1-2) sup_ge1 filter_def) have "(x \ z) \ (y \ z) = z" using 1 sup_absorb2 sup_inf_distrib2 by fastforce thus ?thesis using 2 by force qed lemma inf_sup_principal_aux: assumes "filter F" and "filter G" and "is_principal_up (filter_sup F G)" and "is_principal_up (F \ G)" shows "is_principal_up F" proof - from assms(3-4) obtain x y where 1: "filter_sup F G = \x \ F \ G = \y" by blast from filter_inf_principal_rep obtain t u where 2: "t\F \ u\G \ y = t \ u" using 1 by meson from filter_sup_principal_rep obtain v w where 3: "v\F \ w\G \ x = v \ w" using 1 by (meson assms(1-2)) have "t \ filter_sup F G \ u \ filter_sup F G" - using 2 inf.cobounded1 inf.cobounded2 by blast + unfolding filter_sup_def using 2 inf.cobounded1 inf.cobounded2 by blast hence "x \ t \ x \ u" using 1 by blast hence 4: "(t \ v) \ (u \ w) = x" using 3 by (simp add: inf.absorb2 inf.assoc inf.left_commute) have "(t \ v) \ (u \ w) \ F \ (t \ v) \ (u \ w) \ G" using 2 3 by (metis (no_types, lifting) assms(1-2) filter_inf_closed sup.cobounded1 sup.cobounded2 filter_def) hence "y \ (t \ v) \ (u \ w)" using 1 Int_iff by blast hence 5: "(t \ v) \ (u \ w) = y" using 2 by (simp add: antisym inf.coboundedI1) have "F = \(t \ v)" proof show "F \ \(t \ v)" proof fix z assume 6: "z \ F" hence "z \ filter_sup F G" - using 2 inf.cobounded1 by blast + unfolding filter_sup_def using 2 inf.cobounded1 by blast hence "x \ z" using 1 by simp hence 7: "(t \ v \ z) \ (u \ w) = x" using 4 by (metis inf.absorb1 inf.assoc inf.commute) have "z \ u \ F \ z \ u \ G \ z \ w \ F \ z \ w \ G" using 2 3 6 by (meson assms(1-2) filter_def sup_ge1 sup_ge2) hence "y \ (z \ u) \ (z \ w)" using 1 Int_iff filter_inf_closed by auto hence 8: "(t \ v \ z) \ (u \ w) = y" using 5 by (metis inf.absorb1 sup.commute sup_inf_distrib2) have "t \ v \ z = t \ v" using 4 5 7 8 relative_equality by blast thus "z \ \(t \ v)" by (simp add: inf.orderI) qed next show "\(t \ v) \ F" proof fix z have 9: "t \ v \ F" using 2 3 by (simp add: assms(1) filter_inf_closed) assume "z \ \(t \ v)" hence "t \ v \ z" by simp thus "z \ F" using assms(1) 9 filter_def by auto qed qed thus ?thesis by blast qed text \ The following result is \cite[Lemma II]{GraetzerSchmidt1958}. If both join and meet of two filters are principal filters, both filters are principal filters. \ lemma inf_sup_principal: assumes "filter F" and "filter G" and "is_principal_up (filter_sup F G)" and "is_principal_up (F \ G)" shows "is_principal_up F \ is_principal_up G" proof - have "filter G \ filter F \ is_principal_up (filter_sup G F) \ is_principal_up (G \ F)" by (simp add: assms Int_commute filter_sup_symmetric) thus ?thesis using assms(3) inf_sup_principal_aux by blast qed lemma filter_sup_absorb_inf: "filter F \ filter G \ filter_sup (F \ G) G = G" by (simp add: filter_inf filter_sup_least_upper_bound filter_sup_left_upper_bound filter_sup_symmetric subset_antisym) lemma filter_inf_absorb_sup: "filter F \ filter G \ filter_sup F G \ G = G" apply (rule subset_antisym) apply simp by (simp add: filter_sup_right_upper_bound) lemma filter_inf_right_dist_sup: assumes "filter F" and "filter G" and "filter H" shows "filter_sup F G \ H = filter_sup (F \ H) (G \ H)" proof - have "filter_sup (F \ H) (G \ H) = filter_sup (F \ H) G \ filter_sup (F \ H) H" by (simp add: assms filter_sup_left_dist_inf filter_inf) also have "... = filter_sup (F \ H) G \ H" using assms(1,3) filter_sup_absorb_inf by simp also have "... = filter_sup F G \ filter_sup G H \ H" using assms filter_sup_left_dist_inf filter_sup_symmetric by simp also have "... = filter_sup F G \ H" by (simp add: assms(2-3) filter_inf_absorb_sup semilattice_inf_class.inf_assoc) finally show ?thesis by simp qed text \ The following result generalises \cite[10.11]{DaveyPriestley2002} to distributive lattices as remarked after that section. \ lemma ultra_filter_prime: assumes "ultra_filter F" shows "prime_filter F" proof - { fix x y assume 1: "x \ y \ F \ x \ F" let ?G = "\{ z . \w\F . x \ w = z }" have 2: "filter ?G" using assms filter_inf_filter by simp have "x \ ?G" using 1 by auto hence 3: "F \ ?G" using 1 by auto have "F \ ?G" using inf_le2 order_trans by blast hence "?G = UNIV" using 2 3 assms by blast then obtain z where 4: "z \ F \ x \ z \ y" by blast hence "y \ z = (x \ y) \ z" by (simp add: inf_sup_distrib2 sup_absorb2) also have "... \ F" using 1 4 assms filter_inf_closed by auto finally have "y \ F" using assms by (simp add: filter_def) } thus ?thesis using assms by blast qed +lemma up_dist_inf_inter: + assumes "is_up_set S" + shows "\(x \ y) \ S = filter_sup (\x \ S) (\y \ S) \ S" +proof + show "\(x \ y) \ S \ filter_sup (\x \ S) (\y \ S) \ S" + proof + fix z + let ?x = "x \ z" + let ?y = "y \ z" + assume "z \ \(x \ y) \ S" + hence 1: "x \ y \ z \ z \ S" + by auto + hence "?x \ (\x \ S) \ ?y \ (\y \ S) \ ?x \ ?y \ z" + using assms sup_absorb2 sup_inf_distrib2 by fastforce + thus "z \ filter_sup (\x \ S) (\y \ S) \ S" + using filter_sup_def 1 by fastforce + qed +next + show "filter_sup (\x \ S) (\y \ S) \ S \ \(x \ y) \ S" + proof + fix z + assume "z \ filter_sup (\x \ S) (\y \ S) \ S" + then obtain u v where 2: "u\\x \ v\\y \ u \ v \ z \ z \ S" + using filter_sup_def by auto + hence "x \ y \ z" + using order.trans inf_mono by blast + thus "z \ \(x \ y) \ S" + using 2 by blast + qed +qed + end context distrib_lattice_bot begin lemma prime_filter: "proper_filter F \ \G . prime_filter G \ F \ G" by (metis ultra_filter ultra_filter_prime) end context distrib_lattice_top begin lemma complemented_filter_inf_principal: assumes "filter_complements F G" shows "is_principal_up (F \ \x)" proof - have 1: "filter F \ filter G" by (simp add: assms) hence 2: "filter (F \ \x) \ filter (G \ \x)" by (simp add: filter_inf) have "(F \ \x) \ (G \ \x) = {top}" using assms Int_assoc Int_insert_left_if1 inf_bot_left inf_sup_aci(3) top_in_upset inf.idem by auto hence 3: "is_principal_up ((F \ \x) \ (G \ \x))" using up_top by blast have "filter_sup (F \ \x) (G \ \x) = filter_sup F G \ \x" using 1 filter_inf_right_dist_sup up_filter by auto also have "... = \x" by (simp add: assms) finally have "is_principal_up (filter_sup (F \ \x) (G \ \x))" by auto thus ?thesis using 1 2 3 inf_sup_principal_aux by blast qed end text \ The set of filters over a distributive lattice with a greatest element forms a bounded distributive lattice. \ instantiation filter :: (distrib_lattice_top) bounded_distrib_lattice begin instance -proof - fix x y z :: "'a filter" - have "Rep_filter (x \ (y \ z)) = filter_sup (Rep_filter x) (Rep_filter (y \ z))" - by (simp add: sup_filter.rep_eq) - also have "... = filter_sup (Rep_filter x) (Rep_filter y \ Rep_filter z)" - by (simp add: inf_filter.rep_eq) - also have "... = filter_sup (Rep_filter x) (Rep_filter y) \ filter_sup (Rep_filter x) (Rep_filter z)" - by (simp add: filter_sup_left_dist_inf) - also have "... = Rep_filter (x \ y) \ Rep_filter (x \ z)" - by (simp add: sup_filter.rep_eq) - also have "... = Rep_filter ((x \ y) \ (x \ z))" - by (simp add: inf_filter.rep_eq) - finally show "x \ (y \ z) = (x \ y) \ (x \ z)" - by (simp add: Rep_filter_inject) -qed + apply intro_classes + apply transfer + by (simp add: filter_sup_left_dist_inf) end end diff --git a/thys/Stone_Algebras/Lattice_Basics.thy b/thys/Stone_Algebras/Lattice_Basics.thy --- a/thys/Stone_Algebras/Lattice_Basics.thy +++ b/thys/Stone_Algebras/Lattice_Basics.thy @@ -1,447 +1,558 @@ (* Title: Lattice Basics Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Lattice Basics\ text \ This theory provides notations, basic definitions and facts of lattice-related structures used throughout the subsequent development. \ theory Lattice_Basics imports Main begin +subsection \General Facts and Notations\ + text \ The following results extend basic Isabelle/HOL facts. \ +lemma imp_as_conj: + assumes "P x \ Q x" + shows "P x \ Q x \ P x" + using assms by auto + lemma if_distrib_2: "f (if c then x else y) (if c then z else w) = (if c then f x z else f y w)" by simp lemma left_invertible_inj: "(\x . g (f x) = x) \ inj f" by (metis injI) lemma invertible_bij: assumes "\x . g (f x) = x" and "\y . f (g y) = y" shows "bij f" by (metis assms bijI') lemma finite_ne_subset_induct [consumes 3, case_names singleton insert]: assumes "finite F" and "F \ {}" and "F \ S" and singleton: "\x . P {x}" and insert: "\x F . finite F \ F \ {} \ F \ S \ x \ S \ x \ F \ P F \ P (insert x F)" shows "P F" using assms(1-3) apply (induct rule: finite_ne_induct) apply (simp add: singleton) by (simp add: insert) lemma finite_set_of_finite_funs_pred: assumes "finite { x::'a . True }" and "finite { y::'b . P y }" shows "finite { f . (\x::'a . P (f x)) }" using assms finite_set_of_finite_funs by force text \ We use the following notations for the join, meet and complement operations. Changing the precedence of the unary complement allows us to write terms like \--x\ instead of \-(-x)\. \ context sup begin notation sup (infixl "\" 65) definition additive :: "('a \ 'a) \ bool" where "additive f \ \x y . f (x \ y) = f x \ f y" end context inf begin notation inf (infixl "\" 67) end context uminus begin no_notation uminus ("- _" [81] 80) notation uminus ("- _" [80] 80) end +subsection \Orders\ + text \ We use the following definition of monotonicity for operations defined in classes. The standard \mono\ places a sort constraint on the target type. We also give basic properties of Galois connections and lift orders to functions. \ context ord begin definition isotone :: "('a \ 'a) \ bool" where "isotone f \ \x y . x \ y \ f x \ f y" definition galois :: "('a \ 'a) \ ('a \ 'a) \ bool" where "galois l u \ \x y . l x \ y \ x \ u y" definition lifted_less_eq :: "('a \ 'a) \ ('a \ 'a) \ bool" ("(_ \\ _)" [51, 51] 50) where "f \\ g \ \x . f x \ g x" end context order begin lemma order_lesseq_imp: "(\z . x \ z \ y \ z) \ y \ x" using order_trans by blast lemma galois_char: "galois l u \ (\x . x \ u (l x)) \ (\x . l (u x) \ x) \ isotone l \ isotone u" apply (rule iffI) apply (metis (full_types) galois_def isotone_def order_refl order_trans) using galois_def isotone_def order_trans by blast lemma galois_closure: "galois l u \ l x = l (u (l x)) \ u x = u (l (u x))" by (simp add: galois_char isotone_def antisym) lemma lifted_reflexive: "f = g \ f \\ g" by (simp add: lifted_less_eq_def) lemma lifted_transitive: "f \\ g \ g \\ h \ f \\ h" using lifted_less_eq_def order_trans by blast lemma lifted_antisymmetric: "f \\ g \ g \\ f \ f = g" by (metis (full_types) antisym ext lifted_less_eq_def) text \ If the image of a finite non-empty set under \f\ is a totally ordered, there is an element that minimises the value of \f\. \ lemma finite_set_minimal: assumes "finite s" and "s \ {}" and "\x\s . \y\s . f x \ f y \ f y \ f x" shows "\m\s . \z\s . f m \ f z" apply (rule finite_ne_subset_induct[where S=s]) apply (rule assms(1)) apply (rule assms(2)) apply simp apply simp by (metis assms(3) insert_iff order_trans subsetD) end +subsection \Semilattices\ + text \ The following are basic facts in semilattices. \ context semilattice_sup begin lemma sup_left_isotone: "x \ y \ x \ z \ y \ z" using sup.mono by blast lemma sup_right_isotone: "x \ y \ z \ x \ z \ y" using sup.mono by blast lemma sup_left_divisibility: "x \ y \ (\z . x \ z = y)" using sup.absorb2 sup.cobounded1 by blast lemma sup_right_divisibility: "x \ y \ (\z . z \ x = y)" by (metis sup.cobounded2 sup.orderE) lemma sup_same_context: "x \ y \ z \ y \ x \ z \ x \ z = y \ z" by (simp add: le_iff_sup sup_left_commute) lemma sup_relative_same_increasing: "x \ y \ x \ z = x \ w \ y \ z = y \ w" using sup.assoc sup_right_divisibility by auto end text \ Every bounded semilattice is a commutative monoid. Finite sums defined in commutative monoids are available via the following sublocale. \ context bounded_semilattice_sup_bot begin sublocale sup_monoid: comm_monoid_add where plus = sup and zero = bot apply unfold_locales apply (simp add: sup_assoc) apply (simp add: sup_commute) by simp end context semilattice_inf begin lemma inf_same_context: "x \ y \ z \ y \ x \ z \ x \ z = y \ z" using antisym by auto end text \ The following class requires only the existence of upper bounds, which is a property common to bounded semilattices and (not necessarily bounded) lattices. We use it in our development of filters. \ class directed_semilattice_inf = semilattice_inf + assumes ub: "\z . x \ z \ y \ z" text \ We extend the \inf\ sublocale, which dualises the order in semilattices, to bounded semilattices. \ context bounded_semilattice_inf_top begin subclass directed_semilattice_inf apply unfold_locales using top_greatest by blast sublocale inf: bounded_semilattice_sup_bot where sup = inf and less_eq = greater_eq and less = greater and bot = top by unfold_locales (simp_all add: less_le_not_le) end +subsection \Lattices\ + context lattice begin subclass directed_semilattice_inf apply unfold_locales using sup_ge1 sup_ge2 by blast definition dual_additive :: "('a \ 'a) \ bool" where "dual_additive f \ \x y . f (x \ y) = f x \ f y" end text \ Not every bounded lattice has complements, but two elements might still be complements of each other as captured in the following definition. In this situation we can apply, for example, the shunting property shown below. We introduce most definitions using the \abbreviation\ command. \ context bounded_lattice begin abbreviation "complement x y \ x \ y = top \ x \ y = bot" lemma complement_symmetric: "complement x y \ complement y x" by (simp add: inf.commute sup.commute) definition conjugate :: "('a \ 'a) \ ('a \ 'a) \ bool" where "conjugate f g \ \x y . f x \ y = bot \ x \ g y = bot" end class dense_lattice = bounded_lattice + assumes bot_meet_irreducible: "x \ y = bot \ x = bot \ y = bot" context distrib_lattice begin lemma relative_equality: "x \ z = y \ z \ x \ z = y \ z \ x = y" by (metis inf.commute inf_sup_absorb inf_sup_distrib2) end text \ Distributive lattices with a greatest element are widely used in the construction theorem for Stone algebras. \ class distrib_lattice_bot = bounded_lattice_bot + distrib_lattice class distrib_lattice_top = bounded_lattice_top + distrib_lattice class bounded_distrib_lattice = bounded_lattice + distrib_lattice begin subclass distrib_lattice_bot .. subclass distrib_lattice_top .. lemma complement_shunting: assumes "complement z w" shows "z \ x \ y \ x \ w \ y" proof assume 1: "z \ x \ y" have "x = (z \ w) \ x" by (simp add: assms) also have "... \ y \ (w \ x)" using 1 sup.commute sup.left_commute inf_sup_distrib2 sup_right_divisibility by fastforce also have "... \ w \ y" by (simp add: inf.coboundedI1) finally show "x \ w \ y" . next assume "x \ w \ y" hence "z \ x \ z \ (w \ y)" using inf.sup_right_isotone by auto also have "... = z \ y" by (simp add: assms inf_sup_distrib1) also have "... \ y" by simp finally show "z \ x \ y" . qed end +subsection \Linear Orders\ + text \ We next consider lattices with a linear order structure. In such lattices, join and meet are selective operations, which give the maximum and the minimum of two elements, respectively. Moreover, the lattice is automatically distributive. \ class bounded_linorder = linorder + order_bot + order_top class linear_lattice = lattice + linorder begin lemma max_sup: "max x y = x \ y" by (metis max.boundedI max.cobounded1 max.cobounded2 sup_unique) lemma min_inf: "min x y = x \ y" by (simp add: inf.absorb1 inf.absorb2 min_def) lemma sup_inf_selective: "(x \ y = x \ x \ y = y) \ (x \ y = y \ x \ y = x)" by (meson inf.absorb1 inf.absorb2 le_cases sup.absorb1 sup.absorb2) lemma sup_selective: "x \ y = x \ x \ y = y" using sup_inf_selective by blast lemma inf_selective: "x \ y = x \ x \ y = y" using sup_inf_selective by blast subclass distrib_lattice apply unfold_locales by (metis inf_selective antisym distrib_sup_le inf.commute inf_le2) lemma sup_less_eq: "x \ y \ z \ x \ y \ x \ z" by (metis le_supI1 le_supI2 sup_selective) lemma inf_less_eq: "x \ y \ z \ x \ z \ y \ z" by (metis inf.coboundedI1 inf.coboundedI2 inf_selective) lemma sup_inf_sup: "x \ y = (x \ y) \ (x \ y)" by (metis sup_commute sup_inf_absorb sup_left_commute) end text \ The following class derives additional properties if the linear order of the lattice has a least and a greatest element. \ class linear_bounded_lattice = bounded_lattice + linorder begin subclass linear_lattice .. subclass bounded_linorder .. subclass bounded_distrib_lattice .. lemma sup_dense: "x \ top \ y \ top \ x \ y \ top" by (metis sup_selective) lemma inf_dense: "x \ bot \ y \ bot \ x \ y \ bot" by (metis inf_selective) lemma sup_not_bot: "x \ bot \ x \ y \ bot" by simp lemma inf_not_top: "x \ top \ x \ y \ top" by simp subclass dense_lattice apply unfold_locales using inf_dense by blast end text \ Every bounded linear order can be expanded to a bounded lattice. Join and meet are maximum and minimum, respectively. \ class linorder_lattice_expansion = bounded_linorder + sup + inf + assumes sup_def [simp]: "x \ y = max x y" assumes inf_def [simp]: "x \ y = min x y" begin subclass linear_bounded_lattice apply unfold_locales by auto end +subsection \Non-trivial Algebras\ + text \ Some results, such as the existence of certain filters, require that the algebras are not trivial. This is not an assumption of the order and lattice classes that come with Isabelle/HOL; for example, \bot = top\ may hold in bounded lattices. \ class non_trivial = assumes consistent: "\x y . x \ y" class non_trivial_order = non_trivial + order class non_trivial_order_bot = non_trivial_order + order_bot class non_trivial_bounded_order = non_trivial_order_bot + order_top begin lemma bot_not_top: "bot \ top" proof - from consistent obtain x y :: 'a where "x \ y" by auto thus ?thesis by (metis bot_less top.extremum_strict) qed end +subsection \Homomorphisms\ + +text \ +This section gives definitions of lattice homomorphisms and isomorphisms and basic properties. +\ + +class sup_inf_top_bot_uminus = sup + inf + top + bot + uminus +class sup_inf_top_bot_uminus_ord = sup_inf_top_bot_uminus + ord + +context boolean_algebra +begin + +subclass sup_inf_top_bot_uminus_ord . + end +abbreviation sup_homomorphism :: "('a::sup \ 'b::sup) \ bool" + where "sup_homomorphism f \ \x y . f (x \ y) = f x \ f y" + +abbreviation inf_homomorphism :: "('a::inf \ 'b::inf) \ bool" + where "inf_homomorphism f \ \x y . f (x \ y) = f x \ f y" + +abbreviation bot_homomorphism :: "('a::bot \ 'b::bot) \ bool" + where "bot_homomorphism f \ f bot = bot" + +abbreviation top_homomorphism :: "('a::top \ 'b::top) \ bool" + where "top_homomorphism f \ f top = top" + +abbreviation minus_homomorphism :: "('a::minus \ 'b::minus) \ bool" + where "minus_homomorphism f \ \x y . f (x - y) = f x - f y" + +abbreviation uminus_homomorphism :: "('a::uminus \ 'b::uminus) \ bool" + where "uminus_homomorphism f \ \x . f (-x) = -f x" + +abbreviation sup_inf_homomorphism :: "('a::{sup,inf} \ 'b::{sup,inf}) \ bool" + where "sup_inf_homomorphism f \ sup_homomorphism f \ inf_homomorphism f" + +abbreviation sup_inf_top_homomorphism :: "('a::{sup,inf,top} \ 'b::{sup,inf,top}) \ bool" + where "sup_inf_top_homomorphism f \ sup_inf_homomorphism f \ top_homomorphism f" + +abbreviation sup_inf_top_bot_homomorphism :: "('a::{sup,inf,top,bot} \ 'b::{sup,inf,top,bot}) \ bool" + where "sup_inf_top_bot_homomorphism f \ sup_inf_top_homomorphism f \ bot_homomorphism f" + +abbreviation bounded_lattice_homomorphism :: "('a::bounded_lattice \ 'b::bounded_lattice) \ bool" + where "bounded_lattice_homomorphism f \ sup_inf_top_bot_homomorphism f" + +abbreviation sup_inf_top_bot_uminus_homomorphism :: "('a::sup_inf_top_bot_uminus \ 'b::sup_inf_top_bot_uminus) \ bool" + where "sup_inf_top_bot_uminus_homomorphism f \ sup_inf_top_bot_homomorphism f \ uminus_homomorphism f" + +abbreviation sup_inf_top_bot_uminus_ord_homomorphism :: "('a::sup_inf_top_bot_uminus_ord \ 'b::sup_inf_top_bot_uminus_ord) \ bool" + where "sup_inf_top_bot_uminus_ord_homomorphism f \ sup_inf_top_bot_uminus_homomorphism f \ (\x y . x \ y \ f x \ f y)" + +abbreviation sup_inf_top_isomorphism :: "('a::{sup,inf,top} \ 'b::{sup,inf,top}) \ bool" + where "sup_inf_top_isomorphism f \ sup_inf_top_homomorphism f \ bij f" + +abbreviation bounded_lattice_top_isomorphism :: "('a::bounded_lattice_top \ 'b::bounded_lattice_top) \ bool" + where "bounded_lattice_top_isomorphism f \ sup_inf_top_isomorphism f" + +abbreviation sup_inf_top_bot_uminus_isomorphism :: "('a::sup_inf_top_bot_uminus \ 'b::sup_inf_top_bot_uminus) \ bool" + where "sup_inf_top_bot_uminus_isomorphism f \ sup_inf_top_bot_uminus_homomorphism f \ bij f" + +abbreviation boolean_algebra_isomorphism :: "('a::boolean_algebra \ 'b::boolean_algebra) \ bool" + where "boolean_algebra_isomorphism f \ sup_inf_top_bot_uminus_isomorphism f \ minus_homomorphism f" + +lemma sup_homomorphism_mono: + "sup_homomorphism (f::'a::semilattice_sup \ 'b::semilattice_sup) \ mono f" + by (metis le_iff_sup monoI) + +lemma sup_isomorphism_ord_isomorphism: + assumes "sup_homomorphism (f::'a::semilattice_sup \ 'b::semilattice_sup)" + and "bij f" + shows "x \ y \ f x \ f y" +proof + assume "x \ y" + thus "f x \ f y" + by (metis assms(1) le_iff_sup) +next + assume "f x \ f y" + hence "f (x \ y) = f y" + by (simp add: assms(1) le_iff_sup) + hence "x \ y = y" + by (metis injD bij_is_inj assms(2)) + thus "x \ y" + by (simp add: le_iff_sup) +qed + +lemma minus_homomorphism_default: + assumes "\x y::'a::{inf,minus,uminus} . x - y = x \ -y" + and "\x y::'b::{inf,minus,uminus} . x - y = x \ -y" + and "inf_homomorphism (f::'a \ 'b)" + and "uminus_homomorphism f" + shows "minus_homomorphism f" + by (simp add: assms) + +end + diff --git a/thys/Stone_Algebras/P_Algebras.thy b/thys/Stone_Algebras/P_Algebras.thy --- a/thys/Stone_Algebras/P_Algebras.thy +++ b/thys/Stone_Algebras/P_Algebras.thy @@ -1,1319 +1,1324 @@ (* Title: Pseudocomplemented Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Pseudocomplemented Algebras\ text \ This theory expands lattices with a pseudocomplement operation. In particular, we consider the following algebraic structures: \begin{itemize} \item pseudocomplemented lattices (p-algebras) \item pseudocomplemented distributive lattices (distributive p-algebras) \item Stone algebras \item Heyting semilattices \item Heyting lattices \item Heyting algebras \item Heyting-Stone algebras \item Brouwer algebras \item Boolean algebras \end{itemize} Most of these structures and many results in this theory are discussed in \cite{BalbesDwinger1974,Birkhoff1967,Blyth2005,Curry1977,Graetzer1971,Maddux1996}. \ theory P_Algebras imports Lattice_Basics begin subsection \P-Algebras\ text \ In this section we add a pseudocomplement operation to lattices and to distributive lattices. \ subsubsection \Pseudocomplemented Lattices\ text \ The pseudocomplement of an element \y\ is the greatest element whose meet with \y\ is the least element of the lattice. \ class p_algebra = bounded_lattice + uminus + assumes pseudo_complement: "x \ y = bot \ x \ -y" begin +subclass sup_inf_top_bot_uminus_ord . + text \ Regular elements and dense elements are frequently used in pseudocomplemented algebras. \ abbreviation "regular x \ x = --x" abbreviation "dense x \ -x = bot" abbreviation "complemented x \ \y . x \ y = bot \ x \ y = top" abbreviation "in_p_image x \ \y . x = -y" abbreviation "selection s x \ s = --s \ x" abbreviation "dense_elements \ { x . dense x }" abbreviation "regular_elements \ { x . in_p_image x }" lemma p_bot [simp]: "-bot = top" using inf_top.left_neutral pseudo_complement top_unique by blast lemma p_top [simp]: "-top = bot" by (metis eq_refl inf_top.comm_neutral pseudo_complement) text \ The pseudocomplement satisfies the following half of the requirements of a complement. \ lemma inf_p [simp]: "x \ -x = bot" using inf.commute pseudo_complement by fastforce lemma p_inf [simp]: "-x \ x = bot" by (simp add: inf_commute) lemma pp_inf_p: "--x \ -x = bot" by simp text \ The double complement is a closure operation. \ lemma pp_increasing: "x \ --x" using inf_p pseudo_complement by blast lemma ppp [simp]: "---x = -x" by (metis antisym inf.commute order_trans pseudo_complement pp_increasing) lemma pp_idempotent: "----x = --x" by simp lemma regular_in_p_image_iff: "regular x \ in_p_image x" by auto lemma pseudo_complement_pp: "x \ y = bot \ --x \ -y" by (metis inf_commute pseudo_complement ppp) lemma p_antitone: "x \ y \ -y \ -x" by (metis inf_commute order_trans pseudo_complement pp_increasing) lemma p_antitone_sup: "-(x \ y) \ -x" by (simp add: p_antitone) lemma p_antitone_inf: "-x \ -(x \ y)" by (simp add: p_antitone) lemma p_antitone_iff: "x \ -y \ y \ -x" using order_lesseq_imp p_antitone pp_increasing by blast lemma pp_isotone: "x \ y \ --x \ --y" by (simp add: p_antitone) lemma pp_isotone_sup: "--x \ --(x \ y)" by (simp add: p_antitone) lemma pp_isotone_inf: "--(x \ y) \ --x" by (simp add: p_antitone) text \ One of De Morgan's laws holds in pseudocomplemented lattices. \ lemma p_dist_sup [simp]: "-(x \ y) = -x \ -y" apply (rule antisym) apply (simp add: p_antitone) using inf_le1 inf_le2 le_sup_iff p_antitone_iff by blast lemma p_supdist_inf: "-x \ -y \ -(x \ y)" by (simp add: p_antitone) lemma pp_dist_pp_sup [simp]: "--(--x \ --y) = --(x \ y)" by simp lemma p_sup_p [simp]: "-(x \ -x) = bot" by simp lemma pp_sup_p [simp]: "--(x \ -x) = top" by simp lemma dense_pp: "dense x \ --x = top" by (metis p_bot p_top ppp) lemma dense_sup_p: "dense (x \ -x)" by simp lemma regular_char: "regular x \ (\y . x = -y)" by auto lemma pp_inf_bot_iff: "x \ y = bot \ --x \ y = bot" by (simp add: pseudo_complement_pp) text \ Weak forms of the shunting property hold. Most require a pseudocomplemented element on the right-hand side. \ lemma p_shunting_swap: "x \ y \ -z \ x \ z \ -y" by (metis inf_assoc inf_commute pseudo_complement) lemma pp_inf_below_iff: "x \ y \ -z \ --x \ y \ -z" by (simp add: inf_commute p_shunting_swap) lemma p_inf_pp [simp]: "-(x \ --y) = -(x \ y)" apply (rule antisym) apply (simp add: inf.coboundedI2 p_antitone pp_increasing) using inf_commute p_antitone_iff pp_inf_below_iff by auto lemma p_inf_pp_pp [simp]: "-(--x \ --y) = -(x \ y)" by (simp add: inf_commute) lemma regular_closed_inf: "regular x \ regular y \ regular (x \ y)" by (metis p_dist_sup ppp) lemma regular_closed_p: "regular (-x)" by simp lemma regular_closed_pp: "regular (--x)" by simp lemma regular_closed_bot: "regular bot" by simp lemma regular_closed_top: "regular top" by simp lemma pp_dist_inf [simp]: "--(x \ y) = --x \ --y" by (metis p_dist_sup p_inf_pp_pp ppp) lemma inf_import_p [simp]: "x \ -(x \ y) = x \ -y" apply (rule antisym) using p_shunting_swap apply fastforce using inf.sup_right_isotone p_antitone by auto text \ Pseudocomplements are unique. \ lemma p_unique: "(\x . x \ y = bot \ x \ z) \ z = -y" using inf.eq_iff pseudo_complement by auto lemma maddux_3_5: "x \ x = x \ -(y \ -y)" by simp lemma shunting_1_pp: "x \ --y \ x \ -y = bot" by (simp add: pseudo_complement) lemma pp_pp_inf_bot_iff: "x \ y = bot \ --x \ --y = bot" by (simp add: pseudo_complement_pp) lemma inf_pp_semi_commute: "x \ --y \ --(x \ y)" using inf.eq_refl p_antitone_iff p_inf_pp by presburger lemma inf_pp_commute: "--(--x \ y) = --x \ --y" by simp lemma sup_pp_semi_commute: "x \ --y \ --(x \ y)" by (simp add: p_antitone_iff) lemma regular_sup: "regular z \ (x \ z \ y \ z \ --(x \ y) \ z)" apply (rule iffI) apply (metis le_supI pp_isotone) using dual_order.trans sup_ge2 pp_increasing pp_isotone_sup by blast lemma dense_closed_inf: "dense x \ dense y \ dense (x \ y)" by (simp add: dense_pp) lemma dense_closed_sup: "dense x \ dense y \ dense (x \ y)" by simp lemma dense_closed_pp: "dense x \ dense (--x)" by simp lemma dense_closed_top: "dense top" by simp lemma dense_up_closed: "dense x \ x \ y \ dense y" using dense_pp top_le pp_isotone by auto lemma regular_dense_top: "regular x \ dense x \ x = top" using p_bot by blast lemma selection_char: "selection s x \ (\y . s = -y \ x)" by (metis inf_import_p inf_commute regular_closed_p) lemma selection_closed_inf: "selection s x \ selection t x \ selection (s \ t) x" by (metis inf_assoc inf_commute inf_idem pp_dist_inf) lemma selection_closed_pp: "regular x \ selection s x \ selection (--s) x" by (metis pp_dist_inf) lemma selection_closed_bot: "selection bot x" by simp lemma selection_closed_id: "selection x x" using inf.le_iff_sup pp_increasing by auto text \ Conjugates are usually studied for Boolean algebras, however, some of their properties generalise to pseudocomplemented algebras. \ lemma conjugate_unique_p: assumes "conjugate f g" and "conjugate f h" shows "uminus \ g = uminus \ h" proof - have "\x y . x \ g y = bot \ x \ h y = bot" using assms conjugate_def inf.commute by simp hence "\x y . x \ -(g y) \ x \ -(h y)" using inf.commute pseudo_complement by simp hence "\y . -(g y) = -(h y)" using eq_iff by blast thus ?thesis by auto qed lemma conjugate_symmetric: "conjugate f g \ conjugate g f" by (simp add: conjugate_def inf_commute) lemma additive_isotone: "additive f \ isotone f" by (metis additive_def isotone_def le_iff_sup) lemma dual_additive_antitone: assumes "dual_additive f" shows "isotone (uminus \ f)" proof - have "\x y . f (x \ y) \ f x" using assms dual_additive_def by simp hence "\x y . x \ y \ f y \ f x" by (metis sup_absorb2) hence "\x y . x \ y \ -(f x) \ -(f y)" by (simp add: p_antitone) thus ?thesis by (simp add: isotone_def) qed lemma conjugate_dual_additive: assumes "conjugate f g" shows "dual_additive (uminus \ f)" proof - have 1: "\x y z . -z \ -(f (x \ y)) \ -z \ -(f x) \ -z \ -(f y)" proof (intro allI) fix x y z have "(-z \ -(f (x \ y))) = (f (x \ y) \ -z = bot)" by (simp add: p_antitone_iff pseudo_complement) also have "... = ((x \ y) \ g(-z) = bot)" using assms conjugate_def by auto also have "... = (x \ y \ -(g(-z)))" by (simp add: pseudo_complement) also have "... = (x \ -(g(-z)) \ y \ -(g(-z)))" by (simp add: le_sup_iff) also have "... = (x \ g(-z) = bot \ y \ g(-z) = bot)" by (simp add: pseudo_complement) also have "... = (f x \ -z = bot \ f y \ -z = bot)" using assms conjugate_def by auto also have "... = (-z \ -(f x) \ -z \ -(f y))" by (simp add: p_antitone_iff pseudo_complement) finally show "-z \ -(f (x \ y)) \ -z \ -(f x) \ -z \ -(f y)" by simp qed have "\x y . -(f (x \ y)) = -(f x) \ -(f y)" proof (intro allI) fix x y have "-(f x) \ -(f y) = --(-(f x) \ -(f y))" by simp hence "-(f x) \ -(f y) \ -(f (x \ y))" using 1 by (metis inf_le1 inf_le2) thus "-(f (x \ y)) = -(f x) \ -(f y)" using 1 antisym by fastforce qed thus ?thesis using dual_additive_def by simp qed lemma conjugate_isotone_pp: "conjugate f g \ isotone (uminus \ uminus \ f)" by (simp add: comp_assoc conjugate_dual_additive dual_additive_antitone) lemma conjugate_char_1_pp: "conjugate f g \ (\x y . f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x)" proof assume 1: "conjugate f g" show "\x y . f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x" proof (intro allI) fix x y have 2: "f(x \ -(g y)) \ -y" using 1 by (simp add: conjugate_def pseudo_complement) have "f(x \ -(g y)) \ --f(x \ -(g y))" by (simp add: pp_increasing) also have "... \ --f x" using 1 conjugate_isotone_pp isotone_def by simp finally have 3: "f(x \ -(g y)) \ --f x \ -y" using 2 by simp have 4: "isotone (uminus \ uminus \ g)" using 1 conjugate_isotone_pp conjugate_symmetric by auto have 5: "g(y \ -(f x)) \ -x" using 1 by (metis conjugate_def inf.cobounded2 inf_commute pseudo_complement) have "g(y \ -(f x)) \ --g(y \ -(f x))" by (simp add: pp_increasing) also have "... \ --g y" using 4 isotone_def by auto finally have "g(y \ -(f x)) \ --g y \ -x" using 5 by simp thus "f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x" using 3 by simp qed next assume 6: "\x y . f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x" hence 7: "\x y . f x \ y = bot \ x \ g y = bot" by (metis inf.le_iff_sup inf.le_sup_iff inf_commute pseudo_complement) have "\x y . x \ g y = bot \ f x \ y = bot" using 6 by (metis inf.le_iff_sup inf.le_sup_iff inf_commute pseudo_complement) thus "conjugate f g" using 7 conjugate_def by auto qed lemma conjugate_char_1_isotone: "conjugate f g \ isotone f \ isotone g \ f(x \ -(g y)) \ f x \ -y \ g(y \ -(f x)) \ g y \ -x" by (simp add: conjugate_char_1_pp ord.isotone_def) lemma dense_lattice_char_1: "(\x y . x \ y = bot \ x = bot \ y = bot) \ (\x . x \ bot \ dense x)" by (metis inf_top.left_neutral p_bot p_inf pp_inf_bot_iff) lemma dense_lattice_char_2: "(\x y . x \ y = bot \ x = bot \ y = bot) \ (\x . regular x \ x = bot \ x = top)" by (metis dense_lattice_char_1 inf_top.left_neutral p_inf regular_closed_p regular_closed_top) lemma restrict_below_Rep_eq: "x \ --y \ z \ x \ y = x \ z \ y" by (metis inf.absorb2 inf.commute inf.left_commute pp_increasing) (* lemma p_inf_sup_below: "-x \ (x \ y) \ y" nitpick [expect=genuine] oops lemma complement_p: "x \ y = bot \ x \ y = top \ -x = y" nitpick [expect=genuine] oops lemma complemented_regular: "complemented x \ regular x" nitpick [expect=genuine] oops *) end text \ The following class gives equational axioms for the pseudocomplement operation. \ class p_algebra_eq = bounded_lattice + uminus + assumes p_bot_eq: "-bot = top" and p_top_eq: "-top = bot" and inf_import_p_eq: "x \ -(x \ y) = x \ -y" begin lemma inf_p_eq: "x \ -x = bot" by (metis inf_bot_right inf_import_p_eq inf_top_right p_top_eq) subclass p_algebra apply unfold_locales apply (rule iffI) apply (metis inf.orderI inf_import_p_eq inf_top.right_neutral p_bot_eq) by (metis (full_types) inf.left_commute inf.orderE inf_bot_right inf_commute inf_p_eq) end subsubsection \Pseudocomplemented Distributive Lattices\ text \ We obtain further properties if we assume that the lattice operations are distributive. \ class pd_algebra = p_algebra + bounded_distrib_lattice begin lemma p_inf_sup_below: "-x \ (x \ y) \ y" by (simp add: inf_sup_distrib1) lemma pp_inf_sup_p [simp]: "--x \ (x \ -x) = x" using inf.absorb2 inf_sup_distrib1 pp_increasing by auto lemma complement_p: "x \ y = bot \ x \ y = top \ -x = y" by (metis pseudo_complement inf.commute inf_top.left_neutral sup.absorb_iff1 sup.commute sup_bot.right_neutral sup_inf_distrib2 p_inf) lemma complemented_regular: "complemented x \ regular x" using complement_p inf.commute sup.commute by fastforce lemma regular_inf_dense: "\y z . regular y \ dense z \ x = y \ z" by (metis pp_inf_sup_p dense_sup_p ppp) lemma maddux_3_12 [simp]: "(x \ -y) \ (x \ y) = x" by (metis p_inf sup_bot_right sup_inf_distrib1) lemma maddux_3_13 [simp]: "(x \ y) \ -x = y \ -x" by (simp add: inf_sup_distrib2) lemma maddux_3_20: "((v \ w) \ (-v \ x)) \ -((v \ y) \ (-v \ z)) = (v \ w \ -y) \ (-v \ x \ -z)" proof - have "v \ w \ -(v \ y) \ -(-v \ z) = v \ w \ -(v \ y)" by (meson inf.cobounded1 inf_absorb1 le_infI1 p_antitone_iff) also have "... = v \ w \ -y" using inf.sup_relative_same_increasing inf_import_p inf_le1 by blast finally have 1: "v \ w \ -(v \ y) \ -(-v \ z) = v \ w \ -y" . have "-v \ x \ -(v \ y) \ -(-v \ z) = -v \ x \ -(-v \ z)" by (simp add: inf.absorb1 le_infI1 p_antitone_inf) also have "... = -v \ x \ -z" by (simp add: inf.assoc inf_left_commute) finally have 2: "-v \ x \ -(v \ y) \ -(-v \ z) = -v \ x \ -z" . have "((v \ w) \ (-v \ x)) \ -((v \ y) \ (-v \ z)) = (v \ w \ -(v \ y) \ -(-v \ z)) \ (-v \ x \ -(v \ y) \ -(-v \ z))" by (simp add: inf_assoc inf_sup_distrib2) also have "... = (v \ w \ -y) \ (-v \ x \ -z)" using 1 2 by simp finally show ?thesis . qed lemma order_char_1: "x \ y \ x \ y \ -x" by (metis inf.sup_left_isotone inf_sup_absorb le_supI1 maddux_3_12 sup_commute) lemma order_char_2: "x \ y \ x \ -x \ y \ -x" using order_char_1 by auto (* lemma pp_dist_sup [simp]: "--(x \ y) = --x \ --y" nitpick [expect=genuine] oops lemma regular_closed_sup: "regular x \ regular y \ regular (x \ y)" nitpick [expect=genuine] oops lemma regular_complemented_iff: "regular x \ complemented x" nitpick [expect=genuine] oops lemma selection_closed_sup: "selection s x \ selection t x \ selection (s \ t) x" nitpick [expect=genuine] oops lemma stone [simp]: "-x \ --x = top" nitpick [expect=genuine] oops *) end subsection \Stone Algebras\ text \ A Stone algebra is a distributive lattice with a pseudocomplement that satisfies the following equation. We thus obtain the other half of the requirements of a complement at least for the regular elements. \ class stone_algebra = pd_algebra + assumes stone [simp]: "-x \ --x = top" begin text \ As a consequence, we obtain both De Morgan's laws for all elements. \ lemma p_dist_inf [simp]: "-(x \ y) = -x \ -y" proof (rule p_unique[THEN sym], rule allI, rule iffI) fix w assume "w \ (x \ y) = bot" hence "w \ --x \ y = bot" using inf_commute inf_left_commute pseudo_complement by auto hence 1: "w \ --x \ -y" by (simp add: pseudo_complement) have "w = (w \ -x) \ (w \ --x)" using distrib_imp2 sup_inf_distrib1 by auto thus "w \ -x \ -y" using 1 by (metis inf_le2 sup.mono) next fix w assume "w \ -x \ -y" thus "w \ (x \ y) = bot" using order_trans p_supdist_inf pseudo_complement by blast qed lemma pp_dist_sup [simp]: "--(x \ y) = --x \ --y" by simp lemma regular_closed_sup: "regular x \ regular y \ regular (x \ y)" by simp text \ The regular elements are precisely the ones having a complement. \ lemma regular_complemented_iff: "regular x \ complemented x" by (metis inf_p stone complemented_regular) lemma selection_closed_sup: "selection s x \ selection t x \ selection (s \ t) x" by (simp add: inf_sup_distrib2) lemma huntington_3_pp [simp]: "-(-x \ -y) \ -(-x \ y) = --x" by (metis p_dist_inf p_inf sup.commute sup_bot_left sup_inf_distrib1) lemma maddux_3_3 [simp]: "-(x \ y) \ -(x \ -y) = -x" by (simp add: sup_commute sup_inf_distrib1) lemma maddux_3_11_pp: "(x \ -y) \ (x \ --y) = x" by (metis inf_sup_distrib1 inf_top_right stone) lemma maddux_3_19_pp: "(-x \ y) \ (--x \ z) = (--x \ y) \ (-x \ z)" proof - have "(--x \ y) \ (-x \ z) = (--x \ z) \ (y \ -x) \ (y \ z)" by (simp add: inf.commute inf_sup_distrib1 sup.assoc) also have "... = (--x \ z) \ (y \ -x) \ (y \ z \ (-x \ --x))" by simp also have "... = (--x \ z) \ ((y \ -x) \ (y \ -x \ z)) \ (y \ z \ --x)" using inf_sup_distrib1 sup_assoc inf_commute inf_assoc by presburger also have "... = (--x \ z) \ (y \ -x) \ (y \ z \ --x)" by simp also have "... = ((--x \ z) \ (--x \ z \ y)) \ (y \ -x)" by (simp add: inf_assoc inf_commute sup.left_commute sup_commute) also have "... = (--x \ z) \ (y \ -x)" by simp finally show ?thesis by (simp add: inf_commute sup_commute) qed lemma compl_inter_eq_pp: "--x \ y = --x \ z \ -x \ y = -x \ z \ y = z" by (metis inf_commute inf_p inf_sup_distrib1 inf_top.right_neutral p_bot p_dist_inf) lemma maddux_3_21_pp [simp]: "--x \ (-x \ y) = --x \ y" by (simp add: sup.commute sup_inf_distrib1) lemma shunting_2_pp: "x \ --y \ -x \ --y = top" by (metis inf_top_left p_bot p_dist_inf pseudo_complement) lemma shunting_p: "x \ y \ -z \ x \ -z \ -y" by (metis inf.assoc p_dist_inf p_shunting_swap pseudo_complement) text \ The following weak shunting property is interesting as it does not require the element \z\ on the right-hand side to be regular. \ lemma shunting_var_p: "x \ -y \ z \ x \ z \ --y" proof assume "x \ -y \ z" hence "z \ --y = --y \ (z \ x \ -y)" by (simp add: sup.absorb1 sup.commute) thus "x \ z \ --y" by (metis inf_commute maddux_3_21_pp sup.commute sup.left_commute sup_left_divisibility) next assume "x \ z \ --y" thus "x \ -y \ z" by (metis inf.mono maddux_3_12 sup_ge2) qed (* Whether conjugate_char_2_pp can be proved in pd_algebra or in p_algebra is unknown. *) lemma conjugate_char_2_pp: "conjugate f g \ f bot = bot \ g bot = bot \ (\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x))))" proof assume 1: "conjugate f g" hence 2: "dual_additive (uminus \ g)" using conjugate_symmetric conjugate_dual_additive by auto show "f bot = bot \ g bot = bot \ (\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x))))" proof (intro conjI) show "f bot = bot" using 1 by (metis conjugate_def inf_idem inf_bot_left) next show "g bot = bot" using 1 by (metis conjugate_def inf_idem inf_bot_right) next show "\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x)))" proof (intro allI) fix x y have 3: "y \ -(f(x \ -(g y)))" using 1 by (simp add: conjugate_def pseudo_complement inf_commute) have 4: "x \ -(g(y \ -(f x)))" using 1 conjugate_def inf.commute pseudo_complement by fastforce have "y \ -(f(x \ --(g y))) = y \ -(f(x \ -(g y))) \ -(f(x \ --(g y)))" using 3 by (simp add: inf.le_iff_sup inf_commute) also have "... = y \ -(f((x \ -(g y)) \ (x \ --(g y))))" using 1 conjugate_dual_additive dual_additive_def inf_assoc by auto also have "... = y \ -(f x)" by (simp add: maddux_3_11_pp) also have "... \ -(f x)" by simp finally have 5: "f x \ y \ --(f(x \ --(g y)))" by (simp add: inf_commute p_shunting_swap) have "x \ -(g(y \ --(f x))) = x \ -(g(y \ -(f x))) \ -(g(y \ --(f x)))" using 4 by (simp add: inf.le_iff_sup inf_commute) also have "... = x \ -(g((y \ -(f x)) \ (y \ --(f x))))" using 2 by (simp add: dual_additive_def inf_assoc) also have "... = x \ -(g y)" by (simp add: maddux_3_11_pp) also have "... \ -(g y)" by simp finally have "g y \ x \ --(g(y \ --(f x)))" by (simp add: inf_commute p_shunting_swap) thus "f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x)))" using 5 by simp qed qed next assume "f bot = bot \ g bot = bot \ (\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x))))" thus "conjugate f g" by (unfold conjugate_def, metis inf_commute le_bot pp_inf_bot_iff regular_closed_bot) qed lemma conjugate_char_2_pp_additive: assumes "conjugate f g" and "additive f" and "additive g" shows "f x \ y \ f(x \ --(g y)) \ g y \ x \ g(y \ --(f x))" proof - have "f x \ y = f ((x \ --g y) \ (x \ -g y)) \ y" by (simp add: sup.commute sup_inf_distrib1) also have "... = (f (x \ --g y) \ y) \ (f (x \ -g y) \ y)" using assms(2) additive_def inf_sup_distrib2 by auto also have "... = f (x \ --g y) \ y" by (metis assms(1) conjugate_def inf_le2 pseudo_complement sup_bot.right_neutral) finally have 2: "f x \ y \ f (x \ --g y)" by simp have "g y \ x = g ((y \ --f x) \ (y \ -f x)) \ x" by (simp add: sup.commute sup_inf_distrib1) also have "... = (g (y \ --f x) \ x) \ (g (y \ -f x) \ x)" using assms(3) additive_def inf_sup_distrib2 by auto also have "... = g (y \ --f x) \ x" by (metis assms(1) conjugate_def inf.cobounded2 pseudo_complement sup_bot.right_neutral inf_commute) finally have "g y \ x \ g (y \ --f x)" by simp thus ?thesis using 2 by simp qed (* lemma compl_le_swap2_iff: "-x \ y \ -y \ x" nitpick [expect=genuine] oops lemma huntington_3: "x = -(-x \ -y) \ -(-x \ y)" nitpick [expect=genuine] oops lemma maddux_3_1: "x \ -x = y \ -y" nitpick [expect=genuine] oops lemma maddux_3_4: "x \ (y \ -y) = z \ -z" nitpick [expect=genuine] oops lemma maddux_3_11: "x = (x \ y) \ (x \ -y)" nitpick [expect=genuine] oops lemma maddux_3_19: "(-x \ y) \ (x \ z) = (x \ y) \ (-x \ z)" nitpick [expect=genuine] oops lemma compl_inter_eq: "x \ y = x \ z \ -x \ y = -x \ z \ y = z" nitpick [expect=genuine] oops lemma maddux_3_21: "x \ y = x \ (-x \ y)" nitpick [expect=genuine] oops lemma shunting_1: "x \ y \ x \ -y = bot" nitpick [expect=genuine] oops lemma shunting_2: "x \ y \ -x \ y = top" nitpick [expect=genuine] oops lemma conjugate_unique: "conjugate f g \ conjugate f h \ g = h" nitpick [expect=genuine] oops lemma conjugate_isotone_pp: "conjugate f g \ isotone f" nitpick [expect=genuine] oops lemma conjugate_char_1: "conjugate f g \ (\x y . f(x \ -(g y)) \ f x \ -y \ g(y \ -(f x)) \ g y \ -x)" nitpick [expect=genuine] oops lemma conjugate_char_2: "conjugate f g \ f bot = bot \ g bot = bot \ (\x y . f x \ y \ f(x \ g y) \ g y \ x \ g(y \ f x))" nitpick [expect=genuine] oops lemma shunting: "x \ y \ z \ x \ z \ -y" nitpick [expect=genuine] oops lemma shunting_var: "x \ -y \ z \ x \ z \ y" nitpick [expect=genuine] oops lemma sup_compl_top: "x \ -x = top" nitpick [expect=genuine] oops lemma selection_closed_p: "selection s x \ selection (-s) x" nitpick [expect=genuine] oops lemma selection_closed_pp: "selection s x \ selection (--s) x" nitpick [expect=genuine] oops *) end +abbreviation stone_algebra_isomorphism :: "('a::stone_algebra \ 'b::stone_algebra) \ bool" + where "stone_algebra_isomorphism f \ sup_inf_top_bot_uminus_isomorphism f" + text \ Every bounded linear order can be expanded to a Stone algebra. The pseudocomplement takes \bot\ to the \top\ and every other element to \bot\. \ class linorder_stone_algebra_expansion = linorder_lattice_expansion + uminus + assumes uminus_def [simp]: "-x = (if x = bot then top else bot)" begin subclass stone_algebra apply unfold_locales using bot_unique min_def top_le by auto text \ The regular elements are the least and greatest elements. All elements except the least element are dense. \ lemma regular_bot_top: "regular x \ x = bot \ x = top" by simp lemma not_bot_dense: "x \ bot \ --x = top" by simp end subsection \Heyting Algebras\ text \ In this section we add a relative pseudocomplement operation to semilattices and to lattices. \ subsubsection \Heyting Semilattices\ text \ The pseudocomplement of an element \y\ relative to an element \z\ is the least element whose meet with \y\ is below \z\. This can be stated as a Galois connection. Specialising \z = bot\ gives (non-relative) pseudocomplements. Many properties can already be shown if the underlying structure is just a semilattice. \ class implies = fixes implies :: "'a \ 'a \ 'a" (infixl "\" 65) class heyting_semilattice = semilattice_inf + implies + assumes implies_galois: "x \ y \ z \ x \ y \ z" begin lemma implies_below_eq [simp]: "y \ (x \ y) = y" using implies_galois inf.absorb_iff1 inf.cobounded1 by blast lemma implies_increasing: "x \ y \ x" by (simp add: inf.orderI) lemma implies_galois_swap: "x \ y \ z \ y \ x \ z" by (metis implies_galois inf_commute) lemma implies_galois_var: "x \ y \ z \ y \ x \ z" by (simp add: implies_galois_swap implies_galois) lemma implies_galois_increasing: "x \ y \ (x \ y)" using implies_galois by blast lemma implies_galois_decreasing: "(y \ x) \ y \ x" using implies_galois by blast lemma implies_mp_below: "x \ (x \ y) \ y" using implies_galois_decreasing inf_commute by auto lemma implies_isotone: "x \ y \ z \ x \ z \ y" using implies_galois order_trans by blast lemma implies_antitone: "x \ y \ y \ z \ x \ z" by (meson implies_galois_swap order_lesseq_imp) lemma implies_isotone_inf: "x \ (y \ z) \ x \ y" by (simp add: implies_isotone) lemma implies_antitone_inf: "x \ z \ (x \ y) \ z" by (simp add: implies_antitone) lemma implies_curry: "x \ (y \ z) = (x \ y) \ z" by (metis implies_galois_decreasing implies_galois inf_assoc antisym) lemma implies_curry_flip: "x \ (y \ z) = y \ (x \ z)" by (simp add: implies_curry inf_commute) lemma triple_implies [simp]: "((x \ y) \ y) \ y = x \ y" using implies_antitone implies_galois_swap eq_iff by auto lemma implies_mp_eq [simp]: "x \ (x \ y) = x \ y" by (metis implies_below_eq implies_mp_below inf_left_commute inf.absorb2) lemma implies_dist_implies: "x \ (y \ z) \ (x \ y) \ (x \ z)" using implies_curry implies_curry_flip by auto lemma implies_import_inf [simp]: "x \ ((x \ y) \ (x \ z)) = x \ (y \ z)" by (metis implies_curry implies_mp_eq inf_commute) lemma implies_dist_inf: "x \ (y \ z) = (x \ y) \ (x \ z)" proof - have "(x \ y) \ (x \ z) \ x \ y \ z" by (simp add: implies_galois) hence "(x \ y) \ (x \ z) \ x \ (y \ z)" using implies_galois by blast thus ?thesis by (simp add: implies_isotone eq_iff) qed lemma implies_itself_top: "y \ x \ x" by (simp add: implies_galois_swap implies_increasing) lemma inf_implies_top: "z \ (x \ y) \ x" using implies_galois_var le_infI1 by blast lemma inf_inf_implies [simp]: "z \ ((x \ y) \ x) = z" by (simp add: inf_implies_top inf_absorb1) lemma le_implies_top: "x \ y \ z \ x \ y" using implies_antitone implies_itself_top order.trans by blast lemma le_iff_le_implies: "x \ y \ x \ x \ y" using implies_galois inf_idem by force lemma implies_inf_isotone: "x \ y \ (x \ z) \ (y \ z)" by (metis implies_curry implies_galois_increasing implies_isotone) lemma implies_transitive: "(x \ y) \ (y \ z) \ x \ z" using implies_dist_implies implies_galois_var implies_increasing order_lesseq_imp by blast lemma implies_inf_absorb [simp]: "x \ (x \ y) = x \ y" using implies_dist_inf implies_itself_top inf.absorb_iff2 by auto lemma implies_implies_absorb [simp]: "x \ (x \ y) = x \ y" by (simp add: implies_curry) lemma implies_inf_identity: "(x \ y) \ y = y" by (simp add: inf_commute) lemma implies_itself_same: "x \ x = y \ y" by (simp add: le_implies_top eq_iff) end text \ The following class gives equational axioms for the relative pseudocomplement operation (inequalities can be written as equations). \ class heyting_semilattice_eq = semilattice_inf + implies + assumes implies_mp_below: "x \ (x \ y) \ y" and implies_galois_increasing: "x \ y \ (x \ y)" and implies_isotone_inf: "x \ (y \ z) \ x \ y" begin subclass heyting_semilattice apply unfold_locales apply (rule iffI) apply (metis implies_galois_increasing implies_isotone_inf inf.absorb2 order_lesseq_imp) by (metis implies_mp_below inf_commute order_trans inf_mono order_refl) end text \ The following class allows us to explicitly give the pseudocomplement of an element relative to itself. \ class bounded_heyting_semilattice = bounded_semilattice_inf_top + heyting_semilattice begin lemma implies_itself [simp]: "x \ x = top" using implies_galois inf_le2 top_le by blast lemma implies_order: "x \ y \ x \ y = top" by (metis implies_galois inf_top.left_neutral top_unique) lemma inf_implies [simp]: "(x \ y) \ x = top" using implies_order inf_le1 by blast lemma top_implies [simp]: "top \ x = x" by (metis implies_mp_eq inf_top.left_neutral) end subsubsection \Heyting Lattices\ text \ We obtain further properties if the underlying structure is a lattice. In particular, the lattice operations are automatically distributive in this case. \ class heyting_lattice = lattice + heyting_semilattice begin lemma sup_distrib_inf_le: "(x \ y) \ (x \ z) \ x \ (y \ z)" proof - have "x \ z \ y \ (x \ (y \ z))" using implies_galois_var implies_increasing sup.bounded_iff sup.cobounded2 by blast hence "x \ y \ (x \ z) \ (x \ (y \ z))" using implies_galois_swap implies_increasing le_sup_iff by blast thus ?thesis by (simp add: implies_galois) qed subclass distrib_lattice apply unfold_locales using distrib_sup_le eq_iff sup_distrib_inf_le by auto lemma implies_isotone_sup: "x \ y \ x \ (y \ z)" by (simp add: implies_isotone) lemma implies_antitone_sup: "(x \ y) \ z \ x \ z" by (simp add: implies_antitone) lemma implies_sup: "x \ z \ (y \ z) \ ((x \ y) \ z)" proof - have "(x \ z) \ (y \ z) \ y \ z" by (simp add: implies_galois) hence "(x \ z) \ (y \ z) \ (x \ y) \ z" using implies_galois_swap implies_galois_var by fastforce thus ?thesis by (simp add: implies_galois) qed lemma implies_dist_sup: "(x \ y) \ z = (x \ z) \ (y \ z)" apply (rule antisym) apply (simp add: implies_antitone) by (simp add: implies_sup implies_galois) lemma implies_antitone_isotone: "(x \ y) \ (x \ y) \ x \ y" by (simp add: implies_antitone_sup implies_dist_inf le_infI2) lemma implies_antisymmetry: "(x \ y) \ (y \ x) = (x \ y) \ (x \ y)" by (metis implies_dist_sup implies_inf_absorb inf.commute) lemma sup_inf_implies [simp]: "(x \ y) \ (x \ y) = y" by (simp add: inf_sup_distrib2 sup.absorb2) lemma implies_subdist_sup: "(x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: implies_isotone) lemma implies_subdist_inf: "(x \ z) \ (y \ z) \ (x \ y) \ z" by (simp add: implies_antitone) lemma implies_sup_absorb: "(x \ y) \ z \ (x \ z) \ (y \ z)" by (metis implies_dist_sup implies_isotone_sup implies_increasing inf_inf_implies le_sup_iff sup_inf_implies) lemma sup_below_implies_implies: "x \ y \ (x \ y) \ y" by (simp add: implies_dist_sup implies_galois_swap implies_increasing) end class bounded_heyting_lattice = bounded_lattice + heyting_lattice begin subclass bounded_heyting_semilattice .. lemma implies_bot [simp]: "bot \ x = top" using implies_galois top_unique by fastforce end subsubsection \Heyting Algebras\ text \ The pseudocomplement operation can be defined in Heyting algebras, but it is typically not part of their signature. We add the definition as an axiom so that we can use the class hierarchy, for example, to inherit results from the class \pd_algebra\. \ class heyting_algebra = bounded_heyting_lattice + uminus + assumes uminus_eq: "-x = x \ bot" begin subclass pd_algebra apply unfold_locales using bot_unique implies_galois uminus_eq by auto lemma boolean_implies_below: "-x \ y \ x \ y" by (simp add: implies_increasing implies_isotone uminus_eq) lemma negation_implies: "-(x \ y) = --x \ -y" proof (rule antisym) show "-(x \ y) \ --x \ -y" using boolean_implies_below p_antitone by auto next have "x \ -y \ (x \ y) = bot" by (metis implies_mp_eq inf_p inf_bot_left inf_commute inf_left_commute) hence "--x \ -y \ (x \ y) = bot" using pp_inf_bot_iff inf_assoc by auto thus "--x \ -y \ -(x \ y)" by (simp add: pseudo_complement) qed lemma double_negation_dist_implies: "--(x \ y) = --x \ --y" apply (rule antisym) apply (metis pp_inf_below_iff implies_galois_decreasing implies_galois negation_implies ppp) by (simp add: p_antitone_iff negation_implies) (* lemma stone: "-x \ --x = top" nitpick [expect=genuine] oops *) end text \ The following class gives equational axioms for Heyting algebras. \ class heyting_algebra_eq = bounded_lattice + implies + uminus + assumes implies_mp_eq: "x \ (x \ y) = x \ y" and implies_import_inf: "x \ ((x \ y) \ (x \ z)) = x \ (y \ z)" and inf_inf_implies: "z \ ((x \ y) \ x) = z" and uminus_eq_eq: "-x = x \ bot" begin subclass heyting_algebra apply unfold_locales apply (rule iffI) apply (metis implies_import_inf inf.sup_left_divisibility inf_inf_implies le_iff_inf) apply (metis implies_mp_eq inf.commute inf.le_sup_iff inf.sup_right_isotone) by (simp add: uminus_eq_eq) end text \ A relative pseudocomplement is not enough to obtain the Stone equation, so we add it in the following class. \ class heyting_stone_algebra = heyting_algebra + assumes heyting_stone: "-x \ --x = top" begin subclass stone_algebra by unfold_locales (simp add: heyting_stone) (* lemma pre_linear: "(x \ y) \ (y \ x) = top" nitpick [expect=genuine] oops *) end subsubsection \Brouwer Algebras\ text \ Brouwer algebras are dual to Heyting algebras. The dual pseudocomplement of an element \y\ relative to an element \x\ is the least element whose join with \y\ is above \x\. We can now use the binary operation provided by Boolean algebras in Isabelle/HOL because it is compatible with dual relative pseudocomplements (not relative pseudocomplements). \ class brouwer_algebra = bounded_lattice + minus + uminus + assumes minus_galois: "x \ y \ z \ x - y \ z" and uminus_eq_minus: "-x = top - x" begin sublocale brouwer: heyting_algebra where inf = sup and less_eq = greater_eq and less = greater and sup = inf and bot = top and top = bot and implies = "\x y . y - x" apply unfold_locales apply simp apply simp apply simp apply simp apply (metis minus_galois sup_commute) by (simp add: uminus_eq_minus) lemma curry_minus: "x - (y \ z) = (x - y) - z" by (simp add: brouwer.implies_curry sup_commute) lemma minus_subdist_sup: "(x - z) \ (y - z) \ (x \ y) - z" by (simp add: brouwer.implies_dist_inf) lemma inf_sup_minus: "(x \ y) \ (x - y) = x" by (simp add: inf.absorb1 brouwer.inf_sup_distrib2) end subsection \Boolean Algebras\ text \ This section integrates Boolean algebras in the above hierarchy. In particular, we strengthen several results shown above. \ context boolean_algebra begin text \ Every Boolean algebra is a Stone algebra, a Heyting algebra and a Brouwer algebra. \ subclass stone_algebra apply unfold_locales apply (rule iffI) apply (metis compl_sup_top inf.orderI inf_bot_right inf_sup_distrib1 inf_top_right sup_inf_absorb) using inf.commute inf.sup_right_divisibility apply fastforce by simp sublocale heyting: heyting_algebra where implies = "\x y . -x \ y" apply unfold_locales apply (rule iffI) using shunting_var_p sup_commute apply fastforce using shunting_var_p sup_commute apply force by simp subclass brouwer_algebra apply unfold_locales apply (simp add: diff_eq shunting_var_p sup.commute) by (simp add: diff_eq) lemma huntington_3 [simp]: "-(-x \ -y) \ -(-x \ y) = x" using huntington_3_pp by auto lemma maddux_3_1: "x \ -x = y \ -y" by simp lemma maddux_3_4: "x \ (y \ -y) = z \ -z" by simp lemma maddux_3_11 [simp]: "(x \ y) \ (x \ -y) = x" using brouwer.maddux_3_12 sup.commute by auto lemma maddux_3_19: "(-x \ y) \ (x \ z) = (x \ y) \ (-x \ z)" using maddux_3_19_pp by auto lemma compl_inter_eq: "x \ y = x \ z \ -x \ y = -x \ z \ y = z" by (metis inf_commute maddux_3_11) lemma maddux_3_21 [simp]: "x \ (-x \ y) = x \ y" by (simp add: sup_inf_distrib1) lemma shunting_1: "x \ y \ x \ -y = bot" by (simp add: pseudo_complement) lemma uminus_involutive: "uminus \ uminus = id" by auto lemma uminus_injective: "uminus \ f = uminus \ g \ f = g" by (metis comp_assoc id_o minus_comp_minus) lemma conjugate_unique: "conjugate f g \ conjugate f h \ g = h" using conjugate_unique_p uminus_injective by blast lemma dual_additive_additive: "dual_additive (uminus \ f) \ additive f" by (metis additive_def compl_eq_compl_iff dual_additive_def p_dist_sup o_def) lemma conjugate_additive: "conjugate f g \ additive f" by (simp add: conjugate_dual_additive dual_additive_additive) lemma conjugate_isotone: "conjugate f g \ isotone f" by (simp add: conjugate_additive additive_isotone) lemma conjugate_char_1: "conjugate f g \ (\x y . f(x \ -(g y)) \ f x \ -y \ g(y \ -(f x)) \ g y \ -x)" by (simp add: conjugate_char_1_pp) lemma conjugate_char_2: "conjugate f g \ f bot = bot \ g bot = bot \ (\x y . f x \ y \ f(x \ g y) \ g y \ x \ g(y \ f x))" by (simp add: conjugate_char_2_pp) lemma shunting: "x \ y \ z \ x \ z \ -y" by (simp add: heyting.implies_galois sup.commute) lemma shunting_var: "x \ -y \ z \ x \ z \ y" by (simp add: shunting) end class non_trivial_stone_algebra = non_trivial_bounded_order + stone_algebra class non_trivial_boolean_algebra = non_trivial_stone_algebra + boolean_algebra end diff --git a/thys/Stone_Algebras/Stone_Construction.thy b/thys/Stone_Algebras/Stone_Construction.thy --- a/thys/Stone_Algebras/Stone_Construction.thy +++ b/thys/Stone_Algebras/Stone_Construction.thy @@ -1,2351 +1,2308 @@ (* Title: Construction of Stone Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Stone Construction\ text \ This theory proves the uniqueness theorem for the triple representation of Stone algebras and the construction theorem of Stone algebras \cite{ChenGraetzer1969,Katrinak1973}. Every Stone algebra $S$ has an associated triple consisting of \begin{itemize} \item the set of regular elements $B(S)$ of $S$, \item the set of dense elements $D(S)$ of $S$, and \item the structure map $\varphi(S) : B(S) \to F(D(S))$ defined by $\varphi(x) = {\uparrow} x \cap D(S)$. \end{itemize} Here $F(X)$ is the set of filters of a partially ordered set $X$. We first show that \begin{itemize} \item $B(S)$ is a Boolean algebra, \item $D(S)$ is a distributive lattice with a greatest element, whence $F(D(S))$ is a bounded distributive lattice, and \item $\varphi(S)$ is a bounded lattice homomorphism. \end{itemize} Next, from a triple $T = (B,D,\varphi)$ such that $B$ is a Boolean algebra, $D$ is a distributive lattice with a greatest element and $\varphi : B \to F(D)$ is a bounded lattice homomorphism, we construct a Stone algebra $S(T)$. The elements of $S(T)$ are pairs taken from $B \times F(D)$ following the construction of \cite{Katrinak1973}. We need to represent $S(T)$ as a type to be able to instantiate the Stone algebra class. Because the pairs must satisfy a condition depending on $\varphi$, this would require dependent types. Since Isabelle/HOL does not have dependent types, we use a function lifting instead. The lifted pairs form a Stone algebra. Next, we specialise the construction to start with the triple associated with a Stone algebra $S$, that is, we construct $S(B(S),D(S),\varphi(S))$. In this case, we can instantiate the lifted pairs to obtain a type of pairs (that no longer implements a dependent type). To achieve this, we construct an embedding of the type of pairs into the lifted pairs, so that we inherit the Stone algebra axioms (using a technique of universal algebra that works for universally quantified equations and equational implications). Next, we show that the Stone algebras $S(B(S),D(S),\varphi(S))$ and $S$ are isomorphic. We give explicit mappings in both directions. This implies the uniqueness theorem for the triple representation of Stone algebras. Finally, we show that the triples $(B(S(T)),D(S(T)),\varphi(S(T)))$ and $T$ are isomorphic. This requires an isomorphism of the Boolean algebras $B$ and $B(S(T))$, an isomorphism of the distributive lattices $D$ and $D(S(T))$, and a proof that they preserve the structure maps. We give explicit mappings of the Boolean algebra isomorphism and the distributive lattice isomorphism in both directions. This implies the construction theorem of Stone algebras. Because $S(T)$ is implemented by lifted pairs, so are $B(S(T))$ and $D(S(T))$; we therefore also lift $B$ and $D$ to establish the isomorphisms. \ theory Stone_Construction imports P_Algebras Filters begin -subsection \Triples\ - -text \ -This section gives definitions of lattice homomorphisms and isomorphisms and basic properties. -It concludes with a locale that represents triples as discussed above. -\ - -class sup_inf_top_bot_uminus = sup + inf + top + bot + uminus -class sup_inf_top_bot_uminus_ord = sup_inf_top_bot_uminus + ord - -context p_algebra -begin - -subclass sup_inf_top_bot_uminus_ord . - -end - -abbreviation sup_homomorphism :: "('a::sup \ 'b::sup) \ bool" - where "sup_homomorphism f \ \x y . f (x \ y) = f x \ f y" - -abbreviation inf_homomorphism :: "('a::inf \ 'b::inf) \ bool" - where "inf_homomorphism f \ \x y . f (x \ y) = f x \ f y" - -abbreviation sup_inf_homomorphism :: "('a::{sup,inf} \ 'b::{sup,inf}) \ bool" - where "sup_inf_homomorphism f \ sup_homomorphism f \ inf_homomorphism f" - -abbreviation sup_inf_top_homomorphism :: "('a::{sup,inf,top} \ 'b::{sup,inf,top}) \ bool" - where "sup_inf_top_homomorphism f \ sup_inf_homomorphism f \ f top = top" - -abbreviation sup_inf_top_bot_homomorphism :: "('a::{sup,inf,top,bot} \ 'b::{sup,inf,top,bot}) \ bool" - where "sup_inf_top_bot_homomorphism f \ sup_inf_top_homomorphism f \ f bot = bot" - -abbreviation bounded_lattice_homomorphism :: "('a::bounded_lattice \ 'b::bounded_lattice) \ bool" - where "bounded_lattice_homomorphism f \ sup_inf_top_bot_homomorphism f" - -abbreviation sup_inf_top_bot_uminus_homomorphism :: "('a::sup_inf_top_bot_uminus \ 'b::sup_inf_top_bot_uminus) \ bool" - where "sup_inf_top_bot_uminus_homomorphism f \ sup_inf_top_bot_homomorphism f \ (\x . f (-x) = -f x)" - -abbreviation sup_inf_top_bot_uminus_ord_homomorphism :: "('a::sup_inf_top_bot_uminus_ord \ 'b::sup_inf_top_bot_uminus_ord) \ bool" - where "sup_inf_top_bot_uminus_ord_homomorphism f \ sup_inf_top_bot_uminus_homomorphism f \ (\x y . x \ y \ f x \ f y)" - -abbreviation sup_inf_top_isomorphism :: "('a::{sup,inf,top} \ 'b::{sup,inf,top}) \ bool" - where "sup_inf_top_isomorphism f \ sup_inf_top_homomorphism f \ bij f" - -abbreviation bounded_lattice_top_isomorphism :: "('a::bounded_lattice_top \ 'b::bounded_lattice_top) \ bool" - where "bounded_lattice_top_isomorphism f \ sup_inf_top_isomorphism f" - -abbreviation sup_inf_top_bot_uminus_isomorphism :: "('a::sup_inf_top_bot_uminus \ 'b::sup_inf_top_bot_uminus) \ bool" - where "sup_inf_top_bot_uminus_isomorphism f \ sup_inf_top_bot_uminus_homomorphism f \ bij f" - -abbreviation stone_algebra_isomorphism :: "('a::stone_algebra \ 'b::stone_algebra) \ bool" - where "stone_algebra_isomorphism f \ sup_inf_top_bot_uminus_isomorphism f" - -abbreviation boolean_algebra_isomorphism :: "('a::boolean_algebra \ 'b::boolean_algebra) \ bool" - where "boolean_algebra_isomorphism f \ sup_inf_top_bot_uminus_isomorphism f" - -lemma sup_homomorphism_mono: - "sup_homomorphism (f::'a::semilattice_sup \ 'b::semilattice_sup) \ mono f" - by (metis le_iff_sup monoI) - -lemma sup_isomorphism_ord_isomorphism: - assumes "sup_homomorphism (f::'a::semilattice_sup \ 'b::semilattice_sup)" - and "bij f" - shows "x \ y \ f x \ f y" -proof - assume "x \ y" - thus "f x \ f y" - by (metis assms(1) le_iff_sup) -next - assume "f x \ f y" - hence "f (x \ y) = f y" - by (simp add: assms(1) le_iff_sup) - hence "x \ y = y" - by (metis injD bij_is_inj assms(2)) - thus "x \ y" - by (simp add: le_iff_sup) -qed - text \ A triple consists of a Boolean algebra, a distributive lattice with a greatest element, and a structure map. The Boolean algebra and the distributive lattice are represented as HOL types. Because both occur in the type of the structure map, the triple is determined simply by the structure map and its HOL type. The structure map needs to be a bounded lattice homomorphism. \ locale triple = fixes phi :: "'a::boolean_algebra \ 'b::distrib_lattice_top filter" assumes hom: "bounded_lattice_homomorphism phi" subsection \The Triple of a Stone Algebra\ text \ In this section we construct the triple associated to a Stone algebra. \ subsubsection \Regular Elements\ text \ The regular elements of a Stone algebra form a Boolean subalgebra. \ typedef (overloaded) 'a regular = "regular_elements::'a::stone_algebra set" by auto lemma simp_regular [simp]: "\y . Rep_regular x = -y" using Rep_regular by simp setup_lifting type_definition_regular instantiation regular :: (stone_algebra) boolean_algebra begin lift_definition sup_regular :: "'a regular \ 'a regular \ 'a regular" is sup by (meson regular_in_p_image_iff regular_closed_sup) lift_definition inf_regular :: "'a regular \ 'a regular \ 'a regular" is inf by (meson regular_in_p_image_iff regular_closed_inf) lift_definition minus_regular :: "'a regular \ 'a regular \ 'a regular" is "\x y . x \ -y" by (meson regular_in_p_image_iff regular_closed_inf) lift_definition uminus_regular :: "'a regular \ 'a regular" is uminus by auto lift_definition bot_regular :: "'a regular" is bot by (meson regular_in_p_image_iff regular_closed_bot) lift_definition top_regular :: "'a regular" is top by (meson regular_in_p_image_iff regular_closed_top) lift_definition less_eq_regular :: "'a regular \ 'a regular \ bool" is less_eq . lift_definition less_regular :: "'a regular \ 'a regular \ bool" is less . instance apply intro_classes - apply (simp add: less_eq_regular.rep_eq less_regular.rep_eq inf.less_le_not_le) - apply (simp add: less_eq_regular.rep_eq) - apply (simp add: less_eq_regular.rep_eq) - apply (simp add: Rep_regular_inject less_eq_regular.rep_eq) - apply (simp add: inf_regular.rep_eq less_eq_regular.rep_eq) - apply (simp add: inf_regular.rep_eq less_eq_regular.rep_eq) - apply (simp add: inf_regular.rep_eq less_eq_regular.rep_eq) - apply (simp add: sup_regular.rep_eq less_eq_regular.rep_eq) - apply (simp add: sup_regular.rep_eq less_eq_regular.rep_eq) - apply (simp add: sup_regular.rep_eq less_eq_regular.rep_eq) - apply (simp add: bot_regular.rep_eq less_eq_regular.rep_eq) - apply (simp add: top_regular.rep_eq less_eq_regular.rep_eq) - apply (metis (mono_tags) Rep_regular_inject inf_regular.rep_eq sup_inf_distrib1 sup_regular.rep_eq) - apply (metis (mono_tags) Rep_regular_inverse bot_regular.abs_eq inf_regular.rep_eq inf_p uminus_regular.rep_eq) - apply (metis (mono_tags) top_regular.abs_eq Rep_regular_inverse simp_regular stone sup_regular.rep_eq uminus_regular.rep_eq) - by (metis (mono_tags) Rep_regular_inject inf_regular.rep_eq minus_regular.rep_eq uminus_regular.rep_eq) + subgoal apply transfer by (simp add: less_le_not_le) + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by (simp add: sup_inf_distrib1) + subgoal apply transfer by simp + subgoal apply transfer by auto + subgoal apply transfer by simp + done end instantiation regular :: (non_trivial_stone_algebra) non_trivial_boolean_algebra begin instance proof (intro_classes, rule ccontr) assume "\(\x y::'a regular . x \ y)" hence "(bot::'a regular) = top" by simp hence "(bot::'a) = top" by (metis bot_regular.rep_eq top_regular.rep_eq) thus False by (simp add: bot_not_top) qed end subsubsection \Dense Elements\ text \ The dense elements of a Stone algebra form a distributive lattice with a greatest element. \ typedef (overloaded) 'a dense = "dense_elements::'a::stone_algebra set" using dense_closed_top by blast lemma simp_dense [simp]: "-Rep_dense x = bot" using Rep_dense by simp setup_lifting type_definition_dense instantiation dense :: (stone_algebra) distrib_lattice_top begin lift_definition sup_dense :: "'a dense \ 'a dense \ 'a dense" is sup by simp lift_definition inf_dense :: "'a dense \ 'a dense \ 'a dense" is inf by simp lift_definition top_dense :: "'a dense" is top by simp lift_definition less_eq_dense :: "'a dense \ 'a dense \ bool" is less_eq . lift_definition less_dense :: "'a dense \ 'a dense \ bool" is less . instance apply intro_classes - apply (simp add: less_eq_dense.rep_eq less_dense.rep_eq inf.less_le_not_le) - apply (simp add: less_eq_dense.rep_eq) - apply (simp add: less_eq_dense.rep_eq) - apply (simp add: Rep_dense_inject less_eq_dense.rep_eq) - apply (simp add: inf_dense.rep_eq less_eq_dense.rep_eq) - apply (simp add: inf_dense.rep_eq less_eq_dense.rep_eq) - apply (simp add: inf_dense.rep_eq less_eq_dense.rep_eq) - apply (simp add: sup_dense.rep_eq less_eq_dense.rep_eq) - apply (simp add: sup_dense.rep_eq less_eq_dense.rep_eq) - apply (simp add: sup_dense.rep_eq less_eq_dense.rep_eq) - apply (simp add: top_dense.rep_eq less_eq_dense.rep_eq) - by (metis (mono_tags, lifting) Rep_dense_inject sup_inf_distrib1 inf_dense.rep_eq sup_dense.rep_eq) + subgoal apply transfer by (simp add: inf.less_le_not_le) + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by (simp add: sup_inf_distrib1) + done end lemma up_filter_dense_antitone_dense: "dense (x \ -x \ y) \ dense (x \ -x \ y \ z)" by simp lemma up_filter_dense_antitone: "up_filter (Abs_dense (x \ -x \ y \ z)) \ up_filter (Abs_dense (x \ -x \ y))" by (unfold up_filter_antitone[THEN sym]) (simp add: Abs_dense_inverse less_eq_dense.rep_eq) text \ The filters of dense elements of a Stone algebra form a bounded distributive lattice. \ type_synonym 'a dense_filter = "'a dense filter" typedef (overloaded) 'a dense_filter_type = "{ x::'a dense_filter . True }" using filter_top by blast setup_lifting type_definition_dense_filter_type instantiation dense_filter_type :: (stone_algebra) bounded_distrib_lattice begin lift_definition sup_dense_filter_type :: "'a dense_filter_type \ 'a dense_filter_type \ 'a dense_filter_type" is sup . lift_definition inf_dense_filter_type :: "'a dense_filter_type \ 'a dense_filter_type \ 'a dense_filter_type" is inf . lift_definition bot_dense_filter_type :: "'a dense_filter_type" is bot .. lift_definition top_dense_filter_type :: "'a dense_filter_type" is top .. lift_definition less_eq_dense_filter_type :: "'a dense_filter_type \ 'a dense_filter_type \ bool" is less_eq . lift_definition less_dense_filter_type :: "'a dense_filter_type \ 'a dense_filter_type \ bool" is less . instance apply intro_classes - apply (simp add: less_eq_dense_filter_type.rep_eq less_dense_filter_type.rep_eq inf.less_le_not_le) - apply (simp add: less_eq_dense_filter_type.rep_eq) - apply (simp add: less_eq_dense_filter_type.rep_eq inf.order_lesseq_imp) - apply (simp add: Rep_dense_filter_type_inject less_eq_dense_filter_type.rep_eq) - apply (simp add: inf_dense_filter_type.rep_eq less_eq_dense_filter_type.rep_eq) - apply (simp add: inf_dense_filter_type.rep_eq less_eq_dense_filter_type.rep_eq) - apply (simp add: inf_dense_filter_type.rep_eq less_eq_dense_filter_type.rep_eq) - apply (simp add: less_eq_dense_filter_type.rep_eq sup_dense_filter_type.rep_eq) - apply (simp add: less_eq_dense_filter_type.rep_eq sup_dense_filter_type.rep_eq) - apply (simp add: less_eq_dense_filter_type.rep_eq sup_dense_filter_type.rep_eq) - apply (simp add: less_eq_dense_filter_type.rep_eq bot_dense_filter_type.rep_eq) - apply (simp add: top_dense_filter_type.rep_eq less_eq_dense_filter_type.rep_eq) - by (metis (mono_tags, lifting) Rep_dense_filter_type_inject sup_inf_distrib1 inf_dense_filter_type.rep_eq sup_dense_filter_type.rep_eq) + subgoal apply transfer by (simp add: inf.less_le_not_le) + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by simp + subgoal apply transfer by (simp add: sup_inf_distrib1) + done end subsubsection \The Structure Map\ text \ The structure map of a Stone algebra is a bounded lattice homomorphism. It maps a regular element \x\ to the set of all dense elements above \-x\. This set is a filter. \ -abbreviation stone_phi_set :: "'a::stone_algebra regular \ 'a dense set" - where "stone_phi_set x \ { y . -Rep_regular x \ Rep_dense y }" +abbreviation stone_phi_base :: "'a::stone_algebra regular \ 'a dense set" + where "stone_phi_base x \ { y . -Rep_regular x \ Rep_dense y }" -lemma stone_phi_set_filter: - "filter (stone_phi_set x)" +lemma stone_phi_base_filter: + "filter (stone_phi_base x)" apply (unfold filter_def, intro conjI) apply (metis Collect_empty_eq top_dense.rep_eq top_greatest) apply (metis inf_dense.rep_eq inf_le2 le_inf_iff mem_Collect_eq) using order_trans less_eq_dense.rep_eq by blast definition stone_phi :: "'a::stone_algebra regular \ 'a dense_filter" - where "stone_phi x = Abs_filter (stone_phi_set x)" + where "stone_phi x = Abs_filter (stone_phi_base x)" text \ To show that we obtain a triple, we only need to prove that \stone_phi\ is a bounded lattice homomorphism. The Boolean algebra and the distributive lattice requirements are taken care of by the type system. \ interpretation stone_phi: triple "stone_phi" proof (unfold_locales, intro conjI) have 1: "Rep_regular (Abs_regular bot) = bot" by (metis bot_regular.rep_eq bot_regular_def) show "stone_phi bot = bot" apply (unfold stone_phi_def bot_regular_def 1 p_bot bot_filter_def) by (metis (mono_tags, lifting) Collect_cong Rep_dense_inject order_refl singleton_conv top.extremum_uniqueI top_dense.rep_eq) next show "stone_phi top = top" by (metis Collect_cong stone_phi_def UNIV_I bot.extremum dense_closed_top top_empty_eq top_filter.abs_eq top_regular.rep_eq top_set_def) next show "\x y::'a regular . stone_phi (x \ y) = stone_phi x \ stone_phi y" proof (intro allI) fix x y :: "'a regular" - have "stone_phi_set (x \ y) = filter_sup (stone_phi_set x) (stone_phi_set y)" + have "stone_phi_base (x \ y) = filter_sup (stone_phi_base x) (stone_phi_base y)" proof (rule set_eqI, rule iffI) fix z - assume 2: "z \ stone_phi_set (x \ y)" + assume 2: "z \ stone_phi_base (x \ y)" let ?t = "-Rep_regular x \ Rep_dense z" let ?u = "-Rep_regular y \ Rep_dense z" let ?v = "Abs_dense ?t" let ?w = "Abs_dense ?u" - have 3: "?v \ stone_phi_set x \ ?w \ stone_phi_set y" + have 3: "?v \ stone_phi_base x \ ?w \ stone_phi_base y" by (simp add: Abs_dense_inverse) have "?v \ ?w = Abs_dense (?t \ ?u)" by (simp add: eq_onp_def inf_dense.abs_eq) also have "... = Abs_dense (-Rep_regular (x \ y) \ Rep_dense z)" by (simp add: distrib(1) sup_commute sup_regular.rep_eq) also have "... = Abs_dense (Rep_dense z)" using 2 by (simp add: le_iff_sup) also have "... = z" by (simp add: Rep_dense_inverse) - finally show "z \ filter_sup (stone_phi_set x) (stone_phi_set y)" - using 3 mem_Collect_eq order_refl by fastforce + finally show "z \ filter_sup (stone_phi_base x) (stone_phi_base y)" + using 3 mem_Collect_eq order_refl filter_sup_def by fastforce next fix z - assume "z \ filter_sup (stone_phi_set x) (stone_phi_set y)" - then obtain v w where 4: "v \ stone_phi_set x \ w \ stone_phi_set y \ v \ w \ z" - by auto + assume "z \ filter_sup (stone_phi_base x) (stone_phi_base y)" + then obtain v w where 4: "v \ stone_phi_base x \ w \ stone_phi_base y \ v \ w \ z" + unfolding filter_sup_def by auto have "-Rep_regular (x \ y) = Rep_regular (-(x \ y))" by (metis uminus_regular.rep_eq) also have "... = -Rep_regular x \ -Rep_regular y" by (simp add: inf_regular.rep_eq uminus_regular.rep_eq) also have "... \ Rep_dense v \ Rep_dense w" using 4 inf_mono mem_Collect_eq by blast also have "... = Rep_dense (v \ w)" by (simp add: inf_dense.rep_eq) also have "... \ Rep_dense z" using 4 by (simp add: less_eq_dense.rep_eq) - finally show "z \ stone_phi_set (x \ y)" + finally show "z \ stone_phi_base (x \ y)" by simp qed thus "stone_phi (x \ y) = stone_phi x \ stone_phi y" - by (simp add: stone_phi_def eq_onp_same_args stone_phi_set_filter sup_filter.abs_eq) + by (simp add: stone_phi_def eq_onp_same_args stone_phi_base_filter sup_filter.abs_eq) qed next show "\x y::'a regular . stone_phi (x \ y) = stone_phi x \ stone_phi y" proof (intro allI) fix x y :: "'a regular" have "\z . -Rep_regular (x \ y) \ Rep_dense z \ -Rep_regular x \ Rep_dense z \ -Rep_regular y \ Rep_dense z" by (simp add: inf_regular.rep_eq) - hence "stone_phi_set (x \ y) = (stone_phi_set x) \ (stone_phi_set y)" + hence "stone_phi_base (x \ y) = (stone_phi_base x) \ (stone_phi_base y)" by auto thus "stone_phi (x \ y) = stone_phi x \ stone_phi y" - by (simp add: stone_phi_def eq_onp_same_args stone_phi_set_filter inf_filter.abs_eq) + by (simp add: stone_phi_def eq_onp_same_args stone_phi_base_filter inf_filter.abs_eq) qed qed subsection \Properties of Triples\ text \ In this section we construct a certain set of pairs from a triple, introduce operations on these pairs and develop their properties. The given set and operations will form a Stone algebra. \ context triple begin lemma phi_bot: "phi bot = Abs_filter {top}" by (metis hom bot_filter_def) lemma phi_top: "phi top = Abs_filter UNIV" by (metis hom top_filter_def) text \ The occurrence of \phi\ in the following definition of the pairs creates a need for dependent types. \ definition pairs :: "('a \ 'b filter) set" where "pairs = { (x,y) . \z . y = phi (-x) \ up_filter z }" text \ Operations on pairs are defined in the following. They will be used to establish that the pairs form a Stone algebra. \ fun pairs_less_eq :: "('a \ 'b filter) \ ('a \ 'b filter) \ bool" where "pairs_less_eq (x,y) (z,w) = (x \ z \ w \ y)" fun pairs_less :: "('a \ 'b filter) \ ('a \ 'b filter) \ bool" where "pairs_less (x,y) (z,w) = (pairs_less_eq (x,y) (z,w) \ \ pairs_less_eq (z,w) (x,y))" fun pairs_sup :: "('a \ 'b filter) \ ('a \ 'b filter) \ ('a \ 'b filter)" where "pairs_sup (x,y) (z,w) = (x \ z,y \ w)" fun pairs_inf :: "('a \ 'b filter) \ ('a \ 'b filter) \ ('a \ 'b filter)" where "pairs_inf (x,y) (z,w) = (x \ z,y \ w)" +fun pairs_minus :: "('a \ 'b filter) \ ('a \ 'b filter) \ ('a \ 'b filter)" + where "pairs_minus (x,y) (z,w) = (x \ -z,y \ phi z)" + fun pairs_uminus :: "('a \ 'b filter) \ ('a \ 'b filter)" where "pairs_uminus (x,y) = (-x,phi x)" abbreviation pairs_bot :: "('a \ 'b filter)" where "pairs_bot \ (bot,Abs_filter UNIV)" abbreviation pairs_top :: "('a \ 'b filter)" where "pairs_top \ (top,Abs_filter {top})" lemma pairs_top_in_set: "(x,y) \ pairs \ top \ Rep_filter y" by simp lemma phi_complemented: "complement (phi x) (phi (-x))" by (metis hom inf_compl_bot sup_compl_top) lemma phi_inf_principal: "\z . up_filter z = phi x \ up_filter y" proof - let ?F = "Rep_filter (phi x)" let ?G = "Rep_filter (phi (-x))" have 1: "eq_onp filter ?F ?F \ eq_onp filter (\y) (\y)" by (simp add: eq_onp_def) have "filter_complements ?F ?G" apply (intro conjI) apply simp apply simp apply (metis (no_types) phi_complemented sup_filter.rep_eq top_filter.rep_eq) by (metis (no_types) phi_complemented inf_filter.rep_eq bot_filter.rep_eq) hence "is_principal_up (?F \ \y)" using complemented_filter_inf_principal by blast then obtain z where "\z = ?F \ \y" by auto hence "up_filter z = Abs_filter (?F \ \y)" by simp also have "... = Abs_filter ?F \ up_filter y" using 1 inf_filter.abs_eq by force also have "... = phi x \ up_filter y" by (simp add: Rep_filter_inverse) finally show ?thesis by auto qed text \ Quite a bit of filter theory is involved in showing that the intersection of \phi x\ with a principal filter is a principal filter, so the following function can extract its least element. \ fun rho :: "'a \ 'b \ 'b" where "rho x y = (SOME z . up_filter z = phi x \ up_filter y)" lemma rho_char: "up_filter (rho x y) = phi x \ up_filter y" by (metis (mono_tags) someI_ex rho.simps phi_inf_principal) text \ The following results show that the pairs are closed under the given operations. \ lemma pairs_sup_closed: assumes "(x,y) \ pairs" and "(z,w) \ pairs" shows "pairs_sup (x,y) (z,w) \ pairs" proof - from assms obtain u v where "y = phi (-x) \ up_filter u \ w = phi (-z) \ up_filter v" using pairs_def by auto hence "pairs_sup (x,y) (z,w) = (x \ z,(phi (-x) \ up_filter u) \ (phi (-z) \ up_filter v))" by simp also have "... = (x \ z,(phi (-x) \ phi (-z)) \ (phi (-x) \ up_filter v) \ (up_filter u \ phi (-z)) \ (up_filter u \ up_filter v))" by (simp add: inf.sup_commute inf_sup_distrib1 sup_commute sup_left_commute) also have "... = (x \ z,phi (-(x \ z)) \ (phi (-x) \ up_filter v) \ (up_filter u \ phi (-z)) \ (up_filter u \ up_filter v))" using hom by simp also have "... = (x \ z,phi (-(x \ z)) \ up_filter (rho (-x) v) \ up_filter (rho (-z) u) \ (up_filter u \ up_filter v))" by (metis inf.sup_commute rho_char) also have "... = (x \ z,phi (-(x \ z)) \ up_filter (rho (-x) v) \ up_filter (rho (-z) u) \ up_filter (u \ v))" by (metis up_filter_dist_sup) also have "... = (x \ z,phi (-(x \ z)) \ up_filter (rho (-x) v \ rho (-z) u \ (u \ v)))" by (simp add: sup_commute sup_left_commute up_filter_dist_inf) finally show ?thesis using pairs_def by auto qed lemma pairs_inf_closed: assumes "(x,y) \ pairs" and "(z,w) \ pairs" shows "pairs_inf (x,y) (z,w) \ pairs" proof - from assms obtain u v where "y = phi (-x) \ up_filter u \ w = phi (-z) \ up_filter v" using pairs_def by auto hence "pairs_inf (x,y) (z,w) = (x \ z,(phi (-x) \ up_filter u) \ (phi (-z) \ up_filter v))" by simp also have "... = (x \ z,(phi (-x) \ phi (-z)) \ (up_filter u \ up_filter v))" by (simp add: sup_commute sup_left_commute) also have "... = (x \ z,phi (-(x \ z)) \ (up_filter u \ up_filter v))" using hom by simp also have "... = (x \ z,phi (-(x \ z)) \ up_filter (u \ v))" by (simp add: up_filter_dist_inf) finally show ?thesis using pairs_def by auto qed lemma pairs_uminus_closed: "pairs_uminus (x,y) \ pairs" proof - have "pairs_uminus (x,y) = (-x,phi (--x) \ bot)" by simp also have "... = (-x,phi (--x) \ up_filter top)" by (simp add: bot_filter.abs_eq) finally show ?thesis by (metis (mono_tags, lifting) mem_Collect_eq old.prod.case pairs_def) qed lemma pairs_bot_closed: "pairs_bot \ pairs" using pairs_def phi_top triple.hom triple_axioms by fastforce lemma pairs_top_closed: "pairs_top \ pairs" by (metis p_bot pairs_uminus.simps pairs_uminus_closed phi_bot) text \ We prove enough properties of the pair operations so that we can later show they form a Stone algebra. \ lemma pairs_sup_dist_inf: "(x,y) \ pairs \ (z,w) \ pairs \ (u,v) \ pairs \ pairs_sup (x,y) (pairs_inf (z,w) (u,v)) = pairs_inf (pairs_sup (x,y) (z,w)) (pairs_sup (x,y) (u,v))" using sup_inf_distrib1 inf_sup_distrib1 by auto lemma pairs_phi_less_eq: "(x,y) \ pairs \ phi (-x) \ y" using pairs_def by auto lemma pairs_uminus_galois: assumes "(x,y) \ pairs" and "(z,w) \ pairs" shows "pairs_inf (x,y) (z,w) = pairs_bot \ pairs_less_eq (x,y) (pairs_uminus (z,w))" proof - have 1: "x \ z = bot \ y \ w = Abs_filter UNIV \ phi z \ y" by (metis (no_types, lifting) assms(1) heyting.implies_inf_absorb hom le_supE pairs_phi_less_eq sup_bot_right) have 2: "x \ -z \ phi z \ y \ y \ w = Abs_filter UNIV" proof assume 3: "x \ -z \ phi z \ y" have "Abs_filter UNIV = phi z \ phi (-z)" using hom phi_complemented phi_top by auto also have "... \ y \ w" using 3 assms(2) sup_mono pairs_phi_less_eq by auto finally show "y \ w = Abs_filter UNIV" using hom phi_top top.extremum_uniqueI by auto qed have "x \ z = bot \ x \ -z" by (simp add: shunting_1) thus ?thesis using 1 2 Pair_inject pairs_inf.simps pairs_less_eq.simps pairs_uminus.simps by auto qed lemma pairs_stone: "(x,y) \ pairs \ pairs_sup (pairs_uminus (x,y)) (pairs_uminus (pairs_uminus (x,y))) = pairs_top" by (metis hom pairs_sup.simps pairs_uminus.simps phi_bot phi_complemented stone) text \ The following results show how the regular elements and the dense elements among the pairs look like. \ abbreviation "dense_pairs \ { (x,y) . (x,y) \ pairs \ pairs_uminus (x,y) = pairs_bot }" abbreviation "regular_pairs \ { (x,y) . (x,y) \ pairs \ pairs_uminus (pairs_uminus (x,y)) = (x,y) }" abbreviation "is_principal_up_filter x \ \y . x = up_filter y" lemma dense_pairs: "dense_pairs = { (x,y) . x = top \ is_principal_up_filter y }" proof - have "dense_pairs = { (x,y) . (x,y) \ pairs \ x = top }" by (metis Pair_inject compl_bot_eq double_compl pairs_uminus.simps phi_top) also have "... = { (x,y) . (\z . y = up_filter z) \ x = top }" using hom pairs_def by auto finally show ?thesis by auto qed lemma regular_pairs: "regular_pairs = { (x,y) . y = phi (-x) }" using pairs_def pairs_uminus_closed by fastforce text \ The following extraction function will be used in defining one direction of the Stone algebra isomorphism. \ fun rho_pair :: "'a \ 'b filter \ 'b" where "rho_pair (x,y) = (SOME z . up_filter z = phi x \ y)" lemma get_rho_pair_char: assumes "(x,y) \ pairs" shows "up_filter (rho_pair (x,y)) = phi x \ y" proof - from assms obtain w where "y = phi (-x) \ up_filter w" using pairs_def by auto hence "phi x \ y = phi x \ up_filter w" by (simp add: inf_sup_distrib1 phi_complemented) thus ?thesis using rho_char by auto qed lemma sa_iso_pair: "(--x,phi (-x) \ up_filter y) \ pairs" using pairs_def by auto end subsection \The Stone Algebra of a Triple\ text \ In this section we prove that the set of pairs constructed in a triple forms a Stone Algebra. The following type captures the parameter \phi\ on which the type of triples depends. This parameter is the structure map that occurs in the definition of the set of pairs. The set of all structure maps is the set of all bounded lattice homomorphisms (of appropriate type). In order to make it a HOL type, we need to show that at least one such structure map exists. To this end we use the ultrafilter lemma: the required bounded lattice homomorphism is essentially the characteristic map of an ultrafilter, but the latter must exist. In particular, the underlying Boolean algebra must contain at least two elements. \ typedef (overloaded) ('a,'b) phi = "{ f::'a::non_trivial_boolean_algebra \ 'b::distrib_lattice_top filter . bounded_lattice_homomorphism f }" proof - from ultra_filter_exists obtain F :: "'a set" where 1: "ultra_filter F" by auto hence 2: "prime_filter F" using ultra_filter_prime by auto let ?f = "\x . if x\F then top else bot::'b filter" have "bounded_lattice_homomorphism ?f" proof (intro conjI) show "?f bot = bot" using 1 by (meson bot.extremum filter_def subset_eq top.extremum_unique) next show "?f top = top" using 1 by simp next show "\x y . ?f (x \ y) = ?f x \ ?f y" proof (intro allI) fix x y show "?f (x \ y) = ?f x \ ?f y" apply (cases "x \ F"; cases "y \ F") using 1 filter_def apply fastforce using 1 filter_def apply fastforce using 1 filter_def apply fastforce using 2 sup_bot_left by auto qed next show "\x y . ?f (x \ y) = ?f x \ ?f y" proof (intro allI) fix x y show "?f (x \ y) = ?f x \ ?f y" apply (cases "x \ F"; cases "y \ F") using 1 apply (simp add: filter_inf_closed) using 1 apply (metis (mono_tags, lifting) brouwer.inf_sup_ord(4) inf_top_left filter_def) using 1 apply (metis (mono_tags, lifting) brouwer.inf_sup_ord(3) inf_top_right filter_def) using 1 filter_def by force qed qed hence "?f \ {f . bounded_lattice_homomorphism f}" by simp thus ?thesis by meson qed lemma simp_phi [simp]: "bounded_lattice_homomorphism (Rep_phi x)" using Rep_phi by simp setup_lifting type_definition_phi text \ The following implements the dependent type of pairs depending on structure maps. It uses functions from structure maps to pairs with the requirement that, for each structure map, the corresponding pair is contained in the set of pairs constructed for a triple with that structure map. If this type could be defined in the locale \triple\ and instantiated to Stone algebras there, there would be no need for the lifting and we could work with triples directly. \ typedef (overloaded) ('a,'b) lifted_pair = "{ pf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) phi \ 'a \ 'b filter . \f . pf f \ triple.pairs (Rep_phi f) }" proof - have "\f::('a,'b) phi . triple.pairs_bot \ triple.pairs (Rep_phi f)" proof fix f :: "('a,'b) phi" have "triple (Rep_phi f)" by (simp add: triple_def) thus "triple.pairs_bot \ triple.pairs (Rep_phi f)" using triple.regular_pairs triple.phi_top by fastforce qed thus ?thesis by auto qed lemma simp_lifted_pair [simp]: "\f . Rep_lifted_pair pf f \ triple.pairs (Rep_phi f)" using Rep_lifted_pair by simp setup_lifting type_definition_lifted_pair text \ The lifted pairs form a Stone algebra. \ instantiation lifted_pair :: (non_trivial_boolean_algebra,distrib_lattice_top) stone_algebra begin text \ All operations are lifted point-wise. \ lift_definition sup_lifted_pair :: "('a,'b) lifted_pair \ ('a,'b) lifted_pair \ ('a,'b) lifted_pair" is "\xf yf f . triple.pairs_sup (xf f) (yf f)" by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_sup_closed prod.collapse) lift_definition inf_lifted_pair :: "('a,'b) lifted_pair \ ('a,'b) lifted_pair \ ('a,'b) lifted_pair" is "\xf yf f . triple.pairs_inf (xf f) (yf f)" by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_inf_closed prod.collapse) lift_definition uminus_lifted_pair :: "('a,'b) lifted_pair \ ('a,'b) lifted_pair" is "\xf f . triple.pairs_uminus (Rep_phi f) (xf f)" by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_uminus_closed prod.collapse) lift_definition bot_lifted_pair :: "('a,'b) lifted_pair" is "\f . triple.pairs_bot" by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_bot_closed) lift_definition top_lifted_pair :: "('a,'b) lifted_pair" is "\f . triple.pairs_top" by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_top_closed) lift_definition less_eq_lifted_pair :: "('a,'b) lifted_pair \ ('a,'b) lifted_pair \ bool" is "\xf yf . \f . triple.pairs_less_eq (xf f) (yf f)" . lift_definition less_lifted_pair :: "('a,'b) lifted_pair \ ('a,'b) lifted_pair \ bool" is "\xf yf . (\f . triple.pairs_less_eq (xf f) (yf f)) \ \ (\f . triple.pairs_less_eq (yf f) (xf f))" . instance proof intro_classes fix xf yf :: "('a,'b) lifted_pair" show "xf < yf \ xf \ yf \ \ yf \ xf" by (simp add: less_lifted_pair.rep_eq less_eq_lifted_pair.rep_eq) next fix xf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" obtain x1 x2 where "(x1,x2) = ?x" using prod.collapse by blast hence "triple.pairs_less_eq ?x ?x" using 1 by (metis triple.pairs_less_eq.simps order_refl) } thus "xf \ xf" by (simp add: less_eq_lifted_pair.rep_eq) next fix xf yf zf :: "('a,'b) lifted_pair" assume 1: "xf \ yf" and 2: "yf \ zf" { fix f :: "('a,'b) phi" have 3: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" let ?z = "Rep_lifted_pair zf f" obtain x1 x2 y1 y2 z1 z2 where 4: "(x1,x2) = ?x \ (y1,y2) = ?y \ (z1,z2) = ?z" using prod.collapse by blast have "triple.pairs_less_eq ?x ?y \ triple.pairs_less_eq ?y ?z" using 1 2 3 less_eq_lifted_pair.rep_eq by simp hence "triple.pairs_less_eq ?x ?z" using 3 4 by (metis (mono_tags, lifting) triple.pairs_less_eq.simps order_trans) } thus "xf \ zf" by (simp add: less_eq_lifted_pair.rep_eq) next fix xf yf :: "('a,'b) lifted_pair" assume 1: "xf \ yf" and 2: "yf \ xf" { fix f :: "('a,'b) phi" have 3: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" obtain x1 x2 y1 y2 where 4: "(x1,x2) = ?x \ (y1,y2) = ?y" using prod.collapse by blast have "triple.pairs_less_eq ?x ?y \ triple.pairs_less_eq ?y ?x" using 1 2 3 less_eq_lifted_pair.rep_eq by simp hence "?x = ?y" using 3 4 by (metis (mono_tags, lifting) triple.pairs_less_eq.simps antisym) } thus "xf = yf" by (metis Rep_lifted_pair_inverse ext) next fix xf yf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" obtain x1 x2 y1 y2 where "(x1,x2) = ?x \ (y1,y2) = ?y" using prod.collapse by blast hence "triple.pairs_less_eq (triple.pairs_inf ?x ?y) ?y" using 1 by (metis (mono_tags, lifting) inf_sup_ord(2) sup.cobounded2 triple.pairs_inf.simps triple.pairs_less_eq.simps inf_lifted_pair.rep_eq) } thus "xf \ yf \ yf" by (simp add: less_eq_lifted_pair.rep_eq inf_lifted_pair.rep_eq) next fix xf yf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" obtain x1 x2 y1 y2 where "(x1,x2) = ?x \ (y1,y2) = ?y" using prod.collapse by blast hence "triple.pairs_less_eq (triple.pairs_inf ?x ?y) ?x" using 1 by (metis (mono_tags, lifting) inf_sup_ord(1) sup.cobounded1 triple.pairs_inf.simps triple.pairs_less_eq.simps inf_lifted_pair.rep_eq) } thus "xf \ yf \ xf" by (simp add: less_eq_lifted_pair.rep_eq inf_lifted_pair.rep_eq) next fix xf yf zf :: "('a,'b) lifted_pair" assume 1: "xf \ yf" and 2: "xf \ zf" { fix f :: "('a,'b) phi" have 3: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" let ?z = "Rep_lifted_pair zf f" obtain x1 x2 y1 y2 z1 z2 where 4: "(x1,x2) = ?x \ (y1,y2) = ?y \ (z1,z2) = ?z" using prod.collapse by blast have "triple.pairs_less_eq ?x ?y \ triple.pairs_less_eq ?x ?z" using 1 2 3 less_eq_lifted_pair.rep_eq by simp hence "triple.pairs_less_eq ?x (triple.pairs_inf ?y ?z)" using 3 4 by (metis (mono_tags, lifting) le_inf_iff sup.bounded_iff triple.pairs_inf.simps triple.pairs_less_eq.simps) } thus "xf \ yf \ zf" by (simp add: less_eq_lifted_pair.rep_eq inf_lifted_pair.rep_eq) next fix xf yf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" obtain x1 x2 y1 y2 where "(x1,x2) = ?x \ (y1,y2) = ?y" using prod.collapse by blast hence "triple.pairs_less_eq ?x (triple.pairs_sup ?x ?y)" using 1 by (metis (no_types, lifting) inf_commute sup.cobounded1 inf.cobounded2 triple.pairs_sup.simps triple.pairs_less_eq.simps sup_lifted_pair.rep_eq) } thus "xf \ xf \ yf" by (simp add: less_eq_lifted_pair.rep_eq sup_lifted_pair.rep_eq) next fix xf yf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" obtain x1 x2 y1 y2 where "(x1,x2) = ?x \ (y1,y2) = ?y" using prod.collapse by blast hence "triple.pairs_less_eq ?y (triple.pairs_sup ?x ?y)" using 1 by (metis (no_types, lifting) sup.cobounded2 inf.cobounded2 triple.pairs_sup.simps triple.pairs_less_eq.simps sup_lifted_pair.rep_eq) } thus "yf \ xf \ yf" by (simp add: less_eq_lifted_pair.rep_eq sup_lifted_pair.rep_eq) next fix xf yf zf :: "('a,'b) lifted_pair" assume 1: "yf \ xf" and 2: "zf \ xf" { fix f :: "('a,'b) phi" have 3: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" let ?z = "Rep_lifted_pair zf f" obtain x1 x2 y1 y2 z1 z2 where 4: "(x1,x2) = ?x \ (y1,y2) = ?y \ (z1,z2) = ?z" using prod.collapse by blast have "triple.pairs_less_eq ?y ?x \ triple.pairs_less_eq ?z ?x" using 1 2 3 less_eq_lifted_pair.rep_eq by simp hence "triple.pairs_less_eq (triple.pairs_sup ?y ?z) ?x" using 3 4 by (metis (mono_tags, lifting) le_inf_iff sup.bounded_iff triple.pairs_sup.simps triple.pairs_less_eq.simps) } thus "yf \ zf \ xf" by (simp add: less_eq_lifted_pair.rep_eq sup_lifted_pair.rep_eq) next fix xf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" obtain x1 x2 where "(x1,x2) = ?x" using prod.collapse by blast hence "triple.pairs_less_eq triple.pairs_bot ?x" using 1 by (metis bot.extremum top_greatest top_filter.abs_eq triple.pairs_less_eq.simps) } thus "bot \ xf" by (simp add: less_eq_lifted_pair.rep_eq bot_lifted_pair.rep_eq) next fix xf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" obtain x1 x2 where "(x1,x2) = ?x" using prod.collapse by blast hence "triple.pairs_less_eq ?x triple.pairs_top" using 1 by (metis top.extremum bot_least bot_filter.abs_eq triple.pairs_less_eq.simps) } thus "xf \ top" by (simp add: less_eq_lifted_pair.rep_eq top_lifted_pair.rep_eq) next fix xf yf zf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" let ?z = "Rep_lifted_pair zf f" obtain x1 x2 y1 y2 z1 z2 where "(x1,x2) = ?x \ (y1,y2) = ?y \ (z1,z2) = ?z" using prod.collapse by blast hence "triple.pairs_sup ?x (triple.pairs_inf ?y ?z) = triple.pairs_inf (triple.pairs_sup ?x ?y) (triple.pairs_sup ?x ?z)" using 1 by (metis (no_types) sup_inf_distrib1 inf_sup_distrib1 triple.pairs_sup.simps triple.pairs_inf.simps) } thus "xf \ (yf \ zf) = (xf \ yf) \ (xf \ zf)" by (metis Rep_lifted_pair_inverse ext sup_lifted_pair.rep_eq inf_lifted_pair.rep_eq) next fix xf yf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" let ?y = "Rep_lifted_pair yf f" obtain x1 x2 y1 y2 where 2: "(x1,x2) = ?x \ (y1,y2) = ?y" using prod.collapse by blast have "?x \ triple.pairs (Rep_phi f) \ ?y \ triple.pairs (Rep_phi f)" by simp hence "(triple.pairs_inf ?x ?y = triple.pairs_bot) \ triple.pairs_less_eq ?x (triple.pairs_uminus (Rep_phi f) ?y)" using 1 2 by (metis triple.pairs_uminus_galois) } hence "\f . (Rep_lifted_pair (xf \ yf) f = Rep_lifted_pair bot f) \ triple.pairs_less_eq (Rep_lifted_pair xf f) (Rep_lifted_pair (-yf) f)" using bot_lifted_pair.rep_eq inf_lifted_pair.rep_eq uminus_lifted_pair.rep_eq by simp hence "(Rep_lifted_pair (xf \ yf) = Rep_lifted_pair bot) \ xf \ -yf" using less_eq_lifted_pair.rep_eq by auto thus "(xf \ yf = bot) \ (xf \ -yf)" by (simp add: Rep_lifted_pair_inject) next fix xf :: "('a,'b) lifted_pair" { fix f :: "('a,'b) phi" have 1: "triple (Rep_phi f)" by (simp add: triple_def) let ?x = "Rep_lifted_pair xf f" obtain x1 x2 where "(x1,x2) = ?x" using prod.collapse by blast hence "triple.pairs_sup (triple.pairs_uminus (Rep_phi f) ?x) (triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) ?x)) = triple.pairs_top" using 1 by (metis simp_lifted_pair triple.pairs_stone) } hence "Rep_lifted_pair (-xf \ --xf) = Rep_lifted_pair top" using sup_lifted_pair.rep_eq uminus_lifted_pair.rep_eq top_lifted_pair.rep_eq by simp thus "-xf \ --xf = top" by (simp add: Rep_lifted_pair_inject) qed end subsection \The Stone Algebra of the Triple of a Stone Algebra\ text \ In this section we specialise the above construction to a particular structure map, namely the one obtained in the triple of a Stone algebra. For this particular structure map (as well as for any other particular structure map) the resulting type is no longer a dependent type. It is just the set of pairs obtained for the given structure map. \ typedef (overloaded) 'a stone_phi_pair = "triple.pairs (stone_phi::'a::stone_algebra regular \ 'a dense_filter)" using stone_phi.pairs_bot_closed by auto setup_lifting type_definition_stone_phi_pair instantiation stone_phi_pair :: (stone_algebra) sup_inf_top_bot_uminus_ord begin lift_definition sup_stone_phi_pair :: "'a stone_phi_pair \ 'a stone_phi_pair \ 'a stone_phi_pair" is triple.pairs_sup using stone_phi.pairs_sup_closed by auto lift_definition inf_stone_phi_pair :: "'a stone_phi_pair \ 'a stone_phi_pair \ 'a stone_phi_pair" is triple.pairs_inf using stone_phi.pairs_inf_closed by auto lift_definition uminus_stone_phi_pair :: "'a stone_phi_pair \ 'a stone_phi_pair" is "triple.pairs_uminus stone_phi" using stone_phi.pairs_uminus_closed by auto lift_definition bot_stone_phi_pair :: "'a stone_phi_pair" is "triple.pairs_bot" by (rule stone_phi.pairs_bot_closed) lift_definition top_stone_phi_pair :: "'a stone_phi_pair" is "triple.pairs_top" by (rule stone_phi.pairs_top_closed) lift_definition less_eq_stone_phi_pair :: "'a stone_phi_pair \ 'a stone_phi_pair \ bool" is triple.pairs_less_eq . -lift_definition less_stone_phi_pair :: "'a stone_phi_pair \ 'a stone_phi_pair \ bool" is "\xf yf . triple.pairs_less_eq xf yf \ \ triple.pairs_less_eq yf xf" . +lift_definition less_stone_phi_pair :: "'a stone_phi_pair \ 'a stone_phi_pair \ bool" is triple.pairs_less . instance .. end +(* +instantiation stone_phi_pair :: (stone_algebra) stone_algebra +begin + +instance + apply intro_classes + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by (metis (no_types, lifting) Pair_inject compl_bot_eq heyting.implies_order stone_phi.pairs_less_eq.elims(3) stone_phi.phi_top stone_phi.triple_axioms sup_top_left top_greatest triple_def) + subgoal apply transfer by (metis (no_types, lifting) Pair_inject stone_phi.pairs_less_eq.elims(3) top.extremum bot_least bot_filter.abs_eq) + subgoal apply transfer using stone_phi.triple_axioms triple.pairs_sup_dist_inf by fastforce + subgoal apply transfer using stone_phi.pairs_uminus_galois by fastforce + subgoal apply transfer using stone_phi.pairs_stone by fastforce + done + +end +*) + text \ The result is a Stone algebra and could be proved so by repeating and specialising the above proof for lifted pairs. We choose a different approach, namely by embedding the type of pairs into the lifted type. The embedding injects a pair \x\ into a function as the value at the given structure map; this makes the embedding injective. The value of the function at any other structure map needs to be carefully chosen so that the resulting function is a Stone algebra homomorphism. We use \--x\, which is essentially a projection to the regular element component of \x\, whence the image has the structure of a Boolean algebra. \ fun stone_phi_embed :: "'a::non_trivial_stone_algebra stone_phi_pair \ ('a regular,'a dense) lifted_pair" where "stone_phi_embed x = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x)))" text \ The following lemma shows that in both cases the value of the function is a valid pair for the given structure map. \ lemma stone_phi_embed_triple_pair: "(if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x))) \ triple.pairs (Rep_phi f)" by (metis (no_types, hide_lams) Rep_stone_phi_pair simp_phi surj_pair triple.pairs_uminus_closed triple_def) text \ The following result shows that the embedding preserves the operations of Stone algebras. Of course, it is not (yet) a Stone algebra homomorphism as we do not know (yet) that the domain of the embedding is a Stone algebra. To establish the latter is the purpose of the embedding. \ lemma stone_phi_embed_homomorphism: "sup_inf_top_bot_uminus_ord_homomorphism stone_phi_embed" proof (intro conjI) let ?p = "\f . triple.pairs_uminus (Rep_phi f)" let ?pp = "\f x . ?p f (?p f x)" let ?q = "\f x . ?pp f (Rep_stone_phi_pair x)" show "\x y::'a stone_phi_pair . stone_phi_embed (x \ y) = stone_phi_embed x \ stone_phi_embed y" proof (intro allI) fix x y :: "'a stone_phi_pair" have 1: "\f . triple.pairs_sup (?q f x) (?q f y) = ?q f (x \ y)" proof fix f :: "('a regular,'a dense) phi" let ?r = "Rep_phi f" obtain x1 x2 y1 y2 where 2: "(x1,x2) = Rep_stone_phi_pair x \ (y1,y2) = Rep_stone_phi_pair y" using prod.collapse by blast hence "triple.pairs_sup (?q f x) (?q f y) = triple.pairs_sup (?pp f (x1,x2)) (?pp f (y1,y2))" by simp also have "... = triple.pairs_sup (--x1,?r (-x1)) (--y1,?r (-y1))" by (simp add: triple.pairs_uminus.simps triple_def) also have "... = (--x1 \ --y1,?r (-x1) \ ?r (-y1))" by simp also have "... = (--(x1 \ y1),?r (-(x1 \ y1)))" by simp also have "... = ?pp f (x1 \ y1,x2 \ y2)" by (simp add: triple.pairs_uminus.simps triple_def) also have "... = ?pp f (triple.pairs_sup (x1,x2) (y1,y2))" by simp also have "... = ?q f (x \ y)" using 2 by (simp add: sup_stone_phi_pair.rep_eq) finally show "triple.pairs_sup (?q f x) (?q f y) = ?q f (x \ y)" . qed have "stone_phi_embed x \ stone_phi_embed y = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) \ Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y)" by simp also have "... = Abs_lifted_pair (\f . triple.pairs_sup (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) (if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y))" by (rule sup_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then triple.pairs_sup (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else triple.pairs_sup (?q f x) (?q f y))" by (simp add: if_distrib_2) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then triple.pairs_sup (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else ?q f (x \ y))" using 1 by meson also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair (x \ y) else ?q f (x \ y))" by (metis sup_stone_phi_pair.rep_eq) also have "... = stone_phi_embed (x \ y)" by simp finally show "stone_phi_embed (x \ y) = stone_phi_embed x \ stone_phi_embed y" by simp qed next let ?p = "\f . triple.pairs_uminus (Rep_phi f)" let ?pp = "\f x . ?p f (?p f x)" let ?q = "\f x . ?pp f (Rep_stone_phi_pair x)" show "\x y::'a stone_phi_pair . stone_phi_embed (x \ y) = stone_phi_embed x \ stone_phi_embed y" proof (intro allI) fix x y :: "'a stone_phi_pair" have 1: "\f . triple.pairs_inf (?q f x) (?q f y) = ?q f (x \ y)" proof fix f :: "('a regular,'a dense) phi" let ?r = "Rep_phi f" obtain x1 x2 y1 y2 where 2: "(x1,x2) = Rep_stone_phi_pair x \ (y1,y2) = Rep_stone_phi_pair y" using prod.collapse by blast hence "triple.pairs_inf (?q f x) (?q f y) = triple.pairs_inf (?pp f (x1,x2)) (?pp f (y1,y2))" by simp also have "... = triple.pairs_inf (--x1,?r (-x1)) (--y1,?r (-y1))" by (simp add: triple.pairs_uminus.simps triple_def) also have "... = (--x1 \ --y1,?r (-x1) \ ?r (-y1))" by simp also have "... = (--(x1 \ y1),?r (-(x1 \ y1)))" by simp also have "... = ?pp f (x1 \ y1,x2 \ y2)" by (simp add: triple.pairs_uminus.simps triple_def) also have "... = ?pp f (triple.pairs_inf (x1,x2) (y1,y2))" by simp also have "... = ?q f (x \ y)" using 2 by (simp add: inf_stone_phi_pair.rep_eq) finally show "triple.pairs_inf (?q f x) (?q f y) = ?q f (x \ y)" . qed have "stone_phi_embed x \ stone_phi_embed y = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) \ Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y)" by simp also have "... = Abs_lifted_pair (\f . triple.pairs_inf (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) (if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y))" by (rule inf_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then triple.pairs_inf (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else triple.pairs_inf (?q f x) (?q f y))" by (simp add: if_distrib_2) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then triple.pairs_inf (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else ?q f (x \ y))" using 1 by meson also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair (x \ y) else ?q f (x \ y))" by (metis inf_stone_phi_pair.rep_eq) also have "... = stone_phi_embed (x \ y)" by simp finally show "stone_phi_embed (x \ y) = stone_phi_embed x \ stone_phi_embed y" by simp qed next have "stone_phi_embed (top::'a stone_phi_pair) = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair top else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair top)))" by simp also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then (top,bot) else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (top,bot)))" by (metis (no_types, hide_lams) bot_filter.abs_eq top_stone_phi_pair.rep_eq) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then (top,bot) else triple.pairs_uminus (Rep_phi f) (bot,top))" by (metis (no_types, hide_lams) dense_closed_top simp_phi triple.pairs_uminus.simps triple_def) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then (top,bot) else (top,bot))" by (metis (no_types, hide_lams) p_bot simp_phi triple.pairs_uminus.simps triple_def) also have "... = Abs_lifted_pair (\f . (top,Abs_filter {top}))" by (simp add: bot_filter.abs_eq) also have "... = top" by (rule top_lifted_pair.abs_eq[THEN sym]) finally show "stone_phi_embed (top::'a stone_phi_pair) = top" . next have "stone_phi_embed (bot::'a stone_phi_pair) = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair bot else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair bot)))" by simp also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then (bot,top) else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (bot,top)))" by (metis (no_types, hide_lams) top_filter.abs_eq bot_stone_phi_pair.rep_eq) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then (bot,top) else triple.pairs_uminus (Rep_phi f) (top,bot))" by (metis (no_types, hide_lams) p_bot simp_phi triple.pairs_uminus.simps triple_def) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then (bot,top) else (bot,top))" by (metis (no_types, hide_lams) p_top simp_phi triple.pairs_uminus.simps triple_def) also have "... = Abs_lifted_pair (\f . (bot,Abs_filter UNIV))" by (simp add: top_filter.abs_eq) also have "... = bot" by (rule bot_lifted_pair.abs_eq[THEN sym]) finally show "stone_phi_embed (bot::'a stone_phi_pair) = bot" . next let ?p = "\f . triple.pairs_uminus (Rep_phi f)" let ?pp = "\f x . ?p f (?p f x)" let ?q = "\f x . ?pp f (Rep_stone_phi_pair x)" show "\x::'a stone_phi_pair . stone_phi_embed (-x) = -stone_phi_embed x" proof (intro allI) fix x :: "'a stone_phi_pair" have 1: "\f . triple.pairs_uminus (Rep_phi f) (?q f x) = ?q f (-x)" proof fix f :: "('a regular,'a dense) phi" let ?r = "Rep_phi f" obtain x1 x2 where 2: "(x1,x2) = Rep_stone_phi_pair x" using prod.collapse by blast hence "triple.pairs_uminus (Rep_phi f) (?q f x) = triple.pairs_uminus (Rep_phi f) (?pp f (x1,x2))" by simp also have "... = triple.pairs_uminus (Rep_phi f) (--x1,?r (-x1))" by (simp add: triple.pairs_uminus.simps triple_def) also have "... = (---x1,?r (--x1))" by (simp add: triple.pairs_uminus.simps triple_def) also have "... = ?pp f (-x1,stone_phi x1)" by (simp add: triple.pairs_uminus.simps triple_def) also have "... = ?pp f (triple.pairs_uminus stone_phi (x1,x2))" by simp also have "... = ?q f (-x)" using 2 by (simp add: uminus_stone_phi_pair.rep_eq) finally show "triple.pairs_uminus (Rep_phi f) (?q f x) = ?q f (-x)" . qed have "-stone_phi_embed x = -Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x)" by simp also have "... = Abs_lifted_pair (\f . triple.pairs_uminus (Rep_phi f) (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x))" by (rule uminus_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x) else triple.pairs_uminus (Rep_phi f) (?q f x))" by (simp add: if_distrib) also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x) else ?q f (-x))" using 1 by meson also have "... = Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair (-x) else ?q f (-x))" by (metis uminus_stone_phi_pair.rep_eq) also have "... = stone_phi_embed (-x)" by simp finally show "stone_phi_embed (-x) = -stone_phi_embed x" by simp qed next let ?p = "\f . triple.pairs_uminus (Rep_phi f)" let ?pp = "\f x . ?p f (?p f x)" let ?q = "\f x . ?pp f (Rep_stone_phi_pair x)" show "\x y::'a stone_phi_pair . x \ y \ stone_phi_embed x \ stone_phi_embed y" proof (intro allI, rule impI) fix x y :: "'a stone_phi_pair" assume 1: "x \ y" have "\f . triple.pairs_less_eq (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) (if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y)" proof fix f :: "('a regular,'a dense) phi" let ?r = "Rep_phi f" obtain x1 x2 y1 y2 where 2: "(x1,x2) = Rep_stone_phi_pair x \ (y1,y2) = Rep_stone_phi_pair y" using prod.collapse by blast have "x1 \ y1" using 1 2 by (metis less_eq_stone_phi_pair.rep_eq stone_phi.pairs_less_eq.simps) hence "--x1 \ --y1 \ ?r (-y1) \ ?r (-x1)" by (metis compl_le_compl_iff le_iff_sup simp_phi) hence "triple.pairs_less_eq (--x1,?r (-x1)) (--y1,?r (-y1))" by simp hence "triple.pairs_less_eq (?pp f (x1,x2)) (?pp f (y1,y2))" by (simp add: triple.pairs_uminus.simps triple_def) hence "triple.pairs_less_eq (?q f x) (?q f y)" using 2 by simp hence "if ?r = stone_phi then triple.pairs_less_eq (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else triple.pairs_less_eq (?q f x) (?q f y)" using 1 by (simp add: less_eq_stone_phi_pair.rep_eq) thus "triple.pairs_less_eq (if ?r = stone_phi then Rep_stone_phi_pair x else ?q f x) (if ?r = stone_phi then Rep_stone_phi_pair y else ?q f y)" by (simp add: if_distrib_2) qed hence "Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) \ Abs_lifted_pair (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y)" by (subst less_eq_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair) thus "stone_phi_embed x \ stone_phi_embed y" by simp qed qed text \ The following lemmas show that the embedding is injective and reflects the order. The latter allows us to easily inherit properties involving inequalities from the target of the embedding, without transforming them to equations. \ lemma stone_phi_embed_injective: "inj stone_phi_embed" proof (rule injI) fix x y :: "'a stone_phi_pair" have 1: "Rep_phi (Abs_phi stone_phi) = stone_phi" by (simp add: Abs_phi_inverse stone_phi.hom) assume 2: "stone_phi_embed x = stone_phi_embed y" have "\x::'a stone_phi_pair . Rep_lifted_pair (stone_phi_embed x) = (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x)))" by (simp add: Abs_lifted_pair_inverse stone_phi_embed_triple_pair) hence "(\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x))) = (\f . if Rep_phi f = stone_phi then Rep_stone_phi_pair y else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair y)))" using 2 by metis hence "Rep_stone_phi_pair x = Rep_stone_phi_pair y" using 1 by metis thus "x = y" by (simp add: Rep_stone_phi_pair_inject) qed lemma stone_phi_embed_order_injective: assumes "stone_phi_embed x \ stone_phi_embed y" shows "x \ y" proof - let ?f = "Abs_phi stone_phi" have "\f . triple.pairs_less_eq (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x))) (if Rep_phi f = stone_phi then Rep_stone_phi_pair y else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair y)))" using assms by (subst less_eq_lifted_pair.abs_eq[THEN sym]) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair) hence "triple.pairs_less_eq (if Rep_phi ?f = (stone_phi::'a regular \ 'a dense_filter) then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi ?f) (triple.pairs_uminus (Rep_phi ?f) (Rep_stone_phi_pair x))) (if Rep_phi ?f = (stone_phi::'a regular \ 'a dense_filter) then Rep_stone_phi_pair y else triple.pairs_uminus (Rep_phi ?f) (triple.pairs_uminus (Rep_phi ?f) (Rep_stone_phi_pair y)))" by simp hence "triple.pairs_less_eq (Rep_stone_phi_pair x) (Rep_stone_phi_pair y)" by (simp add: Abs_phi_inverse stone_phi.hom) thus "x \ y" by (simp add: less_eq_stone_phi_pair.rep_eq) qed +lemma stone_phi_embed_strict_order_isomorphism: + "x < y \ stone_phi_embed x < stone_phi_embed y" + by (smt less_eq_stone_phi_pair.rep_eq less_le_not_le less_stone_phi_pair.rep_eq stone_phi.pairs_less.elims(2,3) stone_phi_embed_homomorphism stone_phi_embed_order_injective) + text \ Now all Stone algebra axioms can be inherited using the embedding. This is due to the fact that the axioms are universally quantified equations or conditional equations (or inequalities); this is called a quasivariety in universal algebra. It would be useful to have this construction available for arbitrary quasivarieties. \ instantiation stone_phi_pair :: (non_trivial_stone_algebra) stone_algebra begin instance apply intro_classes - apply (simp add: less_stone_phi_pair.rep_eq less_eq_stone_phi_pair.rep_eq) + apply (metis (mono_tags, lifting) stone_phi_embed_homomorphism stone_phi_embed_strict_order_isomorphism stone_phi_embed_order_injective less_le_not_le) apply (simp add: stone_phi_embed_order_injective) apply (meson order.trans stone_phi_embed_homomorphism stone_phi_embed_order_injective) apply (meson stone_phi_embed_homomorphism antisym stone_phi_embed_injective injD) apply (metis inf.sup_ge1 stone_phi_embed_homomorphism stone_phi_embed_order_injective) apply (metis inf.sup_ge2 stone_phi_embed_homomorphism stone_phi_embed_order_injective) apply (metis inf_greatest stone_phi_embed_homomorphism stone_phi_embed_order_injective) apply (metis stone_phi_embed_homomorphism stone_phi_embed_order_injective sup_ge1) apply (metis stone_phi_embed_homomorphism stone_phi_embed_order_injective sup.cobounded2) apply (metis stone_phi_embed_homomorphism stone_phi_embed_order_injective sup_least) apply (metis bot.extremum stone_phi_embed_homomorphism stone_phi_embed_order_injective) apply (metis stone_phi_embed_homomorphism stone_phi_embed_order_injective top_greatest) apply (metis (mono_tags, lifting) stone_phi_embed_homomorphism sup_inf_distrib1 stone_phi_embed_injective injD) apply (metis stone_phi_embed_homomorphism stone_phi_embed_injective injD stone_phi_embed_order_injective pseudo_complement) by (metis injD stone_phi_embed_homomorphism stone_phi_embed_injective stone) end subsection \Stone Algebra Isomorphism\ text \ In this section we prove that the Stone algebra of the triple of a Stone algebra is isomorphic to the original Stone algebra. The following two definitions give the isomorphism. \ abbreviation sa_iso_inv :: "'a::non_trivial_stone_algebra stone_phi_pair \ 'a" where "sa_iso_inv \ \p . Rep_regular (fst (Rep_stone_phi_pair p)) \ Rep_dense (triple.rho_pair stone_phi (Rep_stone_phi_pair p))" abbreviation sa_iso :: "'a::non_trivial_stone_algebra \ 'a stone_phi_pair" where "sa_iso \ \x . Abs_stone_phi_pair (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x)))" lemma sa_iso_triple_pair: "(Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ triple.pairs stone_phi" by (metis (mono_tags, lifting) double_compl eq_onp_same_args stone_phi.sa_iso_pair uminus_regular.abs_eq) lemma stone_phi_inf_dense: "stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (y \ -y)) \ up_filter (Abs_dense (y \ -y \ x))" proof - have "Rep_filter (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (y \ -y))) \ \(Abs_dense (y \ -y \ x))" proof fix z :: "'a dense" let ?r = "Rep_dense z" assume "z \ Rep_filter (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (y \ -y)))" also have "... = Rep_filter (stone_phi (Abs_regular (-x))) \ Rep_filter (up_filter (Abs_dense (y \ -y)))" by (simp add: inf_filter.rep_eq) - also have "... = stone_phi_set (Abs_regular (-x)) \ \(Abs_dense (y \ -y))" - by (metis Abs_filter_inverse mem_Collect_eq up_filter stone_phi_set_filter stone_phi_def) + also have "... = stone_phi_base (Abs_regular (-x)) \ \(Abs_dense (y \ -y))" + by (metis Abs_filter_inverse mem_Collect_eq up_filter stone_phi_base_filter stone_phi_def) finally have "--x \ ?r \ Abs_dense (y \ -y) \ z" by (metis (mono_tags, lifting) Abs_regular_inverse Int_Collect mem_Collect_eq) hence "--x \ ?r \ y \ -y \ ?r" by (simp add: Abs_dense_inverse less_eq_dense.rep_eq) hence "y \ -y \ x \ ?r" using order_trans pp_increasing by auto hence "Abs_dense (y \ -y \ x) \ Abs_dense ?r" by (subst less_eq_dense.abs_eq) (simp_all add: eq_onp_same_args) thus "z \ \(Abs_dense (y \ -y \ x))" by (simp add: Rep_dense_inverse) qed hence "Abs_filter (Rep_filter (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (y \ -y)))) \ up_filter (Abs_dense (y \ -y \ x))" by (simp add: eq_onp_same_args less_eq_filter.abs_eq) thus ?thesis by (simp add: Rep_filter_inverse) qed lemma stone_phi_complement: "complement (stone_phi (Abs_regular (-x))) (stone_phi (Abs_regular (--x)))" by (metis (mono_tags, lifting) eq_onp_same_args stone_phi.phi_complemented uminus_regular.abs_eq) lemma up_dense_stone_phi: "up_filter (Abs_dense (x \ -x)) \ stone_phi (Abs_regular (--x))" proof - - have "\(Abs_dense (x \ -x)) \ stone_phi_set (Abs_regular (--x))" + have "\(Abs_dense (x \ -x)) \ stone_phi_base (Abs_regular (--x))" proof fix z :: "'a dense" let ?r = "Rep_dense z" assume "z \ \(Abs_dense (x \ -x))" hence "---x \ ?r" by (simp add: Abs_dense_inverse less_eq_dense.rep_eq) hence "-Rep_regular (Abs_regular (--x)) \ ?r" by (metis (mono_tags, lifting) Abs_regular_inverse mem_Collect_eq) - thus "z \ stone_phi_set (Abs_regular (--x))" + thus "z \ stone_phi_base (Abs_regular (--x))" by simp qed thus ?thesis - by (unfold stone_phi_def, subst less_eq_filter.abs_eq, simp_all add: eq_onp_same_args stone_phi_set_filter) + by (unfold stone_phi_def, subst less_eq_filter.abs_eq, simp_all add: eq_onp_same_args stone_phi_base_filter) qed text \ The following two results prove that the isomorphisms are mutually inverse. \ lemma sa_iso_left_invertible: "sa_iso_inv (sa_iso x) = x" proof - have "up_filter (triple.rho_pair stone_phi (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x)))) = stone_phi (Abs_regular (--x)) \ (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x)))" using sa_iso_triple_pair stone_phi.get_rho_pair_char by blast also have "... = stone_phi (Abs_regular (--x)) \ up_filter (Abs_dense (x \ -x))" by (simp add: inf.sup_commute inf_sup_distrib1 stone_phi_complement) also have "... = up_filter (Abs_dense (x \ -x))" using up_dense_stone_phi inf.absorb2 by auto finally have 1: "triple.rho_pair stone_phi (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) = Abs_dense (x \ -x)" using up_filter_injective by auto have "sa_iso_inv (sa_iso x) = (\p . Rep_regular (fst p) \ Rep_dense (triple.rho_pair stone_phi p)) (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x)))" by (simp add: Abs_stone_phi_pair_inverse sa_iso_triple_pair) also have "... = Rep_regular (Abs_regular (--x)) \ Rep_dense (triple.rho_pair stone_phi (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))))" by simp also have "... = --x \ Rep_dense (Abs_dense (x \ -x))" using 1 by (subst Abs_regular_inverse) auto also have "... = --x \ (x \ -x)" by (subst Abs_dense_inverse) simp_all also have "... = x" by simp finally show ?thesis by auto qed lemma sa_iso_right_invertible: "sa_iso (sa_iso_inv p) = p" proof - obtain x y where 1: "(x,y) = Rep_stone_phi_pair p" using prod.collapse by blast hence 2: "(x,y) \ triple.pairs stone_phi" by (simp add: Rep_stone_phi_pair) hence 3: "stone_phi (-x) \ y" by (simp add: stone_phi.pairs_phi_less_eq) have 4: "\z . z \ Rep_filter (stone_phi x \ y) \ -Rep_regular x \ Rep_dense z" proof (rule allI, rule impI) fix z :: "'a dense" let ?r = "Rep_dense z" assume "z \ Rep_filter (stone_phi x \ y)" hence "z \ Rep_filter (stone_phi x)" by (simp add: inf_filter.rep_eq) - also have "... = stone_phi_set x" - by (simp add: stone_phi_def Abs_filter_inverse stone_phi_set_filter) + also have "... = stone_phi_base x" + by (simp add: stone_phi_def Abs_filter_inverse stone_phi_base_filter) finally show "-Rep_regular x \ ?r" by simp qed have "triple.rho_pair stone_phi (x,y) \ \(triple.rho_pair stone_phi (x,y))" by simp also have "... = Rep_filter (Abs_filter (\(triple.rho_pair stone_phi (x,y))))" by (simp add: Abs_filter_inverse) also have "... = Rep_filter (stone_phi x \ y)" using 2 stone_phi.get_rho_pair_char by fastforce finally have "triple.rho_pair stone_phi (x,y) \ Rep_filter (stone_phi x \ y)" by simp hence 5: "-Rep_regular x \ Rep_dense (triple.rho_pair stone_phi (x,y))" using 4 by simp have 6: "sa_iso_inv p = Rep_regular x \ Rep_dense (triple.rho_pair stone_phi (x,y))" using 1 by (metis fstI) hence "-sa_iso_inv p = -Rep_regular x" by simp hence "sa_iso (sa_iso_inv p) = Abs_stone_phi_pair (Abs_regular (--Rep_regular x),stone_phi (Abs_regular (-Rep_regular x)) \ up_filter (Abs_dense ((Rep_regular x \ Rep_dense (triple.rho_pair stone_phi (x,y))) \ -Rep_regular x)))" using 6 by simp also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \ up_filter (Abs_dense ((Rep_regular x \ Rep_dense (triple.rho_pair stone_phi (x,y))) \ -Rep_regular x)))" by (metis (mono_tags, lifting) Rep_regular_inverse double_compl uminus_regular.rep_eq) also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \ up_filter (Abs_dense (Rep_dense (triple.rho_pair stone_phi (x,y)) \ -Rep_regular x)))" by (metis inf_sup_aci(5) maddux_3_21_pp simp_regular) also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \ up_filter (Abs_dense (Rep_dense (triple.rho_pair stone_phi (x,y)))))" using 5 by (simp add: sup.absorb1) also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \ up_filter (triple.rho_pair stone_phi (x,y)))" by (simp add: Rep_dense_inverse) also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \ (stone_phi x \ y))" using 2 stone_phi.get_rho_pair_char by fastforce also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \ y)" by (simp add: stone_phi.phi_complemented sup.commute sup_inf_distrib1) also have "... = Abs_stone_phi_pair (x,y)" using 3 by (simp add: le_iff_sup) also have "... = p" using 1 by (simp add: Rep_stone_phi_pair_inverse) finally show ?thesis . qed text \ It remains to show the homomorphism properties, which is done in the following result. \ lemma sa_iso: "stone_algebra_isomorphism sa_iso" proof (intro conjI) have "Abs_stone_phi_pair (Abs_regular (--bot),stone_phi (Abs_regular (-bot)) \ up_filter (Abs_dense (bot \ -bot))) = Abs_stone_phi_pair (bot,stone_phi top \ up_filter top)" by (simp add: bot_regular.abs_eq top_regular.abs_eq top_dense.abs_eq) also have "... = Abs_stone_phi_pair (bot,stone_phi top)" by (simp add: stone_phi.hom) also have "... = bot" by (simp add: bot_stone_phi_pair_def stone_phi.phi_top) finally show "sa_iso bot = bot" . next have "Abs_stone_phi_pair (Abs_regular (--top),stone_phi (Abs_regular (-top)) \ up_filter (Abs_dense (top \ -top))) = Abs_stone_phi_pair (top,stone_phi bot \ up_filter top)" by (simp add: bot_regular.abs_eq top_regular.abs_eq top_dense.abs_eq) also have "... = top" by (simp add: stone_phi.phi_bot top_stone_phi_pair_def) finally show "sa_iso top = top" . next have 1: "\x y::'a . dense (x \ -x \ y)" by simp have 2: "\x y::'a . up_filter (Abs_dense (x \ -x \ y)) \ (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y)))" proof (intro allI) fix x y :: 'a let ?u = "Abs_dense (x \ -x \ --y)" let ?v = "Abs_dense (y \ -y)" have "\(Abs_dense (x \ -x \ y)) \ Rep_filter (stone_phi (Abs_regular (-y)) \ up_filter ?v)" proof fix z assume "z \ \(Abs_dense (x \ -x \ y))" hence "Abs_dense (x \ -x \ y) \ z" by simp hence 3: "x \ -x \ y \ Rep_dense z" by (simp add: Abs_dense_inverse less_eq_dense.rep_eq) have "y \ x \ -x \ --y" by (simp add: le_supI2 pp_increasing) hence "(x \ -x \ --y) \ (y \ -y) = y \ ((x \ -x \ --y) \ -y)" by (simp add: le_iff_sup sup_inf_distrib1) also have "... = y \ ((x \ -x) \ -y)" by (simp add: inf_commute inf_sup_distrib1) also have "... \ Rep_dense z" using 3 by (meson le_infI1 sup.bounded_iff) finally have "Abs_dense ((x \ -x \ --y) \ (y \ -y)) \ z" by (simp add: Abs_dense_inverse less_eq_dense.rep_eq) hence 4: "?u \ ?v \ z" by (simp add: eq_onp_same_args inf_dense.abs_eq) have "-Rep_regular (Abs_regular (-y)) = --y" by (metis (mono_tags, lifting) mem_Collect_eq Abs_regular_inverse) also have "... \ Rep_dense ?u" by (simp add: Abs_dense_inverse) - finally have "?u \ stone_phi_set (Abs_regular (-y))" + finally have "?u \ stone_phi_base (Abs_regular (-y))" by simp hence 5: "?u \ Rep_filter (stone_phi (Abs_regular (-y)))" - by (metis mem_Collect_eq stone_phi_def stone_phi_set_filter Abs_filter_inverse) + by (metis mem_Collect_eq stone_phi_def stone_phi_base_filter Abs_filter_inverse) have "?v \ \?v" by simp hence "?v \ Rep_filter (up_filter ?v)" by (metis Abs_filter_inverse mem_Collect_eq up_filter) thus "z \ Rep_filter (stone_phi (Abs_regular (-y)) \ up_filter ?v)" - using 4 5 sup_filter.rep_eq by blast + using 4 5 sup_filter.rep_eq filter_sup_def by blast qed hence "up_filter (Abs_dense (x \ -x \ y)) \ Abs_filter (Rep_filter (stone_phi (Abs_regular (-y)) \ up_filter ?v))" by (simp add: eq_onp_same_args less_eq_filter.abs_eq) also have "... = stone_phi (Abs_regular (-y)) \ up_filter ?v" by (simp add: Rep_filter_inverse) finally show "up_filter (Abs_dense (x \ -x \ y)) \ (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y)))" by (metis le_infI le_supI2 sup_bot.right_neutral up_filter_dense_antitone) qed have 6: "\x::'a . in_p_image (-x)" by auto show "\x y::'a . sa_iso (x \ y) = sa_iso x \ sa_iso y" proof (intro allI) fix x y :: 'a have 7: "up_filter (Abs_dense (x \ -x)) \ up_filter (Abs_dense (y \ -y)) \ up_filter (Abs_dense (y \ -y \ x))" proof - have "up_filter (Abs_dense (x \ -x)) \ up_filter (Abs_dense (y \ -y)) = up_filter (Abs_dense (x \ -x) \ Abs_dense (y \ -y))" by (metis up_filter_dist_sup) also have "... = up_filter (Abs_dense (x \ -x \ (y \ -y)))" by (subst sup_dense.abs_eq) (simp_all add: eq_onp_same_args) also have "... = up_filter (Abs_dense (y \ -y \ x \ -x))" by (simp add: sup_commute sup_left_commute) also have "... \ up_filter (Abs_dense (y \ -y \ x))" using up_filter_dense_antitone by auto finally show ?thesis . qed have "Abs_dense (x \ y \ -(x \ y)) = Abs_dense ((x \ -x \ y) \ (y \ -y \ x))" by (simp add: sup_commute sup_inf_distrib1 sup_left_commute) also have "... = Abs_dense (x \ -x \ y) \ Abs_dense (y \ -y \ x)" using 1 by (metis (mono_tags, lifting) Abs_dense_inverse Rep_dense_inverse inf_dense.rep_eq mem_Collect_eq) finally have 8: "up_filter (Abs_dense (x \ y \ -(x \ y))) = up_filter (Abs_dense (x \ -x \ y)) \ up_filter (Abs_dense (y \ -y \ x))" by (simp add: up_filter_dist_inf) also have "... \ (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y)))" using 2 by (simp add: inf.sup_commute le_sup_iff) finally have 9: "(stone_phi (Abs_regular (-x)) \ stone_phi (Abs_regular (-y))) \ up_filter (Abs_dense (x \ y \ -(x \ y))) \ ..." by (simp add: le_supI1) have "... = (stone_phi (Abs_regular (-x)) \ stone_phi (Abs_regular (-y))) \ (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (y \ -y))) \ ((up_filter (Abs_dense (x \ -x)) \ stone_phi (Abs_regular (-y))) \ (up_filter (Abs_dense (x \ -x)) \ up_filter (Abs_dense (y \ -y))))" by (metis (no_types) inf_sup_distrib1 inf_sup_distrib2) also have "... \ (stone_phi (Abs_regular (-x)) \ stone_phi (Abs_regular (-y))) \ up_filter (Abs_dense (y \ -y \ x)) \ ((up_filter (Abs_dense (x \ -x)) \ stone_phi (Abs_regular (-y))) \ (up_filter (Abs_dense (x \ -x)) \ up_filter (Abs_dense (y \ -y))))" by (meson sup_left_isotone sup_right_isotone stone_phi_inf_dense) also have "... \ (stone_phi (Abs_regular (-x)) \ stone_phi (Abs_regular (-y))) \ up_filter (Abs_dense (y \ -y \ x)) \ (up_filter (Abs_dense (x \ -x \ y)) \ (up_filter (Abs_dense (x \ -x)) \ up_filter (Abs_dense (y \ -y))))" by (metis inf.commute sup_left_isotone sup_right_isotone stone_phi_inf_dense) also have "... \ (stone_phi (Abs_regular (-x)) \ stone_phi (Abs_regular (-y))) \ up_filter (Abs_dense (y \ -y \ x)) \ up_filter (Abs_dense (x \ -x \ y))" using 7 by (simp add: sup.absorb1 sup_commute sup_left_commute) also have "... = (stone_phi (Abs_regular (-x)) \ stone_phi (Abs_regular (-y))) \ up_filter (Abs_dense (x \ y \ -(x \ y)))" using 8 by (simp add: sup.commute sup.left_commute) finally have "(stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y))) = ..." using 9 using antisym by blast also have "... = stone_phi (Abs_regular (-x) \ Abs_regular (-y)) \ up_filter (Abs_dense (x \ y \ -(x \ y)))" by (simp add: stone_phi.hom) also have "... = stone_phi (Abs_regular (-(x \ y))) \ up_filter (Abs_dense (x \ y \ -(x \ y)))" using 6 by (subst inf_regular.abs_eq) (simp_all add: eq_onp_same_args) finally have 10: "stone_phi (Abs_regular (-(x \ y))) \ up_filter (Abs_dense (x \ y \ -(x \ y))) = (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y)))" by simp have "Abs_regular (--(x \ y)) = Abs_regular (--x) \ Abs_regular (--y)" using 6 by (subst sup_regular.abs_eq) (simp_all add: eq_onp_same_args) hence "Abs_stone_phi_pair (Abs_regular (--(x \ y)),stone_phi (Abs_regular (-(x \ y))) \ up_filter (Abs_dense (x \ y \ -(x \ y)))) = Abs_stone_phi_pair (triple.pairs_sup (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) (Abs_regular (--y),stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y))))" using 10 by auto also have "... = Abs_stone_phi_pair (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ Abs_stone_phi_pair (Abs_regular (--y),stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y)))" by (rule sup_stone_phi_pair.abs_eq[THEN sym]) (simp_all add: eq_onp_same_args sa_iso_triple_pair) finally show "sa_iso (x \ y) = sa_iso x \ sa_iso y" . qed next have 1: "\x y::'a . dense (x \ -x \ y)" by simp have 2: "\x::'a . in_p_image (-x)" by auto have 3: "\x y::'a . stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (x \ -x)) = stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (x \ -x \ -y))" proof (intro allI) fix x y :: 'a have 4: "up_filter (Abs_dense (x \ -x)) \ stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (x \ -x \ -y))" by (metis (no_types, lifting) complement_shunting stone_phi_inf_dense stone_phi_complement complement_symmetric) have "up_filter (Abs_dense (x \ -x \ -y)) \ up_filter (Abs_dense (x \ -x))" by (metis sup_idem up_filter_dense_antitone) thus "stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (x \ -x)) = stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (x \ -x \ -y))" using 4 by (simp add: le_iff_sup sup_commute sup_left_commute) qed show "\x y::'a . sa_iso (x \ y) = sa_iso x \ sa_iso y" proof (intro allI) fix x y :: 'a have "Abs_dense ((x \ y) \ -(x \ y)) = Abs_dense ((x \ -x \ -y) \ (y \ -y \ -x))" by (simp add: sup_commute sup_inf_distrib1 sup_left_commute) also have "... = Abs_dense (x \ -x \ -y) \ Abs_dense (y \ -y \ -x)" using 1 by (metis (mono_tags, lifting) Abs_dense_inverse Rep_dense_inverse inf_dense.rep_eq mem_Collect_eq) finally have 5: "up_filter (Abs_dense ((x \ y) \ -(x \ y))) = up_filter (Abs_dense (x \ -x \ -y)) \ up_filter (Abs_dense (y \ -y \ -x))" by (simp add: up_filter_dist_inf) have "(stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y))) = (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (x \ -x))) \ (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (y \ -y)))" by (simp add: inf_sup_aci(6) sup_left_commute) also have "... = (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (x \ -x \ -y))) \ (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (y \ -y \ -x)))" using 3 by simp also have "... = (stone_phi (Abs_regular (-x)) \ stone_phi (Abs_regular (-y))) \ (up_filter (Abs_dense (x \ -x \ -y)) \ up_filter (Abs_dense (y \ -y \ -x)))" by (simp add: inf_sup_aci(6) sup_left_commute) also have "... = (stone_phi (Abs_regular (-x)) \ stone_phi (Abs_regular (-y))) \ up_filter (Abs_dense ((x \ y) \ -(x \ y)))" using 5 by (simp add: sup.commute sup.left_commute) finally have "(stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y))) = ..." by simp also have "... = stone_phi (Abs_regular (-x) \ Abs_regular (-y)) \ up_filter (Abs_dense ((x \ y) \ -(x \ y)))" by (simp add: stone_phi.hom) also have "... = stone_phi (Abs_regular (-(x \ y))) \ up_filter (Abs_dense ((x \ y) \ -(x \ y)))" using 2 by (subst sup_regular.abs_eq) (simp_all add: eq_onp_same_args) finally have 6: "stone_phi (Abs_regular (-(x \ y))) \ up_filter (Abs_dense ((x \ y) \ -(x \ y))) = (stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ (stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y)))" by simp have "Abs_regular (--(x \ y)) = Abs_regular (--x) \ Abs_regular (--y)" using 2 by (subst inf_regular.abs_eq) (simp_all add: eq_onp_same_args) hence "Abs_stone_phi_pair (Abs_regular (--(x \ y)),stone_phi (Abs_regular (-(x \ y))) \ up_filter (Abs_dense ((x \ y) \ -(x \ y)))) = Abs_stone_phi_pair (triple.pairs_inf (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) (Abs_regular (--y),stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y))))" using 6 by auto also have "... = Abs_stone_phi_pair (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))) \ Abs_stone_phi_pair (Abs_regular (--y),stone_phi (Abs_regular (-y)) \ up_filter (Abs_dense (y \ -y)))" by (rule inf_stone_phi_pair.abs_eq[THEN sym]) (simp_all add: eq_onp_same_args sa_iso_triple_pair) finally show "sa_iso (x \ y) = sa_iso x \ sa_iso y" . qed next show "\x::'a . sa_iso (-x) = -sa_iso x" proof fix x :: 'a have "sa_iso (-x) = Abs_stone_phi_pair (Abs_regular (---x),stone_phi (Abs_regular (--x)) \ up_filter top)" by (simp add: top_dense_def) also have "... = Abs_stone_phi_pair (Abs_regular (---x),stone_phi (Abs_regular (--x)))" by (metis bot_filter.abs_eq sup_bot.right_neutral up_top) also have "... = Abs_stone_phi_pair (triple.pairs_uminus stone_phi (Abs_regular (--x),stone_phi (Abs_regular (-x)) \ up_filter (Abs_dense (x \ -x))))" by (subst uminus_regular.abs_eq[THEN sym], unfold eq_onp_same_args) auto also have "... = -sa_iso x" by (simp add: eq_onp_def sa_iso_triple_pair uminus_stone_phi_pair.abs_eq) finally show "sa_iso (-x) = -sa_iso x" by simp qed next show "bij sa_iso" by (metis (mono_tags, lifting) sa_iso_left_invertible sa_iso_right_invertible invertible_bij[where g=sa_iso_inv]) qed subsection \Triple Isomorphism\ text \ In this section we prove that the triple of the Stone algebra of a triple is isomorphic to the original triple. The notion of isomorphism for triples is described in \cite{ChenGraetzer1969}. It amounts to an isomorphism of Boolean algebras, an isomorphism of distributive lattices with a greatest element, and a commuting diagram involving the structure maps. \ subsubsection \Boolean Algebra Isomorphism\ text \ We first define and prove the isomorphism of Boolean algebras. Because the Stone algebra of a triple is implemented as a lifted pair, we also lift the Boolean algebra. \ typedef (overloaded) ('a,'b) lifted_boolean_algebra = "{ xf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) phi \ 'a . True }" by simp setup_lifting type_definition_lifted_boolean_algebra instantiation lifted_boolean_algebra :: (non_trivial_boolean_algebra,distrib_lattice_top) boolean_algebra begin lift_definition sup_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \ ('a,'b) lifted_boolean_algebra \ ('a,'b) lifted_boolean_algebra" is "\xf yf f . sup (xf f) (yf f)" . lift_definition inf_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \ ('a,'b) lifted_boolean_algebra \ ('a,'b) lifted_boolean_algebra" is "\xf yf f . inf (xf f) (yf f)" . lift_definition minus_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \ ('a,'b) lifted_boolean_algebra \ ('a,'b) lifted_boolean_algebra" is "\xf yf f . minus (xf f) (yf f)" . lift_definition uminus_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \ ('a,'b) lifted_boolean_algebra" is "\xf f . uminus (xf f)" . lift_definition bot_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra" is "\f . bot" .. lift_definition top_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra" is "\f . top" .. lift_definition less_eq_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \ ('a,'b) lifted_boolean_algebra \ bool" is "\xf yf . \f . less_eq (xf f) (yf f)" . lift_definition less_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \ ('a,'b) lifted_boolean_algebra \ bool" is "\xf yf . (\f . less_eq (xf f) (yf f)) \ \ (\f . less_eq (yf f) (xf f))" . instance apply intro_classes - apply (simp add: less_eq_lifted_boolean_algebra.rep_eq less_lifted_boolean_algebra.rep_eq) - apply (simp add: less_eq_lifted_boolean_algebra.rep_eq) - using less_eq_lifted_boolean_algebra.rep_eq order_trans apply fastforce - apply (metis less_eq_lifted_boolean_algebra.rep_eq antisym ext Rep_lifted_boolean_algebra_inject) - apply (simp add: inf_lifted_boolean_algebra.rep_eq less_eq_lifted_boolean_algebra.rep_eq) - apply (simp add: inf_lifted_boolean_algebra.rep_eq less_eq_lifted_boolean_algebra.rep_eq) - apply (simp add: inf_lifted_boolean_algebra.rep_eq less_eq_lifted_boolean_algebra.rep_eq) - apply (simp add: sup_lifted_boolean_algebra.rep_eq less_eq_lifted_boolean_algebra.rep_eq) - apply (simp add: less_eq_lifted_boolean_algebra.rep_eq sup_lifted_boolean_algebra.rep_eq) - apply (simp add: less_eq_lifted_boolean_algebra.rep_eq sup_lifted_boolean_algebra.rep_eq) - apply (simp add: bot_lifted_boolean_algebra.rep_eq less_eq_lifted_boolean_algebra.rep_eq) - apply (simp add: less_eq_lifted_boolean_algebra.rep_eq top_lifted_boolean_algebra.rep_eq) - apply (unfold Rep_lifted_boolean_algebra_inject[THEN sym] sup_lifted_boolean_algebra.rep_eq inf_lifted_boolean_algebra.rep_eq, simp add: sup_inf_distrib1) - apply (unfold Rep_lifted_boolean_algebra_inject[THEN sym] inf_lifted_boolean_algebra.rep_eq uminus_lifted_boolean_algebra.rep_eq bot_lifted_boolean_algebra.rep_eq, simp) - apply (unfold Rep_lifted_boolean_algebra_inject[THEN sym] sup_lifted_boolean_algebra.rep_eq uminus_lifted_boolean_algebra.rep_eq top_lifted_boolean_algebra.rep_eq, simp) - by (unfold Rep_lifted_boolean_algebra_inject[THEN sym] inf_lifted_boolean_algebra.rep_eq uminus_lifted_boolean_algebra.rep_eq minus_lifted_boolean_algebra.rep_eq, simp add: diff_eq) + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer using order_trans by blast + subgoal apply transfer using antisym ext by blast + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by (simp add: sup_inf_distrib1) + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by (simp add: diff_eq) + done end text \ The following two definitions give the Boolean algebra isomorphism. \ abbreviation ba_iso_inv :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_boolean_algebra \ ('a,'b) lifted_pair regular" where "ba_iso_inv \ \xf . Abs_regular (Abs_lifted_pair (\f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f))))" abbreviation ba_iso :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair regular \ ('a,'b) lifted_boolean_algebra" where "ba_iso \ \pf . Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular pf) f))" lemma ba_iso_inv_lifted_pair: "(Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f)) \ triple.pairs (Rep_phi f)" by (metis (no_types, hide_lams) double_compl simp_phi triple.pairs_uminus.simps triple_def triple.pairs_uminus_closed) lemma ba_iso_inv_regular: "regular (Abs_lifted_pair (\f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f))))" proof - have "\f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f)) = triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f)))" by (simp add: triple.pairs_uminus.simps triple_def) hence "Abs_lifted_pair (\f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f))) = --Abs_lifted_pair (\f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f)))" by (simp add: triple.pairs_uminus_closed triple_def eq_onp_def uminus_lifted_pair.abs_eq ba_iso_inv_lifted_pair) thus ?thesis by simp qed text \ The following two results prove that the isomorphisms are mutually inverse. \ lemma ba_iso_left_invertible: "ba_iso_inv (ba_iso pf) = pf" proof - have 1: "\f . snd (Rep_lifted_pair (Rep_regular pf) f) = Rep_phi f (-fst (Rep_lifted_pair (Rep_regular pf) f))" proof fix f :: "('a,'b) phi" let ?r = "Rep_phi f" have "triple ?r" by (simp add: triple_def) hence 2: "\p . triple.pairs_uminus ?r p = (-fst p,?r (fst p))" by (metis prod.collapse triple.pairs_uminus.simps) have 3: "Rep_regular pf = --Rep_regular pf" by (simp add: regular_in_p_image_iff) show "snd (Rep_lifted_pair (Rep_regular pf) f) = ?r (-fst (Rep_lifted_pair (Rep_regular pf) f))" using 2 3 by (metis fstI sndI uminus_lifted_pair.rep_eq) qed have "ba_iso_inv (ba_iso pf) = Abs_regular (Abs_lifted_pair (\f . (fst (Rep_lifted_pair (Rep_regular pf) f),Rep_phi f (-fst (Rep_lifted_pair (Rep_regular pf) f)))))" by (simp add: Abs_lifted_boolean_algebra_inverse) also have "... = Abs_regular (Abs_lifted_pair (Rep_lifted_pair (Rep_regular pf)))" using 1 by (metis prod.collapse) also have "... = pf" by (simp add: Rep_regular_inverse Rep_lifted_pair_inverse) finally show ?thesis . qed lemma ba_iso_right_invertible: "ba_iso (ba_iso_inv xf) = xf" proof - let ?rf = "Rep_lifted_boolean_algebra xf" have 1: "\f . (-?rf f,Rep_phi f (?rf f)) \ triple.pairs (Rep_phi f) \ (?rf f,Rep_phi f (-?rf f)) \ triple.pairs (Rep_phi f)" proof fix f have "up_filter top = bot" by (simp add: bot_filter.abs_eq) hence "(\z . Rep_phi f (?rf f) = Rep_phi f (?rf f) \ up_filter z) \ (\z . Rep_phi f (-?rf f) = Rep_phi f (-?rf f) \ up_filter z)" by (metis sup_bot_right) thus "(-?rf f,Rep_phi f (?rf f)) \ triple.pairs (Rep_phi f) \ (?rf f,Rep_phi f (-?rf f)) \ triple.pairs (Rep_phi f)" by (simp add: triple_def triple.pairs_def) qed have "regular (Abs_lifted_pair (\f . (?rf f,Rep_phi f (-?rf f))))" proof - have "--Abs_lifted_pair (\f . (?rf f,Rep_phi f (-?rf f))) = -Abs_lifted_pair (\f . triple.pairs_uminus (Rep_phi f) (?rf f,Rep_phi f (-?rf f)))" using 1 by (simp add: eq_onp_same_args uminus_lifted_pair.abs_eq) also have "... = -Abs_lifted_pair (\f . (-?rf f,Rep_phi f (?rf f)))" by (metis (no_types, lifting) simp_phi triple_def triple.pairs_uminus.simps) also have "... = Abs_lifted_pair (\f . triple.pairs_uminus (Rep_phi f) (-?rf f,Rep_phi f (?rf f)))" using 1 by (simp add: eq_onp_same_args uminus_lifted_pair.abs_eq) also have "... = Abs_lifted_pair (\f . (?rf f,Rep_phi f (-?rf f)))" by (metis (no_types, lifting) simp_phi triple_def triple.pairs_uminus.simps double_compl) finally show ?thesis by simp qed hence "in_p_image (Abs_lifted_pair (\f . (?rf f,Rep_phi f (-?rf f))))" by blast thus ?thesis using 1 by (simp add: Rep_lifted_boolean_algebra_inverse Abs_lifted_pair_inverse Abs_regular_inverse) qed text \ The isomorphism is established by proving the remaining Boolean algebra homomorphism properties. \ lemma ba_iso: "boolean_algebra_isomorphism ba_iso" proof (intro conjI) show "Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular bot) f)) = bot" by (simp add: bot_lifted_boolean_algebra_def bot_regular.rep_eq bot_lifted_pair.rep_eq) -next show "Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular top) f)) = top" by (simp add: top_lifted_boolean_algebra_def top_regular.rep_eq top_lifted_pair.rep_eq) -next show "\pf qf . Abs_lifted_boolean_algebra (\f::('a,'b) phi . fst (Rep_lifted_pair (Rep_regular (pf \ qf)) f)) = Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular pf) f)) \ Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular qf) f))" proof (intro allI) fix pf qf :: "('a,'b) lifted_pair regular" { fix f obtain x y z w where 1: "(x,y) = Rep_lifted_pair (Rep_regular pf) f \ (z,w) = Rep_lifted_pair (Rep_regular qf) f" using prod.collapse by blast have "triple (Rep_phi f)" by (simp add: triple_def) hence "fst (triple.pairs_sup (x,y) (z,w)) = fst (x,y) \ fst (z,w)" using triple.pairs_sup.simps by force hence "fst (triple.pairs_sup (Rep_lifted_pair (Rep_regular pf) f) (Rep_lifted_pair (Rep_regular qf) f)) = fst (Rep_lifted_pair (Rep_regular pf) f) \ fst (Rep_lifted_pair (Rep_regular qf) f)" using 1 by simp hence "fst (Rep_lifted_pair (Rep_regular (pf \ qf)) f) = fst (Rep_lifted_pair (Rep_regular pf) f) \ fst (Rep_lifted_pair (Rep_regular qf) f)" by (unfold sup_regular.rep_eq sup_lifted_pair.rep_eq) simp } thus "Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular (pf \ qf)) f)) = Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular pf) f)) \ Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular qf) f))" by (simp add: eq_onp_same_args sup_lifted_boolean_algebra.abs_eq sup_regular.rep_eq sup_lifted_boolean_algebra.rep_eq) qed -next - show "\pf qf . Abs_lifted_boolean_algebra (\f::('a,'b) phi . fst (Rep_lifted_pair (Rep_regular (pf \ qf)) f)) = Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular pf) f)) \ Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular qf) f))" + show 1: "\pf qf . Abs_lifted_boolean_algebra (\f::('a,'b) phi . fst (Rep_lifted_pair (Rep_regular (pf \ qf)) f)) = Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular pf) f)) \ Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular qf) f))" proof (intro allI) fix pf qf :: "('a,'b) lifted_pair regular" { fix f obtain x y z w where 1: "(x,y) = Rep_lifted_pair (Rep_regular pf) f \ (z,w) = Rep_lifted_pair (Rep_regular qf) f" using prod.collapse by blast have "triple (Rep_phi f)" by (simp add: triple_def) hence "fst (triple.pairs_inf (x,y) (z,w)) = fst (x,y) \ fst (z,w)" using triple.pairs_inf.simps by force hence "fst (triple.pairs_inf (Rep_lifted_pair (Rep_regular pf) f) (Rep_lifted_pair (Rep_regular qf) f)) = fst (Rep_lifted_pair (Rep_regular pf) f) \ fst (Rep_lifted_pair (Rep_regular qf) f)" using 1 by simp hence "fst (Rep_lifted_pair (Rep_regular (pf \ qf)) f) = fst (Rep_lifted_pair (Rep_regular pf) f) \ fst (Rep_lifted_pair (Rep_regular qf) f)" by (unfold inf_regular.rep_eq inf_lifted_pair.rep_eq) simp } thus "Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular (pf \ qf)) f)) = Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular pf) f)) \ Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular qf) f))" by (simp add: eq_onp_same_args inf_lifted_boolean_algebra.abs_eq inf_regular.rep_eq inf_lifted_boolean_algebra.rep_eq) qed -next show "\pf . Abs_lifted_boolean_algebra (\f::('a,'b) phi . fst (Rep_lifted_pair (Rep_regular (-pf)) f)) = -Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular pf) f))" proof fix pf :: "('a,'b) lifted_pair regular" { fix f obtain x y where 1: "(x,y) = Rep_lifted_pair (Rep_regular pf) f" using prod.collapse by blast have "triple (Rep_phi f)" by (simp add: triple_def) hence "fst (triple.pairs_uminus (Rep_phi f) (x,y)) = -fst (x,y)" using triple.pairs_uminus.simps by force hence "fst (triple.pairs_uminus (Rep_phi f) (Rep_lifted_pair (Rep_regular pf) f)) = -fst (Rep_lifted_pair (Rep_regular pf) f)" using 1 by simp hence "fst (Rep_lifted_pair (Rep_regular (-pf)) f) = -fst (Rep_lifted_pair (Rep_regular pf) f)" by (unfold uminus_regular.rep_eq uminus_lifted_pair.rep_eq) simp } thus "Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular (-pf)) f)) = -Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular pf) f))" by (simp add: eq_onp_same_args uminus_lifted_boolean_algebra.abs_eq uminus_regular.rep_eq uminus_lifted_boolean_algebra.rep_eq) qed -next + thus "\pf qf . Abs_lifted_boolean_algebra (\f::('a,'b) phi . fst (Rep_lifted_pair (Rep_regular (pf - qf)) f)) = Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular pf) f)) - Abs_lifted_boolean_algebra (\f . fst (Rep_lifted_pair (Rep_regular qf) f))" + using 1 by (simp add: diff_eq) show "bij ba_iso" by (rule invertible_bij[where g=ba_iso_inv]) (simp_all add: ba_iso_left_invertible ba_iso_right_invertible) qed subsubsection \Distributive Lattice Isomorphism\ text \ We carry out a similar development for the isomorphism of distributive lattices. Again, the original distributive lattice with a greatest element needs to be lifted to match the lifted pairs. \ typedef (overloaded) ('a,'b) lifted_distrib_lattice_top = "{ xf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) phi \ 'b . True }" by simp setup_lifting type_definition_lifted_distrib_lattice_top instantiation lifted_distrib_lattice_top :: (non_trivial_boolean_algebra,distrib_lattice_top) distrib_lattice_top begin lift_definition sup_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top \ ('a,'b) lifted_distrib_lattice_top \ ('a,'b) lifted_distrib_lattice_top" is "\xf yf f . sup (xf f) (yf f)" . lift_definition inf_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top \ ('a,'b) lifted_distrib_lattice_top \ ('a,'b) lifted_distrib_lattice_top" is "\xf yf f . inf (xf f) (yf f)" . lift_definition top_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top" is "\f . top" .. lift_definition less_eq_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top \ ('a,'b) lifted_distrib_lattice_top \ bool" is "\xf yf . \f . less_eq (xf f) (yf f)" . lift_definition less_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top \ ('a,'b) lifted_distrib_lattice_top \ bool" is "\xf yf . (\f . less_eq (xf f) (yf f)) \ \ (\f . less_eq (yf f) (xf f))" . instance apply intro_classes - apply (simp add: less_eq_lifted_distrib_lattice_top.rep_eq less_lifted_distrib_lattice_top.rep_eq) - apply (simp add: less_eq_lifted_distrib_lattice_top.rep_eq) - using less_eq_lifted_distrib_lattice_top.rep_eq order_trans apply fastforce - apply (metis less_eq_lifted_distrib_lattice_top.rep_eq antisym ext Rep_lifted_distrib_lattice_top_inject) - apply (simp add: inf_lifted_distrib_lattice_top.rep_eq less_eq_lifted_distrib_lattice_top.rep_eq) - apply (simp add: inf_lifted_distrib_lattice_top.rep_eq less_eq_lifted_distrib_lattice_top.rep_eq) - apply (simp add: inf_lifted_distrib_lattice_top.rep_eq less_eq_lifted_distrib_lattice_top.rep_eq) - apply (simp add: sup_lifted_distrib_lattice_top.rep_eq less_eq_lifted_distrib_lattice_top.rep_eq) - apply (simp add: less_eq_lifted_distrib_lattice_top.rep_eq sup_lifted_distrib_lattice_top.rep_eq) - apply (simp add: less_eq_lifted_distrib_lattice_top.rep_eq sup_lifted_distrib_lattice_top.rep_eq) - apply (simp add: less_eq_lifted_distrib_lattice_top.rep_eq top_lifted_distrib_lattice_top.rep_eq) - by (unfold Rep_lifted_distrib_lattice_top_inject[THEN sym] sup_lifted_distrib_lattice_top.rep_eq inf_lifted_distrib_lattice_top.rep_eq, simp add: sup_inf_distrib1) + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer using order_trans by blast + subgoal apply transfer using antisym ext by blast + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by auto + subgoal apply transfer by (simp add: sup_inf_distrib1) + done end text \ The following function extracts the least element of the filter of a dense pair, which turns out to be a principal filter. It is used to define one of the isomorphisms below. \ fun get_dense :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense \ ('a,'b) phi \ 'b" where "get_dense pf f = (SOME z . Rep_lifted_pair (Rep_dense pf) f = (top,up_filter z))" lemma get_dense_char: "Rep_lifted_pair (Rep_dense pf) f = (top,up_filter (get_dense pf f))" proof - obtain x y where 1: "(x,y) = Rep_lifted_pair (Rep_dense pf) f \ (x,y) \ triple.pairs (Rep_phi f) \ triple.pairs_uminus (Rep_phi f) (x,y) = triple.pairs_bot" by (metis bot_lifted_pair.rep_eq prod.collapse simp_dense simp_lifted_pair uminus_lifted_pair.rep_eq) hence 2: "x = top" by (simp add: triple.intro triple.pairs_uminus.simps dense_pp) have "triple (Rep_phi f)" by (simp add: triple_def) hence "\z. y = Rep_phi f (-x) \ up_filter z" using 1 triple.pairs_def by blast then obtain z where "y = up_filter z" using 2 by auto hence "Rep_lifted_pair (Rep_dense pf) f = (top,up_filter z)" using 1 2 by simp thus ?thesis by (metis (mono_tags, lifting) tfl_some get_dense.simps) qed text \ The following two definitions give the distributive lattice isomorphism. \ abbreviation dl_iso_inv :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_distrib_lattice_top \ ('a,'b) lifted_pair dense" where "dl_iso_inv \ \xf . Abs_dense (Abs_lifted_pair (\f . (top,up_filter (Rep_lifted_distrib_lattice_top xf f))))" abbreviation dl_iso :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense \ ('a,'b) lifted_distrib_lattice_top" where "dl_iso \ \pf . Abs_lifted_distrib_lattice_top (get_dense pf)" lemma dl_iso_inv_lifted_pair: "(top,up_filter (Rep_lifted_distrib_lattice_top xf f)) \ triple.pairs (Rep_phi f)" by (metis (no_types, hide_lams) compl_bot_eq double_compl simp_phi sup_bot.left_neutral triple.sa_iso_pair triple_def) lemma dl_iso_inv_dense: "dense (Abs_lifted_pair (\f . (top,up_filter (Rep_lifted_distrib_lattice_top xf f))))" proof - have "\f . triple.pairs_uminus (Rep_phi f) (top,up_filter (Rep_lifted_distrib_lattice_top xf f)) = triple.pairs_bot" by (simp add: top_filter.abs_eq triple.pairs_uminus.simps triple_def) hence "bot = -Abs_lifted_pair (\f . (top,up_filter (Rep_lifted_distrib_lattice_top xf f)))" by (simp add: eq_onp_def uminus_lifted_pair.abs_eq dl_iso_inv_lifted_pair bot_lifted_pair_def) thus ?thesis by simp qed text \ The following two results prove that the isomorphisms are mutually inverse. \ lemma dl_iso_left_invertible: "dl_iso_inv (dl_iso pf) = pf" proof - have "dl_iso_inv (dl_iso pf) = Abs_dense (Abs_lifted_pair (\f . (top,up_filter (get_dense pf f))))" by (metis Abs_lifted_distrib_lattice_top_inverse UNIV_I UNIV_def) also have "... = Abs_dense (Abs_lifted_pair (Rep_lifted_pair (Rep_dense pf)))" by (metis get_dense_char) also have "... = pf" by (simp add: Rep_dense_inverse Rep_lifted_pair_inverse) finally show ?thesis . qed lemma dl_iso_right_invertible: "dl_iso (dl_iso_inv xf) = xf" proof - let ?rf = "Rep_lifted_distrib_lattice_top xf" let ?pf = "Abs_dense (Abs_lifted_pair (\f . (top,up_filter (?rf f))))" have 1: "\f . (top,up_filter (?rf f)) \ triple.pairs (Rep_phi f)" proof fix f :: "('a,'b) phi" have "triple (Rep_phi f)" by (simp add: triple_def) thus "(top,up_filter (?rf f)) \ triple.pairs (Rep_phi f)" using triple.pairs_def by force qed have 2: "dense (Abs_lifted_pair (\f . (top,up_filter (?rf f))))" proof - have "-Abs_lifted_pair (\f . (top,up_filter (?rf f))) = Abs_lifted_pair (\f . triple.pairs_uminus (Rep_phi f) (top,up_filter (?rf f)))" using 1 by (simp add: eq_onp_same_args uminus_lifted_pair.abs_eq) also have "... = Abs_lifted_pair (\f . (bot,Rep_phi f top))" by (simp add: triple.pairs_uminus.simps triple_def) also have "... = Abs_lifted_pair (\f . triple.pairs_bot)" by (metis (no_types, hide_lams) simp_phi triple.phi_top triple_def) also have "... = bot" by (simp add: bot_lifted_pair_def) finally show ?thesis by simp qed have "get_dense ?pf = ?rf" proof fix f have "(top,up_filter (get_dense ?pf f)) = Rep_lifted_pair (Rep_dense ?pf) f" by (metis get_dense_char) also have "... = Rep_lifted_pair (Abs_lifted_pair (\f . (top,up_filter (?rf f)))) f" using Abs_dense_inverse 2 by force also have "... = (top,up_filter (?rf f))" using 1 by (simp add: Abs_lifted_pair_inverse) finally show "get_dense ?pf f = ?rf f" using up_filter_injective by auto qed thus ?thesis by (simp add: Rep_lifted_distrib_lattice_top_inverse) qed text \ To obtain the isomorphism, it remains to show the homomorphism properties of lattices with a greatest element. \ lemma dl_iso: "bounded_lattice_top_isomorphism dl_iso" proof (intro conjI) have "get_dense top = (\f::('a,'b) phi . top)" proof fix f :: "('a,'b) phi" have "Rep_lifted_pair (Rep_dense top) f = (top,Abs_filter {top})" by (simp add: top_dense.rep_eq top_lifted_pair.rep_eq) hence "up_filter (get_dense top f) = Abs_filter {top}" by (metis prod.inject get_dense_char) hence "Rep_filter (up_filter (get_dense top f)) = {top}" by (metis bot_filter.abs_eq bot_filter.rep_eq) thus "get_dense top f = top" by (metis self_in_upset singletonD Abs_filter_inverse mem_Collect_eq up_filter) qed thus "Abs_lifted_distrib_lattice_top (get_dense top::('a,'b) phi \ 'b) = top" by (metis top_lifted_distrib_lattice_top_def) next show "\pf qf :: ('a,'b) lifted_pair dense . Abs_lifted_distrib_lattice_top (get_dense (pf \ qf)) = Abs_lifted_distrib_lattice_top (get_dense pf) \ Abs_lifted_distrib_lattice_top (get_dense qf)" proof (intro allI) fix pf qf :: "('a,'b) lifted_pair dense" have 1: "Abs_lifted_distrib_lattice_top (get_dense pf) \ Abs_lifted_distrib_lattice_top (get_dense qf) = Abs_lifted_distrib_lattice_top (\f . get_dense pf f \ get_dense qf f)" by (simp add: eq_onp_same_args sup_lifted_distrib_lattice_top.abs_eq) have "(\f . get_dense (pf \ qf) f) = (\f . get_dense pf f \ get_dense qf f)" proof fix f have "(top,up_filter (get_dense (pf \ qf) f)) = Rep_lifted_pair (Rep_dense (pf \ qf)) f" by (metis get_dense_char) also have "... = triple.pairs_sup (Rep_lifted_pair (Rep_dense pf) f) (Rep_lifted_pair (Rep_dense qf) f)" by (simp add: sup_lifted_pair.rep_eq sup_dense.rep_eq) also have "... = triple.pairs_sup (top,up_filter (get_dense pf f)) (top,up_filter (get_dense qf f))" by (metis get_dense_char) also have "... = (top,up_filter (get_dense pf f) \ up_filter (get_dense qf f))" by (metis (no_types, lifting) calculation prod.simps(1) simp_phi triple.pairs_sup.simps triple_def) also have "... = (top,up_filter (get_dense pf f \ get_dense qf f))" by (metis up_filter_dist_sup) finally show "get_dense (pf \ qf) f = get_dense pf f \ get_dense qf f" using up_filter_injective by blast qed thus "Abs_lifted_distrib_lattice_top (get_dense (pf \ qf)) = Abs_lifted_distrib_lattice_top (get_dense pf) \ Abs_lifted_distrib_lattice_top (get_dense qf)" using 1 by metis qed next show "\pf qf :: ('a,'b) lifted_pair dense . Abs_lifted_distrib_lattice_top (get_dense (pf \ qf)) = Abs_lifted_distrib_lattice_top (get_dense pf) \ Abs_lifted_distrib_lattice_top (get_dense qf)" proof (intro allI) fix pf qf :: "('a,'b) lifted_pair dense" have 1: "Abs_lifted_distrib_lattice_top (get_dense pf) \ Abs_lifted_distrib_lattice_top (get_dense qf) = Abs_lifted_distrib_lattice_top (\f . get_dense pf f \ get_dense qf f)" by (simp add: eq_onp_same_args inf_lifted_distrib_lattice_top.abs_eq) have "(\f . get_dense (pf \ qf) f) = (\f . get_dense pf f \ get_dense qf f)" proof fix f have "(top,up_filter (get_dense (pf \ qf) f)) = Rep_lifted_pair (Rep_dense (pf \ qf)) f" by (metis get_dense_char) also have "... = triple.pairs_inf (Rep_lifted_pair (Rep_dense pf) f) (Rep_lifted_pair (Rep_dense qf) f)" by (simp add: inf_lifted_pair.rep_eq inf_dense.rep_eq) also have "... = triple.pairs_inf (top,up_filter (get_dense pf f)) (top,up_filter (get_dense qf f))" by (metis get_dense_char) also have "... = (top,up_filter (get_dense pf f) \ up_filter (get_dense qf f))" by (metis (no_types, lifting) calculation prod.simps(1) simp_phi triple.pairs_inf.simps triple_def) also have "... = (top,up_filter (get_dense pf f \ get_dense qf f))" by (metis up_filter_dist_inf) finally show "get_dense (pf \ qf) f = get_dense pf f \ get_dense qf f" using up_filter_injective by blast qed thus "Abs_lifted_distrib_lattice_top (get_dense (pf \ qf)) = Abs_lifted_distrib_lattice_top (get_dense pf) \ Abs_lifted_distrib_lattice_top (get_dense qf)" using 1 by metis qed next show "bij dl_iso" by (rule invertible_bij[where g=dl_iso_inv]) (simp_all add: dl_iso_left_invertible dl_iso_right_invertible) qed subsubsection \Structure Map Preservation\ text \ We finally show that the isomorphisms are compatible with the structure maps. This involves lifting the distributive lattice isomorphism to filters of distributive lattices (as these are the targets of the structure maps). To this end, we first show that the lifted isomorphism preserves filters. \ lemma phi_iso_filter: "filter ((\qf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense . Rep_lifted_distrib_lattice_top (dl_iso qf) f) ` Rep_filter (stone_phi pf))" proof (rule filter_map_filter) show "mono (\qf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense . Rep_lifted_distrib_lattice_top (dl_iso qf) f)" by (metis (no_types, lifting) mono_def dl_iso le_iff_sup sup_lifted_distrib_lattice_top.rep_eq) next show "\qf y . Rep_lifted_distrib_lattice_top (dl_iso qf) f \ y \ (\rf . qf \ rf \ y = Rep_lifted_distrib_lattice_top (dl_iso rf) f)" proof (intro allI, rule impI) fix qf :: "('a,'b) lifted_pair dense" fix y :: 'b assume 1: "Rep_lifted_distrib_lattice_top (dl_iso qf) f \ y" let ?rf = "Abs_dense (Abs_lifted_pair (\g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g))" have 2: "\g . (if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g) \ triple.pairs (Rep_phi g)" by (metis Abs_lifted_distrib_lattice_top_inverse dl_iso_inv_lifted_pair mem_Collect_eq simp_lifted_pair) hence "-Abs_lifted_pair (\g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g) = Abs_lifted_pair (\g . triple.pairs_uminus (Rep_phi g) (if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g))" by (simp add: eq_onp_def uminus_lifted_pair.abs_eq) also have "... = Abs_lifted_pair (\g . if g = f then triple.pairs_uminus (Rep_phi g) (top,up_filter y) else triple.pairs_uminus (Rep_phi g) (Rep_lifted_pair (Rep_dense qf) g))" by (simp add: if_distrib) also have "... = Abs_lifted_pair (\g . if g = f then (bot,top) else triple.pairs_uminus (Rep_phi g) (Rep_lifted_pair (Rep_dense qf) g))" by (subst triple.pairs_uminus.simps, simp add: triple_def, metis compl_top_eq simp_phi) also have "... = Abs_lifted_pair (\g . if g = f then (bot,top) else (bot,top))" by (metis bot_lifted_pair.rep_eq simp_dense top_filter.abs_eq uminus_lifted_pair.rep_eq) also have "... = bot" by (simp add: bot_lifted_pair.abs_eq top_filter.abs_eq) finally have 3: "Abs_lifted_pair (\g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g) \ dense_elements" by blast hence "(top,up_filter (get_dense (Abs_dense (Abs_lifted_pair (\g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g))) f)) = Rep_lifted_pair (Rep_dense (Abs_dense (Abs_lifted_pair (\g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g)))) f" by (metis (mono_tags, lifting) get_dense_char) also have "... = Rep_lifted_pair (Abs_lifted_pair (\g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g)) f" using 3 by (simp add: Abs_dense_inverse) also have "... = (top,up_filter y)" using 2 by (simp add: Abs_lifted_pair_inverse) finally have "get_dense (Abs_dense (Abs_lifted_pair (\g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g))) f = y" using up_filter_injective by blast hence 4: "Rep_lifted_distrib_lattice_top (dl_iso ?rf) f = y" by (simp add: Abs_lifted_distrib_lattice_top_inverse) { fix g have "Rep_lifted_distrib_lattice_top (dl_iso qf) g \ Rep_lifted_distrib_lattice_top (dl_iso ?rf) g" proof (cases "g = f") assume "g = f" thus ?thesis using 1 4 by simp next assume 5: "g \ f" have "(top,up_filter (get_dense ?rf g)) = Rep_lifted_pair (Rep_dense (Abs_dense (Abs_lifted_pair (\g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g)))) g" by (metis (mono_tags, lifting) get_dense_char) also have "... = Rep_lifted_pair (Abs_lifted_pair (\g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g)) g" using 3 by (simp add: Abs_dense_inverse) also have "... = Rep_lifted_pair (Rep_dense qf) g" using 2 5 by (simp add: Abs_lifted_pair_inverse) also have "... = (top,up_filter (get_dense qf g))" using get_dense_char by auto finally have "get_dense ?rf g = get_dense qf g" using up_filter_injective by blast thus "Rep_lifted_distrib_lattice_top (dl_iso qf) g \ Rep_lifted_distrib_lattice_top (dl_iso ?rf) g" by (simp add: Abs_lifted_distrib_lattice_top_inverse) qed } hence "Rep_lifted_distrib_lattice_top (dl_iso qf) \ Rep_lifted_distrib_lattice_top (dl_iso ?rf)" by (simp add: le_funI) hence 6: "dl_iso qf \ dl_iso ?rf" by (simp add: le_funD less_eq_lifted_distrib_lattice_top.rep_eq) hence "qf \ ?rf" by (metis (no_types, lifting) dl_iso sup_isomorphism_ord_isomorphism) thus "\rf . qf \ rf \ y = Rep_lifted_distrib_lattice_top (dl_iso rf) f" using 4 by auto qed qed text \ The commutativity property states that the same result is obtained in two ways by starting with a regular lifted pair \pf\: \begin{itemize} \item apply the Boolean algebra isomorphism to the pair; then apply a structure map \f\ to obtain a filter of dense elements; or, \item apply the structure map \stone_phi\ to the pair; then apply the distributive lattice isomorphism lifted to the resulting filter. \end{itemize} \ lemma phi_iso: "Rep_phi f (Rep_lifted_boolean_algebra (ba_iso pf) f) = filter_map (\qf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense . Rep_lifted_distrib_lattice_top (dl_iso qf) f) (stone_phi pf)" proof - let ?r = "Rep_phi f" let ?ppf = "\g . triple.pairs_uminus (Rep_phi g) (Rep_lifted_pair (Rep_regular pf) g)" have 1: "triple ?r" by (simp add: triple_def) have 2: "Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f))) \ { z . \qf . -Rep_regular pf \ Rep_dense qf \ z = get_dense qf f }" proof fix z obtain x where 3: "x = fst (Rep_lifted_pair (Rep_regular pf) f)" by simp assume "z \ Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f)))" hence "\z \ Rep_filter (?r x)" using 3 filter_def by fastforce hence 4: "up_filter z \ ?r x" by (metis Rep_filter_cases Rep_filter_inverse less_eq_filter.rep_eq mem_Collect_eq up_filter) have 5: "\g . ?ppf g \ triple.pairs (Rep_phi g)" by (metis (no_types) simp_lifted_pair uminus_lifted_pair.rep_eq) let ?zf = "\g . if g = f then (top,up_filter z) else triple.pairs_top" have 6: "\g . ?zf g \ triple.pairs (Rep_phi g)" proof fix g :: "('a,'b) phi" have "triple (Rep_phi g)" by (simp add: triple_def) hence "(top,up_filter z) \ triple.pairs (Rep_phi g)" using triple.pairs_def by force thus "?zf g \ triple.pairs (Rep_phi g)" by (metis simp_lifted_pair top_lifted_pair.rep_eq) qed hence "-Abs_lifted_pair ?zf = Abs_lifted_pair (\g . triple.pairs_uminus (Rep_phi g) (?zf g))" by (subst uminus_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args) also have "... = Abs_lifted_pair (\g . if g = f then triple.pairs_uminus (Rep_phi g) (top,up_filter z) else triple.pairs_uminus (Rep_phi g) triple.pairs_top)" by (rule arg_cong[where f=Abs_lifted_pair]) auto also have "... = Abs_lifted_pair (\g . triple.pairs_bot)" using 1 by (metis bot_lifted_pair.rep_eq dense_closed_top top_lifted_pair.rep_eq triple.pairs_uminus.simps uminus_lifted_pair.rep_eq) finally have 7: "Abs_lifted_pair ?zf \ dense_elements" by (simp add: bot_lifted_pair.abs_eq) let ?qf = "Abs_dense (Abs_lifted_pair ?zf)" have "\g . triple.pairs_less_eq (?ppf g) (?zf g)" proof fix g show "triple.pairs_less_eq (?ppf g) (?zf g)" proof (cases "g = f") assume 8: "g = f" hence 9: "?ppf g = (-x,?r x)" using 1 3 by (metis prod.collapse triple.pairs_uminus.simps) have "triple.pairs_less_eq (-x,?r x) (top,up_filter z)" using 1 4 by (meson inf.bot_least triple.pairs_less_eq.simps) thus ?thesis using 8 9 by simp next assume 10: "g \ f" have "triple.pairs_less_eq (?ppf g) triple.pairs_top" using 1 by (metis (no_types, hide_lams) bot.extremum top_greatest prod.collapse triple_def triple.pairs_less_eq.simps triple.phi_bot) thus ?thesis using 10 by simp qed qed hence "Abs_lifted_pair ?ppf \ Abs_lifted_pair ?zf" using 5 6 by (subst less_eq_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args) hence 11: "-Rep_regular pf \ Rep_dense ?qf" using 7 by (simp add: uminus_lifted_pair_def Abs_dense_inverse) have "(top,up_filter (get_dense ?qf f)) = Rep_lifted_pair (Rep_dense ?qf) f" by (metis get_dense_char) also have "... = (top,up_filter z)" using 6 7 Abs_dense_inverse Abs_lifted_pair_inverse by force finally have "z = get_dense ?qf f" using up_filter_injective by force thus "z \ { z . \qf . -Rep_regular pf \ Rep_dense qf \ z = get_dense qf f }" using 11 by auto qed have 12: "Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f))) \ { z . \qf . -Rep_regular pf \ Rep_dense qf \ z = get_dense qf f }" proof fix z assume "z \ { z . \qf . -Rep_regular pf \ Rep_dense qf \ z = get_dense qf f }" hence "\qf . -Rep_regular pf \ Rep_dense qf \ z = get_dense qf f" by auto hence "triple.pairs_less_eq (Rep_lifted_pair (-Rep_regular pf) f) (top,up_filter z)" by (metis less_eq_lifted_pair.rep_eq get_dense_char) hence "up_filter z \ snd (Rep_lifted_pair (-Rep_regular pf) f)" using 1 by (metis (no_types, hide_lams) prod.collapse triple.pairs_less_eq.simps) also have "... = snd (?ppf f)" by (metis uminus_lifted_pair.rep_eq) also have "... = ?r (fst (Rep_lifted_pair (Rep_regular pf) f))" using 1 by (metis (no_types) prod.collapse prod.inject triple.pairs_uminus.simps) finally have "Rep_filter (up_filter z) \ Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f)))" by (simp add: less_eq_filter.rep_eq) hence "\z \ Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f)))" by (metis Abs_filter_inverse mem_Collect_eq up_filter) thus "z \ Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f)))" by blast qed have 13: "\qf\Rep_filter (stone_phi pf) . Rep_lifted_distrib_lattice_top (Abs_lifted_distrib_lattice_top (get_dense qf)) f = get_dense qf f" by (metis Abs_lifted_distrib_lattice_top_inverse UNIV_I UNIV_def) - have "Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f))) = { z . \qf\stone_phi_set pf . z = get_dense qf f }" + have "Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f))) = { z . \qf\stone_phi_base pf . z = get_dense qf f }" using 2 12 by simp - hence "?r (fst (Rep_lifted_pair (Rep_regular pf) f)) = Abs_filter { z . \qf\stone_phi_set pf . z = get_dense qf f }" + hence "?r (fst (Rep_lifted_pair (Rep_regular pf) f)) = Abs_filter { z . \qf\stone_phi_base pf . z = get_dense qf f }" by (metis Rep_filter_inverse) hence "?r (Rep_lifted_boolean_algebra (ba_iso pf) f) = Abs_filter { z . \qf\Rep_filter (stone_phi pf) . z = Rep_lifted_distrib_lattice_top (dl_iso qf) f }" - using 13 by (simp add: Abs_filter_inverse stone_phi_set_filter stone_phi_def Abs_lifted_boolean_algebra_inverse) + using 13 by (simp add: Abs_filter_inverse stone_phi_base_filter stone_phi_def Abs_lifted_boolean_algebra_inverse) thus ?thesis by (simp add: image_def) qed end diff --git a/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy b/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy --- a/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy +++ b/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy @@ -1,3173 +1,3173 @@ (* Title: Kleene Relation Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Kleene Relation Algebras\ text \ This theory combines Kleene algebras with Stone relation algebras. Relation algebras with transitive closure have been studied by \cite{Ng1984}. The weakening to Stone relation algebras allows us to talk about reachability in weighted graphs, for example. Many results in this theory are used in the correctness proof of Prim's minimum spanning tree algorithm. In particular, they are concerned with the exchange property, preservation of parts of the invariant and with establishing parts of the postcondition. \ theory Kleene_Relation_Algebras imports Stone_Relation_Algebras.Relation_Algebras Kleene_Algebras begin text \ We first note that bounded distributive lattices can be expanded to Kleene algebras by reusing some of the operations. \ sublocale bounded_distrib_lattice < comp_inf: bounded_kleene_algebra where star = "\x . top" and one = top and times = inf apply unfold_locales apply (simp add: inf.assoc) apply simp apply simp apply (simp add: le_infI2) apply (simp add: inf_sup_distrib2) apply simp apply simp apply simp apply simp apply simp apply (simp add: inf_sup_distrib1) apply simp apply simp by (simp add: inf_assoc) text \ We add the Kleene star operation to each of bounded distributive allegories, pseudocomplemented distributive allegories and Stone relation algebras. We start with single-object bounded distributive allegories. \ class bounded_distrib_kleene_allegory = bounded_distrib_allegory + kleene_algebra begin subclass bounded_kleene_algebra .. lemma conv_star_conv: "x\<^sup>\ \ x\<^sup>T\<^sup>\\<^sup>T" proof - have "x\<^sup>T\<^sup>\ * x\<^sup>T \ x\<^sup>T\<^sup>\" by (simp add: star.right_plus_below_circ) hence 1: "x * x\<^sup>T\<^sup>\\<^sup>T \ x\<^sup>T\<^sup>\\<^sup>T" using conv_dist_comp conv_isotone by fastforce have "1 \ x\<^sup>T\<^sup>\\<^sup>T" by (simp add: reflexive_conv_closed star.circ_reflexive) hence "1 \ x * x\<^sup>T\<^sup>\\<^sup>T \ x\<^sup>T\<^sup>\\<^sup>T" using 1 by simp thus ?thesis using star_left_induct by fastforce qed text \ It follows that star and converse commute. \ lemma conv_star_commute: "x\<^sup>\\<^sup>T = x\<^sup>T\<^sup>\" proof (rule antisym) show "x\<^sup>\\<^sup>T \ x\<^sup>T\<^sup>\" using conv_star_conv conv_isotone by fastforce next show "x\<^sup>T\<^sup>\ \ x\<^sup>\\<^sup>T" by (metis conv_star_conv conv_involutive) qed lemma conv_plus_commute: "x\<^sup>+\<^sup>T = x\<^sup>T\<^sup>+" by (simp add: conv_dist_comp conv_star_commute star_plus) text \ The following results are variants of a separation lemma of Kleene algebras. \ lemma cancel_separate_2: assumes "x * y \ 1" shows "((w \ x) \ (z \ y))\<^sup>\ = (z \ y)\<^sup>\ * (w \ x)\<^sup>\" proof - have "(w \ x) * (z \ y) \ 1" by (meson assms comp_isotone order.trans inf.cobounded2) thus ?thesis using cancel_separate_1 sup_commute by simp qed lemma cancel_separate_3: assumes "x * y \ 1" shows "(w \ x)\<^sup>\ * (z \ y)\<^sup>\ = (w \ x)\<^sup>\ \ (z \ y)\<^sup>\" proof - have "(w \ x) * (z \ y) \ 1" by (meson assms comp_isotone order.trans inf.cobounded2) thus ?thesis by (simp add: cancel_separate_eq) qed lemma cancel_separate_4: assumes "z * y \ 1" and "w \ y \ z" and "x \ y \ z" shows "w\<^sup>\ * x\<^sup>\ = (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\) * (x \ z)\<^sup>\" proof - have "w\<^sup>\ * x\<^sup>\ = ((w \ y) \ (w \ z))\<^sup>\ * ((x \ y) \ (x \ z))\<^sup>\" by (metis assms(2,3) inf.orderE inf_sup_distrib1) also have "... = (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ * (x \ y)\<^sup>\) * (x \ z)\<^sup>\" by (metis assms(1) cancel_separate_2 sup_commute mult_assoc) finally show ?thesis by (simp add: assms(1) cancel_separate_3) qed lemma cancel_separate_5: assumes "w * z\<^sup>T \ 1" shows "w \ x * (y \ z) \ y" proof - have "w \ x * (y \ z) \ (x \ w * (y \ z)\<^sup>T) * (y \ z)" by (metis dedekind_2 inf_commute) also have "... \ w * z\<^sup>T * (y \ z)" by (simp add: conv_dist_inf inf.coboundedI2 mult_left_isotone mult_right_isotone) also have "... \ y \ z" by (metis assms mult_1_left mult_left_isotone) finally show ?thesis by simp qed lemma cancel_separate_6: assumes "z * y \ 1" and "w \ y \ z" and "x \ y \ z" and "v * z\<^sup>T \ 1" and "v \ y\<^sup>\ = bot" shows "v \ w\<^sup>\ * x\<^sup>\ \ x \ w" proof - have "v \ (w \ y)\<^sup>\ * (x \ y)\<^sup>\ \ v \ y\<^sup>\ * (x \ y)\<^sup>\" using comp_inf.mult_right_isotone mult_left_isotone star_isotone by simp also have "... \ v \ y\<^sup>\" by (simp add: inf.coboundedI2 star.circ_increasing star.circ_mult_upper_bound star_right_induct_mult) finally have 1: "v \ (w \ y)\<^sup>\ * (x \ y)\<^sup>\ = bot" using assms(5) le_bot by simp have "v \ w\<^sup>\ * x\<^sup>\ = v \ (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\) * (x \ z)\<^sup>\" using assms(1-3) cancel_separate_4 by simp also have "... = (v \ (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\) * (x \ z)\<^sup>\ * (x \ z)) \ (v \ (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\))" by (metis inf_sup_distrib1 star.circ_back_loop_fixpoint) also have "... \ x \ (v \ (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\))" using assms(4) cancel_separate_5 semiring.add_right_mono by simp also have "... = x \ (v \ (w \ y)\<^sup>\ * (w \ z)\<^sup>\)" using 1 by (simp add: inf_sup_distrib1 mult_left_dist_sup sup_monoid.add_assoc) also have "... = x \ (v \ (w \ y)\<^sup>\ * (w \ z)\<^sup>\ * (w \ z)) \ (v \ (w \ y)\<^sup>\)" by (metis comp_inf.semiring.distrib_left star.circ_back_loop_fixpoint sup_assoc) also have "... \ x \ w \ (v \ (w \ y)\<^sup>\)" using assms(4) cancel_separate_5 sup_left_isotone sup_right_isotone by simp also have "... \ x \ w \ (v \ y\<^sup>\)" using comp_inf.mult_right_isotone star_isotone sup_right_isotone by simp finally show ?thesis using assms(5) le_bot by simp qed text \ We show several results about the interaction of vectors and the Kleene star. \ lemma vector_star_1: assumes "vector x" shows "x\<^sup>T * (x * x\<^sup>T)\<^sup>\ \ x\<^sup>T" proof - have "x\<^sup>T * (x * x\<^sup>T)\<^sup>\ = (x\<^sup>T * x)\<^sup>\ * x\<^sup>T" by (simp add: star_slide) also have "... \ top * x\<^sup>T" by (simp add: mult_left_isotone) also have "... = x\<^sup>T" using assms vector_conv_covector by auto finally show ?thesis . qed lemma vector_star_2: "vector x \ x\<^sup>T * (x * x\<^sup>T)\<^sup>\ \ x\<^sup>T * bot\<^sup>\" by (simp add: star_absorb vector_star_1) lemma vector_vector_star: "vector v \ (v * v\<^sup>T)\<^sup>\ = 1 \ v * v\<^sup>T" by (simp add: transitive_star vv_transitive) text \ The following equivalence relation characterises the component trees of a forest. This is a special case of undirected reachability in a directed graph. \ abbreviation "forest_components f \ f\<^sup>T\<^sup>\ * f\<^sup>\" lemma forest_components_equivalence: "injective x \ equivalence (forest_components x)" apply (intro conjI) apply (simp add: reflexive_mult_closed star.circ_reflexive) apply (metis cancel_separate_1 eq_iff star.circ_transitive_equal) by (simp add: conv_dist_comp conv_star_commute) lemma forest_components_increasing: "x \ forest_components x" by (metis order.trans mult_left_isotone mult_left_one star.circ_increasing star.circ_reflexive) lemma forest_components_isotone: "x \ y \ forest_components x \ forest_components y" by (simp add: comp_isotone conv_isotone star_isotone) lemma forest_components_idempotent: "injective x \ forest_components (forest_components x) = forest_components x" by (metis forest_components_equivalence cancel_separate_1 star.circ_transitive_equal star_involutive) lemma forest_components_star: "injective x \ (forest_components x)\<^sup>\ = forest_components x" using forest_components_equivalence forest_components_idempotent star.circ_transitive_equal by simp text \ The following lemma shows that the nodes reachable in the graph can be reached by only using edges between reachable nodes. \ lemma reachable_restrict: assumes "vector r" shows "r\<^sup>T * g\<^sup>\ = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\" proof - have 1: "r\<^sup>T \ r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\" using mult_right_isotone mult_1_right star.circ_reflexive by fastforce have 2: "covector (r\<^sup>T * g\<^sup>\)" using assms covector_mult_closed vector_conv_covector by auto have "r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * g \ r\<^sup>T * g\<^sup>\ * g" by (simp add: mult_left_isotone mult_right_isotone star_isotone) also have "... \ r\<^sup>T * g\<^sup>\" by (simp add: mult_assoc mult_right_isotone star.left_plus_below_circ star_plus) finally have "r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * g = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * g \ r\<^sup>T * g\<^sup>\" by (simp add: le_iff_inf) also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * (g \ r\<^sup>T * g\<^sup>\)" using assms covector_comp_inf covector_mult_closed vector_conv_covector by auto also have "... = (r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ \ r\<^sup>T * g\<^sup>\) * (g \ r\<^sup>T * g\<^sup>\)" by (simp add: inf.absorb2 inf_commute mult_right_isotone star_isotone) also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * (g \ r\<^sup>T * g\<^sup>\ \ (r\<^sup>T * g\<^sup>\)\<^sup>T)" using 2 by (metis comp_inf_vector_1) also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * ((r\<^sup>T * g\<^sup>\)\<^sup>T \ r\<^sup>T * g\<^sup>\ \ g)" using inf_commute inf_assoc by simp also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)" using 2 by (metis covector_conv_vector inf_top.right_neutral vector_inf_comp) also have "... \ r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\" by (simp add: mult_assoc mult_right_isotone star.left_plus_below_circ star_plus) finally have "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\" using 1 star_right_induct by auto thus ?thesis by (simp add: inf.eq_iff mult_right_isotone star_isotone) qed lemma kruskal_acyclic_inv_1: assumes "injective f" and "e * forest_components f * e = bot" shows "(f \ top * e * f\<^sup>T\<^sup>\)\<^sup>T * f\<^sup>\ * e = bot" proof - let ?q = "top * e * f\<^sup>T\<^sup>\" let ?F = "forest_components f" have "(f \ ?q)\<^sup>T * f\<^sup>\ * e = ?q\<^sup>T \ f\<^sup>T * f\<^sup>\ * e" by (metis (mono_tags) comp_associative conv_dist_inf covector_conv_vector inf_vector_comp vector_top_closed) also have "... \ ?q\<^sup>T \ ?F * e" using comp_inf.mult_right_isotone mult_left_isotone star.circ_increasing by simp also have "... = f\<^sup>\ * e\<^sup>T * top \ ?F * e" by (simp add: conv_dist_comp conv_star_commute mult_assoc) also have "... \ ?F * e\<^sup>T * top \ ?F * e" by (metis conv_dist_comp conv_star_commute conv_top inf.sup_left_isotone star.circ_right_top star_outer_increasing mult_assoc) also have "... = ?F * (e\<^sup>T * top \ ?F * e)" by (metis assms(1) forest_components_equivalence equivalence_comp_dist_inf mult_assoc) also have "... = (?F \ top * e) * ?F * e" by (simp add: comp_associative comp_inf_vector_1 conv_dist_comp inf_vector_comp) also have "... \ top * e * ?F * e" by (simp add: mult_left_isotone) also have "... = bot" using assms(2) mult_assoc by simp finally show ?thesis by (simp add: bot_unique) qed lemma kruskal_forest_components_inf_1: assumes "f \ w \ w\<^sup>T" and "injective w" and "f \ forest_components g" shows "f * forest_components (forest_components g \ w) \ forest_components (forest_components g \ w)" proof - let ?f = "forest_components g" let ?w = "forest_components (?f \ w)" have "f * ?w = (f \ (w \ w\<^sup>T)) * ?w" by (simp add: assms(1) inf.absorb1) also have "... = (f \ w) * ?w \ (f \ w\<^sup>T) * ?w" by (simp add: inf_sup_distrib1 semiring.distrib_right) also have "... \ (?f \ w) * ?w \ (f \ w\<^sup>T) * ?w" using assms(3) inf.sup_left_isotone mult_left_isotone sup_left_isotone by simp also have "... \ (?f \ w) * ?w \ (?f \ w\<^sup>T) * ?w" using assms(3) inf.sup_left_isotone mult_left_isotone sup_right_isotone by simp also have "... = (?f \ w) * ?w \ (?f \ w)\<^sup>T * ?w" by (simp add: conv_dist_comp conv_dist_inf conv_star_commute) also have "... \ (?f \ w) * ?w \ ?w" by (metis star.circ_loop_fixpoint sup_ge1 sup_right_isotone) also have "... = ?w \ (?f \ w) * (?f \ w)\<^sup>\ \ (?f \ w) * (?f \ w)\<^sup>T\<^sup>+ * (?f \ w)\<^sup>\" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute sup_assoc) also have "... \ ?w \ (?f \ w)\<^sup>\ \ (?f \ w) * (?f \ w)\<^sup>T\<^sup>+ * (?f \ w)\<^sup>\" using star.left_plus_below_circ sup_left_isotone sup_right_isotone by auto also have "... = ?w \ (?f \ w) * (?f \ w)\<^sup>T\<^sup>+ * (?f \ w)\<^sup>\" by (metis star.circ_loop_fixpoint sup.right_idem) also have "... \ ?w \ w * w\<^sup>T * ?w" using comp_associative conv_dist_inf mult_isotone sup_right_isotone by simp also have "... = ?w" by (metis assms(2) coreflexive_comp_top_inf inf.cobounded2 sup.orderE) finally show ?thesis by simp qed lemma kruskal_forest_components_inf: assumes "f \ w \ w\<^sup>T" and "injective w" shows "forest_components f \ forest_components (forest_components f \ w)" proof - let ?f = "forest_components f" let ?w = "forest_components (?f \ w)" have 1: "1 \ ?w" by (simp add: reflexive_mult_closed star.circ_reflexive) have "f * ?w \ ?w" using assms forest_components_increasing kruskal_forest_components_inf_1 by simp hence 2: "f\<^sup>\ \ ?w" using 1 star_left_induct by fastforce have "f\<^sup>T * ?w \ ?w" apply (rule kruskal_forest_components_inf_1) apply (metis assms(1) conv_dist_sup conv_involutive conv_isotone sup_commute) apply (simp add: assms(2)) by (metis le_supI2 star.circ_back_loop_fixpoint star.circ_increasing) thus "?f \ ?w" using 2 star_left_induct by simp qed end text \ We next add the Kleene star to single-object pseudocomplemented distributive allegories. \ class pd_kleene_allegory = pd_allegory + bounded_distrib_kleene_allegory begin text \ The following definitions and results concern acyclic graphs and forests. \ abbreviation acyclic :: "'a \ bool" where "acyclic x \ x\<^sup>+ \ -1" abbreviation forest :: "'a \ bool" where "forest x \ injective x \ acyclic x" lemma forest_bot: "forest bot" by simp lemma acyclic_star_below_complement: "acyclic w \ w\<^sup>T\<^sup>\ \ -w" by (simp add: conv_star_commute schroeder_4_p) lemma acyclic_star_below_complement_1: "acyclic w \ w\<^sup>\ \ w\<^sup>T = bot" using pseudo_complement schroeder_5_p by force lemma acyclic_star_inf_conv: assumes "acyclic w" shows "w\<^sup>\ \ w\<^sup>T\<^sup>\ = 1" proof - have "w\<^sup>+ \ w\<^sup>T\<^sup>\ \ (w \ w\<^sup>T\<^sup>\) * w\<^sup>\" by (metis conv_star_commute dedekind_2 star.circ_transitive_equal) also have "... = bot" by (metis assms conv_star_commute p_antitone_iff pseudo_complement schroeder_4_p semiring.mult_not_zero star.circ_circ_mult star_involutive star_one) finally have "w\<^sup>\ \ w\<^sup>T\<^sup>\ \ 1" by (metis eq_iff le_bot mult_left_zero star.circ_plus_one star.circ_zero star_left_unfold_equal sup_inf_distrib1) thus ?thesis by (simp add: inf.antisym star.circ_reflexive) qed lemma acyclic_asymmetric: "acyclic w \ asymmetric w" by (simp add: local.dual_order.trans local.pseudo_complement local.schroeder_5_p local.star.circ_increasing) lemma forest_separate: assumes "forest x" shows "x\<^sup>\ * x\<^sup>T\<^sup>\ \ x\<^sup>T * x \ 1" proof - have "x\<^sup>\ * 1 \ -x\<^sup>T" using assms schroeder_5_p by force hence 1: "x\<^sup>\ \ x\<^sup>T = bot" by (simp add: pseudo_complement) have "x\<^sup>\ \ x\<^sup>T * x = (1 \ x\<^sup>\ * x) \ x\<^sup>T * x" using star.circ_right_unfold_1 by simp also have "... = (1 \ x\<^sup>T * x) \ (x\<^sup>\ * x \ x\<^sup>T * x)" by (simp add: inf_sup_distrib2) also have "... \ 1 \ (x\<^sup>\ * x \ x\<^sup>T * x)" using sup_left_isotone by simp also have "... = 1 \ (x\<^sup>\ \ x\<^sup>T) * x" by (simp add: assms injective_comp_right_dist_inf) also have "... = 1" using 1 by simp finally have 2: "x\<^sup>\ \ x\<^sup>T * x \ 1" . hence 3: "x\<^sup>T\<^sup>\ \ x\<^sup>T * x \ 1" by (metis (mono_tags, lifting) conv_star_commute conv_dist_comp conv_dist_inf conv_involutive coreflexive_symmetric) have "x\<^sup>\ * x\<^sup>T\<^sup>\ \ x\<^sup>T * x \ (x\<^sup>\ \ x\<^sup>T\<^sup>\) \ x\<^sup>T * x" using assms cancel_separate inf.sup_left_isotone by simp also have "... \ 1" using 2 3 by (simp add: inf_sup_distrib2) finally show ?thesis . qed text \ The following definition captures the components of undirected weighted graphs. \ abbreviation "components g \ (--g)\<^sup>\" lemma components_equivalence: "symmetric x \ equivalence (components x)" by (simp add: conv_star_commute conv_complement star.circ_reflexive star.circ_transitive_equal) lemma components_increasing: "x \ components x" using order_trans pp_increasing star.circ_increasing by blast lemma components_isotone: "x \ y \ components x \ components y" by (simp add: pp_isotone star_isotone) lemma cut_reachable: assumes "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" shows "v * -v\<^sup>T \ g \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\)" proof - have "v * -v\<^sup>T \ g \ v * top \ g" using inf.sup_left_isotone mult_right_isotone top_greatest by blast also have "... = (r\<^sup>T * t\<^sup>\)\<^sup>T * top \ g" by (metis assms(1) conv_involutive) also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * top \ g" using assms(2) conv_isotone inf.sup_left_isotone mult_left_isotone mult_right_isotone star_isotone by auto also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * ((r\<^sup>T * g\<^sup>\) * g)" by (metis conv_involutive dedekind_1 inf_top.left_neutral) also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\)" by (simp add: mult_assoc mult_right_isotone star.left_plus_below_circ star_plus) finally show ?thesis . qed text \ The following lemma shows that the predecessors of visited nodes in the minimum spanning tree extending the current tree have all been visited. \ lemma predecessors_reachable: assumes "vector r" and "injective r" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "forest w" and "t \ w" and "w \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" shows "w * v \ v" proof - have "w * r \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) * r" using assms(6) mult_left_isotone by auto also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * top" by (simp add: mult_assoc mult_right_isotone) also have "... = (r\<^sup>T * g\<^sup>\)\<^sup>T" by (simp add: assms(1) comp_associative conv_dist_comp) also have "... \ (r\<^sup>T * w\<^sup>\)\<^sup>T" by (simp add: assms(7) conv_isotone) also have "... = w\<^sup>T\<^sup>\ * r" by (simp add: conv_dist_comp conv_star_commute) also have "... \ -w * r" using assms(4) by (simp add: mult_left_isotone acyclic_star_below_complement) also have "... \ -(w * r)" by (simp add: assms(2) comp_injective_below_complement) finally have 1: "w * r = bot" by (simp add: le_iff_inf) have "v = t\<^sup>T\<^sup>\ * r" by (metis assms(3) conv_dist_comp conv_involutive conv_star_commute) also have "... = t\<^sup>T * v \ r" by (simp add: calculation star.circ_loop_fixpoint) also have "... \ w\<^sup>T * v \ r" using assms(5) comp_isotone conv_isotone semiring.add_right_mono by auto finally have "w * v \ w * w\<^sup>T * v \ w * r" by (simp add: comp_left_dist_sup mult_assoc mult_right_isotone) also have "... = w * w\<^sup>T * v" using 1 by simp also have "... \ v" using assms(4) by (simp add: star_left_induct_mult_iff star_sub_one) finally show ?thesis . qed subsection \Prim's Algorithm\ text \ The following results are used for proving the correctness of Prim's minimum spanning tree algorithm. \ subsubsection \Preservation of Invariant\ text \ We first treat the preservation of the invariant. The following lemma shows that the while-loop preserves that \v\ represents the nodes of the constructed tree. The remaining lemmas in this section show that \t\ is a spanning tree. The exchange property is treated in the following two sections. \ lemma reachable_inv: assumes "vector v" and "e \ v * -v\<^sup>T" and "e * t = bot" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" shows "(v \ e\<^sup>T * top)\<^sup>T = r\<^sup>T * (t \ e)\<^sup>\" proof - have 1: "v\<^sup>T \ r\<^sup>T * (t \ e)\<^sup>\" by (simp add: assms(4) mult_right_isotone star.circ_sub_dist) have 2: "(e\<^sup>T * top)\<^sup>T = top * e" by (simp add: conv_dist_comp) also have "... = top * (v * -v\<^sup>T \ e)" by (simp add: assms(2) inf_absorb2) also have "... \ top * (v * top \ e)" using inf.sup_left_isotone mult_right_isotone top_greatest by blast also have "... = top * v\<^sup>T * e" by (simp add: comp_inf_vector inf.sup_monoid.add_commute) also have "... = v\<^sup>T * e" using assms(1) vector_conv_covector by auto also have "... \ r\<^sup>T * (t \ e)\<^sup>\ * e" using 1 by (simp add: mult_left_isotone) also have "... \ r\<^sup>T * (t \ e)\<^sup>\ * (t \ e)" by (simp add: mult_right_isotone) also have "... \ r\<^sup>T * (t \ e)\<^sup>\" by (simp add: comp_associative mult_right_isotone star.right_plus_below_circ) finally have 3: "(v \ e\<^sup>T * top)\<^sup>T \ r\<^sup>T * (t \ e)\<^sup>\" using 1 by (simp add: conv_dist_sup) have "r\<^sup>T \ r\<^sup>T * t\<^sup>\" using sup.bounded_iff star.circ_back_loop_prefixpoint by blast also have "... \ (v \ e\<^sup>T * top)\<^sup>T" by (metis assms(4) conv_isotone sup_ge1) finally have 4: "r\<^sup>T \ (v \ e\<^sup>T * top)\<^sup>T" . have "(v \ e\<^sup>T * top)\<^sup>T * (t \ e) = (v \ e\<^sup>T * top)\<^sup>T * t \ (v \ e\<^sup>T * top)\<^sup>T * e" by (simp add: mult_left_dist_sup) also have "... \ (v \ e\<^sup>T * top)\<^sup>T * t \ top * e" using comp_isotone semiring.add_left_mono by auto also have "... = v\<^sup>T * t \ top * e * t \ top * e" using 2 by (simp add: conv_dist_sup mult_right_dist_sup) also have "... = v\<^sup>T * t \ top * e" by (simp add: assms(3) comp_associative) also have "... \ r\<^sup>T * t\<^sup>\ \ top * e" by (metis assms(4) star.circ_back_loop_fixpoint sup_ge1 sup_left_isotone) also have "... = v\<^sup>T \ top * e" by (simp add: assms(4)) finally have 5: "(v \ e\<^sup>T * top)\<^sup>T * (t \ e) \ (v \ e\<^sup>T * top)\<^sup>T" using 2 by (simp add: conv_dist_sup) have "r\<^sup>T * (t \ e)\<^sup>\ \ (v \ e\<^sup>T * top)\<^sup>T * (t \ e)\<^sup>\" using 4 by (simp add: mult_left_isotone) also have "... \ (v \ e\<^sup>T * top)\<^sup>T" using 5 by (simp add: star_right_induct_mult) finally show ?thesis using 3 by (simp add: inf.eq_iff) qed text \ The next result is used to show that the while-loop preserves acyclicity of the constructed tree. \ lemma acyclic_inv: assumes "acyclic t" and "vector v" and "e \ v * -v\<^sup>T" and "t \ v * v\<^sup>T" shows "acyclic (t \ e)" proof - have "t\<^sup>+ * e \ t\<^sup>+ * v * -v\<^sup>T" by (simp add: assms(3) comp_associative mult_right_isotone) also have "... \ v * v\<^sup>T * t\<^sup>\ * v * -v\<^sup>T" by (simp add: assms(4) mult_left_isotone) also have "... \ v * top * -v\<^sup>T" by (metis mult_assoc mult_left_isotone mult_right_isotone top_greatest) also have "... = v * -v\<^sup>T" by (simp add: assms(2)) also have "... \ -1" by (simp add: pp_increasing schroeder_3_p) finally have 1: "t\<^sup>+ * e \ -1" . have 2: "e * t\<^sup>\ = e" using assms(2-4) et(1) star_absorb by blast have "e\<^sup>\ = 1 \ e \ e * e * e\<^sup>\" by (metis star.circ_loop_fixpoint star_square_2 sup_commute) also have "... = 1 \ e" using assms(2,3) ee comp_left_zero bot_least sup_absorb1 by simp finally have 3: "e\<^sup>\ = 1 \ e" . have "e \ v * -v\<^sup>T" by (simp add: assms(3)) also have "... \ -1" by (simp add: pp_increasing schroeder_3_p) finally have 4: "t\<^sup>+ * e \ e \ -1" using 1 by simp have "(t \ e)\<^sup>+ = (t \ e) * t\<^sup>\ * (e * t\<^sup>\)\<^sup>\" using star_sup_1 mult_assoc by simp also have "... = (t \ e) * t\<^sup>\ * (1 \ e)" using 2 3 by simp also have "... = t\<^sup>+ * (1 \ e) \ e * t\<^sup>\ * (1 \ e)" by (simp add: comp_right_dist_sup) also have "... = t\<^sup>+ * (1 \ e) \ e * (1 \ e)" using 2 by simp also have "... = t\<^sup>+ * (1 \ e) \ e" using 3 by (metis star_absorb assms(2,3) ee) also have "... = t\<^sup>+ \ t\<^sup>+ * e \ e" by (simp add: mult_left_dist_sup) also have "... \ -1" using 4 by (metis assms(1) sup.absorb1 sup.orderI sup_assoc) finally show ?thesis . qed text \ The following lemma shows that the extended tree is in the component reachable from the root. \ lemma mst_subgraph_inv_2: assumes "regular (v * v\<^sup>T)" and "t \ v * v\<^sup>T \ --g" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "e \ v * -v\<^sup>T \ --g" and "vector v" and "regular ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T)" shows "t \ e \ (r\<^sup>T * (--((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g))\<^sup>\)\<^sup>T * (r\<^sup>T * (--((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g))\<^sup>\)" proof - let ?v = "v \ e\<^sup>T * top" let ?G = "?v * ?v\<^sup>T \ g" let ?c = "r\<^sup>T * (--?G)\<^sup>\" have "v\<^sup>T \ r\<^sup>T * (--(v * v\<^sup>T \ g))\<^sup>\" using assms(1-3) inf_pp_commute mult_right_isotone star_isotone by auto also have "... \ ?c" using comp_inf.mult_right_isotone comp_isotone conv_isotone inf.commute mult_right_isotone pp_isotone star_isotone sup.cobounded1 by presburger finally have 2: "v\<^sup>T \ ?c \ v \ ?c\<^sup>T" by (metis conv_isotone conv_involutive) have "t \ v * v\<^sup>T" using assms(2) by auto hence 3: "t \ ?c\<^sup>T * ?c" using 2 order_trans mult_isotone by blast have "e \ v * top \ --g" by (metis assms(4,5) inf.bounded_iff inf.sup_left_divisibility mult_right_isotone top.extremum) hence "e \ v * top \ top * e \ --g" by (simp add: top_left_mult_increasing inf.boundedI) hence "e \ v * top * e \ --g" by (metis comp_inf_covector inf.absorb2 mult_assoc top.extremum) hence "t \ e \ (v * v\<^sup>T \ --g) \ (v * top * e \ --g)" using assms(2) sup_mono by blast also have "... = v * ?v\<^sup>T \ --g" by (simp add: inf_sup_distrib2 mult_assoc mult_left_dist_sup conv_dist_comp conv_dist_sup) also have "... \ --?G" using assms(6) comp_left_increasing_sup inf.sup_left_isotone pp_dist_inf by auto finally have 4: "t \ e \ --?G" . have "e \ e * e\<^sup>T * e" by (simp add: ex231c) also have "... \ v * -v\<^sup>T * -v * v\<^sup>T * e" by (metis assms(4) mult_left_isotone conv_isotone conv_dist_comp mult_assoc mult_isotone conv_involutive conv_complement inf.boundedE) also have "... \ v * top * v\<^sup>T * e" by (metis mult_assoc mult_left_isotone mult_right_isotone top.extremum) also have "... = v * r\<^sup>T * t\<^sup>\ * e" using assms(3,5) by (simp add: mult_assoc) also have "... \ v * r\<^sup>T * (t \ e)\<^sup>\" by (simp add: comp_associative mult_right_isotone star.circ_mult_upper_bound star.circ_sub_dist_1 star_isotone sup_commute) also have "... \ v * ?c" using 4 by (simp add: mult_assoc mult_right_isotone star_isotone) also have "... \ ?c\<^sup>T * ?c" using 2 by (simp add: mult_left_isotone) finally show ?thesis using 3 by simp qed lemma span_inv: assumes "e \ v * -v\<^sup>T" and "vector v" and "arc e" and "t \ (v * v\<^sup>T) \ g" and "g\<^sup>T = g" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "injective r" and "r\<^sup>T \ v\<^sup>T" and "r\<^sup>T * ((v * v\<^sup>T) \ g)\<^sup>\ \ r\<^sup>T * t\<^sup>\" shows "r\<^sup>T * (((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g)\<^sup>\ \ r\<^sup>T * (t \ e)\<^sup>\" proof - let ?d = "(v * v\<^sup>T) \ g" have 1: "(v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T = v * v\<^sup>T \ v * v\<^sup>T * e \ e\<^sup>T * v * v\<^sup>T \ e\<^sup>T * e" using assms(1-3) ve_dist by simp have "t\<^sup>T \ ?d\<^sup>T" using assms(4) conv_isotone by simp also have "... = (v * v\<^sup>T) \ g\<^sup>T" by (simp add: conv_dist_comp conv_dist_inf) also have "... = ?d" by (simp add: assms(5)) finally have 2: "t\<^sup>T \ ?d" . have "v * v\<^sup>T = (r\<^sup>T * t\<^sup>\)\<^sup>T * (r\<^sup>T * t\<^sup>\)" by (metis assms(6) conv_involutive) also have "... = t\<^sup>T\<^sup>\ * (r * r\<^sup>T) * t\<^sup>\" by (simp add: comp_associative conv_dist_comp conv_star_commute) also have "... \ t\<^sup>T\<^sup>\ * 1 * t\<^sup>\" by (simp add: assms(7) mult_left_isotone star_right_induct_mult_iff star_sub_one) also have "... = t\<^sup>T\<^sup>\ * t\<^sup>\" by simp also have "... \ ?d\<^sup>\ * t\<^sup>\" using 2 by (simp add: comp_left_isotone star.circ_isotone) also have "... \ ?d\<^sup>\ * ?d\<^sup>\" using assms(4) mult_right_isotone star_isotone by simp also have 3: "... = ?d\<^sup>\" by (simp add: star.circ_transitive_equal) finally have 4: "v * v\<^sup>T \ ?d\<^sup>\" . have 5: "r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\" by (simp add: comp_associative mult_right_isotone star.circ_plus_same star.left_plus_below_circ) have "r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * v * v\<^sup>T * e" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * ?d\<^sup>\ * e" using 3 4 by (metis comp_associative comp_isotone eq_refl) finally have 6: "r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e" . have 7: "\x . r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * x = bot" proof fix x have "r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * x \ r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * top" by (simp add: mult_right_isotone) also have "... = r\<^sup>T * e\<^sup>T * top \ r\<^sup>T * v * v\<^sup>T * e\<^sup>T * top" by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup) also have "... = r\<^sup>T * e\<^sup>T * top" by (metis assms(1,2) mult_assoc mult_right_dist_sup mult_right_zero sup_bot_right vTeT) also have "... \ v\<^sup>T * e\<^sup>T * top" by (simp add: assms(8) comp_isotone) also have "... = bot" using vTeT assms(1,2) by simp finally show "r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * x = bot" by (simp add: le_bot) qed have "r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e\<^sup>T * v * v\<^sup>T" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * v * v\<^sup>T" by (metis assms(2) star.circ_isotone vector_vector_star inf_le1 comp_associative comp_right_isotone comp_left_isotone) also have "... = bot" using 7 by simp finally have 8: "r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * v * v\<^sup>T \ g) = bot" by (simp add: le_bot) have "r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e\<^sup>T * e" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * e" by (metis assms(2) star.circ_isotone vector_vector_star inf_le1 comp_associative comp_right_isotone comp_left_isotone) also have "... = bot" using 7 by simp finally have 9: "r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * e \ g) = bot" by (simp add: le_bot) have "r\<^sup>T * ?d\<^sup>\ * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) = r\<^sup>T * ?d\<^sup>\ * ((v * v\<^sup>T \ v * v\<^sup>T * e \ e\<^sup>T * v * v\<^sup>T \ e\<^sup>T * e) \ g)" using 1 by simp also have "... = r\<^sup>T * ?d\<^sup>\ * ((v * v\<^sup>T \ g) \ (v * v\<^sup>T * e \ g) \ (e\<^sup>T * v * v\<^sup>T \ g) \ (e\<^sup>T * e \ g))" by (simp add: inf_sup_distrib2) also have "... = r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * e \ g)" by (simp add: comp_left_dist_sup) also have "... = r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T * e \ g)" using 8 9 by simp also have "... \ r\<^sup>T * ?d\<^sup>\ \ r\<^sup>T * ?d\<^sup>\ * e" using 5 6 sup.mono by simp also have "... = r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by (simp add: mult_left_dist_sup) finally have 10: "r\<^sup>T * ?d\<^sup>\ * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by simp have "r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * v * v\<^sup>T" by (simp add: comp_associative comp_right_isotone) also have "... = bot" by (metis assms(1,2) comp_associative comp_right_zero ev comp_left_zero) finally have 11: "r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T \ g) = bot" by (simp add: le_bot) have "r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e * v * v\<^sup>T * e" by (simp add: comp_associative comp_right_isotone) also have "... = bot" by (metis assms(1,2) comp_associative comp_right_zero ev comp_left_zero) finally have 12: "r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T * e \ g) = bot" by (simp add: le_bot) have "r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * e\<^sup>T * v * v\<^sup>T" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * ?d\<^sup>\ * 1 * v * v\<^sup>T" by (metis assms(3) arc_injective comp_associative comp_left_isotone comp_right_isotone) also have "... = r\<^sup>T * ?d\<^sup>\ * v * v\<^sup>T" by simp also have "... \ r\<^sup>T * ?d\<^sup>\ * ?d\<^sup>\" using 4 by (simp add: mult_right_isotone mult_assoc) also have "... = r\<^sup>T * ?d\<^sup>\" by (simp add: star.circ_transitive_equal comp_associative) finally have 13: "r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\" . have "r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e * e\<^sup>T * e" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * ?d\<^sup>\ * 1 * e" by (metis assms(3) arc_injective comp_associative comp_left_isotone comp_right_isotone) also have "... = r\<^sup>T * ?d\<^sup>\ * e" by simp finally have 14: "r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e" . have "r\<^sup>T * ?d\<^sup>\ * e * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) = r\<^sup>T * ?d\<^sup>\ * e * ((v * v\<^sup>T \ v * v\<^sup>T * e \ e\<^sup>T * v * v\<^sup>T \ e\<^sup>T * e) \ g)" using 1 by simp also have "... = r\<^sup>T * ?d\<^sup>\ * e * ((v * v\<^sup>T \ g) \ (v * v\<^sup>T * e \ g) \ (e\<^sup>T * v * v\<^sup>T \ g) \ (e\<^sup>T * e \ g))" by (simp add: inf_sup_distrib2) also have "... = r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * e \ g)" by (simp add: comp_left_dist_sup) also have "... = r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * e \ g)" using 11 12 by simp also have "... \ r\<^sup>T * ?d\<^sup>\ \ r\<^sup>T * ?d\<^sup>\ * e" using 13 14 sup_mono by simp also have "... = r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by (simp add: mult_left_dist_sup) finally have 15: "r\<^sup>T * ?d\<^sup>\ * e * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by simp have "r\<^sup>T \ r\<^sup>T * ?d\<^sup>\" using mult_right_isotone star.circ_reflexive by fastforce also have "... \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by (simp add: semiring.distrib_left) finally have 16: "r\<^sup>T \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" . have "r\<^sup>T * ?d\<^sup>\ * (1 \ e) * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) = r\<^sup>T * ?d\<^sup>\ * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g)" by (simp add: semiring.distrib_left semiring.distrib_right) also have "... \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" using 10 15 le_supI by simp finally have "r\<^sup>T * ?d\<^sup>\ * (1 \ e) * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" . hence "r\<^sup>T \ r\<^sup>T * ?d\<^sup>\ * (1 \ e) * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" using 16 sup_least by simp hence "r\<^sup>T * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g)\<^sup>\ \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by (simp add: star_right_induct) also have "... \ r\<^sup>T * t\<^sup>\ * (1 \ e)" by (simp add: assms(9) mult_left_isotone) also have "... \ r\<^sup>T * (t \ e)\<^sup>\" by (simp add: star_one_sup_below) finally show ?thesis . qed subsubsection \Exchange gives Spanning Trees\ text \ The following abbreviations are used in the spanning tree application using Prim's algorithm to construct the new tree for the exchange property. It is obtained by replacing an edge with one that has minimal weight and reversing the path connecting these edges. Here, w represents a weighted graph, v represents a set of nodes and e represents an edge. \ abbreviation prim_E :: "'a \ 'a \ 'a \ 'a" where "prim_E w v e \ w \ --v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" abbreviation prim_P :: "'a \ 'a \ 'a \ 'a" where "prim_P w v e \ w \ -v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" abbreviation prim_EP :: "'a \ 'a \ 'a \ 'a" where "prim_EP w v e \ w \ -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" abbreviation prim_W :: "'a \ 'a \ 'a \ 'a" where "prim_W w v e \ (w \ -(prim_EP w v e)) \ (prim_P w v e)\<^sup>T \ e" text \ The lemmas in this section are used to show that the relation after exchange represents a spanning tree. The results in the next section are used to show that it is a minimum spanning tree. \ lemma exchange_injective_3: assumes "e \ v * -v\<^sup>T" and "vector v" shows "(w \ -(prim_EP w v e)) * e\<^sup>T = bot" proof - have 1: "top * e \ -v\<^sup>T" by (simp add: assms schroeder_4_p vTeT) have "top * e \ top * e * w\<^sup>T\<^sup>\" using sup_right_divisibility star.circ_back_loop_fixpoint by blast hence "top * e \ -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" using 1 by simp hence "top * e \ -(w \ -prim_EP w v e)" by (metis inf.assoc inf_import_p le_infI2 p_antitone p_antitone_iff) hence "(w \ -(prim_EP w v e)) * e\<^sup>T \ bot" using p_top schroeder_4_p by blast thus ?thesis using le_bot by simp qed lemma exchange_injective_6: assumes "arc e" and "forest w" shows "(prim_P w v e)\<^sup>T * e\<^sup>T = bot" proof - have "e\<^sup>T * top * e \ --1" by (simp add: assms(1) local.p_antitone local.p_antitone_iff local.point_injective) hence 1: "e * -1 * e\<^sup>T \ bot" by (metis conv_involutive p_top triple_schroeder_p) have "(prim_P w v e)\<^sup>T * e\<^sup>T \ (w \ top * e * w\<^sup>T\<^sup>\)\<^sup>T * e\<^sup>T" using comp_inf.mult_left_isotone conv_dist_inf mult_left_isotone by simp also have "... = (w\<^sup>T \ w\<^sup>T\<^sup>\\<^sup>T * e\<^sup>T * top) * e\<^sup>T" by (simp add: comp_associative conv_dist_comp conv_dist_inf) also have "... = w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T * e\<^sup>T" by (simp add: conv_star_commute inf_vector_comp) also have "... \ (w\<^sup>T \ w\<^sup>\ * e\<^sup>T * top * e) * (e\<^sup>T \ w\<^sup>+ * e\<^sup>T * top)" by (metis dedekind mult_assoc conv_involutive inf_commute) also have "... \ (w\<^sup>\ * e\<^sup>T * top * e) * (w\<^sup>+ * e\<^sup>T * top)" by (simp add: mult_isotone) also have "... \ (top * e) * (w\<^sup>+ * e\<^sup>T * top)" by (simp add: mult_left_isotone) also have "... = top * e * w\<^sup>+ * e\<^sup>T * top" using mult_assoc by simp also have "... \ top * e * -1 * e\<^sup>T * top" using assms(2) mult_left_isotone mult_right_isotone by simp also have "... \ bot" using 1 by (metis le_bot semiring.mult_not_zero mult_assoc) finally show ?thesis using le_bot by simp qed text \ The graph after exchanging is injective. \ lemma exchange_injective: assumes "arc e" and "e \ v * -v\<^sup>T" and "forest w" and "vector v" shows "injective (prim_W w v e)" proof - have 1: "(w \ -(prim_EP w v e)) * (w \ -(prim_EP w v e))\<^sup>T \ 1" proof - have "(w \ -(prim_EP w v e)) * (w \ -(prim_EP w v e))\<^sup>T \ w * w\<^sup>T" by (simp add: comp_isotone conv_isotone) also have "... \ 1" by (simp add: assms(3)) finally show ?thesis . qed have 2: "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>T \ 1" proof - have "top * (prim_P w v e)\<^sup>T = top * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>T\<^sup>\\<^sup>T * e\<^sup>T * top)" by (simp add: comp_associative conv_complement conv_dist_comp conv_dist_inf) also have "... = top * e * w\<^sup>T\<^sup>\ * (w\<^sup>T \ -v * -v\<^sup>T)" by (metis comp_inf_vector conv_dist_comp conv_involutive inf_top_left mult_assoc) also have "... \ top * e * w\<^sup>T\<^sup>\ * (w\<^sup>T \ top * -v\<^sup>T)" using comp_inf.mult_right_isotone mult_left_isotone mult_right_isotone by simp also have "... = top * e * w\<^sup>T\<^sup>\ * w\<^sup>T \ -v\<^sup>T" by (metis assms(4) comp_inf_covector vector_conv_compl) also have "... \ -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" by (simp add: comp_associative comp_isotone inf.coboundedI1 star.circ_plus_same star.left_plus_below_circ) finally have "top * (prim_P w v e)\<^sup>T \ -(w \ -prim_EP w v e)" by (metis inf.assoc inf_import_p le_infI2 p_antitone p_antitone_iff) hence "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>T \ bot" using p_top schroeder_4_p by blast thus ?thesis by (simp add: bot_unique) qed have 3: "(w \ -(prim_EP w v e)) * e\<^sup>T \ 1" by (metis assms(2,4) exchange_injective_3 bot_least) have 4: "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>T \ 1" using 2 conv_dist_comp coreflexive_symmetric by fastforce have 5: "(prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \ 1" proof - have "(prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \ (top * e * w\<^sup>T\<^sup>\)\<^sup>T * (top * e * w\<^sup>T\<^sup>\)" by (simp add: conv_dist_inf mult_isotone) also have "... = w\<^sup>\ * e\<^sup>T * top * top * e * w\<^sup>T\<^sup>\" using conv_star_commute conv_dist_comp conv_involutive conv_top mult_assoc by presburger also have "... = w\<^sup>\ * e\<^sup>T * top * e * w\<^sup>T\<^sup>\" by (simp add: comp_associative) also have "... \ w\<^sup>\ * 1 * w\<^sup>T\<^sup>\" by (metis comp_left_isotone comp_right_isotone mult_assoc assms(1) point_injective) finally have "(prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \ w\<^sup>\ * w\<^sup>T\<^sup>\ \ w\<^sup>T * w" by (simp add: conv_isotone inf.left_commute inf.sup_monoid.add_commute mult_isotone) also have "... \ 1" by (simp add: assms(3) forest_separate) finally show ?thesis . qed have 6: "(prim_P w v e)\<^sup>T * e\<^sup>T \ 1" using assms exchange_injective_6 bot_least by simp have 7: "e * (w \ -(prim_EP w v e))\<^sup>T \ 1" using 3 by (metis conv_dist_comp conv_involutive coreflexive_symmetric) have 8: "e * (prim_P w v e)\<^sup>T\<^sup>T \ 1" using 6 conv_dist_comp coreflexive_symmetric by fastforce have 9: "e * e\<^sup>T \ 1" by (simp add: assms(1) arc_injective) have "(prim_W w v e) * (prim_W w v e)\<^sup>T = (w \ -(prim_EP w v e)) * (w \ -(prim_EP w v e))\<^sup>T \ (w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>T \ (w \ -(prim_EP w v e)) * e\<^sup>T \ (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>T \ (prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \ (prim_P w v e)\<^sup>T * e\<^sup>T \ e * (w \ -(prim_EP w v e))\<^sup>T \ e * (prim_P w v e)\<^sup>T\<^sup>T \ e * e\<^sup>T" using comp_left_dist_sup comp_right_dist_sup conv_dist_sup sup.assoc by simp also have "... \ 1" using 1 2 3 4 5 6 7 8 9 by simp finally show ?thesis . qed lemma pv: assumes "vector v" shows "(prim_P w v e)\<^sup>T * v = bot" proof - have "(prim_P w v e)\<^sup>T * v \ (-v * -v\<^sup>T)\<^sup>T * v" by (meson conv_isotone inf_le1 inf_le2 mult_left_isotone order_trans) also have "... = -v * -v\<^sup>T * v" by (simp add: conv_complement conv_dist_comp) also have "... = bot" by (simp add: assms covector_vector_comp mult_assoc) finally show ?thesis by (simp add: antisym) qed lemma vector_pred_inv: assumes "arc e" and "e \ v * -v\<^sup>T" and "forest w" and "vector v" and "w * v \ v" shows "(prim_W w v e) * (v \ e\<^sup>T * top) \ v \ e\<^sup>T * top" proof - have "(prim_W w v e) * e\<^sup>T * top = (w \ -(prim_EP w v e)) * e\<^sup>T * top \ (prim_P w v e)\<^sup>T * e\<^sup>T * top \ e * e\<^sup>T * top" by (simp add: mult_right_dist_sup) also have "... = e * e\<^sup>T * top" using assms exchange_injective_3 exchange_injective_6 comp_left_zero by simp also have "... \ v * -v\<^sup>T * e\<^sup>T * top" by (simp add: assms(2) comp_isotone) also have "... \ v * top" by (simp add: comp_associative mult_right_isotone) also have "... = v" by (simp add: assms(4)) finally have 1: "(prim_W w v e) * e\<^sup>T * top \ v" . have "(prim_W w v e) * v = (w \ -(prim_EP w v e)) * v \ (prim_P w v e)\<^sup>T * v \ e * v" by (simp add: mult_right_dist_sup) also have "... = (w \ -(prim_EP w v e)) * v" by (metis assms(2,4) pv ev sup_bot_right) also have "... \ w * v" by (simp add: mult_left_isotone) finally have 2: "(prim_W w v e) * v \ v" using assms(5) order_trans by blast have "(prim_W w v e) * (v \ e\<^sup>T * top) = (prim_W w v e) * v \ (prim_W w v e) * e\<^sup>T * top" by (simp add: semiring.distrib_left mult_assoc) also have "... \ v" using 1 2 by simp also have "... \ v \ e\<^sup>T * top" by simp finally show ?thesis . qed text \ The graph after exchanging is acyclic. \ lemma exchange_acyclic: assumes "vector v" and "e \ v * -v\<^sup>T" and "w * v \ v" and "acyclic w" shows "acyclic (prim_W w v e)" proof - have 1: "(prim_P w v e)\<^sup>T * e = bot" proof - have "(prim_P w v e)\<^sup>T * e \ (-v * -v\<^sup>T)\<^sup>T * e" by (meson conv_order dual_order.trans inf.cobounded1 inf.cobounded2 mult_left_isotone) also have "... = -v * -v\<^sup>T * e" by (simp add: conv_complement conv_dist_comp) also have "... \ -v * -v\<^sup>T * v * -v\<^sup>T" by (simp add: assms(2) comp_associative mult_right_isotone) also have "... = bot" by (simp add: assms(1) covector_vector_comp mult_assoc) finally show ?thesis by (simp add: bot_unique) qed have 2: "e * e = bot" using assms(1,2) ee by auto have 3: "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T = bot" proof - have "top * (prim_P w v e) \ top * (-v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\)" using comp_inf.mult_semi_associative mult_right_isotone by auto also have "... \ top * -v * -v\<^sup>T \ top * top * e * w\<^sup>T\<^sup>\" by (simp add: comp_inf_covector mult_assoc) also have "... \ top * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" using mult_left_isotone top.extremum local.inf_mono by presburger also have "... = -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" by (simp add: assms(1) vector_conv_compl) finally have "top * (prim_P w v e) \ -(w \ -prim_EP w v e)" by (metis inf.assoc inf_import_p le_infI2 p_antitone p_antitone_iff) hence "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T \ bot" using p_top schroeder_4_p by blast thus ?thesis using bot_unique by blast qed hence 4: "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>\ = w \ -(prim_EP w v e)" using star_absorb by blast hence 5: "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\ = (w \ -(prim_EP w v e))\<^sup>+" by (metis star_plus mult_assoc) hence 6: "(w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ = (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>\" by (metis star.circ_loop_fixpoint mult_assoc) have 7: "(w \ -(prim_EP w v e))\<^sup>+ * e \ v * top" proof - have "e \ v * top" using assms(2) dual_order.trans mult_right_isotone top_greatest by blast hence 8: "e \ w * v * top \ v * top" by (simp add: assms(1,3) comp_associative) have "(w \ -(prim_EP w v e))\<^sup>+ * e \ w\<^sup>+ * e" by (simp add: comp_isotone star_isotone) also have "... \ w\<^sup>\ * e" by (simp add: mult_left_isotone star.left_plus_below_circ) also have "... \ v * top" using 8 by (simp add: comp_associative star_left_induct) finally show ?thesis . qed have 9: "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ * e = bot" proof - have "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ * e \ (prim_P w v e)\<^sup>T * v * top" using 7 by (simp add: mult_assoc mult_right_isotone) also have "... = bot" by (simp add: assms(1) pv) finally show ?thesis using bot_unique by blast qed have 10: "e * (w \ -(prim_EP w v e))\<^sup>+ * e = bot" proof - have "e * (w \ -(prim_EP w v e))\<^sup>+ * e \ e * v * top" using 7 by (simp add: mult_assoc mult_right_isotone) also have "... \ v * -v\<^sup>T * v * top" by (simp add: assms(2) mult_left_isotone) also have "... = bot" by (simp add: assms(1) covector_vector_comp mult_assoc) finally show ?thesis using bot_unique by blast qed have 11: "e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ v * -v\<^sup>T" proof - have 12: "-v\<^sup>T * w \ -v\<^sup>T" by (metis assms(3) conv_complement order_lesseq_imp pp_increasing schroeder_6_p) have "v * -v\<^sup>T * (w \ -(prim_EP w v e)) \ v * -v\<^sup>T * w" by (simp add: comp_isotone star_isotone) also have "... \ v * -v\<^sup>T" using 12 by (simp add: comp_isotone comp_associative) finally have 13: "v * -v\<^sup>T * (w \ -(prim_EP w v e)) \ v * -v\<^sup>T" . have 14: "(prim_P w v e)\<^sup>T \ -v * -v\<^sup>T" by (metis conv_complement conv_dist_comp conv_involutive conv_order inf_le1 inf_le2 order_trans) have "e * (prim_P w v e)\<^sup>T\<^sup>\ \ v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: assms(2) mult_left_isotone) also have "... = v * -v\<^sup>T \ v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>+" by (metis mult_assoc star.circ_back_loop_fixpoint star_plus sup_commute) also have "... = v * -v\<^sup>T \ v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>\ * (prim_P w v e)\<^sup>T" by (simp add: mult_assoc star_plus) also have "... \ v * -v\<^sup>T \ v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>\ * -v * -v\<^sup>T" using 14 mult_assoc mult_right_isotone sup_right_isotone by simp also have "... \ v * -v\<^sup>T \ v * top * -v\<^sup>T" by (metis top_greatest mult_right_isotone mult_left_isotone mult_assoc sup_right_isotone) also have "... = v * -v\<^sup>T" by (simp add: assms(1)) finally have "e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ v * -v\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_left_isotone) also have "... \ v * -v\<^sup>T" using 13 by (simp add: star_right_induct_mult) finally show ?thesis . qed have 15: "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ = (w \ -(prim_EP w v e))\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\" using 5 by simp also have "... = (w \ -(prim_EP w v e))\<^sup>+" by (simp add: mult_assoc star.circ_transitive_equal) also have "... \ w\<^sup>+" by (simp add: comp_isotone star_isotone) finally show ?thesis using assms(4) by simp qed have 16: "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>+ \ (w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: mult_right_isotone star.left_plus_below_circ) also have "... = (w \ -(prim_EP w v e))\<^sup>+" using 5 by simp also have "... \ w\<^sup>+" by (simp add: comp_isotone star_isotone) finally have "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>+ \ -1" using assms(4) by simp hence 17: "(prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>+ \ -1" by (simp add: comp_commute_below_diversity) have "(prim_P w v e)\<^sup>T\<^sup>+ \ w\<^sup>T\<^sup>+" by (simp add: comp_isotone conv_dist_inf inf.left_commute inf.sup_monoid.add_commute star_isotone) also have "... = w\<^sup>+\<^sup>T" by (simp add: conv_dist_comp conv_star_commute star_plus) also have "... \ -1" using assms(4) conv_complement conv_isotone by force finally have 18: "(prim_P w v e)\<^sup>T\<^sup>+ \ -1" . have "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ = (prim_P w v e)\<^sup>T * ((w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>\) * (w \ -(prim_EP w v e))\<^sup>\" using 6 by (simp add: comp_associative) also have "... = (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\ \ (prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_left_dist_sup mult_right_dist_sup) also have "... = (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_assoc star.circ_transitive_equal) also have "... = (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+ * (1 \ (w \ -(prim_EP w v e))\<^sup>+)" using star_left_unfold_equal by simp also have "... = (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+" by (simp add: mult_left_dist_sup sup.left_commute sup_commute) also have "... = ((prim_P w v e)\<^sup>T \ (prim_P w v e)\<^sup>T\<^sup>+) * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+" by (simp add: mult_right_dist_sup) also have "... = (prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+" using star.circ_mult_increasing by (simp add: le_iff_sup) also have "... \ -1" using 17 18 by simp finally show ?thesis . qed have 19: "e * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "e * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ = e * ((w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>\) * (w \ -(prim_EP w v e))\<^sup>\" using 6 by (simp add: mult_assoc) also have "... = e * (w \ -(prim_EP w v e))\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_left_dist_sup mult_right_dist_sup) also have "... = e * (w \ -(prim_EP w v e))\<^sup>+ \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_assoc star.circ_transitive_equal) also have "... \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>+ \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (metis mult_right_sub_dist_sup_right semiring.add_right_mono star.circ_back_loop_fixpoint) also have "... \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" using mult_right_isotone star.left_plus_below_circ by auto also have "... \ v * -v\<^sup>T" using 11 by simp also have "... \ -1" by (simp add: pp_increasing schroeder_3_p) finally show ?thesis . qed have 20: "(prim_W w v e) * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" using 15 16 19 by (simp add: comp_right_dist_sup) have 21: "(w \ -(prim_EP w v e))\<^sup>+ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "(w \ -(prim_EP w v e)) * v * -v\<^sup>T \ w * v * -v\<^sup>T" by (simp add: comp_isotone star_isotone) also have "... \ v * -v\<^sup>T" by (simp add: assms(3) mult_left_isotone) finally have 22: "(w \ -(prim_EP w v e)) * v * -v\<^sup>T \ v * -v\<^sup>T" . have "(w \ -(prim_EP w v e))\<^sup>+ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ (w \ -(prim_EP w v e))\<^sup>+ * v * -v\<^sup>T" using 11 by (simp add: mult_right_isotone mult_assoc) also have "... \ (w \ -(prim_EP w v e))\<^sup>\ * v * -v\<^sup>T" using mult_left_isotone star.left_plus_below_circ by blast also have "... \ v * -v\<^sup>T" using 22 by (simp add: star_left_induct_mult mult_assoc) also have "... \ -1" by (simp add: pp_increasing schroeder_3_p) finally show ?thesis . qed have 23: "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\ * e = (prim_P w v e)\<^sup>T * e \ (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ * e" using comp_left_dist_sup mult_assoc star.circ_loop_fixpoint sup_commute by auto also have "... = bot" using 1 9 by simp finally show ?thesis by simp qed have 24: "e * (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "e * (w \ -(prim_EP w v e))\<^sup>\ * e = e * e \ e * (w \ -(prim_EP w v e))\<^sup>+ * e" using comp_left_dist_sup mult_assoc star.circ_loop_fixpoint sup_commute by auto also have "... = bot" using 2 10 by simp finally show ?thesis by simp qed have 25: "(prim_W w v e) * (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" using 21 23 24 by (simp add: comp_right_dist_sup) have "(prim_W w v e)\<^sup>\ = ((prim_P w v e)\<^sup>T \ e)\<^sup>\ * ((w \ -(prim_EP w v e)) * ((prim_P w v e)\<^sup>T \ e)\<^sup>\)\<^sup>\" by (metis star_sup_1 sup.left_commute sup_commute) also have "... = ((prim_P w v e)\<^sup>T\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\) * ((w \ -(prim_EP w v e)) * ((prim_P w v e)\<^sup>T\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\))\<^sup>\" using 1 2 star_separate by auto also have "... = ((prim_P w v e)\<^sup>T\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\) * ((w \ -(prim_EP w v e)) * (1 \ e * (prim_P w v e)\<^sup>T\<^sup>\))\<^sup>\" using 4 mult_left_dist_sup by auto also have "... = (w \ -(prim_EP w v e))\<^sup>\ * ((prim_P w v e)\<^sup>T\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\) * (w \ -(prim_EP w v e))\<^sup>\" using 3 9 10 star_separate_2 by blast also have "... = (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: semiring.distrib_left semiring.distrib_right mult_assoc) finally have "(prim_W w v e)\<^sup>+ = (prim_W w v e) * ((w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\)" by simp also have "... = (prim_W w v e) * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ (prim_W w v e) * (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: comp_left_dist_sup comp_associative) also have "... \ -1" using 20 25 by simp finally show ?thesis . qed text \ The following lemma shows that an edge across the cut between visited nodes and unvisited nodes does not leave the component of visited nodes. \ lemma mst_subgraph_inv: assumes "e \ v * -v\<^sup>T \ g" and "t \ g" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" shows "e \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g" proof - have "e \ v * -v\<^sup>T \ g" by (rule assms(1)) also have "... \ v * (-v\<^sup>T \ v\<^sup>T * g) \ g" by (simp add: dedekind_1) also have "... \ v * v\<^sup>T * g \ g" by (simp add: comp_associative comp_right_isotone inf_commute le_infI2) also have "... = v * (r\<^sup>T * t\<^sup>\) * g \ g" by (simp add: assms(3)) also have "... = (r\<^sup>T * t\<^sup>\)\<^sup>T * (r\<^sup>T * t\<^sup>\) * g \ g" by (metis assms(3) conv_involutive) also have "... \ (r\<^sup>T * t\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) * g \ g" using assms(2) comp_inf.mult_left_isotone comp_isotone star_isotone by auto also have "... \ (r\<^sup>T * t\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g" using inf.sup_right_isotone inf_commute mult_assoc mult_right_isotone star.left_plus_below_circ star_plus by presburger also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g" using assms(2) comp_inf.mult_left_isotone conv_dist_comp conv_isotone mult_left_isotone star_isotone by auto finally show ?thesis . qed text \ The following lemmas show that the tree after exchanging contains the currently constructed and tree and its extension by the chosen edge. \ lemma mst_extends_old_tree: assumes "t \ w" and "t \ v * v\<^sup>T" and "vector v" shows "t \ prim_W w v e" proof - have "t \ prim_EP w v e \ t \ -v\<^sup>T" by (simp add: inf.coboundedI2 inf.sup_monoid.add_assoc) also have "... \ v * v\<^sup>T \ -v\<^sup>T" by (simp add: assms(2) inf.coboundedI1) also have "... \ bot" by (simp add: assms(3) covector_vector_comp eq_refl schroeder_2) finally have "t \ -(prim_EP w v e)" using le_bot pseudo_complement by blast hence "t \ w \ -(prim_EP w v e)" using assms(1) by simp thus ?thesis using local.le_supI1 by blast 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 end text \ We finally add the Kleene star to Stone relation algebras. Kleene star and the relational operations are reasonably independent. The only additional axiom we need in the generalisation to Stone-Kleene relation algebras is that star distributes over double complement. \ class stone_kleene_relation_algebra = stone_relation_algebra + pd_kleene_allegory + assumes pp_dist_star: "--(x\<^sup>\) = (--x)\<^sup>\" begin lemma regular_closed_star: "regular x \ regular (x\<^sup>\)" by (simp add: pp_dist_star) lemma components_idempotent: "components (components x) = components x" using pp_dist_star star_involutive by auto text \ The following lemma shows that the nodes reachable in the tree after exchange contain the nodes reachable in the tree before exchange. \ lemma mst_reachable_inv: assumes "regular (prim_EP w v e)" and "vector r" and "e \ v * -v\<^sup>T" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ w" and "t \ v * v\<^sup>T" and "w * v \ v" shows "r\<^sup>T * w\<^sup>\ \ r\<^sup>T * (prim_W w v e)\<^sup>\" proof - have 1: "r\<^sup>T \ r\<^sup>T * (prim_W w v e)\<^sup>\" using sup.bounded_iff star.circ_back_loop_prefixpoint by blast have "top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ * w\<^sup>T \ -v\<^sup>T = top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T)" by (simp add: assms(4) covector_comp_inf vector_conv_compl) also have "... \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\" by (simp add: comp_isotone mult_assoc star.circ_plus_same star.left_plus_below_circ) finally have 2: "top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ * w\<^sup>T \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" by (simp add: shunting_var_p) have 3: "--v\<^sup>T * w\<^sup>T \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" by (metis assms(8) conv_dist_comp conv_order mult_assoc order.trans pp_comp_semi_commute pp_isotone sup.coboundedI1 sup_commute) have 4: "top * e \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" using sup_right_divisibility star.circ_back_loop_fixpoint le_supI1 by blast have "(top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T) * w\<^sup>T = top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ * w\<^sup>T \ --v\<^sup>T * w\<^sup>T" by (simp add: comp_right_dist_sup) also have "... \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" using 2 3 by simp finally have "top * e \ (top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T) * w\<^sup>T \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" using 4 by simp hence 5: "top * e * w\<^sup>T\<^sup>\ \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" by (simp add: star_right_induct) have 6: "top * e \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\" using sup_right_divisibility star.circ_back_loop_fixpoint by blast have "(top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T \ (top * e * w\<^sup>T\<^sup>\)\<^sup>T" by (simp add: star_isotone mult_right_isotone conv_isotone inf_assoc) also have "... = w\<^sup>\ * e\<^sup>T * top" by (simp add: conv_dist_comp conv_star_commute mult_assoc) finally have 7: "(top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T \ w\<^sup>\ * e\<^sup>T * top" . have "(top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T \ (top * e * (-v * -v\<^sup>T)\<^sup>\)\<^sup>T" by (simp add: conv_isotone inf_commute mult_right_isotone star_isotone le_infI2) also have "... \ (top * v * -v\<^sup>T * (-v * -v\<^sup>T)\<^sup>\)\<^sup>T" by (metis assms(3) conv_isotone mult_left_isotone mult_right_isotone mult_assoc) also have "... = (top * v * (-v\<^sup>T * -v)\<^sup>\ * -v\<^sup>T)\<^sup>T" by (simp add: mult_assoc star_slide) also have "... \ (top * -v\<^sup>T)\<^sup>T" using conv_order mult_left_isotone by auto also have "... = -v" by (simp add: assms(4) conv_complement vector_conv_compl) finally have 8: "(top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T \ w\<^sup>\ * e\<^sup>T * top \ -v" using 7 by simp have "covector (top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)" by (simp add: covector_mult_closed) hence "top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T) = top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T \ (top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T)" by (metis comp_inf_vector_1 inf.idem) also have "... \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top \ -v)" using 8 mult_right_isotone inf.sup_right_isotone inf_assoc by simp also have "... = top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ (-v \ -v\<^sup>T) \ w\<^sup>\ * e\<^sup>T * top)" using inf_assoc inf_commute by (simp add: inf_assoc) also have "... = top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)" using assms(4) conv_complement vector_complement_closed vector_covector by fastforce also have "... \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\" by (simp add: comp_associative comp_isotone star.circ_plus_same star.left_plus_below_circ) finally have 9: "top * e \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T) \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\" using 6 by simp have "prim_EP w v e \ -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" using inf.sup_left_isotone by auto also have "... \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\" using 5 by (metis inf_commute shunting_var_p) also have "... \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\" using 9 by (simp add: star_right_induct) finally have 10: "prim_EP w v e \ top * e * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: conv_complement conv_dist_comp conv_dist_inf conv_star_commute mult_assoc) have "top * e = top * (v * -v\<^sup>T \ e)" by (simp add: assms(3) inf.absorb2) also have "... \ top * (v * top \ e)" using inf.sup_right_isotone inf_commute mult_right_isotone top_greatest by presburger also have "... = (top \ (v * top)\<^sup>T) * e" using assms(4) covector_inf_comp_3 by presburger also have "... = top * v\<^sup>T * e" by (simp add: conv_dist_comp) also have "... = top * r\<^sup>T * t\<^sup>\ * e" by (simp add: assms(5) comp_associative) also have "... \ top * r\<^sup>T * (prim_W w v e)\<^sup>\ * e" by (metis assms(4,6,7) mst_extends_old_tree star_isotone mult_left_isotone mult_right_isotone) finally have 11: "top * e \ top * r\<^sup>T * (prim_W w v e)\<^sup>\ * e" . have "r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e) \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (top * e * (prim_P w v e)\<^sup>T\<^sup>\)" using 10 mult_right_isotone by blast also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * top * e * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: mult_assoc) also have "... \ top * e * (prim_P w v e)\<^sup>T\<^sup>\" by (metis comp_associative comp_inf_covector inf.idem inf.sup_right_divisibility) also have "... \ top * r\<^sup>T * (prim_W w v e)\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\" using 11 by (simp add: mult_left_isotone) also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\" using assms(2) vector_conv_covector by auto also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_W w v e) * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: mult_left_isotone mult_right_isotone) also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_W w v e) * (prim_W w v e)\<^sup>\" by (meson dual_order.trans mult_right_isotone star_isotone sup_ge1 sup_ge2) also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\" by (metis mult_assoc mult_right_isotone star.circ_transitive_equal star.left_plus_below_circ) finally have 12: "r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e) \ r\<^sup>T * (prim_W w v e)\<^sup>\" . have "r\<^sup>T * (prim_W w v e)\<^sup>\ * w \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (w \ prim_EP w v e)" by (simp add: inf_assoc) also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * ((w \ prim_EP w v e) \ (-(prim_EP w v e) \ prim_EP w v e))" by (metis assms(1) inf_top_right stone) also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * ((w \ -(prim_EP w v e)) \ prim_EP w v e)" by (simp add: sup_inf_distrib2) also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * (w \ -(prim_EP w v e)) \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e)" by (simp add: comp_left_dist_sup) also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_W w v e) \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e)" using mult_right_isotone sup_left_isotone by auto also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\ \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e)" using mult_assoc mult_right_isotone star.circ_plus_same star.left_plus_below_circ sup_left_isotone by auto also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\" using 12 sup.absorb1 by blast finally have "r\<^sup>T \ r\<^sup>T * (prim_W w v e)\<^sup>\ * w \ r\<^sup>T * (prim_W w v e)\<^sup>\" using 1 by simp thus ?thesis by (simp add: star_right_induct) qed text \ Some of the following lemmas already hold in pseudocomplemented distributive Kleene allegories. \ subsubsection \Exchange gives Minimum Spanning Trees\ text \ The lemmas in this section are used to show that the after exchange we obtain a minimum spanning tree. The following lemmas show various interactions between the three constituents of the tree after exchange. \ lemma epm_1: "vector v \ prim_E w v e \ prim_P w v e = prim_EP w v e" by (metis inf_commute inf_sup_distrib1 mult_assoc mult_right_dist_sup regular_closed_p regular_complement_top vector_conv_compl) lemma epm_2: assumes "regular (prim_EP w v e)" and "vector v" shows "(w \ -(prim_EP w v e)) \ prim_P w v e \ prim_E w v e = w" proof - have "(w \ -(prim_EP w v e)) \ prim_P w v e \ prim_E w v e = (w \ -(prim_EP w v e)) \ prim_EP w v e" using epm_1 sup_assoc sup_commute assms(2) by (simp add: inf_sup_distrib1) also have "... = w \ prim_EP w v e" by (metis assms(1) inf_top.right_neutral regular_complement_top sup_inf_distrib2) also have "... = w" by (simp add: sup_inf_distrib1) finally show ?thesis . qed lemma epm_4: assumes "e \ w" and "injective w" and "w * v \ v" and "e \ v * -v\<^sup>T" shows "top * e * w\<^sup>T\<^sup>+ \ top * v\<^sup>T" proof - have "w\<^sup>\ * v \ v" by (simp add: assms(3) star_left_induct_mult) hence 1: "v\<^sup>T * w\<^sup>T\<^sup>\ \ v\<^sup>T" using conv_star_commute conv_dist_comp conv_isotone by fastforce have "e * w\<^sup>T \ w * w\<^sup>T \ e * w\<^sup>T" by (simp add: assms(1) mult_left_isotone) also have "... \ 1 \ e * w\<^sup>T" using assms(2) inf.sup_left_isotone by auto also have "... = 1 \ w * e\<^sup>T" using calculation conv_dist_comp conv_involutive coreflexive_symmetric by fastforce also have "... \ w * e\<^sup>T" by simp also have "... \ w * -v * v\<^sup>T" by (metis assms(4) conv_complement conv_dist_comp conv_involutive conv_order mult_assoc mult_right_isotone) also have "... \ top * v\<^sup>T" by (simp add: mult_left_isotone) finally have "top * e * w\<^sup>T\<^sup>+ \ top * v\<^sup>T * w\<^sup>T\<^sup>\" by (metis antisym comp_associative comp_isotone dense_top_closed mult_left_isotone transitive_top_closed) also have "... \ top * v\<^sup>T" using 1 by (simp add: mult_assoc mult_right_isotone) finally show ?thesis . qed lemma epm_5: assumes "e \ w" and "injective w" and "w * v \ v" and "e \ v * -v\<^sup>T" and "vector v" shows "prim_P w v e = bot" proof - have 1: "e = w \ top * e" by (simp add: assms(1,2) epm_3) have 2: "top * e * w\<^sup>T\<^sup>+ \ top * v\<^sup>T" by (simp add: assms(1-4) epm_4) have 3: "-v * -v\<^sup>T \ top * v\<^sup>T = bot" by (simp add: assms(5) comp_associative covector_vector_comp inf.sup_monoid.add_commute schroeder_2) have "prim_P w v e = (w \ -v * -v\<^sup>T \ top * e) \ (w \ -v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>+)" by (metis inf_sup_distrib1 mult_assoc star.circ_back_loop_fixpoint star_plus sup_commute) also have "... \ (e \ -v * -v\<^sup>T) \ (w \ -v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>+)" using 1 by (metis comp_inf.mult_semi_associative inf.sup_monoid.add_commute semiring.add_right_mono) also have "... \ (e \ -v * -v\<^sup>T) \ (w \ -v * -v\<^sup>T \ top * v\<^sup>T)" using 2 by (metis sup_right_isotone inf.sup_right_isotone) also have "... \ (e \ -v * -v\<^sup>T) \ (-v * -v\<^sup>T \ top * v\<^sup>T)" using inf.assoc le_infI2 by auto also have "... \ v * -v\<^sup>T \ -v * -v\<^sup>T" using 3 assms(4) inf.sup_left_isotone by auto also have "... \ v * top \ -v * top" using inf.sup_mono mult_right_isotone top_greatest by blast also have "... = bot" using assms(5) inf_compl_bot vector_complement_closed by auto finally show ?thesis by (simp add: le_iff_inf) qed lemma epm_6: assumes "e \ w" and "injective w" and "w * v \ v" and "e \ v * -v\<^sup>T" and "vector v" shows "prim_E w v e = e" proof - have 1: "e \ --v * -v\<^sup>T" using assms(4) mult_isotone order_lesseq_imp pp_increasing by blast have 2: "top * e * w\<^sup>T\<^sup>+ \ top * v\<^sup>T" by (simp add: assms(1-4) epm_4) have 3: "e = w \ top * e" by (simp add: assms(1,2) epm_3) hence "e \ top * e * w\<^sup>T\<^sup>\" by (metis le_infI2 star.circ_back_loop_fixpoint sup.commute sup_ge1) hence 4: "e \ prim_E w v e" using 1 by (simp add: assms(1)) have 5: "--v * -v\<^sup>T \ top * v\<^sup>T = bot" by (simp add: assms(5) comp_associative covector_vector_comp inf.sup_monoid.add_commute schroeder_2) have "prim_E w v e = (w \ --v * -v\<^sup>T \ top * e) \ (w \ --v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>+)" by (metis inf_sup_distrib1 mult_assoc star.circ_back_loop_fixpoint star_plus sup_commute) also have "... \ (e \ --v * -v\<^sup>T) \ (w \ --v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>+)" using 3 by (metis comp_inf.mult_semi_associative inf.sup_monoid.add_commute semiring.add_right_mono) also have "... \ (e \ --v * -v\<^sup>T) \ (w \ --v * -v\<^sup>T \ top * v\<^sup>T)" using 2 by (metis sup_right_isotone inf.sup_right_isotone) also have "... \ (e \ --v * -v\<^sup>T) \ (--v * -v\<^sup>T \ top * v\<^sup>T)" using inf.assoc le_infI2 by auto also have "... \ e" by (simp add: "5") finally show ?thesis using 4 by (simp add: antisym) qed lemma epm_7: "regular (prim_EP w v e) \ e \ w \ injective w \ w * v \ v \ e \ v * -v\<^sup>T \ vector v \ prim_W w v e = w" by (metis conv_bot epm_2 epm_5 epm_6) lemma epm_8: assumes "acyclic w" shows "(w \ -(prim_EP w v e)) \ (prim_P w v e)\<^sup>T = bot" proof - have "(w \ -(prim_EP w v e)) \ (prim_P w v e)\<^sup>T \ w \ w\<^sup>T" by (meson conv_isotone inf_le1 inf_mono order_trans) thus ?thesis by (metis assms acyclic_asymmetric inf.commute le_bot) qed lemma epm_9: assumes "e \ v * -v\<^sup>T" and "vector v" shows "(w \ -(prim_EP w v e)) \ e = bot" proof - have 1: "e \ -v\<^sup>T" by (metis assms complement_conv_sub vector_conv_covector ev p_antitone_iff p_bot) have "(w \ -(prim_EP w v e)) \ e = (w \ --v\<^sup>T \ e) \ (w \ -(top * e * w\<^sup>T\<^sup>\) \ e)" by (simp add: inf_commute inf_sup_distrib1) also have "... \ (--v\<^sup>T \ e) \ (-(top * e * w\<^sup>T\<^sup>\) \ e)" using comp_inf.mult_left_isotone inf.cobounded2 semiring.add_mono by blast also have "... = -(top * e * w\<^sup>T\<^sup>\) \ e" using 1 by (metis inf.sup_relative_same_increasing inf_commute inf_sup_distrib1 maddux_3_13 regular_closed_p) also have "... = bot" by (metis inf.sup_relative_same_increasing inf_bot_right inf_commute inf_p mult_left_isotone star_outer_increasing top_greatest) finally show ?thesis by (simp add: le_iff_inf) qed lemma epm_10: assumes "e \ v * -v\<^sup>T" and "vector v" shows "(prim_P w v e)\<^sup>T \ e = bot" proof - have "(prim_P w v e)\<^sup>T \ -v * -v\<^sup>T" by (simp add: conv_complement conv_dist_comp conv_dist_inf inf.absorb_iff1 inf.left_commute inf_commute) hence "(prim_P w v e)\<^sup>T \ e \ -v * -v\<^sup>T \ v * -v\<^sup>T" using assms(1) inf_mono by blast also have "... \ -v * top \ v * top" using inf.sup_mono mult_right_isotone top_greatest by blast also have "... = bot" using assms(2) inf_compl_bot vector_complement_closed by auto finally show ?thesis by (simp add: le_iff_inf) qed lemma epm_11: assumes "vector v" shows "(w \ -(prim_EP w v e)) \ prim_P w v e = bot" proof - have "prim_P w v e \ prim_EP w v e" by (metis assms comp_isotone inf.sup_left_isotone inf.sup_right_isotone order.refl top_greatest vector_conv_compl) thus ?thesis using inf_le2 order_trans p_antitone pseudo_complement by blast qed lemma epm_12: assumes "vector v" shows "(w \ -(prim_EP w v e)) \ prim_E w v e = bot" proof - have "prim_E w v e \ prim_EP w v e" by (metis assms comp_isotone inf.sup_left_isotone inf.sup_right_isotone order.refl top_greatest vector_conv_compl) thus ?thesis using inf_le2 order_trans p_antitone pseudo_complement by blast qed lemma epm_13: assumes "vector v" shows "prim_P w v e \ prim_E w v e = bot" proof - have "prim_P w v e \ prim_E w v e \ -v * -v\<^sup>T \ --v * -v\<^sup>T" by (meson dual_order.trans inf.cobounded1 inf.sup_mono inf_le2) also have "... \ -v * top \ --v * top" using inf.sup_mono mult_right_isotone top_greatest by blast also have "... = bot" using assms inf_compl_bot vector_complement_closed by auto finally show ?thesis by (simp add: le_iff_inf) qed text \ The following lemmas show that the relation characterising the edge across the cut is an arc. \ lemma arc_edge_1: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" shows "top * e \ v\<^sup>T * w\<^sup>\" proof - have "top * e \ top * (v * -v\<^sup>T \ g)" using assms(1) mult_right_isotone by auto also have "... \ top * (v * top \ g)" using inf.sup_right_isotone inf_commute mult_right_isotone top_greatest by presburger also have "... = v\<^sup>T * g" by (metis assms(2) covector_inf_comp_3 inf_top.left_neutral) also have "... = r\<^sup>T * t\<^sup>\ * g" by (simp add: assms(3)) also have "... \ r\<^sup>T * g\<^sup>\ * g" by (simp add: assms(4) mult_left_isotone mult_right_isotone star_isotone) also have "... \ r\<^sup>T * g\<^sup>\" by (simp add: mult_assoc mult_right_isotone star.right_plus_below_circ) also have "... \ r\<^sup>T * w\<^sup>\" by (simp add: assms(5)) also have "... \ v\<^sup>T * w\<^sup>\" by (metis assms(3) mult_left_isotone mult_right_isotone mult_1_right star.circ_reflexive) finally show ?thesis . qed lemma arc_edge_2: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" and "w * v \ v" and "injective w" shows "top * e * w\<^sup>T\<^sup>\ \ v\<^sup>T * w\<^sup>\" proof - have 1: "top * e \ v\<^sup>T * w\<^sup>\" using assms(1-5) arc_edge_1 by blast have "v\<^sup>T * w\<^sup>\ * w\<^sup>T = v\<^sup>T * w\<^sup>T \ v\<^sup>T * w\<^sup>+ * w\<^sup>T" by (metis mult_assoc mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... \ v\<^sup>T \ v\<^sup>T * w\<^sup>+ * w\<^sup>T" by (metis assms(6) conv_dist_comp conv_isotone sup_left_isotone) also have "... = v\<^sup>T \ v\<^sup>T * w\<^sup>\ * (w * w\<^sup>T)" by (metis mult_assoc star_plus) also have "... \ v\<^sup>T \ v\<^sup>T * w\<^sup>\" by (metis assms(7) mult_right_isotone mult_1_right sup_right_isotone) also have "... = v\<^sup>T * w\<^sup>\" by (metis star.circ_back_loop_fixpoint sup_absorb2 sup_ge2) finally show ?thesis using 1 star_right_induct by auto qed lemma arc_edge_3: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" and "w * v \ v" and "injective w" and "prim_E w v e = bot" shows "e = bot" proof - have "bot = prim_E w v e" by (simp add: assms(8)) also have "... = w \ --v * top \ top * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" by (metis assms(2) comp_inf_covector inf.assoc inf_top.left_neutral vector_conv_compl) also have "... = w \ top * e * w\<^sup>T\<^sup>\ \ -v\<^sup>T \ --v" using assms(2) inf.assoc inf.commute vector_conv_compl vector_complement_closed by (simp add: inf_assoc) finally have 1: "w \ top * e * w\<^sup>T\<^sup>\ \ -v\<^sup>T \ -v" using shunting_1_pp by force have "w\<^sup>\ * e\<^sup>T * top = (top * e * w\<^sup>T\<^sup>\)\<^sup>T" by (simp add: conv_star_commute comp_associative conv_dist_comp) also have "... \ (v\<^sup>T * w\<^sup>\)\<^sup>T" using assms(1-7) arc_edge_2 by (simp add: conv_isotone) also have "... = w\<^sup>T\<^sup>\ * v" by (simp add: conv_star_commute conv_dist_comp) finally have 2: "w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T\<^sup>\ * v" . have "(w\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>T * -v = (w \ top * e * w\<^sup>T\<^sup>\) * -v" by (simp add: conv_dist_comp conv_dist_inf conv_star_commute mult_assoc) also have "... = (w \ top * e * w\<^sup>T\<^sup>\ \ -v\<^sup>T) * top" by (metis assms(2) conv_complement covector_inf_comp_3 inf_top.right_neutral vector_complement_closed) also have "... \ -v * top" using 1 by (simp add: comp_isotone) also have "... = -v" using assms(2) vector_complement_closed by auto finally have "(w\<^sup>T \ w\<^sup>\ * e\<^sup>T * top) * --v \ --v" using p_antitone_iff schroeder_3_p by auto hence "w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T * --v \ --v" by (simp add: inf_vector_comp) hence 3: "w\<^sup>T * --v \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" by (simp add: inf.commute shunting_p) have "w\<^sup>T * -(w\<^sup>\ * e\<^sup>T * top) \ -(w\<^sup>\ * e\<^sup>T * top)" by (metis mult_assoc p_antitone p_antitone_iff schroeder_3_p star.circ_loop_fixpoint sup_commute sup_right_divisibility) also have "... \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" by simp finally have "w\<^sup>T * (--v \ -(w\<^sup>\ * e\<^sup>T * top)) \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" using 3 by (simp add: mult_left_dist_sup) hence "w\<^sup>T\<^sup>\ * (--v \ -(w\<^sup>\ * e\<^sup>T * top)) \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" using star_left_induct_mult_iff by blast hence "w\<^sup>T\<^sup>\ * --v \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" by (simp add: semiring.distrib_left) hence "w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T\<^sup>\ * --v \ --v" by (simp add: inf_commute shunting_p) hence "w\<^sup>\ * e\<^sup>T * top \ --v" using 2 by (metis inf.absorb1 p_antitone_iff p_comp_pp vector_export_comp) hence 4: "e\<^sup>T * top \ --v" by (metis mult_assoc star.circ_loop_fixpoint sup.bounded_iff) have "e\<^sup>T * top \ (v * -v\<^sup>T)\<^sup>T * top" using assms(1) comp_isotone conv_isotone by auto also have "... \ -v * top" by (simp add: conv_complement conv_dist_comp mult_assoc mult_right_isotone) also have "... = -v" using assms(2) vector_complement_closed by auto finally have "e\<^sup>T * top \ bot" using 4 shunting_1_pp by auto hence "e\<^sup>T = bot" using antisym bot_least top_right_mult_increasing by blast thus ?thesis using conv_bot by fastforce qed lemma arc_edge_4: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" and "arc e" shows "top * prim_E w v e * top = top" proof - have "--v\<^sup>T * w = (--v\<^sup>T * w \ -v\<^sup>T) \ (--v\<^sup>T * w \ --v\<^sup>T)" by (simp add: maddux_3_11_pp) also have "... \ (--v\<^sup>T * w \ -v\<^sup>T) \ --v\<^sup>T" using sup_right_isotone by auto also have "... = --v\<^sup>T * (w \ -v\<^sup>T) \ --v\<^sup>T" using assms(2) covector_comp_inf covector_complement_closed vector_conv_covector by auto also have "... \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" by (metis star.circ_back_loop_fixpoint sup.cobounded2 sup_left_isotone) finally have 1: "--v\<^sup>T * w \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" . have "--v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ * w \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" by (simp add: le_supI1 mult_assoc mult_right_isotone star.circ_plus_same star.left_plus_below_circ) hence 2: "(--v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T) * w \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" using 1 by (simp add: inf.orderE mult_right_dist_sup) have "v\<^sup>T \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" by (simp add: pp_increasing sup.coboundedI2) hence "v\<^sup>T * w\<^sup>\ \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" using 2 by (simp add: star_right_induct) hence 3: "-v\<^sup>T \ v\<^sup>T * w\<^sup>\ \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\" by (metis inf_commute shunting_var_p) have "top * e = top * e \ v\<^sup>T * w\<^sup>\" by (meson assms(1-5) arc_edge_1 inf.orderE) also have "... \ top * v * -v\<^sup>T \ v\<^sup>T * w\<^sup>\" using assms(1) inf.sup_left_isotone mult_assoc mult_right_isotone by auto also have "... \ top * -v\<^sup>T \ v\<^sup>T * w\<^sup>\" using inf.sup_left_isotone mult_left_isotone top_greatest by blast also have "... = -v\<^sup>T \ v\<^sup>T * w\<^sup>\" by (simp add: assms(2) vector_conv_compl) also have "... \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\" using 3 by simp also have "... = (top \ (--v)\<^sup>T) * (w \ -v\<^sup>T) * w\<^sup>\" by (simp add: conv_complement) also have "... = top * (w \ --v \ -v\<^sup>T) * w\<^sup>\" using assms(2) covector_inf_comp_3 inf_assoc inf_left_commute vector_complement_closed by presburger also have "... = top * (w \ --v * -v\<^sup>T) * w\<^sup>\" by (metis assms(2) vector_complement_closed conv_complement inf_assoc vector_covector) finally have "top * (e\<^sup>T * top)\<^sup>T \ top * (w \ --v * -v\<^sup>T) * w\<^sup>\" by (metis conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top) hence "top \ top * (w \ --v * -v\<^sup>T) * w\<^sup>\ * (e\<^sup>T * top)" using assms(6) shunt_bijective by blast also have "... = top * (w \ --v * -v\<^sup>T) * (top * e * w\<^sup>\\<^sup>T)\<^sup>T" by (simp add: conv_dist_comp mult_assoc) also have "... = top * (w \ --v * -v\<^sup>T \ top * e * w\<^sup>\\<^sup>T) * top" by (simp add: comp_inf_vector_1 mult_assoc) finally show ?thesis by (simp add: conv_star_commute top_le) qed lemma arc_edge_5: assumes "vector v" and "w * v \ v" and "injective w" and "arc e" shows "(prim_E w v e)\<^sup>T * top * prim_E w v e \ 1" proof - have 1: "e\<^sup>T * top * e \ 1" by (simp add: assms(4) point_injective) have "prim_E w v e \ --v * top" by (simp add: inf_commute le_infI2 mult_right_isotone) hence 2: "prim_E w v e \ --v" by (simp add: assms(1) vector_complement_closed) have 3: "w * --v \ --v" by (simp add: assms(2) p_antitone p_antitone_iff) have "w \ top * prim_E w v e \ w * (prim_E w v e)\<^sup>T * prim_E w v e" by (metis dedekind_2 inf.commute inf_top.left_neutral) also have "... \ w * w\<^sup>T * prim_E w v e" by (simp add: conv_isotone le_infI1 mult_left_isotone mult_right_isotone) also have "... \ prim_E w v e" by (metis assms(3) mult_left_isotone mult_left_one) finally have 4: "w \ top * prim_E w v e \ prim_E w v e" . have "w\<^sup>+ \ top * prim_E w v e = w\<^sup>\ * (w \ top * prim_E w v e)" by (simp add: comp_inf_covector star_plus) also have "... \ w\<^sup>\ * prim_E w v e" using 4 by (simp add: mult_right_isotone) also have "... \ --v" using 2 3 star_left_induct sup.bounded_iff by blast finally have 5: "w\<^sup>+ \ top * prim_E w v e \ -v = bot" using shunting_1_pp by blast hence 6: "w\<^sup>+\<^sup>T \ (prim_E w v e)\<^sup>T * top \ -v\<^sup>T = bot" using conv_complement conv_dist_comp conv_dist_inf conv_top conv_bot by force have "(prim_E w v e)\<^sup>T * top * prim_E w v e \ (top * e * w\<^sup>T\<^sup>\)\<^sup>T * top * (top * e * w\<^sup>T\<^sup>\)" by (simp add: conv_isotone mult_isotone) also have "... = w\<^sup>\ * e\<^sup>T * top * e * w\<^sup>T\<^sup>\" by (metis conv_star_commute conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top) also have "... \ w\<^sup>\ * w\<^sup>T\<^sup>\" using 1 by (metis mult_assoc mult_1_right mult_right_isotone mult_left_isotone) also have "... = w\<^sup>\ \ w\<^sup>T\<^sup>\" by (metis assms(3) cancel_separate inf.eq_iff star.circ_sup_sub_sup_one_1 star.circ_plus_one star_involutive) also have "... = w\<^sup>+ \ w\<^sup>T\<^sup>+ \ 1" by (metis star.circ_plus_one star_left_unfold_equal sup.assoc sup.commute) finally have 7: "(prim_E w v e)\<^sup>T * top * prim_E w v e \ w\<^sup>+ \ w\<^sup>T\<^sup>+ \ 1" . have "prim_E w v e \ --v * -v\<^sup>T" by (simp add: le_infI1) also have "... \ top * -v\<^sup>T" by (simp add: mult_left_isotone) also have "... = -v\<^sup>T" by (simp add: assms(1) vector_conv_compl) finally have 8: "prim_E w v e \ -v\<^sup>T" . hence 9: "(prim_E w v e)\<^sup>T \ -v" by (metis conv_complement conv_involutive conv_isotone) have "(prim_E w v e)\<^sup>T * top * prim_E w v e = (w\<^sup>+ \ w\<^sup>T\<^sup>+ \ 1) \ (prim_E w v e)\<^sup>T * top * prim_E w v e" using 7 by (simp add: inf.absorb_iff2) also have "... = (1 \ (prim_E w v e)\<^sup>T * top * prim_E w v e) \ (w\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e)" using comp_inf.mult_right_dist_sup sup_assoc sup_commute by auto also have "... \ 1 \ (w\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e)" using inf_le1 sup_left_isotone by blast also have "... \ 1 \ (w\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top * -v\<^sup>T)" using 8 inf.sup_right_isotone mult_right_isotone sup_right_isotone by blast also have "... \ 1 \ (w\<^sup>+ \ -v * top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top * -v\<^sup>T)" using 9 by (metis inf.sup_right_isotone mult_left_isotone sup.commute sup_right_isotone) also have "... = 1 \ (w\<^sup>+ \ -v * top \ top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top \ top * -v\<^sup>T)" by (metis (no_types) vector_export_comp inf_top_right inf_assoc) also have "... = 1 \ (w\<^sup>+ \ -v \ top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top \ -v\<^sup>T)" using assms(1) vector_complement_closed vector_conv_compl by auto also have "... = 1" using 5 6 by (simp add: conv_star_commute conv_dist_comp inf.commute inf_assoc star.circ_plus_same) finally show ?thesis . qed lemma arc_edge_6: assumes "vector v" and "w * v \ v" and "injective w" and "arc e" shows "prim_E w v e * top * (prim_E w v e)\<^sup>T \ 1" proof - have "prim_E w v e * 1 * (prim_E w v e)\<^sup>T \ w * w\<^sup>T" using comp_isotone conv_order inf.coboundedI1 mult_one_associative by auto also have "... \ 1" by (simp add: assms(3)) finally have 1: "prim_E w v e * 1 * (prim_E w v e)\<^sup>T \ 1" . have "(prim_E w v e)\<^sup>T * top * prim_E w v e \ 1" by (simp add: assms arc_edge_5) also have "... \ --1" by (simp add: pp_increasing) finally have 2: "prim_E w v e * -1 * (prim_E w v e)\<^sup>T \ bot" by (metis conv_involutive regular_closed_bot regular_dense_top triple_schroeder_p) have "prim_E w v e * top * (prim_E w v e)\<^sup>T = prim_E w v e * 1 * (prim_E w v e)\<^sup>T \ prim_E w v e * -1 * (prim_E w v e)\<^sup>T" by (metis mult_left_dist_sup mult_right_dist_sup regular_complement_top regular_one_closed) also have "... \ 1" using 1 2 by (simp add: bot_unique) finally show ?thesis . qed lemma arc_edge: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" and "w * v \ v" and "injective w" and "arc e" shows "arc (prim_E w v e)" proof (intro conjI) have "prim_E w v e * top * (prim_E w v e)\<^sup>T \ 1" using assms(2,6-8) arc_edge_6 by simp thus "injective (prim_E w v e * top)" by (metis conv_dist_comp conv_top mult_assoc top_mult_top) next show "surjective (prim_E w v e * top)" using assms(1-5,8) arc_edge_4 mult_assoc by simp next have "(prim_E w v e)\<^sup>T * top * prim_E w v e \ 1" using assms(2,6-8) arc_edge_5 by simp thus "injective ((prim_E w v e)\<^sup>T * top)" by (metis conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top) next have "top * prim_E w v e * top = top" using assms(1-5,8) arc_edge_4 by simp thus "surjective ((prim_E w v e)\<^sup>T * top)" by (metis mult_assoc conv_dist_comp conv_top) qed subsubsection \Invariant implies Postcondition\ text \ The lemmas in this section are used to show that the invariant implies the postcondition at the end of the algorithm. The following lemma shows that the nodes reachable in the graph are the same as those reachable in the constructed tree. \ lemma span_post: assumes "regular v" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "v * -v\<^sup>T \ g = bot" and "t \ v * v\<^sup>T \ g" and "r\<^sup>T * (v * v\<^sup>T \ g)\<^sup>\ \ r\<^sup>T * t\<^sup>\" shows "v\<^sup>T = r\<^sup>T * g\<^sup>\" proof - let ?vv = "v * v\<^sup>T \ g" have 1: "r\<^sup>T \ v\<^sup>T" using assms(3) mult_right_isotone mult_1_right star.circ_reflexive by fastforce have "v * top \ g = (v * v\<^sup>T \ v * -v\<^sup>T) \ g" by (metis assms(1) conv_complement mult_left_dist_sup regular_complement_top) also have "... = ?vv \ (v * -v\<^sup>T \ g)" by (simp add: inf_sup_distrib2) also have "... = ?vv" by (simp add: assms(4)) finally have 2: "v * top \ g = ?vv" by simp have "r\<^sup>T * ?vv\<^sup>\ \ v\<^sup>T * ?vv\<^sup>\" using 1 by (simp add: comp_left_isotone) also have "... \ v\<^sup>T * (v * v\<^sup>T)\<^sup>\" by (simp add: comp_right_isotone star.circ_isotone) also have "... \ v\<^sup>T" by (simp add: assms(2) vector_star_1) finally have "r\<^sup>T * ?vv\<^sup>\ \ v\<^sup>T" by simp hence "r\<^sup>T * ?vv\<^sup>\ * g = (r\<^sup>T * ?vv\<^sup>\ \ v\<^sup>T) * g" by (simp add: inf.absorb1) also have "... = r\<^sup>T * ?vv\<^sup>\ * (v * top \ g)" by (simp add: assms(2) covector_inf_comp_3) also have "... = r\<^sup>T * ?vv\<^sup>\ * ?vv" using 2 by simp also have "... \ r\<^sup>T * ?vv\<^sup>\" by (simp add: comp_associative comp_right_isotone star.left_plus_below_circ star_plus) finally have "r\<^sup>T \ r\<^sup>T * ?vv\<^sup>\ * g \ r\<^sup>T * ?vv\<^sup>\" using star.circ_back_loop_prefixpoint by auto hence "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * ?vv\<^sup>\" using star_right_induct by blast hence "r\<^sup>T * g\<^sup>\ = r\<^sup>T * ?vv\<^sup>\" by (simp add: antisym mult_right_isotone star_isotone) also have "... = r\<^sup>T * t\<^sup>\" using assms(5,6) antisym mult_right_isotone star_isotone by auto also have "... = v\<^sup>T" by (simp add: assms(3)) finally show ?thesis by simp qed text \ The following lemma shows that the minimum spanning tree extending a tree is the same as the tree at the end of the algorithm. \ lemma mst_post: assumes "vector r" and "injective r" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "forest w" and "t \ w" and "w \ v * v\<^sup>T" shows "w = t" proof - have 1: "vector v" using assms(1,3) covector_mult_closed vector_conv_covector by auto have "w * v \ v * v\<^sup>T * v" by (simp add: assms(6) mult_left_isotone) also have "... \ v" using 1 by (metis mult_assoc mult_right_isotone top_greatest) finally have 2: "w * v \ v" . have 3: "r \ v" by (metis assms(3) conv_order mult_right_isotone mult_1_right star.circ_reflexive) have 4: "v \ -r = t\<^sup>T\<^sup>\ * r \ -r" by (metis assms(3) conv_dist_comp conv_involutive conv_star_commute) also have "... = (r \ t\<^sup>T\<^sup>+ * r) \ -r" using mult_assoc star.circ_loop_fixpoint sup_commute by auto also have "... \ t\<^sup>T\<^sup>+ * r" by (simp add: shunting) also have "... \ t\<^sup>T * top" by (simp add: comp_isotone mult_assoc) finally have "1 \ (v \ -r) * (v \ -r)\<^sup>T \ 1 \ t\<^sup>T * top * (t\<^sup>T * top)\<^sup>T" using conv_order inf.sup_right_isotone mult_isotone by auto also have "... = 1 \ t\<^sup>T * top * t" by (metis conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top) also have "... \ t\<^sup>T * (top * t \ t * 1)" by (metis conv_involutive dedekind_1 inf.commute mult_assoc) also have "... \ t\<^sup>T * t" by (simp add: mult_right_isotone) finally have 5: "1 \ (v \ -r) * (v \ -r)\<^sup>T \ t\<^sup>T * t" . have "w * w\<^sup>+ \ -1" by (metis assms(4) mult_right_isotone order_trans star.circ_increasing star.left_plus_circ) hence 6: "w\<^sup>T\<^sup>+ \ -w" by (metis conv_star_commute mult_assoc mult_1_left triple_schroeder_p) have "w * r \ w\<^sup>T\<^sup>+ * r = (w \ w\<^sup>T\<^sup>+) * r" using assms(2) by (simp add: injective_comp_right_dist_inf) also have "... = bot" using 6 p_antitone pseudo_complement_pp semiring.mult_not_zero by blast finally have 7: "w * r \ w\<^sup>T\<^sup>+ * r = bot" . have "-1 * r \ -r" using assms(2) local.dual_order.trans local.pp_increasing local.schroeder_4_p by blast hence "-1 * r * top \ -r" by (simp add: assms(1) comp_associative) hence 8: "r\<^sup>T * -1 * r \ bot" by (simp add: mult_assoc schroeder_6_p) have "r\<^sup>T * w * r \ r\<^sup>T * w\<^sup>+ * r" by (simp add: mult_left_isotone mult_right_isotone star.circ_mult_increasing) also have "... \ r\<^sup>T * -1 * r" by (simp add: assms(4) comp_isotone) finally have "r\<^sup>T * w * r \ bot" using 8 by simp hence "w * r * top \ -r" by (simp add: mult_assoc schroeder_6_p) hence "w * r \ -r" by (simp add: assms(1) comp_associative) hence "w * r \ -r \ w * v" using 3 by (simp add: mult_right_isotone) also have "... \ -r \ v" using 2 by (simp add: le_infI2) also have "... = -r \ t\<^sup>T\<^sup>\ * r" using 4 by (simp add: inf_commute) also have "... \ -r \ w\<^sup>T\<^sup>\ * r" using assms(5) comp_inf.mult_right_isotone conv_isotone mult_left_isotone star_isotone by auto also have "... = -r \ (r \ w\<^sup>T\<^sup>+ * r)" using mult_assoc star.circ_loop_fixpoint sup_commute by auto also have "... \ w\<^sup>T\<^sup>+ * r" using inf.commute maddux_3_13 by auto finally have "w * r = bot" using 7 by (simp add: le_iff_inf) hence "w = w \ top * -r\<^sup>T" by (metis complement_conv_sub conv_dist_comp conv_involutive conv_bot inf.assoc inf.orderE regular_closed_bot regular_dense_top top_left_mult_increasing) also have "... = w \ v * v\<^sup>T \ top * -r\<^sup>T" by (simp add: assms(6) inf_absorb1) also have "... \ w \ top * v\<^sup>T \ top * -r\<^sup>T" using comp_inf.mult_left_isotone comp_inf.mult_right_isotone mult_left_isotone by auto also have "... = w \ top * (v\<^sup>T \ -r\<^sup>T)" using 1 assms(1) covector_inf_closed inf_assoc vector_conv_compl vector_conv_covector by auto also have "... = w * (1 \ (v \ -r) * top)" by (simp add: comp_inf_vector conv_complement conv_dist_inf) also have "... = w * (1 \ (v \ -r) * (v \ -r)\<^sup>T)" by (metis conv_top dedekind_eq inf_commute inf_top_left mult_1_left mult_1_right) also have "... \ w * t\<^sup>T * t" using 5 by (simp add: comp_isotone mult_assoc) also have "... \ w * w\<^sup>T * t" by (simp add: assms(5) comp_isotone conv_isotone) also have "... \ t" using assms(4) mult_left_isotone mult_1_left by fastforce finally show ?thesis by (simp add: assms(5) antisym) qed subsection \Kruskal's Algorithm\ text \ The following results are used for proving the correctness of Kruskal's minimum spanning tree algorithm. \ subsubsection \Preservation of Invariant\ text \ We first treat the preservation of the invariant. The following lemmas show conditions necessary for preserving that \f\ is a forest. \ lemma kruskal_injective_inv_2: assumes "arc e" and "acyclic f" shows "top * e * f\<^sup>T\<^sup>\ * f\<^sup>T \ -e" proof - have "f \ -f\<^sup>T\<^sup>\" using assms(2) acyclic_star_below_complement p_antitone_iff by simp hence "e * f \ top * e * -f\<^sup>T\<^sup>\" by (simp add: comp_isotone top_left_mult_increasing) also have "... = -(top * e * f\<^sup>T\<^sup>\)" by (metis assms(1) comp_mapping_complement conv_dist_comp conv_involutive conv_top) finally show ?thesis using schroeder_4_p by simp qed lemma kruskal_injective_inv_3: assumes "arc e" and "forest f" shows "(top * e * f\<^sup>T\<^sup>\)\<^sup>T * (top * e * f\<^sup>T\<^sup>\) \ f\<^sup>T * f \ 1" proof - have "(top * e * f\<^sup>T\<^sup>\)\<^sup>T * (top * e * f\<^sup>T\<^sup>\) = f\<^sup>\ * e\<^sup>T * top * e * f\<^sup>T\<^sup>\" by (metis conv_dist_comp conv_involutive conv_star_commute conv_top vector_top_closed mult_assoc) also have "... \ f\<^sup>\ * f\<^sup>T\<^sup>\" by (metis assms(1) arc_expanded mult_left_isotone mult_right_isotone mult_1_left mult_assoc) finally have "(top * e * f\<^sup>T\<^sup>\)\<^sup>T * (top * e * f\<^sup>T\<^sup>\) \ f\<^sup>T * f \ f\<^sup>\ * f\<^sup>T\<^sup>\ \ f\<^sup>T * f" using inf.sup_left_isotone by simp also have "... \ 1" using assms(2) forest_separate by simp finally show ?thesis by simp qed lemma kruskal_acyclic_inv: assumes "acyclic f" and "covector q" and "(f \ q)\<^sup>T * f\<^sup>\ * e = bot" and "e * f\<^sup>\ * e = bot" and "f\<^sup>T\<^sup>\ * f\<^sup>\ \ -e" shows "acyclic ((f \ -q) \ (f \ q)\<^sup>T \ e)" proof - have "(f \ -q) * (f \ q)\<^sup>T = (f \ -q) * (f\<^sup>T \ q\<^sup>T)" by (simp add: conv_dist_inf) hence 1: "(f \ -q) * (f \ q)\<^sup>T = bot" by (metis assms(2) comp_inf.semiring.mult_zero_right comp_inf_vector_1 conv_bot covector_bot_closed inf.sup_monoid.add_assoc p_inf) hence 2: "(f \ -q)\<^sup>\ * (f \ q)\<^sup>T = (f \ q)\<^sup>T" using mult_right_zero star_absorb star_simulation_right_equal by fastforce hence 3: "((f \ -q) \ (f \ q)\<^sup>T)\<^sup>+ = (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+" by (simp add: plus_sup) have 4: "((f \ -q) \ (f \ q)\<^sup>T)\<^sup>\ = (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\" using 2 by (simp add: star.circ_sup_9) have "(f \ q)\<^sup>T * (f \ -q)\<^sup>\ * e \ (f \ q)\<^sup>T * f\<^sup>\ * e" by (simp add: mult_left_isotone mult_right_isotone star_isotone) hence "(f \ q)\<^sup>T * (f \ -q)\<^sup>\ * e = bot" using assms(3) le_bot by simp hence 5: "(f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ * e = (f \ -q)\<^sup>\ * e" by (metis comp_associative conv_bot conv_dist_comp conv_involutive conv_star_commute star_absorb) have "e * (f \ -q)\<^sup>\ * e \ e * f\<^sup>\ * e" by (simp add: mult_left_isotone mult_right_isotone star_isotone) hence "e * (f \ -q)\<^sup>\ * e = bot" using assms(4) le_bot by simp hence 6: "((f \ -q)\<^sup>\ * e)\<^sup>+ = (f \ -q)\<^sup>\ * e" by (simp add: comp_associative star_absorb) have "f\<^sup>T\<^sup>\ * 1 * f\<^sup>T\<^sup>\ * f\<^sup>\ \ -e" by (simp add: assms(5) star.circ_transitive_equal) hence 7: "f\<^sup>\ * e * f\<^sup>T\<^sup>\ * f\<^sup>\ \ -1" by (metis comp_right_one conv_involutive conv_one conv_star_commute triple_schroeder_p) have "(f \ -q)\<^sup>+ * (f \ q)\<^sup>T\<^sup>+ \ -1" using 1 2 by (metis forest_bot mult_left_zero mult_assoc) hence 8: "(f \ q)\<^sup>T\<^sup>+ * (f \ -q)\<^sup>+ \ -1" using comp_commute_below_diversity by simp have 9: "f\<^sup>T\<^sup>+ \ -1" using assms(1) acyclic_star_below_complement schroeder_5_p by force have "((f \ -q) \ (f \ q)\<^sup>T \ e)\<^sup>+ = (((f \ -q) \ (f \ q)\<^sup>T)\<^sup>\ * e)\<^sup>\ * ((f \ -q) \ (f \ q)\<^sup>T)\<^sup>+ \ (((f \ -q) \ (f \ q)\<^sup>T)\<^sup>\ * e)\<^sup>+" by (simp add: plus_sup) also have "... = ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ * e)\<^sup>\ * ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+) \ ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ * e)\<^sup>+" using 3 4 by simp also have "... = ((f \ -q)\<^sup>\ * e)\<^sup>\ * ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+) \ ((f \ -q)\<^sup>\ * e)\<^sup>+" using 5 by simp also have "... = ((f \ -q)\<^sup>\ * e \ 1) * ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+) \ (f \ -q)\<^sup>\ * e" using 6 by (metis star_left_unfold_equal sup_monoid.add_commute) also have "... = (f \ -q)\<^sup>\ * e \ (f \ -q)\<^sup>\ * e * (f \ q)\<^sup>T\<^sup>+ \ (f \ -q)\<^sup>\ * e * (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+" using comp_associative mult_left_dist_sup mult_right_dist_sup sup_assoc sup_commute by simp also have "... = (f \ -q)\<^sup>\ * e * (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ \ (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+" by (metis star.circ_back_loop_fixpoint star_plus sup_monoid.add_commute mult_assoc) also have "... \ f\<^sup>\ * e * f\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ \ (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+" using mult_left_isotone mult_right_isotone star_isotone sup_left_isotone conv_isotone order_trans inf_le1 by meson also have "... \ f\<^sup>\ * e * f\<^sup>T\<^sup>\ * f\<^sup>\ \ (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ f\<^sup>T\<^sup>+" using mult_left_isotone mult_right_isotone star_isotone sup_left_isotone sup_right_isotone conv_isotone order_trans inf_le1 by meson also have "... = f\<^sup>\ * e * f\<^sup>T\<^sup>\ * f\<^sup>\ \ (f \ q)\<^sup>T\<^sup>+ * (f \ -q)\<^sup>+ \ (f \ -q)\<^sup>+ \ f\<^sup>T\<^sup>+" by (simp add: star.circ_loop_fixpoint sup_monoid.add_assoc mult_assoc) also have "... \ f\<^sup>\ * e * f\<^sup>T\<^sup>\ * f\<^sup>\ \ (f \ q)\<^sup>T\<^sup>+ * (f \ -q)\<^sup>+ \ f\<^sup>+ \ f\<^sup>T\<^sup>+" using mult_left_isotone mult_right_isotone star_isotone sup_left_isotone sup_right_isotone order_trans inf_le1 by meson also have "... \ -1" using 7 8 9 assms(1) by simp finally show ?thesis by simp qed lemma kruskal_exchange_acyclic_inv_1: assumes "acyclic f" and "covector q" shows "acyclic ((f \ -q) \ (f \ q)\<^sup>T)" using kruskal_acyclic_inv[where e=bot] by (simp add: assms) lemma kruskal_exchange_acyclic_inv_2: assumes "acyclic w" and "injective w" and "d \ w" and "bijective (d\<^sup>T * top)" and "bijective (e * top)" and "d \ top * e\<^sup>T * w\<^sup>T\<^sup>\" and "w * e\<^sup>T * top = bot" shows "acyclic ((w \ -d) \ e)" proof - let ?v = "w \ -d" let ?w = "?v \ e" have "d\<^sup>T * top \ w\<^sup>\ * e * top" by (metis assms(6) comp_associative comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb comp_isotone conv_dist_comp conv_involutive conv_order conv_star_commute conv_top inf.cobounded1 vector_top_closed) hence 1: "e * top \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" by (metis assms(4,5) bijective_reverse comp_associative conv_star_commute) have 2: "?v * d\<^sup>T * top = bot" by (simp add: assms(2,3) kruskal_exchange_acyclic_inv_3) have "?v * w\<^sup>T\<^sup>+ * d\<^sup>T * top \ w * w\<^sup>T\<^sup>+ * d\<^sup>T * top" by (simp add: mult_left_isotone) also have "... \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" by (metis assms(2) mult_left_isotone mult_1_left mult_assoc) finally have "?v * w\<^sup>T\<^sup>\ * d\<^sup>T * top \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" using 2 by (metis bot_least comp_associative mult_right_dist_sup star.circ_back_loop_fixpoint star.circ_plus_same sup_least) hence 3: "?v\<^sup>\ * e * top \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" using 1 by (simp add: comp_associative star_left_induct sup_least) have "d * e\<^sup>T \ bot" by (metis assms(3,7) conv_bot conv_dist_comp conv_involutive conv_top order.trans inf.absorb2 inf.cobounded2 inf_commute le_bot p_antitone_iff p_top schroeder_4_p top_left_mult_increasing) hence 4: "e\<^sup>T * top \ -(d\<^sup>T * top)" by (metis (no_types) comp_associative inf.cobounded2 le_bot p_antitone_iff schroeder_3_p semiring.mult_zero_left) have "?v\<^sup>T * -(d\<^sup>T * top) \ -(d\<^sup>T * top)" using schroeder_3_p mult_assoc 2 by simp hence "?v\<^sup>T\<^sup>\ * e\<^sup>T * top \ -(d\<^sup>T * top)" using 4 by (simp add: comp_associative star_left_induct sup_least) hence 5: "d\<^sup>T * top \ -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" by (simp add: p_antitone_iff) have "w * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top = w * e\<^sup>T * top \ w * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top" by (metis star_left_unfold_equal mult_right_dist_sup mult_left_dist_sup mult_1_right mult_assoc) also have "... = w * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top" using assms(7) by simp also have "... \ w * w\<^sup>T * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (simp add: comp_associative conv_isotone mult_left_isotone mult_right_isotone) also have "... \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (metis assms(2) mult_1_left mult_left_isotone) finally have "w * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top \ --(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" by (simp add: p_antitone p_antitone_iff) hence "w\<^sup>T * -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top) \ -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" using comp_associative schroeder_3_p by simp hence 6: "w\<^sup>T\<^sup>\ * d\<^sup>T * top \ -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" using 5 by (simp add: comp_associative star_left_induct sup_least) have "e * ?v\<^sup>\ * e \ e * ?v\<^sup>\ * e * top" by (simp add: top_right_mult_increasing) also have "... \ e * w\<^sup>T\<^sup>\ * d\<^sup>T * top" using 3 by (simp add: comp_associative mult_right_isotone) also have "... \ e * -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" using 6 by (simp add: comp_associative mult_right_isotone) also have "... \ bot" by (metis conv_complement_sub_leq conv_dist_comp conv_involutive conv_star_commute le_bot mult_right_sub_dist_sup_right p_bot regular_closed_bot star.circ_back_loop_fixpoint) finally have 7: "e * ?v\<^sup>\ * e = bot" by (simp add: antisym) hence "?v\<^sup>\ * e \ -1" by (metis bot_least comp_associative comp_commute_below_diversity ex231d order_lesseq_imp semiring.mult_zero_left star.circ_left_top) hence 8: "?v\<^sup>\ * e * ?v\<^sup>\ \ -1" by (metis comp_associative comp_commute_below_diversity star.circ_transitive_equal) have "1 \ ?w\<^sup>+ = 1 \ ?w * ?v\<^sup>\ * (e * ?v\<^sup>\)\<^sup>\" by (simp add: star_sup_1 mult_assoc) also have "... = 1 \ ?w * ?v\<^sup>\ * (e * ?v\<^sup>\ \ 1)" using 7 by (metis star.circ_mult_1 star_absorb sup_monoid.add_commute mult_assoc) also have "... = 1 \ (?v\<^sup>+ * e * ?v\<^sup>\ \ ?v\<^sup>+ \ e * ?v\<^sup>\ * e * ?v\<^sup>\ \ e * ?v\<^sup>\)" by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup sup_assoc sup_commute sup_left_commute) also have "... = 1 \ (?v\<^sup>+ * e * ?v\<^sup>\ \ ?v\<^sup>+ \ e * ?v\<^sup>\)" using 7 by simp also have "... = 1 \ (?v\<^sup>\ * e * ?v\<^sup>\ \ ?v\<^sup>+)" by (metis (mono_tags, hide_lams) comp_associative star.circ_loop_fixpoint sup_assoc sup_commute) also have "... \ 1 \ (?v\<^sup>\ * e * ?v\<^sup>\ \ w\<^sup>+)" using comp_inf.mult_right_isotone comp_isotone semiring.add_right_mono star_isotone sup_commute by simp also have "... = (1 \ ?v\<^sup>\ * e * ?v\<^sup>\) \ (1 \ w\<^sup>+)" by (simp add: inf_sup_distrib1) also have "... = 1 \ ?v\<^sup>\ * e * ?v\<^sup>\" by (metis assms(1) inf_commute pseudo_complement sup_bot_right) also have "... = bot" using 8 p_antitone_iff pseudo_complement by simp finally show ?thesis using le_bot p_antitone_iff pseudo_complement by auto qed subsubsection \Exchange gives Spanning Trees\ text \ The lemmas in this section are used to show that the relation after exchange represents a spanning tree. \ lemma inf_star_import: assumes "x \ z" and "univalent z" and "reflexive y" and "regular z" shows "x\<^sup>\ * y \ z\<^sup>\ \ x\<^sup>\ * (y \ z\<^sup>\)" proof - have 1: "y \ x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\" by (metis assms(4) pp_dist_star shunting_var_p star.circ_loop_fixpoint sup.cobounded2) have "x * -z\<^sup>\ \ z\<^sup>+ \ x * (-z\<^sup>\ \ x\<^sup>T * z\<^sup>+)" by (simp add: dedekind_1) also have "... \ x * (-z\<^sup>\ \ z\<^sup>T * z\<^sup>+)" using assms(1) comp_inf.mult_right_isotone conv_isotone mult_left_isotone mult_right_isotone by simp also have "... \ x * (-z\<^sup>\ \ 1 * z\<^sup>\)" by (metis assms(2) comp_associative comp_inf.mult_right_isotone mult_left_isotone mult_right_isotone) finally have 2: "x * -z\<^sup>\ \ z\<^sup>+ = bot" by (simp add: antisym) have "x * -z\<^sup>\ \ z\<^sup>\ = (x * -z\<^sup>\ \ z\<^sup>+) \ (x * -z\<^sup>\ \ 1)" by (metis comp_inf.semiring.distrib_left star_left_unfold_equal sup_commute) also have "... \ x\<^sup>\ * (y \ z\<^sup>\)" using 2 by (simp add: assms(3) inf.coboundedI2 reflexive_mult_closed star.circ_reflexive) finally have "x * -z\<^sup>\ \ x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\" by (metis assms(4) pp_dist_star shunting_var_p) hence "x * (x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\) \ x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\" by (metis le_supE le_supI mult_left_dist_sup star.circ_loop_fixpoint sup.cobounded1) hence "x\<^sup>\ * y \ x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\" using 1 by (simp add: star_left_induct) hence "x\<^sup>\ * y \ --z\<^sup>\ \ x\<^sup>\ * (y \ z\<^sup>\)" using shunting_var_p by simp thus ?thesis using order.trans inf.sup_right_isotone pp_increasing by blast qed lemma kruskal_exchange_forest_components_inv: assumes "injective ((w \ -d) \ e)" and "regular d" and "e * top * e = e" and "d \ top * e\<^sup>T * w\<^sup>T\<^sup>\" and "w * e\<^sup>T * top = bot" and "injective w" and "d \ w" and "d \ (w \ -d)\<^sup>T\<^sup>\ * e\<^sup>T * top" shows "forest_components w \ forest_components ((w \ -d) \ e)" proof - let ?v = "w \ -d" let ?w = "?v \ e" let ?f = "forest_components ?w" have 1: "?v * d\<^sup>T * top = bot" by (simp add: assms(6,7) kruskal_exchange_acyclic_inv_3) have 2: "d * e\<^sup>T \ bot" by (metis assms(5,7) conv_bot conv_dist_comp conv_involutive conv_top order.trans inf.absorb2 inf.cobounded2 inf_commute le_bot p_antitone_iff p_top schroeder_4_p top_left_mult_increasing) have "w\<^sup>\ * e\<^sup>T * top = e\<^sup>T * top" by (metis assms(5) conv_bot conv_dist_comp conv_involutive conv_star_commute star.circ_top star_absorb) hence "w\<^sup>\ * e\<^sup>T * top \ -(d\<^sup>T * top)" using 2 by (metis (no_types) comp_associative inf.cobounded2 le_bot p_antitone_iff schroeder_3_p semiring.mult_zero_left) hence 3: "e\<^sup>T * top \ -(w\<^sup>T\<^sup>\ * d\<^sup>T * top)" by (metis conv_star_commute p_antitone_iff schroeder_3_p mult_assoc) have "?v * w\<^sup>T\<^sup>\ * d\<^sup>T * top = ?v * d\<^sup>T * top \ ?v * w\<^sup>T\<^sup>+ * d\<^sup>T * top" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... \ w * w\<^sup>T\<^sup>+ * d\<^sup>T * top" using 1 by (simp add: mult_left_isotone) also have "... \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" by (metis assms(6) mult_assoc mult_1_left mult_left_isotone) finally have "?v * w\<^sup>T\<^sup>\ * d\<^sup>T * top \ --(w\<^sup>T\<^sup>\ * d\<^sup>T * top)" using p_antitone p_antitone_iff by auto hence 4: "?v\<^sup>T * -(w\<^sup>T\<^sup>\ * d\<^sup>T * top) \ -(w\<^sup>T\<^sup>\ * d\<^sup>T * top)" using comp_associative schroeder_3_p by simp have 5: "injective ?v" using assms(1) conv_dist_sup mult_left_dist_sup mult_right_dist_sup by simp have "?v * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top = ?v * e\<^sup>T * top \ ?v * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... \ w * e\<^sup>T * top \ ?v * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top" using mult_left_isotone sup_left_isotone by simp also have "... \ w * e\<^sup>T * top \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" using 5 by (metis mult_assoc mult_1_left mult_left_isotone sup_right_isotone) finally have "?v * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (simp add: assms(5)) hence "?v\<^sup>\ * d * top \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (metis assms(8) star_left_induct sup_least comp_associative mult_right_sub_dist_sup_right sup.orderE vector_top_closed) also have "... \ -(w\<^sup>T\<^sup>\ * d\<^sup>T * top)" using 3 4 by (simp add: comp_associative star_left_induct) also have "... \ -(d\<^sup>T * top)" by (metis p_antitone star.circ_left_top star_outer_increasing mult_assoc) finally have 6: "?v\<^sup>\ * d * top \ -(d\<^sup>T * top)" by simp have "d\<^sup>T * top \ w\<^sup>\ * e * top" by (metis assms(4) comp_associative comp_inf.star.circ_sup_2 comp_isotone conv_dist_comp conv_involutive conv_order conv_star_commute conv_top vector_top_closed) also have "... \ (?v \ d)\<^sup>\ * e * top" by (metis assms(2) comp_inf.semiring.distrib_left maddux_3_11_pp mult_left_isotone star_isotone sup.cobounded2 sup_commute sup_inf_distrib1) also have "... = ?v\<^sup>\ * (d * ?v\<^sup>\)\<^sup>\ * e * top" by (simp add: star_sup_1) also have "... = ?v\<^sup>\ * e * top \ ?v\<^sup>\ * d * ?v\<^sup>\ * (d * ?v\<^sup>\)\<^sup>\ * e * top" by (metis semiring.distrib_right star.circ_unfold_sum star_decompose_1 star_decompose_3 mult_assoc) also have "... \ ?v\<^sup>\ * e * top \ ?v\<^sup>\ * d * top" by (metis comp_associative comp_isotone le_supI mult_left_dist_sup mult_right_dist_sup mult_right_isotone star.circ_decompose_5 star_decompose_3 sup.cobounded1 sup_commute top.extremum) finally have "d\<^sup>T * top \ ?v\<^sup>\ * e * top \ (d\<^sup>T * top \ ?v\<^sup>\ * d * top)" using sup_inf_distrib2 sup_monoid.add_commute by simp hence "d\<^sup>T * top \ ?v\<^sup>\ * e * top" using 6 by (metis inf_commute pseudo_complement sup_monoid.add_0_right) hence 7: "d \ top * e\<^sup>T * ?v\<^sup>T\<^sup>\" by (metis comp_associative conv_dist_comp conv_involutive conv_isotone conv_star_commute conv_top order.trans top_right_mult_increasing) have 8: "?v \ ?f" using forest_components_increasing le_supE by blast have "d \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top \ top * e\<^sup>T * ?v\<^sup>T\<^sup>\" using 7 assms(8) by simp also have "... = ?v\<^sup>T\<^sup>\ * e\<^sup>T * top * e\<^sup>T * ?v\<^sup>T\<^sup>\" by (metis inf_top_right vector_inf_comp vector_top_closed mult_assoc) also have "... = ?v\<^sup>T\<^sup>\ * e\<^sup>T * ?v\<^sup>T\<^sup>\" by (metis assms(3) comp_associative conv_dist_comp conv_top) also have "... \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * ?f" using 8 by (metis assms(1) forest_components_equivalence cancel_separate_1 conv_dist_comp conv_order mult_left_isotone star_involutive star_isotone) also have "... \ ?v\<^sup>T\<^sup>\ * ?f * ?f" by (metis assms(1) forest_components_equivalence forest_components_increasing conv_isotone le_supE mult_left_isotone mult_right_isotone) also have "... \ ?f * ?f * ?f" by (metis comp_associative comp_isotone conv_dist_sup star.circ_loop_fixpoint star_isotone sup.cobounded1 sup.cobounded2) also have "... = ?f" by (simp add: assms(1) forest_components_equivalence preorder_idempotent) finally have "w \ ?f" using 8 by (metis assms(2) shunting_var_p sup.orderE) thus ?thesis using assms(1) forest_components_idempotent forest_components_isotone by fastforce qed lemma kruskal_spanning_inv: assumes "injective ((f \ -q) \ (f \ q)\<^sup>T \ e)" and "regular q" and "regular e" and "(-h \ --g)\<^sup>\ \ forest_components f" shows "components (-(h \ -e \ -e\<^sup>T) \ g) \ forest_components ((f \ -q) \ (f \ q)\<^sup>T \ e)" proof - let ?f = "(f \ -q) \ (f \ q)\<^sup>T \ e" let ?h = "h \ -e \ -e\<^sup>T" let ?F = "forest_components f" let ?FF = "forest_components ?f" have 1: "equivalence ?FF" using assms(1) forest_components_equivalence by simp hence 2: "?f * ?FF \ ?FF" using order.trans forest_components_increasing mult_left_isotone by blast have 3: "?f\<^sup>T * ?FF \ ?FF" using 1 by (metis forest_components_increasing mult_left_isotone conv_isotone preorder_idempotent) have "(f \ q) * ?FF \ ?f\<^sup>T * ?FF" using conv_dist_sup conv_involutive sup_assoc sup_left_commute mult_left_isotone by simp hence 4: "(f \ q) * ?FF \ ?FF" using 3 order.trans by blast have "(f \ -q) * ?FF \ ?f * ?FF" using le_supI1 mult_left_isotone by simp hence "(f \ -q) * ?FF \ ?FF" using 2 order.trans by blast hence "((f \ q) \ (f \ -q)) * ?FF \ ?FF" using 4 mult_right_dist_sup by simp hence "f * ?FF \ ?FF" by (metis assms(2) maddux_3_11_pp) hence 5: "f\<^sup>\ * ?FF \ ?FF" using star_left_induct_mult_iff by simp have "(f \ -q)\<^sup>T * ?FF \ ?f\<^sup>T * ?FF" by (meson conv_isotone order.trans mult_left_isotone sup.cobounded1) hence 6: "(f \ -q)\<^sup>T * ?FF \ ?FF" using 3 order.trans by blast have "(f \ q)\<^sup>T * ?FF \ ?f * ?FF" by (simp add: mult_left_isotone sup.left_commute sup_assoc) hence "(f \ q)\<^sup>T * ?FF \ ?FF" using 2 order.trans by blast hence "((f \ -q)\<^sup>T \ (f \ q)\<^sup>T) * ?FF \ ?FF" using 6 mult_right_dist_sup by simp hence "f\<^sup>T * ?FF \ ?FF" by (metis assms(2) conv_dist_sup maddux_3_11_pp) hence 7: "?F * ?FF \ ?FF" using 5 star_left_induct mult_assoc by simp have 8: "e * ?FF \ ?FF" using 2 by (simp add: mult_right_dist_sup mult_left_isotone) have "e\<^sup>T * ?FF \ ?f\<^sup>T * ?FF" by (simp add: mult_left_isotone conv_isotone) also have "... \ ?FF * ?FF" using 1 by (metis forest_components_increasing mult_left_isotone conv_isotone) finally have "e\<^sup>T * ?FF \ ?FF" using 1 preorder_idempotent by auto hence 9: "(?F \ e \ e\<^sup>T) * ?FF \ ?FF" using 7 8 mult_right_dist_sup by simp have "components (-?h \ g) \ ((-h \ --g) \ e \ e\<^sup>T)\<^sup>\" by (metis assms(3) comp_inf.mult_left_sub_dist_sup_left conv_complement p_dist_inf pp_dist_inf regular_closed_p star_isotone sup_inf_distrib2 sup_monoid.add_assoc) also have "... \ ((-h \ --g)\<^sup>\ \ e \ e\<^sup>T)\<^sup>\" using star.circ_increasing star_isotone sup_left_isotone by simp also have "... \ (?F \ e \ e\<^sup>T)\<^sup>\" using assms(4) sup_left_isotone star_isotone by simp also have "... \ ?FF" using 1 9 star_left_induct by force finally show ?thesis by simp qed lemma kruskal_exchange_spanning_inv_1: assumes "injective ((w \ -q) \ (w \ q)\<^sup>T)" and "regular (w \ q)" and "components g \ forest_components w" shows "components g \ forest_components ((w \ -q) \ (w \ q)\<^sup>T)" proof - let ?p = "w \ q" let ?w = "(w \ -q) \ ?p\<^sup>T" have 1: "w \ -?p \ forest_components ?w" by (metis forest_components_increasing inf_import_p le_supE) have "w \ ?p \ ?w\<^sup>T" by (simp add: conv_dist_sup) also have "... \ forest_components ?w" by (metis assms(1) conv_isotone forest_components_equivalence forest_components_increasing) finally have "w \ (?p \ -?p) \ forest_components ?w" using 1 inf_sup_distrib1 by simp hence "w \ forest_components ?w" by (metis assms(2) inf_top_right stone) hence 2: "w\<^sup>\ \ forest_components ?w" using assms(1) star_isotone forest_components_star by force hence 3: "w\<^sup>T\<^sup>\ \ forest_components ?w" using assms(1) conv_isotone conv_star_commute forest_components_equivalence by force have "components g \ forest_components w" using assms(3) by simp also have "... \ forest_components ?w * forest_components ?w" using 2 3 mult_isotone by simp also have "... = forest_components ?w" using assms(1) forest_components_equivalence preorder_idempotent by simp finally show ?thesis by simp qed lemma kruskal_exchange_spanning_inv_2: assumes "injective w" and "w\<^sup>\ * e\<^sup>T = e\<^sup>T" and "f \ f\<^sup>T \ (w \ -d \ -d\<^sup>T) \ (w\<^sup>T \ -d \ -d\<^sup>T)" and "d \ forest_components f * e\<^sup>T * top" shows "d \ (w \ -d)\<^sup>T\<^sup>\ * e\<^sup>T * top" proof - have 1: "(w \ -d \ -d\<^sup>T) * (w\<^sup>T \ -d \ -d\<^sup>T) \ 1" using assms(1) comp_isotone order.trans inf.cobounded1 by blast have "d \ forest_components f * e\<^sup>T * top" using assms(4) by simp also have "... \ (f \ f\<^sup>T)\<^sup>\ * (f \ f\<^sup>T)\<^sup>\ * e\<^sup>T * top" by (simp add: comp_isotone star_isotone) also have "... = (f \ f\<^sup>T)\<^sup>\ * e\<^sup>T * top" by (simp add: star.circ_transitive_equal) also have "... \ ((w \ -d \ -d\<^sup>T) \ (w\<^sup>T \ -d \ -d\<^sup>T))\<^sup>\ * e\<^sup>T * top" using assms(3) by (simp add: comp_isotone star_isotone) also have "... = (w\<^sup>T \ -d \ -d\<^sup>T)\<^sup>\ * (w \ -d \ -d\<^sup>T)\<^sup>\ * e\<^sup>T * top" using 1 cancel_separate_1 by simp also have "... \ (w\<^sup>T \ -d \ -d\<^sup>T)\<^sup>\ * w\<^sup>\ * e\<^sup>T * top" by (simp add: inf_assoc mult_left_isotone mult_right_isotone star_isotone) also have "... = (w\<^sup>T \ -d \ -d\<^sup>T)\<^sup>\ * e\<^sup>T * top" using assms(2) mult_assoc by simp also have "... \ (w\<^sup>T \ -d\<^sup>T)\<^sup>\ * e\<^sup>T * top" using mult_left_isotone conv_isotone star_isotone comp_inf.mult_right_isotone inf.cobounded2 inf.left_commute inf.sup_monoid.add_commute by presburger also have "... = (w \ -d)\<^sup>T\<^sup>\ * e\<^sup>T * top" using conv_complement conv_dist_inf by presburger finally show ?thesis by simp qed lemma kruskal_spanning_inv_1: assumes "e \ F" and "regular e" and "components (-h \ g) \ F" and "equivalence F" shows "components (-(h \ -e \ -e\<^sup>T) \ g) \ F" proof - have 1: "F * F \ F" using assms(4) by simp hence 2: "e * F \ F" using assms(1) mult_left_isotone order_lesseq_imp by blast have "e\<^sup>T * F \ F" by (metis assms(1,4) conv_isotone mult_left_isotone preorder_idempotent) hence 3: "(F \ e \ e\<^sup>T) * F \ F" using 1 2 mult_right_dist_sup by simp have "components (-(h \ -e \ -e\<^sup>T) \ g) \ ((-h \ --g) \ e \ e\<^sup>T)\<^sup>\" by (metis assms(2) comp_inf.mult_left_sub_dist_sup_left conv_complement p_dist_inf pp_dist_inf regular_closed_p star_isotone sup_inf_distrib2 sup_monoid.add_assoc) also have "... \ ((-h \ --g)\<^sup>\ \ e \ e\<^sup>T)\<^sup>\" using sup_left_isotone star.circ_increasing star_isotone by simp also have "... \ (F \ e \ e\<^sup>T)\<^sup>\" using assms(3) sup_left_isotone star_isotone by simp also have "... \ F" using 3 assms(4) star_left_induct by force finally show ?thesis by simp qed lemma kruskal_reroot_edge: assumes "injective (e\<^sup>T * top)" and "acyclic w" shows "((w \ -(top * e * w\<^sup>T\<^sup>\)) \ (w \ top * e * w\<^sup>T\<^sup>\)\<^sup>T) * e\<^sup>T = bot" proof - let ?q = "top * e * w\<^sup>T\<^sup>\" let ?p = "w \ ?q" let ?w = "(w \ -?q) \ ?p\<^sup>T" have "(w \ -?q) * e\<^sup>T * top = w * (e\<^sup>T * top \ -?q\<^sup>T)" by (metis comp_associative comp_inf_vector_1 conv_complement covector_complement_closed vector_top_closed) also have "... = w * (e\<^sup>T * top \ -(w\<^sup>\ * e\<^sup>T * top))" by (simp add: conv_dist_comp conv_star_commute mult_assoc) also have "... = bot" by (metis comp_associative comp_inf.semiring.mult_not_zero inf.sup_relative_same_increasing inf_p mult_right_zero star.circ_loop_fixpoint sup_commute sup_left_divisibility) finally have 1: "(w \ -?q) * e\<^sup>T * top = bot" by simp have "?p\<^sup>T * e\<^sup>T * top = (w\<^sup>T \ w\<^sup>\ * e\<^sup>T * top) * e\<^sup>T * top" by (simp add: conv_dist_comp conv_star_commute mult_assoc conv_dist_inf) also have "... = w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T * e\<^sup>T * top" by (simp add: inf_vector_comp vector_export_comp) also have "... = (w\<^sup>\ \ w\<^sup>T) * e\<^sup>T * top" using assms(1) injective_comp_right_dist_inf mult_assoc by simp also have "... = bot" using assms(2) acyclic_star_below_complement_1 semiring.mult_not_zero by blast finally have "?w * e\<^sup>T * top = bot" using 1 mult_right_dist_sup by simp thus ?thesis by (metis star.circ_top star_absorb) qed subsubsection \Exchange gives Minimum Spanning Trees\ text \ The lemmas in this section are used to show that the after exchange we obtain a minimum spanning tree. The following lemmas show that the relation characterising the edge across the cut is an arc. \ lemma kruskal_edge_arc: assumes "equivalence F" and "forest w" and "arc e" and "regular F" and "F \ forest_components (F \ w)" and "regular w" and "w * e\<^sup>T = bot" and "e * F * e = bot" and "e\<^sup>T \ w\<^sup>\" shows "arc (w \ top * e\<^sup>T * w\<^sup>T\<^sup>\ \ F * e\<^sup>T * top \ top * e * -F)" proof (unfold arc_expanded, intro conjI) let ?E = "top * e\<^sup>T * w\<^sup>T\<^sup>\" let ?F = "F * e\<^sup>T * top" let ?G = "top * e * -F" let ?FF = "F * e\<^sup>T * e * F" let ?GG = "-F * e\<^sup>T * e * -F" let ?w = "forest_components (F \ w)" have "F \ w\<^sup>T\<^sup>\ \ forest_components (F \ w) \ w\<^sup>T\<^sup>\" by (simp add: assms(5) inf.coboundedI1) also have "... \ (F \ w)\<^sup>T\<^sup>\ * ((F \ w)\<^sup>\ \ w\<^sup>T\<^sup>\)" apply (rule inf_star_import) apply (simp add: conv_isotone) apply (simp add: assms(2)) apply (simp add: star.circ_reflexive) by (metis assms(6) conv_complement) also have "... \ (F \ w)\<^sup>T\<^sup>\ * (w\<^sup>\ \ w\<^sup>T\<^sup>\)" using comp_inf.mult_left_isotone mult_right_isotone star_isotone by simp also have "... = (F \ w)\<^sup>T\<^sup>\" by (simp add: assms(2) acyclic_star_inf_conv) finally have "w * (F \ w\<^sup>T\<^sup>\) * e\<^sup>T * e \ w * (F \ w)\<^sup>T\<^sup>\ * e\<^sup>T * e" by (simp add: mult_left_isotone mult_right_isotone) also have "... = w * e\<^sup>T * e \ w * (F \ w)\<^sup>T\<^sup>+ * e\<^sup>T * e" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... = w * (F \ w)\<^sup>T\<^sup>+ * e\<^sup>T * e" by (simp add: assms(7)) also have "... \ w * (F \ w)\<^sup>T\<^sup>+" by (metis assms(3) arc_univalent mult_assoc mult_1_right mult_right_isotone) also have "... \ w * w\<^sup>T * (F \ w)\<^sup>T\<^sup>\" by (simp add: comp_associative conv_isotone mult_left_isotone mult_right_isotone) also have "... \ (F \ w)\<^sup>T\<^sup>\" using assms(2) coreflexive_comp_top_inf inf.sup_right_divisibility by auto also have "... \ F\<^sup>T\<^sup>\" by (simp add: conv_dist_inf star_isotone) finally have 1: "w * (F \ w\<^sup>T\<^sup>\) * e\<^sup>T * e \ F" by (metis assms(1) antisym mult_1_left mult_left_isotone star.circ_plus_same star.circ_reflexive star.left_plus_below_circ star_left_induct_mult_iff) have "F * e\<^sup>T * e \ forest_components (F \ w) * e\<^sup>T * e" by (simp add: assms(5) mult_left_isotone) also have "... \ forest_components w * e\<^sup>T * e" by (simp add: comp_isotone conv_dist_inf star_isotone) also have "... = w\<^sup>T\<^sup>\ * e\<^sup>T * e" by (metis (no_types) assms(7) comp_associative conv_bot conv_dist_comp conv_involutive conv_star_commute star_absorb) also have "... \ w\<^sup>T\<^sup>\" by (metis assms(3) arc_univalent mult_assoc mult_1_right mult_right_isotone) finally have 2: "F * e\<^sup>T * e \ w\<^sup>T\<^sup>\" by simp have "w * F * e\<^sup>T * e \ w * F * e\<^sup>T * e * e\<^sup>T * e" using comp_associative ex231c mult_right_isotone by simp also have "... = w * (F * e\<^sup>T * e \ w\<^sup>T\<^sup>\) * e\<^sup>T * e" using 2 by (simp add: comp_associative inf.absorb1) also have "... \ w * (F \ w\<^sup>T\<^sup>\) * e\<^sup>T * e" by (metis assms(3) arc_univalent mult_assoc mult_1_right mult_right_isotone mult_left_isotone inf.sup_left_isotone) also have "... \ F" using 1 by simp finally have 3: "w * F * e\<^sup>T * e \ F" by simp hence "e\<^sup>T * e * F * w\<^sup>T \ F" by (metis assms(1) conv_dist_comp conv_dist_inf conv_involutive inf.absorb_iff1 mult_assoc) hence "e\<^sup>T * e * F * w\<^sup>T \ e\<^sup>T * top \ F" by (simp add: comp_associative mult_right_isotone) also have "... \ e\<^sup>T * e * F" by (metis conv_involutive dedekind_1 inf_top_left mult_assoc) finally have 4: "e\<^sup>T * e * F * w\<^sup>T \ e\<^sup>T * e * F" by simp have "(top * e)\<^sup>T * (?F \ w\<^sup>T\<^sup>\) = e\<^sup>T * top * e * F * w\<^sup>T\<^sup>\" by (metis assms(1) comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb conv_dist_comp conv_involutive conv_top covector_inf_comp_3 vector_top_closed mult_assoc) also have "... = e\<^sup>T * e * F * w\<^sup>T\<^sup>\" by (simp add: assms(3) arc_top_edge) also have "... \ e\<^sup>T * e * F" using 4 star_right_induct_mult by simp also have "... \ F" by (metis assms(3) arc_injective conv_involutive mult_1_left mult_left_isotone) finally have 5: "(top * e)\<^sup>T * (?F \ w\<^sup>T\<^sup>\) \ F" by simp have "(?F \ w) * w\<^sup>T\<^sup>+ = ?F \ w * w\<^sup>T\<^sup>+" by (simp add: vector_export_comp) also have "... \ ?F \ w\<^sup>T\<^sup>\" by (metis assms(2) comp_associative inf.sup_right_isotone mult_left_isotone star.circ_transitive_equal star_left_unfold_equal sup.absorb_iff2 sup_monoid.add_assoc) also have 6: "... \ top * e * F" using 5 by (metis assms(3) shunt_mapping conv_dist_comp conv_involutive conv_top) finally have 7: "(?F \ w) * w\<^sup>T\<^sup>+ \ top * e * F" by simp have "e\<^sup>T * top * e \ 1" by (simp add: assms(3) point_injective) also have "... \ F" by (simp add: assms(1)) finally have 8: "e * -F * e\<^sup>T \ bot" by (metis p_antitone p_antitone_iff p_bot regular_closed_bot schroeder_3_p schroeder_4_p mult_assoc) have "?FF \ w * (w\<^sup>T\<^sup>+ \ ?GG) * w\<^sup>T \ ?F \ w * (w\<^sup>T\<^sup>+ \ ?GG) * w\<^sup>T" using comp_inf.mult_left_isotone mult_isotone mult_assoc by simp also have "... \ ?F \ w * (w\<^sup>T\<^sup>+ \ ?G) * w\<^sup>T" by (metis assms(3) arc_top_edge comp_inf.star.circ_decompose_9 comp_inf_covector inf.sup_right_isotone inf_le2 mult_left_isotone mult_right_isotone vector_top_closed mult_assoc) also have "... = (?F \ w) * (w\<^sup>T\<^sup>+ \ ?G) * w\<^sup>T" by (simp add: vector_export_comp) also have "... = (?F \ w) * w\<^sup>T\<^sup>+ * (?G\<^sup>T \ w\<^sup>T)" by (simp add: covector_comp_inf covector_comp_inf_1 covector_mult_closed) also have "... \ top * e * F * (?G\<^sup>T \ w\<^sup>T)" using 7 mult_left_isotone by simp also have "... \ top * e * F * ?G\<^sup>T" by (simp add: mult_right_isotone) also have "... = top * e * -F * e\<^sup>T * top" by (metis assms(1) conv_complement conv_dist_comp conv_top equivalence_comp_left_complement mult_assoc) finally have 9: "?FF \ w * (w\<^sup>T\<^sup>+ \ ?GG) * w\<^sup>T = bot" using 8 by (metis comp_associative covector_bot_closed le_bot vector_bot_closed) hence 10: "?FF \ w * (w\<^sup>+ \ ?GG) * w\<^sup>T = bot" using assms(1) comp_associative conv_bot conv_complement conv_dist_comp conv_dist_inf conv_star_commute star.circ_plus_same by fastforce have "(w \ ?E \ ?F \ ?G) * top * (w \ ?E \ ?F \ ?G)\<^sup>T = (?F \ (w \ ?E \ ?G)) * top * ((w \ ?E \ ?G)\<^sup>T \ ?F\<^sup>T)" by (simp add: conv_dist_inf inf_commute inf_left_commute) also have "... = (?F \ (w \ ?E \ ?G)) * top * (w \ ?E \ ?G)\<^sup>T \ ?F\<^sup>T" using covector_comp_inf vector_conv_covector vector_mult_closed vector_top_closed by simp also have "... = ?F \ (w \ ?E \ ?G) * top * (w \ ?E \ ?G)\<^sup>T \ ?F\<^sup>T" by (simp add: vector_export_comp) also have "... = ?F \ top * e * F \ (w \ ?E \ ?G) * top * (w \ ?E \ ?G)\<^sup>T" by (simp add: assms(1) conv_dist_comp inf_assoc inf_commute mult_assoc) also have "... = ?F * e * F \ (w \ ?E \ ?G) * top * (w \ ?E \ ?G)\<^sup>T" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... = ?FF \ (w \ ?E \ ?G) * (top * (w \ ?E \ ?G)\<^sup>T)" using assms(3) arc_top_edge comp_associative by simp also have "... = ?FF \ (w \ ?E \ ?G) * (top * (?G\<^sup>T \ (?E\<^sup>T \ w\<^sup>T)))" by (simp add: conv_dist_inf inf_assoc inf_commute inf_left_commute) also have "... = ?FF \ (w \ ?E \ ?G) * (?G * (?E\<^sup>T \ w\<^sup>T))" by (metis covector_comp_inf_1 covector_top_closed covector_mult_closed inf_top_left) also have "... = ?FF \ (w \ ?E \ ?G) * (?G \ ?E) * w\<^sup>T" by (metis covector_comp_inf_1 covector_top_closed mult_assoc) also have "... = ?FF \ (w \ ?E) * (?G\<^sup>T \ ?G \ ?E) * w\<^sup>T" by (metis covector_comp_inf_1 covector_mult_closed inf.sup_monoid.add_assoc vector_top_closed) also have "... = ?FF \ w * (?E\<^sup>T \ ?G\<^sup>T \ ?G \ ?E) * w\<^sup>T" by (metis covector_comp_inf_1 covector_mult_closed inf.sup_monoid.add_assoc vector_top_closed) also have "... = ?FF \ w * (?E\<^sup>T \ ?E \ (?G\<^sup>T \ ?G)) * w\<^sup>T" by (simp add: inf_commute inf_left_commute) also have "... = ?FF \ w * (?E\<^sup>T \ ?E \ (-F * e\<^sup>T * top \ ?G)) * w\<^sup>T" by (simp add: assms(1) conv_complement conv_dist_comp mult_assoc) also have "... = ?FF \ w * (?E\<^sup>T \ ?E \ (-F * e\<^sup>T * ?G)) * w\<^sup>T" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... = ?FF \ w * (?E\<^sup>T \ ?E \ ?GG) * w\<^sup>T" by (metis assms(3) arc_top_edge comp_associative) also have "... = ?FF \ w * (w\<^sup>\ * e * top \ ?E \ ?GG) * w\<^sup>T" by (simp add: comp_associative conv_dist_comp conv_star_commute) also have "... = ?FF \ w * (w\<^sup>\ * e * ?E \ ?GG) * w\<^sup>T" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... \ ?FF \ w * (w\<^sup>\ * w\<^sup>T\<^sup>\ \ ?GG) * w\<^sup>T" by (metis assms(3) mult_assoc mult_1_right mult_left_isotone mult_right_isotone inf.sup_left_isotone inf.sup_right_isotone arc_expanded) also have "... = ?FF \ w * ((w\<^sup>+ \ 1 \ w\<^sup>T\<^sup>\) \ ?GG) * w\<^sup>T" by (simp add: assms(2) cancel_separate_eq star_left_unfold_equal sup_monoid.add_commute) also have "... = ?FF \ w * ((w\<^sup>+ \ 1 \ w\<^sup>T\<^sup>+) \ ?GG) * w\<^sup>T" using star.circ_plus_one star_left_unfold_equal sup_assoc by presburger also have "... = (?FF \ w * (w\<^sup>+ \ ?GG) * w\<^sup>T) \ (?FF \ w * (1 \ ?GG) * w\<^sup>T) \ (?FF \ w * (w\<^sup>T\<^sup>+ \ ?GG) * w\<^sup>T)" by (simp add: inf_sup_distrib1 inf_sup_distrib2 semiring.distrib_left semiring.distrib_right) also have "... \ w * (1 \ ?GG) * w\<^sup>T" using 9 10 by simp also have "... \ w * w\<^sup>T" by (metis inf.cobounded1 mult_1_right mult_left_isotone mult_right_isotone) also have "... \ 1" by (simp add: assms(2)) finally show "(w \ ?E \ ?F \ ?G) * top * (w \ ?E \ ?F \ ?G)\<^sup>T \ 1" by simp have "w\<^sup>T\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w \ w\<^sup>T\<^sup>+ \ ?G \ w\<^sup>T * F * e\<^sup>T * e * F * w" using top_greatest inf.sup_left_isotone inf.sup_right_isotone mult_left_isotone by simp also have "... \ w\<^sup>T\<^sup>+ \ ?G \ w\<^sup>T * ?F" using comp_associative inf.sup_right_isotone mult_right_isotone top.extremum by presburger also have "... = w\<^sup>T * (w\<^sup>T\<^sup>\ \ ?F) \ ?G" using assms(2) inf_assoc inf_commute inf_left_commute univalent_comp_left_dist_inf by simp also have "... \ w\<^sup>T * (top * e * F) \ ?G" using 6 by (metis inf.sup_monoid.add_commute inf.sup_right_isotone mult_right_isotone) also have "... \ top * e * F \ ?G" by (metis comp_associative comp_inf_covector mult_left_isotone top.extremum) also have "... = bot" by (metis assms(3) conv_dist_comp conv_involutive conv_top inf_p mult_right_zero univalent_comp_left_dist_inf) finally have 11: "w\<^sup>T\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w = bot" by (simp add: antisym) hence 12: "w\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w = bot" using assms(1) comp_associative conv_bot conv_complement conv_dist_comp conv_dist_inf conv_star_commute star.circ_plus_same by fastforce have "(w \ ?E \ ?F \ ?G)\<^sup>T * top * (w \ ?E \ ?F \ ?G) = ((w \ ?E \ ?G)\<^sup>T \ ?F\<^sup>T) * top * (?F \ (w \ ?E \ ?G))" by (simp add: conv_dist_inf inf_commute inf_left_commute) also have "... = (w \ ?E \ ?G)\<^sup>T * ?F * (?F \ (w \ ?E \ ?G))" by (simp add: covector_inf_comp_3 vector_mult_closed) also have "... = (w \ ?E \ ?G)\<^sup>T * (?F \ ?F\<^sup>T) * (w \ ?E \ ?G)" using covector_comp_inf covector_inf_comp_3 vector_conv_covector vector_mult_closed by simp also have "... = (w \ ?E \ ?G)\<^sup>T * (?F \ ?F\<^sup>T) * (w \ ?E) \ ?G" by (simp add: comp_associative comp_inf_covector) also have "... = (w \ ?E \ ?G)\<^sup>T * (?F \ ?F\<^sup>T) * w \ ?E \ ?G" by (simp add: comp_associative comp_inf_covector) also have "... = (?G\<^sup>T \ (?E\<^sup>T \ w\<^sup>T)) * (?F \ ?F\<^sup>T) * w \ ?E \ ?G" by (simp add: conv_dist_inf inf.left_commute inf.sup_monoid.add_commute) also have "... = ?G\<^sup>T \ (?E\<^sup>T \ w\<^sup>T) * (?F \ ?F\<^sup>T) * w \ ?E \ ?G" by (metis (no_types) comp_associative conv_dist_comp conv_top vector_export_comp) also have "... = ?G\<^sup>T \ ?E\<^sup>T \ w\<^sup>T * (?F \ ?F\<^sup>T) * w \ ?E \ ?G" by (metis (no_types) comp_associative inf_assoc conv_dist_comp conv_top vector_export_comp) also have "... = ?E\<^sup>T \ ?E \ (?G\<^sup>T \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (simp add: inf_assoc inf.left_commute inf.sup_monoid.add_commute) also have "... = w\<^sup>\ * e * top \ ?E \ (?G\<^sup>T \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (simp add: comp_associative conv_dist_comp conv_star_commute) also have "... = w\<^sup>\ * e * ?E \ (?G\<^sup>T \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... \ w\<^sup>\ * w\<^sup>T\<^sup>\ \ (?G\<^sup>T \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (metis assms(3) mult_assoc mult_1_right mult_left_isotone mult_right_isotone inf.sup_left_isotone arc_expanded) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ (-F * e\<^sup>T * top \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (simp add: assms(1) conv_complement conv_dist_comp mult_assoc) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * ?G \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (metis assms(3) arc_top_edge mult_assoc) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * (?F \ top * e * F) * w" by (simp add: assms(1) conv_dist_comp mult_assoc) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * (?F * e * F) * w" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w" by (metis assms(3) arc_top_edge mult_assoc) also have "... = (w\<^sup>+ \ 1 \ w\<^sup>T\<^sup>\) \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w" by (simp add: assms(2) cancel_separate_eq star_left_unfold_equal sup_monoid.add_commute) also have "... = (w\<^sup>+ \ 1 \ w\<^sup>T\<^sup>+) \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w" using star.circ_plus_one star_left_unfold_equal sup_assoc by presburger also have "... = (w\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w) \ (1 \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w) \ (w\<^sup>T\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w)" by (simp add: inf_sup_distrib2) also have "... \ 1" using 11 12 by (simp add: inf.coboundedI1) finally show "(w \ ?E \ ?F \ ?G)\<^sup>T * top * (w \ ?E \ ?F \ ?G) \ 1" by simp have "(w \ -F) * (F \ w\<^sup>T) \ w * w\<^sup>T \ -F * F" by (simp add: mult_isotone) also have "... \ 1 \ -F" using assms(1,2) comp_inf.comp_isotone equivalence_comp_right_complement by auto also have "... = bot" using assms(1) bot_unique pp_isotone pseudo_complement_pp by blast finally have 13: "(w \ -F) * (F \ w\<^sup>T) = bot" by (simp add: antisym) have "w \ ?G \ F * (w \ ?G)" by (metis assms(1) mult_1_left mult_right_dist_sup sup.absorb_iff2) also have "... \ F * (w \ ?G) * w\<^sup>\" by (metis eq_refl le_supE star.circ_back_loop_fixpoint) finally have 14: "w \ ?G \ F * (w \ ?G) * w\<^sup>\" by simp have "w \ top * e * F \ w * (e * F)\<^sup>T * e * F" by (metis (no_types) comp_inf.star_slide dedekind_2 inf_left_commute inf_top_right mult_assoc) also have "... \ F" using 3 assms(1) by (metis comp_associative conv_dist_comp mult_left_isotone preorder_idempotent) finally have "w \ -F \ -(top * e * F)" using order.trans p_shunting_swap pp_increasing by blast also have "... = ?G" by (metis assms(3) comp_mapping_complement conv_dist_comp conv_involutive conv_top) finally have "(w \ -F) * F * (w \ ?G) = (w \ -F \ ?G) * F * (w \ ?G)" by (simp add: inf.absorb1) also have "... \ (w \ -F \ ?G) * F * w" by (simp add: comp_isotone) also have "... \ (w \ -F \ ?G) * forest_components (F \ w) * w" by (simp add: assms(5) mult_left_isotone mult_right_isotone) also have "... \ (w \ -F \ ?G) * (F \ w)\<^sup>T\<^sup>\ * w\<^sup>\ * w" by (simp add: mult_left_isotone mult_right_isotone star_isotone mult_assoc) also have "... \ (w \ -F \ ?G) * (F \ w)\<^sup>T\<^sup>\ * w\<^sup>\" by (simp add: comp_associative mult_right_isotone star.circ_plus_same star.left_plus_below_circ) also have "... = (w \ -F \ ?G) * w\<^sup>\ \ (w \ -F \ ?G) * (F \ w)\<^sup>T\<^sup>+ * w\<^sup>\" by (metis comp_associative inf.sup_monoid.add_assoc mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... \ (w \ -F \ ?G) * w\<^sup>\ \ (w \ -F \ ?G) * (F \ w)\<^sup>T * top" by (metis mult_assoc top_greatest mult_right_isotone sup_right_isotone) also have "... \ (w \ -F \ ?G) * w\<^sup>\ \ (w \ -F) * (F \ w)\<^sup>T * top" using inf.cobounded1 mult_left_isotone sup_right_isotone by blast also have "... \ (w \ ?G) * w\<^sup>\ \ (w \ -F) * (F \ w)\<^sup>T * top" using inf.sup_monoid.add_assoc inf.sup_right_isotone mult_left_isotone sup_commute sup_right_isotone by simp also have "... = (w \ ?G) * w\<^sup>\ \ (w \ -F) * (F \ w\<^sup>T) * top" by (simp add: assms(1) conv_dist_inf) also have "... \ 1 * (w \ ?G) * w\<^sup>\" using 13 by simp also have "... \ F * (w \ ?G) * w\<^sup>\" using assms(1) mult_left_isotone by blast finally have 15: "(w \ -F) * F * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" by simp have "(w \ F) * F * (w \ ?G) \ F * F * (w \ ?G)" by (simp add: mult_left_isotone) also have "... = F * (w \ ?G)" by (simp add: assms(1) preorder_idempotent) also have "... \ F * (w \ ?G) * w\<^sup>\" by (metis eq_refl le_supE star.circ_back_loop_fixpoint) finally have "(w \ F) * F * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" by simp hence "((w \ F) \ (w \ -F)) * F * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" using 15 by (simp add: semiring.distrib_right) hence "w * F * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" by (metis assms(4) maddux_3_11_pp) hence "w * F * (w \ ?G) * w\<^sup>\ \ F * (w \ ?G) * w\<^sup>\" by (metis (full_types) comp_associative mult_left_isotone star.circ_transitive_equal) hence "w\<^sup>\ * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" using 14 by (simp add: mult_assoc star_left_induct) hence 16: "w\<^sup>+ \ ?G \ F * (w \ ?G) * w\<^sup>\" by (simp add: covector_comp_inf covector_mult_closed star.circ_plus_same) have 17: "e\<^sup>T * top * e\<^sup>T \ -F" using assms(8) le_bot triple_schroeder_p by simp hence "(top * e)\<^sup>T * e\<^sup>T \ -F" by (simp add: conv_dist_comp) hence 18: "e\<^sup>T \ ?G" by (metis assms(3) shunt_mapping conv_dist_comp conv_involutive conv_top) have "e\<^sup>T \ -F" using 17 by (simp add: assms(3) arc_top_arc) also have "... \ -1" by (simp add: assms(1) p_antitone) finally have "e\<^sup>T \ w\<^sup>\ \ -1" using assms(9) by simp also have "... \ w\<^sup>+" using shunting_var_p star_left_unfold_equal sup_commute by simp finally have "e\<^sup>T \ w\<^sup>+ \ ?G" using 18 by simp hence "e\<^sup>T \ F * (w \ ?G) * w\<^sup>\" using 16 order_trans by blast also have "... = (F * w \ ?G) * w\<^sup>\" by (simp add: comp_associative comp_inf_covector) finally have "e\<^sup>T * top * e\<^sup>T \ (F * w \ ?G) * w\<^sup>\" by (simp add: assms(3) arc_top_arc) hence "e\<^sup>T * top * (e * top)\<^sup>T \ (F * w \ ?G) * w\<^sup>\" by (metis conv_dist_comp conv_top vector_top_closed mult_assoc) hence "e\<^sup>T * top \ (F * w \ ?G) * w\<^sup>\ * e * top" by (metis assms(3) shunt_bijective mult_assoc) hence "(top * e)\<^sup>T * top \ (F * w \ ?G) * w\<^sup>\ * e * top" by (simp add: conv_dist_comp mult_assoc) hence "top \ top * e * (F * w \ ?G) * w\<^sup>\ * e * top" by (metis assms(3) shunt_mapping conv_dist_comp conv_involutive conv_top mult_assoc) also have "... = top * e * F * w * (w\<^sup>\ * e * top \ ?G\<^sup>T)" by (metis comp_associative comp_inf_vector_1) also have "... = top * (w \ (top * e * F)\<^sup>T) * (w\<^sup>\ * e * top \ ?G\<^sup>T)" by (metis comp_inf_vector_1 inf_top.left_neutral) also have "... = top * (w \ ?F) * (w\<^sup>\ * e * top \ ?G\<^sup>T)" by (simp add: assms(1) conv_dist_comp mult_assoc) also have "... = top * (w \ ?F) * (?E\<^sup>T \ ?G\<^sup>T)" by (simp add: comp_associative conv_dist_comp conv_star_commute) also have "... = top * (w \ ?F \ ?G) * ?E\<^sup>T" by (simp add: comp_associative comp_inf_vector_1) also have "... = top * (w \ ?F \ ?G \ ?E) * top" using comp_inf_vector_1 mult_assoc by simp finally show "top * (w \ ?E \ ?F \ ?G) * top = top" by (simp add: inf_commute inf_left_commute top_le) qed lemma kruskal_edge_arc_1: assumes "e \ --h" and "h \ g" - and "symmetric h" + 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" - using assms(2,3) inf.sup_right_isotone pp_isotone by simp + by (metis assms(2,3) inf.sup_right_isotone pp_isotone conv_isotone) also have "... \ e\<^sup>T * top \ components g" using inf.sup_right_isotone star.circ_increasing by simp also have "... \ e\<^sup>T * top \ forest_components w" using assms(4) comp_inf.mult_right_isotone by simp also have "... = (e\<^sup>T * top \ w\<^sup>T\<^sup>\) * w\<^sup>\" by (simp add: inf_assoc vector_export_comp) also have "... = (e\<^sup>T * top \ 1) * w\<^sup>\ \ (e\<^sup>T * top \ w\<^sup>T\<^sup>+) * w\<^sup>\" by (metis inf_sup_distrib1 semiring.distrib_right star_left_unfold_equal) also have "... \ w\<^sup>\ \ (e\<^sup>T * top \ w\<^sup>T\<^sup>+) * w\<^sup>\" by (metis inf_le2 mult_1_left mult_left_isotone sup_left_isotone) also have "... \ w\<^sup>\ \ (e\<^sup>T * top \ w\<^sup>T) * top" using comp_associative comp_inf.mult_right_isotone sup_right_isotone mult_right_isotone top.extremum vector_export_comp by presburger also have "... = w\<^sup>\" using 1 inf.sup_monoid.add_commute inf_vector_comp by simp finally show ?thesis by simp qed lemma kruskal_edge_between_components_1: assumes "equivalence F" and "mapping (top * e)" shows "F \ -(w \ top * e\<^sup>T * w\<^sup>T\<^sup>\ \ F * e\<^sup>T * top \ top * e * -F)" proof - let ?d = "w \ top * e\<^sup>T * w\<^sup>T\<^sup>\ \ F * e\<^sup>T * top \ top * e * -F" have "?d \ F \ F * e\<^sup>T * top \ F" by (meson inf_le1 inf_le2 le_infI order_trans) also have "... \ (F * e\<^sup>T * top)\<^sup>T * F" by (simp add: mult_assoc vector_restrict_comp_conv) also have "... = top * e * F * F" by (simp add: assms(1) comp_associative conv_dist_comp conv_star_commute) also have "... = top * e * F" using assms(1) preorder_idempotent mult_assoc by fastforce finally have "?d \ F \ top * e * F \ top * e * -F" by (simp add: le_infI1) also have "... = top * e * F \ -(top * e * F)" using assms(2) conv_dist_comp total_conv_surjective comp_mapping_complement by simp finally show ?thesis by (metis inf_p le_bot p_antitone_iff pseudo_complement) qed lemma kruskal_edge_between_components_2: assumes "forest_components f \ -d" and "injective f" and "f \ f\<^sup>T \ w \ w\<^sup>T" shows "f \ f\<^sup>T \ (w \ -d \ -d\<^sup>T) \ (w\<^sup>T \ -d \ -d\<^sup>T)" proof - let ?F = "forest_components f" have "?F\<^sup>T \ -d\<^sup>T" using assms(1) conv_complement conv_order by fastforce hence 1: "?F \ -d\<^sup>T" by (simp add: conv_dist_comp conv_star_commute) have "equivalence ?F" using assms(2) forest_components_equivalence by simp hence "f \ f\<^sup>T \ ?F" by (metis conv_dist_inf forest_components_increasing inf.absorb_iff2 sup.boundedI) also have "... \ -d \ -d\<^sup>T" using 1 assms(1) by simp finally have "f \ f\<^sup>T \ -d \ -d\<^sup>T" by simp thus ?thesis by (metis assms(3) inf_sup_distrib2 le_inf_iff) qed end subsection \Related Structures\ text \ Stone algebras can be expanded to Stone-Kleene relation algebras by reusing some operations. \ sublocale stone_algebra < comp_inf: stone_kleene_relation_algebra where star = "\x . top" and one = top and times = inf and conv = id apply unfold_locales by simp text \ Every bounded linear order can be expanded to a Stone algebra, which can be expanded to a Stone relation algebra, which can be expanded to a Stone-Kleene relation algebra. \ class linorder_stone_kleene_relation_algebra_expansion = linorder_stone_relation_algebra_expansion + star + assumes star_def [simp]: "x\<^sup>\ = top" begin subclass kleene_algebra apply unfold_locales apply simp apply (simp add: min.coboundedI1 min.commute) by (simp add: min.coboundedI1) subclass stone_kleene_relation_algebra apply unfold_locales by simp end text \ A Kleene relation algebra is based on a relation algebra. \ class kleene_relation_algebra = relation_algebra + stone_kleene_relation_algebra end