diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relation_Functions.thy b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relation_Functions.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relation_Functions.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relation_Functions.thy @@ -1,283 +1,292 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Binary Relations\ subsection \Basic Functions\ theory Binary_Relation_Functions imports HOL_Basics_Base begin paragraph \Summary\ text \Basic functions on binary relations.\ definition "rel_comp R S x y \ \z. R x z \ S z y" bundle rel_comp_syntax begin notation rel_comp (infixl "\\" 55) end bundle no_rel_comp_syntax begin no_notation rel_comp (infixl "\\" 55) end unbundle rel_comp_syntax lemma rel_compI [intro]: assumes "R x y" and "S y z" shows "(R \\ S) x z" using assms unfolding rel_comp_def by blast lemma rel_compE [elim]: assumes "(R \\ S) x y" obtains z where "R x z" "S z y" using assms unfolding rel_comp_def by blast lemma rel_comp_assoc: "R \\ (S \\ T) = (R \\ S) \\ T" by (intro ext) blast definition rel_inv :: "('a \ 'b \ bool) \ 'b \ 'a \ bool" where "rel_inv R x y \ R y x" bundle rel_inv_syntax begin notation rel_inv ("(_\)" [1000]) end bundle no_rel_inv_syntax begin no_notation rel_inv ("(_\)" [1000]) end unbundle rel_inv_syntax lemma rel_invI [intro]: assumes "R x y" shows "R\ y x" using assms unfolding rel_inv_def . lemma rel_invD [dest]: assumes "R\ x y" shows "R y x" using assms unfolding rel_inv_def . lemma rel_inv_iff_rel [simp]: "R\ x y \ R y x" by blast lemma rel_inv_comp_eq [simp]: "(R \\ S)\ = S\ \\ R\" by (intro ext) blast lemma rel_inv_inv_eq_self [simp]: "R\\ = R" by blast lemma rel_inv_eq_iff_eq [iff]: "R\ = S\ \ R = S" by (blast dest: fun_cong) definition "in_dom R x \ \y. R x y" lemma in_domI [intro]: assumes "R x y" shows "in_dom R x" using assms unfolding in_dom_def by blast lemma in_domE [elim]: assumes "in_dom R x" obtains y where "R x y" using assms unfolding in_dom_def by blast lemma in_dom_if_in_dom_rel_comp: assumes "in_dom (R \\ S) x" shows "in_dom R x" using assms by blast definition "in_codom R y \ \x. R x y" lemma in_codomI [intro]: assumes "R x y" shows "in_codom R y" using assms unfolding in_codom_def by blast lemma in_codomE [elim]: assumes "in_codom R y" obtains x where "R x y" using assms unfolding in_codom_def by blast lemma in_codom_if_in_codom_rel_comp: assumes "in_codom (R \\ S) y" shows "in_codom S y" using assms by blast lemma in_codom_rel_inv_eq_in_dom [simp]: "in_codom (R\) = in_dom R" by (intro ext) blast lemma in_dom_rel_inv_eq_in_codom [simp]: "in_dom (R\) = in_codom R" by (intro ext) blast definition "in_field R x \ in_dom R x \ in_codom R x" lemma in_field_if_in_dom: assumes "in_dom R x" shows "in_field R x" unfolding in_field_def using assms by blast lemma in_field_if_in_codom: assumes "in_codom R x" shows "in_field R x" unfolding in_field_def using assms by blast lemma in_fieldE [elim]: assumes "in_field R x" obtains (in_dom) x' where "R x x'" | (in_codom) x' where "R x' x" using assms unfolding in_field_def by blast lemma in_fieldE': assumes "in_field R x" obtains (in_dom) "in_dom R x" | (in_codom) "in_codom R x" using assms by blast lemma in_fieldI [intro]: assumes "R x y" shows "in_field R x" "in_field R y" using assms by (auto intro: in_field_if_in_dom in_field_if_in_codom) lemma in_field_iff_in_dom_or_in_codom: "in_field L x \ in_dom L x \ in_codom L x" by blast lemma in_field_rel_inv_eq [simp]: "in_field R\ = in_field R" by (intro ext) auto lemma in_field_compE [elim]: assumes "in_field (R \\ S) x" obtains (in_dom) "in_dom R x" | (in_codom) "in_codom S x" using assms by blast lemma in_field_eq_in_dom_if_in_codom_eq_in_dom: assumes "in_codom R = in_dom R" shows "in_field R = in_dom R" using assms by (intro ext) (auto elim: in_fieldE') definition "rel_if B R x y \ B \ R x y" bundle rel_if_syntax begin notation (output) rel_if (infixl "\" 50) end bundle no_rel_if_syntax begin no_notation (output) rel_if (infixl "\" 50) end unbundle rel_if_syntax lemma rel_if_eq_rel_if_pred [simp]: assumes "B" shows "(rel_if B R) = R" unfolding rel_if_def using assms by blast lemma rel_if_eq_top_if_not_pred [simp]: assumes "\B" shows "(rel_if B R) = (\_ _. True)" unfolding rel_if_def using assms by blast lemma rel_if_if_impI [intro]: assumes "B \ R x y" shows "(rel_if B R) x y" unfolding rel_if_def using assms by blast lemma rel_ifE [elim]: assumes "(rel_if B R) x y" obtains "\B" | "B" "R x y" using assms unfolding rel_if_def by blast lemma rel_ifD: assumes "(rel_if B R) x y" and "B" shows "R x y" using assms by blast -consts restrict_left :: "('a \ 'b \ bool) \ 'c \ 'a \ 'b \ bool" - -definition "restrict_right R P \ (restrict_left R\ P)\" +consts restrict_left :: "'a \ 'b \ 'a" +consts restrict_right :: "'a \ 'b \ 'a" overloading - restrict_left_pred \ "restrict_left :: ('a \ 'b \ bool) \ ('a \ bool) \ 'a \ 'b \ bool" + bin_rel_restrict_left_pred \ + "restrict_left :: ('a \ 'b \ bool) \ ('a \ bool) \ 'a \ 'b \ bool" + bin_rel_restrict_right_pred \ + "restrict_right :: ('a \ 'b \ bool) \ ('b \ bool) \ 'a \ 'b \ bool" begin - definition "restrict_left_pred R P x y \ P x \ R x y" + definition "bin_rel_restrict_left_pred R P x y \ P x \ R x y" + definition "bin_rel_restrict_right_pred R P x y \ P y \ R x y" end bundle restrict_syntax begin notation restrict_left ("(_)\(\<^bsub>_\<^esub>)" [1000]) notation restrict_right ("(_)\(\<^bsub>_\<^esub>)" [1000]) end bundle no_restrict_syntax begin no_notation restrict_left ("(_)\(\<^bsub>_\<^esub>)" [1000]) no_notation restrict_right ("(_)\(\<^bsub>_\<^esub>)" [1000]) end unbundle restrict_syntax -lemma restrict_leftI [intro]: +lemma bin_rel_restrict_leftI [intro]: assumes "R x y" and "P x" shows "R\\<^bsub>P\<^esub> x y" - using assms unfolding restrict_left_pred_def by blast + using assms unfolding bin_rel_restrict_left_pred_def by blast -lemma restrict_leftE [elim]: +lemma bin_rel_restrict_leftE [elim]: assumes "R\\<^bsub>P\<^esub> x y" obtains "P x" "R x y" - using assms unfolding restrict_left_pred_def by blast - -lemma restrict_right_eq: "R\\<^bsub>P\<^esub> = ((R\)\\<^bsub>P\<^esub>)\" - unfolding restrict_right_def .. + using assms unfolding bin_rel_restrict_left_pred_def by blast -lemma rel_inv_restrict_right_rel_inv_eq_restrict_left [simp]: "((R\)\\<^bsub>P\<^esub>)\ = R\\<^bsub>P\<^esub>" - by (simp add: restrict_right_eq) - -lemma restrict_right_iff_restrict_left: "R\\<^bsub>P\<^esub> x y = (R\)\\<^bsub>P\<^esub> y x" - unfolding restrict_right_eq by simp - -lemma restrict_rightI [intro]: +lemma bin_rel_restrict_rightI [intro]: assumes "R x y" and "P y" shows "R\\<^bsub>P\<^esub> x y" - using assms by (auto iff: restrict_right_iff_restrict_left) + using assms unfolding bin_rel_restrict_right_pred_def by blast -lemma restrict_rightE [elim]: +lemma bin_rel_restrict_rightE [elim]: assumes "R\\<^bsub>P\<^esub> x y" obtains "P y" "R x y" - using assms by (auto iff: restrict_right_iff_restrict_left) + using assms unfolding bin_rel_restrict_right_pred_def by blast -lemma rel_inv_restrict_left_inv_restrict_left_eq: +context + fixes R :: "'a \ 'b \ bool" +begin + +lemma bin_rel_restrict_right_eq: "R\\<^bsub>P :: 'b \ bool\<^esub> = ((R\)\\<^bsub>P\<^esub>)\" + by blast + +lemma rel_inv_restrict_right_rel_inv_eq_restrict_left [simp]: "((R\)\\<^bsub>P :: 'a \ bool\<^esub>)\ = R\\<^bsub>P\<^esub>" + by blast + +lemma bin_rel_restrict_right_iff_restrict_left: "R\\<^bsub>P :: 'b \ bool\<^esub> x y \ (R\)\\<^bsub>P\<^esub> y x" + unfolding bin_rel_restrict_right_eq by simp + +end + +lemma rel_inv_bin_rel_restrict_left_inv_bin_rel_restrict_left_eq: fixes R :: "'a \ 'b \ bool" and P :: "'a \ bool" and Q :: "'b \ bool" shows "(((R\\<^bsub>P\<^esub>)\)\\<^bsub>Q\<^esub>)\ = (((R\)\\<^bsub>Q\<^esub>)\)\\<^bsub>P\<^esub>" - by (intro ext iffI restrict_leftI rel_invI) auto + by (intro ext iffI bin_rel_restrict_leftI rel_invI) auto -lemma restrict_left_right_eq_restrict_right_left: +lemma bin_rel_restrict_left_right_eq_restrict_right_left: fixes R :: "'a \ 'b \ bool" and P :: "'a \ bool" and Q :: "'b \ bool" shows "R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub> = R\\<^bsub>Q\<^esub>\\<^bsub>P\<^esub>" - unfolding restrict_right_eq - by (fact rel_inv_restrict_left_inv_restrict_left_eq) + unfolding bin_rel_restrict_right_eq + by (fact rel_inv_bin_rel_restrict_left_inv_bin_rel_restrict_left_eq) -lemma in_dom_restrict_leftI [intro]: +lemma in_dom_bin_rel_restrict_leftI [intro]: assumes "R x y" and "P x" shows "in_dom R\\<^bsub>P\<^esub> x" using assms by blast -lemma in_dom_restrict_left_if_in_dom: +lemma in_dom_bin_rel_restrict_left_if_in_dom: assumes "in_dom R x" and "P x" shows "in_dom R\\<^bsub>P\<^esub> x" using assms by blast -lemma in_dom_restrict_leftE [elim]: +lemma in_dom_bin_rel_restrict_leftE [elim]: assumes "in_dom R\\<^bsub>P\<^esub> x" obtains y where "P x" "R x y" using assms by blast -lemma in_codom_restrict_leftI [intro]: +lemma in_codom_bin_rel_restrict_leftI [intro]: assumes "R x y" and "P x" shows "in_codom R\\<^bsub>P\<^esub> y" using assms by blast -lemma in_codom_restrict_leftE [elim]: +lemma in_codom_bin_rel_restrict_leftE [elim]: assumes "in_codom R\\<^bsub>P\<^esub> y" obtains x where "P x" "R x y" using assms by blast definition "rel_bimap f g (R :: 'a \ 'b \ bool) x y \ R (f x) (g y)" lemma rel_bimap_eq [simp]: "rel_bimap f g R x y = R (f x) (g y)" unfolding rel_bimap_def by simp definition "rel_map f R \ rel_bimap f f R" lemma rel_bimap_self_eq_rel_map [simp]: "rel_bimap f f R = rel_map f R" unfolding rel_map_def by simp lemma rel_map_eq [simp]: "rel_map f R x y = R (f x) (f y)" by (simp only: rel_bimap_self_eq_rel_map[symmetric] rel_bimap_eq) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Lattice.thy b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Lattice.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Lattice.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Lattice.thy @@ -1,90 +1,92 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Lattice\ theory Binary_Relations_Lattice imports Binary_Relations_Order_Base HOL.Boolean_Algebras begin paragraph \Summary\ text \Basic results about the lattice structure on binary relations.\ lemma rel_infI [intro]: assumes "R x y" and "S x y" shows "(R \ S) x y" using assms by (rule inf2I) lemma rel_infE [elim]: assumes "(R \ S) x y" obtains "R x y" "S x y" using assms by (rule inf2E) lemma rel_infD: assumes "(R \ S) x y" shows "R x y" and "S x y" using assms by auto lemma in_dom_rel_infI [intro]: assumes "R x y" and "S x y" shows "in_dom (R \ S) x" using assms by blast lemma in_dom_rel_infE [elim]: assumes "in_dom (R \ S) x" obtains y where "R x y" "S x y" using assms by blast lemma in_codom_rel_infI [intro]: assumes "R x y" and "S x y" shows "in_codom (R \ S) y" using assms by blast lemma in_codom_rel_infE [elim]: assumes "in_codom (R \ S) y" obtains x where "R x y" "S x y" using assms by blast lemma in_field_eq_in_dom_sup_in_codom: "in_field L = (in_dom L \ in_codom L)" by (intro ext) (simp add: in_field_iff_in_dom_or_in_codom) -lemma in_dom_restrict_left_eq [simp]: "in_dom R\\<^bsub>P\<^esub> = (in_dom R \ P)" - by (intro ext) auto - -lemma in_codom_restrict_left_eq [simp]: "in_codom R\\<^bsub>P\<^esub> = (in_codom R \ P)" +lemma in_dom_bin_rel_restrict_left_eq [simp]: "in_dom R\\<^bsub>P\<^esub> = (in_dom R \ P)" by (intro ext) auto -lemma restrict_left_restrict_left_eq [simp]: - fixes R :: "'a \ _" and P Q :: "'a \ bool" +lemma in_codom_bin_rel_restrict_left_eq [simp]: "in_codom R\\<^bsub>P\<^esub> = (in_codom R \ P)" + by (intro ext) auto + +lemma bin_rel_restrict_left_restrict_left_eq [simp]: + fixes R :: "'a \ 'b \ bool" and P Q :: "'a \ bool" shows "R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub>" - by (intro ext iffI restrict_leftI) auto + by (intro ext iffI bin_rel_restrict_leftI) auto -lemma restrict_left_restrict_right_eq [simp]: +lemma bin_rel_restrict_left_restrict_right_eq [simp]: fixes R :: "'a \ 'b \ bool" and P :: "'a \ bool" and Q :: "'b \ bool" shows "R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub>" - by (intro ext iffI restrict_leftI restrict_rightI) auto + by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI) auto -lemma restrict_right_restrict_left_eq [simp]: +lemma bin_rel_restrict_right_restrict_left_eq [simp]: fixes R :: "'a \ 'b \ bool" and P :: "'b \ bool" and Q :: "'a \ bool" shows "R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub>" - by (intro ext iffI restrict_leftI restrict_rightI) auto + by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI) auto -lemma restrict_right_restrict_right_eq [simp]: +lemma bin_rel_restrict_right_restrict_right_eq [simp]: fixes R :: "'a \ 'b \ bool" and P Q :: "'b \ bool" shows "R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub>" by (intro ext iffI) auto -lemma restrict_left_sup_eq [simp]: "(R :: 'a \ _)\\<^bsub>((P :: 'a \ bool) \ Q)\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub> " - by (intro antisym le_relI) (auto elim!: restrict_leftE) +lemma bin_rel_restrict_left_sup_eq [simp]: + "(R :: 'a \ 'b \ bool)\\<^bsub>((P :: 'a \ bool) \ Q)\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub> " + by (intro antisym le_relI) (auto elim!: bin_rel_restrict_leftE) -lemma restrict_left_inf_eq [simp]: "(R :: 'a \ _)\\<^bsub>((P :: 'a \ bool) \ Q)\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub> " - by (intro antisym le_relI) (auto elim!: restrict_leftE) +lemma bin_rel_restrict_left_inf_eq [simp]: + "(R :: 'a \ 'b \ bool)\\<^bsub>((P :: 'a \ bool) \ Q)\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub> " + by (intro antisym le_relI) (auto elim!: bin_rel_restrict_leftE) lemma inf_rel_bimap_and_eq_restrict_left_restrict_right: "R \ (rel_bimap P Q (\)) = R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub>" by (intro ext) auto end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order_Base.thy b/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order_Base.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order_Base.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order_Base.thy @@ -1,33 +1,42 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Order\ theory Binary_Relations_Order_Base imports Binary_Relation_Functions HOL.Orderings begin lemma le_relI [intro]: assumes "\x y. R x y \ S x y" shows "R \ S" using assms by (rule predicate2I) lemma le_relD [dest]: assumes "R \ S" and "R x y" shows "S x y" using assms by (rule predicate2D) lemma le_relE: assumes "R \ S" and "R x y" obtains "S x y" using assms by blast lemma rel_inv_le_rel_inv_iff [iff]: "R\ \ S\ \ R \ S" by blast -lemma restrict_left_top_eq [simp]: "(R :: 'a \ _)\\<^bsub>(\ :: 'a \ bool)\<^esub> = R" +lemma bin_rel_restrict_left_le_self: "(R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'a \ bool)\<^esub> \ R" + by blast + +lemma bin_rel_restrict_right_le_self: "(R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'b \ bool)\<^esub> \ R" + by blast + +lemma bin_rel_restrict_left_top_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>(\ :: 'a \ bool)\<^esub> = R" + by (intro ext) auto + +lemma bin_rel_restrict_right_top_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>(\ :: 'b \ bool)\<^esub> = R" by (intro ext) auto end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Antisymmetric.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Antisymmetric.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Antisymmetric.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Antisymmetric.thy @@ -1,65 +1,65 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Antisymmetric\ theory Binary_Relations_Antisymmetric imports Binary_Relation_Functions HOL_Syntax_Bundles_Lattices begin consts antisymmetric_on :: "'a \ ('b \ 'b \ bool) \ bool" overloading antisymmetric_on_pred \ "antisymmetric_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool" begin definition "antisymmetric_on_pred P R \ \x y. P x \ P y \ R x y \ R y x \ x = y" end lemma antisymmetric_onI [intro]: assumes "\x y. P x \ P y \ R x y \ R y x \ x = y" shows "antisymmetric_on P R" unfolding antisymmetric_on_pred_def using assms by blast lemma antisymmetric_onD: assumes "antisymmetric_on P R" and "P x" "P y" and "R x y" "R y x" shows "x = y" using assms unfolding antisymmetric_on_pred_def by blast -definition "antisymmetric (R :: 'a \ _) \ antisymmetric_on (\ :: 'a \ bool) R" +definition "(antisymmetric :: ('a \ _) \ _) \ antisymmetric_on (\ :: 'a \ bool)" lemma antisymmetric_eq_antisymmetric_on: - "antisymmetric (R :: 'a \ _) = antisymmetric_on (\ :: 'a \ bool) R" + "(antisymmetric :: ('a \ _) \ _) = antisymmetric_on (\ :: 'a \ bool)" unfolding antisymmetric_def .. lemma antisymmetricI [intro]: assumes "\x y. R x y \ R y x \ x = y" shows "antisymmetric R" unfolding antisymmetric_eq_antisymmetric_on using assms by (intro antisymmetric_onI) lemma antisymmetricD: assumes "antisymmetric R" and "R x y" "R y x" shows "x = y" using assms unfolding antisymmetric_eq_antisymmetric_on by (auto dest: antisymmetric_onD) lemma antisymmetric_on_if_antisymmetric: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "antisymmetric R" shows "antisymmetric_on P R" using assms by (intro antisymmetric_onI) (blast dest: antisymmetricD) lemma antisymmetric_if_antisymmetric_on_in_field: assumes "antisymmetric_on (in_field R) R" shows "antisymmetric R" using assms by (intro antisymmetricI) (blast dest: antisymmetric_onD) corollary antisymmetric_on_in_field_iff_antisymmetric [simp]: "antisymmetric_on (in_field R) R \ antisymmetric R" using antisymmetric_if_antisymmetric_on_in_field antisymmetric_on_if_antisymmetric by blast end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Injective.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Injective.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Injective.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Injective.thy @@ -1,106 +1,106 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Injective\ theory Binary_Relations_Injective imports Binary_Relation_Functions HOL_Syntax_Bundles_Lattices ML_Unification.ML_Unification_HOL_Setup begin consts rel_injective_on :: "'a \ ('b \ 'c \ bool) \ bool" overloading rel_injective_on_pred \ "rel_injective_on :: ('a \ bool) \ ('a \ 'b \ bool) \ bool" begin definition "rel_injective_on_pred P R \ \x x' y. P x \ P x' \ R x y \ R x' y \ x = x'" end lemma rel_injective_onI [intro]: assumes "\x x' y. P x \ P x' \ R x y \ R x' y \ x = x'" shows "rel_injective_on P R" unfolding rel_injective_on_pred_def using assms by blast lemma rel_injective_onD: assumes "rel_injective_on P R" and "P x" "P x'" and "R x y" "R x' y" shows "x = x'" using assms unfolding rel_injective_on_pred_def by blast consts rel_injective_at :: "'a \ ('b \ 'c \ bool) \ bool" overloading rel_injective_at_pred \ "rel_injective_at :: ('a \ bool) \ ('b \ 'a \ bool) \ bool" begin definition "rel_injective_at_pred P R \ \x x' y. P y \ R x y \ R x' y \ x = x'" end lemma rel_injective_atI [intro]: assumes "\x x' y. P y \ R x y \ R x' y \ x = x'" shows "rel_injective_at P R" unfolding rel_injective_at_pred_def using assms by blast lemma rel_injective_atD: assumes "rel_injective_at P R" and "P y" and "R x y" "R x' y" shows "x = x'" using assms unfolding rel_injective_at_pred_def by blast -definition "rel_injective (R :: 'a \ _) \ rel_injective_on (\ :: 'a \ bool) R" +definition "(rel_injective :: ('a \ _) \ _) \ rel_injective_on (\ :: 'a \ bool)" lemma rel_injective_eq_rel_injective_on: - "rel_injective (R :: 'a \ _) = rel_injective_on (\ :: 'a \ bool) R" + "(rel_injective :: ('a \ _) \ _) = rel_injective_on (\ :: 'a \ bool)" unfolding rel_injective_def .. lemma rel_injectiveI [intro]: assumes "\x x' y. R x y \ R x' y \ x = x'" shows "rel_injective R" unfolding rel_injective_eq_rel_injective_on using assms by blast lemma rel_injectiveD: assumes "rel_injective R" and "R x y" "R x' y" shows "x = x'" using assms unfolding rel_injective_eq_rel_injective_on by (auto dest: rel_injective_onD) lemma rel_injective_eq_rel_injective_at: "rel_injective (R :: 'a \ 'b \ bool) = rel_injective_at (\ :: 'b \ bool) R" by (intro iffI rel_injectiveI) (auto dest: rel_injective_atD rel_injectiveD) lemma rel_injective_on_if_rel_injective: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "rel_injective R" shows "rel_injective_on P R" using assms by (blast dest: rel_injectiveD) lemma rel_injective_at_if_rel_injective: fixes P :: "'a \ bool" and R :: "'b \ 'a \ bool" assumes "rel_injective R" shows "rel_injective_at P R" using assms by (blast dest: rel_injectiveD) lemma rel_injective_if_rel_injective_on_in_dom: assumes "rel_injective_on (in_dom R) R" shows "rel_injective R" using assms by (blast dest: rel_injective_onD) lemma rel_injective_if_rel_injective_at_in_codom: assumes "rel_injective_at (in_codom R) R" shows "rel_injective R" using assms by (blast dest: rel_injective_atD) corollary rel_injective_on_in_dom_iff_rel_injective [simp]: "rel_injective_on (in_dom R) R \ rel_injective R" using rel_injective_if_rel_injective_on_in_dom rel_injective_on_if_rel_injective by blast corollary rel_injective_at_in_codom_iff_rel_injective [iff]: "rel_injective_at (in_codom R) R \ rel_injective R" using rel_injective_if_rel_injective_at_in_codom rel_injective_at_if_rel_injective by blast end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Irreflexive.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Irreflexive.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Irreflexive.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Irreflexive.thy @@ -1,51 +1,51 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Irreflexive\ theory Binary_Relations_Irreflexive imports Binary_Relation_Functions HOL_Syntax_Bundles_Lattices begin consts irreflexive_on :: "'a \ ('b \ 'b \ bool) \ bool" overloading irreflexive_on_pred \ "irreflexive_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool" begin definition "irreflexive_on_pred P R \ \x. P x \ \(R x x)" end lemma irreflexive_onI [intro]: assumes "\x. P x \ \(R x x)" shows "irreflexive_on P R" using assms unfolding irreflexive_on_pred_def by blast lemma irreflexive_onD [dest]: assumes "irreflexive_on P R" and "P x" shows "\(R x x)" using assms unfolding irreflexive_on_pred_def by blast -definition "irreflexive (R :: 'a \ _) \ irreflexive_on (\ :: 'a \ bool) R" +definition "(irreflexive :: ('a \ _) \ _) \ irreflexive_on (\ :: 'a \ bool)" lemma irreflexive_eq_irreflexive_on: - "irreflexive (R :: 'a \ _) = irreflexive_on (\ :: 'a \ bool) R" + "(irreflexive :: ('a \ _) \ _) = irreflexive_on (\ :: 'a \ bool)" unfolding irreflexive_def .. lemma irreflexiveI [intro]: assumes "\x. \(R x x)" shows "irreflexive R" unfolding irreflexive_eq_irreflexive_on using assms by (intro irreflexive_onI) lemma irreflexiveD: assumes "irreflexive R" shows "\(R x x)" using assms unfolding irreflexive_eq_irreflexive_on by auto lemma irreflexive_on_if_irreflexive: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "irreflexive R" shows "irreflexive_on P R" using assms by (intro irreflexive_onI) (blast dest: irreflexiveD) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Left_Total.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Left_Total.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Left_Total.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Left_Total.thy @@ -1,62 +1,61 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Left Total\ theory Binary_Relations_Left_Total imports Binary_Relation_Functions HOL_Syntax_Bundles_Lattices begin consts left_total_on :: "'a \ ('b \ 'c \ bool) \ bool" overloading left_total_on_pred \ "left_total_on :: ('a \ bool) \ ('a \ 'b \ bool) \ bool" begin definition "left_total_on_pred P R \ \x. P x \ in_dom R x" end lemma left_total_onI [intro]: assumes "\x. P x \ in_dom R x" shows "left_total_on P R" unfolding left_total_on_pred_def using assms by blast lemma left_total_onE [elim]: assumes "left_total_on P R" and "P x" obtains y where "R x y" using assms unfolding left_total_on_pred_def by blast lemma in_dom_if_left_total_on: assumes "left_total_on P R" and "P x" shows "in_dom R x" using assms by blast -definition "left_total (R :: 'a \ _) \ left_total_on (\ :: 'a \ bool) R" +definition "(left_total :: ('a \ _) \ _) \ left_total_on (\ :: 'a \ bool)" -lemma left_total_eq_left_total_on: - "left_total (R :: 'a \ _) = left_total_on (\ :: 'a \ bool) R" +lemma left_total_eq_left_total_on: "(left_total :: ('a \ _) \ _) = left_total_on (\ :: 'a \ bool)" unfolding left_total_def .. lemma left_totalI [intro]: assumes "\x. in_dom R x" shows "left_total R" unfolding left_total_eq_left_total_on using assms by (intro left_total_onI) lemma left_totalE: assumes "left_total R" obtains y where "R x y" using assms unfolding left_total_eq_left_total_on by (blast intro: top1I) lemma in_dom_if_left_total: assumes "left_total R" shows "in_dom R x" using assms by (blast elim: left_totalE) lemma left_total_on_if_left_total: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "left_total R" shows "left_total_on P R" using assms by (intro left_total_onI) (blast dest: in_dom_if_left_total) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Reflexive.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Reflexive.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Reflexive.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Reflexive.thy @@ -1,105 +1,104 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Reflexive\ theory Binary_Relations_Reflexive imports Functions_Monotone begin consts reflexive_on :: "'a \ ('b \ 'b \ bool) \ bool" overloading reflexive_on_pred \ "reflexive_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool" begin definition "reflexive_on_pred P R \ \x. P x \ R x x" end lemma reflexive_onI [intro]: assumes "\x. P x \ R x x" shows "reflexive_on P R" using assms unfolding reflexive_on_pred_def by blast lemma reflexive_onD [dest]: assumes "reflexive_on P R" and "P x" shows "R x x" using assms unfolding reflexive_on_pred_def by blast lemma le_in_dom_if_reflexive_on: assumes "reflexive_on P R" shows "P \ in_dom R" using assms by blast lemma le_in_codom_if_reflexive_on: assumes "reflexive_on P R" shows "P \ in_codom R" using assms by blast lemma in_codom_eq_in_dom_if_reflexive_on_in_field: assumes "reflexive_on (in_field R) R" shows "in_codom R = in_dom R" using assms by blast lemma reflexive_on_rel_inv_iff_reflexive_on [iff]: "reflexive_on P R\ \ reflexive_on (P :: 'a \ bool) (R :: 'a \ _)" by blast lemma antimono_reflexive_on [iff]: "antimono (\(P :: 'a \ bool). reflexive_on P (R :: 'a \ _))" by (intro antimonoI) auto lemma reflexive_on_if_le_pred_if_reflexive_on: fixes P P' :: "'a \ bool" and R :: "'a \ _" assumes "reflexive_on P R" and "P' \ P" shows "reflexive_on P' R" using assms by blast lemma reflexive_on_sup_eq [simp]: "(reflexive_on :: ('a \ bool) \ ('a \ _) \ _) ((P :: 'a \ bool) \ Q) = reflexive_on P \ reflexive_on Q" by (intro ext iffI reflexive_onI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on) -lemma reflexive_on_iff_eq_restrict_left_le: +lemma reflexive_on_iff_eq_restrict_le: "reflexive_on (P :: 'a \ bool) (R :: 'a \ _) \ ((=)\\<^bsub>P\<^esub> \ R)" by blast -definition "reflexive (R :: 'a \ _) \ reflexive_on (\ :: 'a \ bool) R" +definition "(reflexive :: ('a \ _) \ _) \ reflexive_on (\ :: 'a \ bool)" -lemma reflexive_eq_reflexive_on: - "reflexive (R :: 'a \ _) = reflexive_on (\ :: 'a \ bool) R" +lemma reflexive_eq_reflexive_on: "(reflexive :: ('a \ _) \ _) = reflexive_on (\ :: 'a \ bool)" unfolding reflexive_def .. lemma reflexiveI [intro]: assumes "\x. R x x" shows "reflexive R" unfolding reflexive_eq_reflexive_on using assms by (intro reflexive_onI) lemma reflexiveD: assumes "reflexive R" shows "R x x" using assms unfolding reflexive_eq_reflexive_on by (blast intro: top1I) lemma reflexive_on_if_reflexive: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "reflexive R" shows "reflexive_on P R" using assms by (intro reflexive_onI) (blast dest: reflexiveD) lemma reflexive_rel_inv_iff_reflexive [iff]: "reflexive R\ \ reflexive R" by (blast dest: reflexiveD) lemma reflexive_iff_eq_le: "reflexive R \ ((=) \ R)" - unfolding reflexive_eq_reflexive_on reflexive_on_iff_eq_restrict_left_le + unfolding reflexive_eq_reflexive_on reflexive_on_iff_eq_restrict_le by simp paragraph \Instantiations\ lemma reflexive_eq: "reflexive (=)" by (rule reflexiveI) (rule refl) lemma reflexive_top: "reflexive \" by (rule reflexiveI) auto end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Right_Unique.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Right_Unique.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Right_Unique.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Right_Unique.thy @@ -1,136 +1,136 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Right Unique\ theory Binary_Relations_Right_Unique imports Binary_Relations_Injective HOL_Syntax_Bundles_Lattices begin consts right_unique_on :: "'a \ ('b \ 'c \ bool) \ bool" overloading right_unique_on_pred \ "right_unique_on :: ('a \ bool) \ ('a \ 'b \ bool) \ bool" begin definition "right_unique_on_pred P R \ \x y y'. P x \ R x y \ R x y' \ y = y'" end lemma right_unique_onI [intro]: assumes "\x y y'. P x \ R x y \ R x y' \ y = y'" shows "right_unique_on P R" using assms unfolding right_unique_on_pred_def by blast lemma right_unique_onD: assumes "right_unique_on P R" and "P x" and "R x y" "R x y'" shows "y = y'" using assms unfolding right_unique_on_pred_def by blast consts right_unique_at :: "'a \ ('b \ 'c \ bool) \ bool" overloading right_unique_at_pred \ "right_unique_at :: ('a \ bool) \ ('b \ 'a \ bool) \ bool" begin definition "right_unique_at_pred P R \ \x y y'. P y \ P y' \ R x y \ R x y' \ y = y'" end lemma right_unique_atI [intro]: assumes "\x y y'. P y \ P y' \ R x y \ R x y' \ y = y'" shows "right_unique_at P R" using assms unfolding right_unique_at_pred_def by blast lemma right_unique_atD: assumes "right_unique_at P R" and "P y" and "P y'" and "R x y" "R x y'" shows "y = y'" using assms unfolding right_unique_at_pred_def by blast lemma right_unique_at_rel_inv_iff_rel_injective_on [iff]: "right_unique_at (P :: 'a \ bool) (R\ :: 'b \ 'a \ bool) \ rel_injective_on P R" by (blast dest: right_unique_atD rel_injective_onD) lemma rel_injective_on_rel_inv_iff_right_unique_at [iff]: "rel_injective_on (P :: 'a \ bool) (R\ :: 'a \ 'b \ bool) \ right_unique_at P R" by (blast dest: right_unique_atD rel_injective_onD) lemma right_unique_on_rel_inv_iff_rel_injective_at [iff]: "right_unique_on (P :: 'a \ bool) (R\ :: 'a \ 'b \ bool) \ rel_injective_at P R" by (blast dest: right_unique_onD rel_injective_atD) lemma rel_injective_at_rel_inv_iff_right_unique_on [iff]: "rel_injective_at (P :: 'b \ bool) (R\ :: 'a \ 'b \ bool) \ right_unique_on P R" by (blast dest: right_unique_onD rel_injective_atD) -definition "right_unique (R :: 'a \ _) \ right_unique_on (\ :: 'a \ bool) R" +definition "(right_unique :: ('a \ _) \ _) \ right_unique_on (\ :: 'a \ bool)" lemma right_unique_eq_right_unique_on: - "right_unique (R :: 'a \ _) = right_unique_on (\ :: 'a \ bool) R" + "(right_unique :: ('a \ _) \ _) = right_unique_on (\ :: 'a \ bool)" unfolding right_unique_def .. lemma right_uniqueI [intro]: assumes "\x y y'. R x y \ R x y' \ y = y'" shows "right_unique R" unfolding right_unique_eq_right_unique_on using assms by blast lemma right_uniqueD: assumes "right_unique R" and "R x y" "R x y'" shows "y = y'" using assms unfolding right_unique_eq_right_unique_on by (auto dest: right_unique_onD) lemma right_unique_eq_right_unique_at: "right_unique (R :: 'a \ 'b \ bool) = right_unique_at (\ :: 'b \ bool) R" by (intro iffI right_uniqueI) (auto dest: right_unique_atD right_uniqueD) lemma right_unique_on_if_right_unique: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "right_unique R" shows "right_unique_on P R" using assms by (blast dest: right_uniqueD) lemma right_unique_at_if_right_unique: fixes P :: "'a \ bool" and R :: "'b \ 'a \ bool" assumes "right_unique R" shows "right_unique_at P R" using assms by (blast dest: right_uniqueD) lemma right_unique_if_right_unique_on_in_dom: assumes "right_unique_on (in_dom R) R" shows "right_unique R" using assms by (blast dest: right_unique_onD) lemma right_unique_if_right_unique_at_in_codom: assumes "right_unique_at (in_codom R) R" shows "right_unique R" using assms by (blast dest: right_unique_atD) corollary right_unique_on_in_dom_iff_right_unique [iff]: "right_unique_on (in_dom R) R \ right_unique R" using right_unique_if_right_unique_on_in_dom right_unique_on_if_right_unique by blast corollary right_unique_at_in_codom_iff_right_unique [iff]: "right_unique_at (in_codom R) R \ right_unique R" using right_unique_if_right_unique_at_in_codom right_unique_at_if_right_unique by blast lemma right_unique_rel_inv_iff_rel_injective [iff]: "right_unique R\ \ rel_injective R" by (blast dest: right_uniqueD rel_injectiveD) lemma rel_injective_rel_inv_iff_right_unique [iff]: "rel_injective R\ \ right_unique R" by (blast dest: right_uniqueD rel_injectiveD) -paragraph \Instantiatiats\ +paragraph \Instantiations\ lemma right_unique_eq: "right_unique (=)" by (rule right_uniqueI) blast end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Surjective.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Surjective.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Surjective.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Surjective.thy @@ -1,79 +1,79 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Surjective\ theory Binary_Relations_Surjective imports Binary_Relations_Left_Total HOL_Syntax_Bundles_Lattices begin consts rel_surjective_at :: "'a \ ('b \ 'c \ bool) \ bool" overloading rel_surjective_at_pred \ "rel_surjective_at :: ('a \ bool) \ ('b \ 'a \ bool) \ bool" begin definition "rel_surjective_at_pred P R \ \y. P y \ in_codom R y" end lemma rel_surjective_atI [intro]: assumes "\y. P y \ in_codom R y" shows "rel_surjective_at P R" unfolding rel_surjective_at_pred_def using assms by blast lemma rel_surjective_atE [elim]: assumes "rel_surjective_at P R" and "P y" obtains x where "R x y" using assms unfolding rel_surjective_at_pred_def by blast -lemma in_codom_if_rel_surjective_at_on: +lemma in_codom_if_rel_surjective_at: assumes "rel_surjective_at P R" and "P y" shows "in_codom R y" using assms by blast lemma rel_surjective_at_rel_inv_iff_left_total_on [iff]: "rel_surjective_at (P :: 'a \ bool) (R\ :: 'b \ 'a \ bool) \ left_total_on P R" by fast lemma left_total_on_rel_inv_iff_rel_surjective_at [iff]: "left_total_on (P :: 'a \ bool) (R\ :: 'a \ 'b \ bool) \ rel_surjective_at P R" by fast -definition "rel_surjective (R :: _ \ 'a \ _) \ rel_surjective_at (\ :: 'a \ bool) R" +definition "(rel_surjective :: (_ \ 'a \ _) \ _) \ rel_surjective_at (\ :: 'a \ bool)" lemma rel_surjective_eq_rel_surjective_at: - "rel_surjective (R :: _ \ 'a \ _) = rel_surjective_at (\ :: 'a \ bool) R" + "(rel_surjective :: (_ \ 'a \ _) \ _) = rel_surjective_at (\ :: 'a \ bool)" unfolding rel_surjective_def .. lemma rel_surjectiveI: assumes "\y. in_codom R y" shows "rel_surjective R" unfolding rel_surjective_eq_rel_surjective_at using assms by (intro rel_surjective_atI) lemma rel_surjectiveE: assumes "rel_surjective R" obtains x where "R x y" using assms unfolding rel_surjective_eq_rel_surjective_at by (blast intro: top1I) -lemma in_codom_if_rel_surjective_at: +lemma in_codom_if_rel_surjective: assumes "rel_surjective R" shows "in_codom R y" using assms by (blast elim: rel_surjectiveE) lemma rel_surjective_rel_inv_iff_left_total [iff]: "rel_surjective R\ \ left_total R" unfolding rel_surjective_eq_rel_surjective_at left_total_eq_left_total_on by simp lemma left_total_rel_inv_iff_rel_surjective [iff]: "left_total R\ \ rel_surjective R" unfolding rel_surjective_eq_rel_surjective_at left_total_eq_left_total_on by simp lemma rel_surjective_at_if_surjective: fixes P :: "'a \ bool" and R :: "_ \ 'a \ _" assumes "rel_surjective R" shows "rel_surjective_at P R" - using assms by (intro rel_surjective_atI) (blast dest: in_codom_if_rel_surjective_at) + using assms by (intro rel_surjective_atI) (blast dest: in_codom_if_rel_surjective) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Symmetric.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Symmetric.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Symmetric.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Symmetric.thy @@ -1,106 +1,106 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Symmetric\ theory Binary_Relations_Symmetric imports Functions_Monotone begin consts symmetric_on :: "'a \ ('b \ 'b \ bool) \ bool" overloading symmetric_on_pred \ "symmetric_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool" begin definition "symmetric_on_pred P R \ \x y. P x \ P y \ R x y \ R y x" end lemma symmetric_onI [intro]: assumes "\x y. P x \ P y \ R x y \ R y x" shows "symmetric_on P R" unfolding symmetric_on_pred_def using assms by blast lemma symmetric_onD: assumes "symmetric_on P R" and "P x" "P y" and "R x y" shows "R y x" using assms unfolding symmetric_on_pred_def by blast lemma symmetric_on_rel_inv_iff_symmetric_on [iff]: "symmetric_on P R\ \ symmetric_on (P :: 'a \ bool) (R :: 'a \ _)" by (blast dest: symmetric_onD) lemma antimono_symmetric_on [iff]: "antimono (\(P :: 'a \ bool). symmetric_on P (R :: 'a \ _))" by (intro antimonoI) (auto dest: symmetric_onD) lemma symmetric_on_if_le_pred_if_symmetric_on: fixes P P' :: "'a \ bool" and R :: "'a \ _" assumes "symmetric_on P R" and "P' \ P" shows "symmetric_on P' R" using assms by (blast dest: symmetric_onD) -definition "symmetric (R :: 'a \ _) \ symmetric_on (\ :: 'a \ bool) R" +definition "(symmetric :: ('a \ _) \ _) \ symmetric_on (\ :: 'a \ bool)" lemma symmetric_eq_symmetric_on: - "symmetric (R :: 'a \ _) = symmetric_on (\ :: 'a \ bool) R" + "(symmetric :: ('a \ _) \ _) = symmetric_on (\ :: 'a \ bool)" unfolding symmetric_def .. lemma symmetricI [intro]: assumes "\x y. R x y \ R y x" shows "symmetric R" unfolding symmetric_eq_symmetric_on using assms by (intro symmetric_onI) lemma symmetricD: assumes "symmetric R" and "R x y" shows "R y x" using assms unfolding symmetric_eq_symmetric_on by (auto dest: symmetric_onD) lemma symmetric_on_if_symmetric: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "symmetric R" shows "symmetric_on P R" using assms by (intro symmetric_onI) (blast dest: symmetricD) lemma symmetric_rel_inv_iff_symmetric [iff]: "symmetric R\ \ symmetric R" by (blast dest: symmetricD) lemma rel_inv_eq_self_if_symmetric [simp]: assumes "symmetric R" shows "R\ = R" using assms by (blast dest: symmetricD) lemma rel_iff_rel_if_symmetric: assumes "symmetric R" shows "R x y \ R y x" using assms by (blast dest: symmetricD) lemma symmetric_if_rel_inv_eq_self: assumes "R\ = R" shows "symmetric R" by (intro symmetricI, subst assms[symmetric]) simp lemma symmetric_iff_rel_inv_eq_self: "symmetric R \ R\ = R" using rel_inv_eq_self_if_symmetric symmetric_if_rel_inv_eq_self by blast lemma symmetric_if_symmetric_on_in_field: assumes "symmetric_on (in_field R) R" shows "symmetric R" using assms by (intro symmetricI) (blast dest: symmetric_onD) corollary symmetric_on_in_field_iff_symmetric [simp]: "symmetric_on (in_field R) R \ symmetric R" using symmetric_if_symmetric_on_in_field symmetric_on_if_symmetric by blast paragraph \Instantiations\ lemma symmetric_eq [iff]: "symmetric (=)" by (rule symmetricI) (rule sym) lemma symmetric_top: "symmetric \" by (rule symmetricI) auto end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Transitive.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Transitive.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Transitive.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Transitive.thy @@ -1,115 +1,114 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Transitive\ theory Binary_Relations_Transitive imports Binary_Relation_Functions Functions_Monotone begin consts transitive_on :: "'a \ ('b \ 'b \ bool) \ bool" overloading transitive_on_pred \ "transitive_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool" begin definition "transitive_on_pred P R \ \x y z. P x \ P y \ P z \ R x y \ R y z \ R x z" end lemma transitive_onI [intro]: assumes "\x y z. P x \ P y \ P z \ R x y \ R y z \ R x z" shows "transitive_on P R" unfolding transitive_on_pred_def using assms by blast lemma transitive_onD: assumes "transitive_on P R" and "P x" "P y" "P z" and "R x y" "R y z" shows "R x z" using assms unfolding transitive_on_pred_def by blast lemma transitive_on_if_rel_comp_self_imp: assumes "\x y. P x \ P y \ (R \\ R) x y \ R x y" shows "transitive_on P R" proof (rule transitive_onI) fix x y z assume "R x y" "R y z" then have "(R \\ R) x z" by (intro rel_compI) moreover assume "P x" "P y" "P z" ultimately show "R x z" by (simp only: assms) qed lemma transitive_on_rel_inv_iff_transitive_on [iff]: "transitive_on P R\ \ transitive_on (P :: 'a \ bool) (R :: 'a \ _)" by (auto intro!: transitive_onI dest: transitive_onD) lemma antimono_transitive_on [iff]: "antimono (\(P :: 'a \ bool). transitive_on P (R :: 'a \ _))" by (intro antimonoI) (auto dest: transitive_onD) lemma transitive_on_if_le_pred_if_transitive_on: fixes P P' :: "'a \ bool" and R :: "'a \ _" assumes "transitive_on P R" and "P' \ P" shows "transitive_on P' R" using assms by (auto dest: transitive_onD) -definition "transitive (R :: 'a \ _) \ transitive_on (\ :: 'a \ bool) R" +definition "(transitive :: ('a \ _) \ _) \ transitive_on (\ :: 'a \ bool)" lemma transitive_eq_transitive_on: - "transitive (R :: 'a \ _) = transitive_on (\ :: 'a \ bool) R" + "(transitive :: ('a \ _) \ _) = transitive_on (\ :: 'a \ bool)" unfolding transitive_def .. lemma transitiveI [intro]: assumes "\x y z. R x y \ R y z \ R x z" shows "transitive R" unfolding transitive_eq_transitive_on using assms by (intro transitive_onI) lemma transitiveD [dest]: assumes "transitive R" and "R x y" "R y z" shows "R x z" using assms unfolding transitive_eq_transitive_on by (auto dest: transitive_onD) lemma transitive_on_if_transitive: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "transitive R" shows "transitive_on P R" using assms by (intro transitive_onI) blast lemma transitive_if_rel_comp_le_self: assumes "R \\ R \ R" shows "transitive R" using assms unfolding transitive_eq_transitive_on by (intro transitive_on_if_rel_comp_self_imp) blast lemma rel_comp_le_self_if_transitive: assumes "transitive R" shows "R \\ R \ R" using assms by blast corollary transitive_iff_rel_comp_le_self: "transitive R \ R \\ R \ R" using transitive_if_rel_comp_le_self rel_comp_le_self_if_transitive by blast lemma transitive_if_transitive_on_in_field: assumes "transitive_on (in_field R) R" shows "transitive R" using assms by (intro transitiveI) (blast dest: transitive_onD) corollary transitive_on_in_field_iff_transitive [simp]: "transitive_on (in_field R) R \ transitive R" using transitive_if_transitive_on_in_field transitive_on_if_transitive by blast -lemma transitive_rel_inv_iff_transitive [iff]: - "transitive R\ \ transitive R" - by (auto intro!: transitiveI) +lemma transitive_rel_inv_iff_transitive [iff]: "transitive R\ \ transitive R" + by fast paragraph \Instantiations\ lemma transitive_eq: "transitive (=)" by (rule transitiveI) (rule trans) lemma transitive_top: "transitive \" by (rule transitiveI) auto end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Restricted_Equality.thy b/thys/Transport/HOL_Basics/Binary_Relations/Restricted_Equality.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Restricted_Equality.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Restricted_Equality.thy @@ -1,98 +1,68 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Restricted Equality\ theory Restricted_Equality imports Binary_Relations_Order_Base Binary_Relation_Functions Equivalence_Relations Partial_Orders begin paragraph \Summary\ -text \Introduces the concept of restricted equalities. +text \Introduces notations and theorems for restricted equalities. An equality @{term "(=)"} can be restricted to only apply to a subset of its -elements. The restriction can be formulated, for example, by a predicate or a -set.\ - -consts eq_restrict :: "'a \ 'b \ 'b \ bool" +elements. The restriction can be formulated, for example, by a predicate or a set.\ bundle eq_restrict_syntax begin syntax - "_eq_restrict" :: "'a \ ('a \ 'a \ bool) \ 'a \ bool" ("(_) =(\<^bsub>_\<^esub>) (_)" [51,51,51] 50) -notation eq_restrict ("'(=(\<^bsub>_\<^esub>)')") + "_eq_restrict_infix" :: "'a \ ('a \ bool) \ 'a \ bool" ("(_) =(\<^bsub>_\<^esub>) (_)" [51,51,51] 50) + "_eq_restrict" :: "('a \ bool) \ 'a \ bool" ("'(=(\<^bsub>_\<^esub>)')") end bundle no_eq_restrict_syntax begin no_syntax - "_eq_restrict" :: "'a \ ('a \ 'a \ bool) \ 'a \ bool" ("(_) =(\<^bsub>_\<^esub>) (_)" [51,51,51] 50) -no_notation eq_restrict ("'(=(\<^bsub>_\<^esub>)')") + "_eq_restrict_infix" :: "'a \ ('a \ bool) \ 'a \ bool" ("(_) =(\<^bsub>_\<^esub>) (_)" [51,51,51] 50) + "_eq_restrict" :: "('a \ bool) \ 'a \ bool" ("'(=(\<^bsub>_\<^esub>)')") end unbundle eq_restrict_syntax translations - "x =\<^bsub>P\<^esub> y" \ "CONST eq_restrict P x y" - -overloading - eq_restrict_pred \ "eq_restrict :: ('a \ bool) \ 'a \ 'a \ bool" -begin - definition "eq_restrict_pred (P :: 'a \ bool) \ ((=) :: 'a \ _)\\<^bsub>P\<^esub>" -end - -lemma eq_restrict_eq_eq_restrict_left: "((=\<^bsub>P :: 'a \ bool\<^esub>) :: 'a \ _) = (=)\\<^bsub>P\<^esub>" - unfolding eq_restrict_pred_def by simp - -lemma eq_restrictI [intro]: - assumes "x = y" - and "P x" - shows "x =\<^bsub>P\<^esub> y" - unfolding eq_restrict_eq_eq_restrict_left using assms by auto - -lemma eq_restrictE [elim]: - assumes "x =\<^bsub>P\<^esub> y" - obtains "P x" "y = x" - using assms unfolding eq_restrict_eq_eq_restrict_left by auto - -lemma eq_restrict_iff: "x =\<^bsub>P\<^esub> y \ y = x \ P x" by auto - -lemma eq_restrict_le_eq: "((=\<^bsub>P :: 'a \ bool\<^esub>) :: 'a \ _) \ (=)" - by (intro le_relI) auto - -lemma eq_restrict_top_eq_eq [simp]: "(=\<^bsub>\ :: 'a \ bool\<^esub>) = ((=) :: 'a \ _)" - unfolding eq_restrict_eq_eq_restrict_left by simp + "(=\<^bsub>P\<^esub>)" \ "CONST restrict_left (=) P" + "x =\<^bsub>P\<^esub> y" \ "CONST restrict_left (=) P x y" lemma in_dom_eq_restrict_eq [simp]: "in_dom (=\<^bsub>P\<^esub>) = P" by auto lemma in_codom_eq_restrict_eq [simp]: "in_codom (=\<^bsub>P\<^esub>) = P" by auto lemma in_field_eq_restrict_eq [simp]: "in_field (=\<^bsub>P\<^esub>) = P" by auto paragraph \Order Properties\ context fixes P :: "'a \ bool" begin context begin lemma reflexive_on_eq_restrict: "reflexive_on P ((=\<^bsub>P\<^esub>) :: 'a \ _)" by auto lemma transitive_eq_restrict: "transitive ((=\<^bsub>P\<^esub>) :: 'a \ _)" by auto lemma symmetric_eq_restrict: "symmetric ((=\<^bsub>P\<^esub>) :: 'a \ _)" by auto lemma antisymmetric_eq_restrict: "antisymmetric ((=\<^bsub>P\<^esub>) :: 'a \ _)" by auto end context begin lemma preorder_on_eq_restrict: "preorder_on P ((=\<^bsub>P\<^esub>) :: 'a \ _)" using reflexive_on_eq_restrict transitive_eq_restrict by auto lemma partial_equivalence_rel_eq_restrict: "partial_equivalence_rel ((=\<^bsub>P\<^esub>) :: 'a \ _)" using symmetric_eq_restrict transitive_eq_restrict by auto end lemma partial_order_on_eq_restrict: "partial_order_on P ((=\<^bsub>P\<^esub>) :: 'a \ _)" using preorder_on_eq_restrict antisymmetric_eq_restrict by auto lemma equivalence_rel_on_eq_restrict: "equivalence_rel_on P ((=\<^bsub>P\<^esub>) :: 'a \ _)" using partial_equivalence_rel_eq_restrict reflexive_on_eq_restrict by blast end end diff --git a/thys/Transport/HOL_Basics/Functions/Function_Relators.thy b/thys/Transport/HOL_Basics/Functions/Function_Relators.thy --- a/thys/Transport/HOL_Basics/Functions/Function_Relators.thy +++ b/thys/Transport/HOL_Basics/Functions/Function_Relators.thy @@ -1,163 +1,163 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Relators\ theory Function_Relators imports Binary_Relation_Functions Functions_Base Predicates_Lattice begin paragraph \Summary\ text \Introduces the concept of function relators. The slogan of function relators is "related functions map related inputs to related outputs".\ definition "Dep_Fun_Rel_rel R S f g \ \x y. R x y \ S x y (f x) (g y)" abbreviation "Fun_Rel_rel R S \ Dep_Fun_Rel_rel R (\_ _. S)" definition "Dep_Fun_Rel_pred P R f g \ \x. P x \ R x (f x) (g x)" abbreviation "Fun_Rel_pred P R \ Dep_Fun_Rel_pred P (\_. R)" bundle Dep_Fun_Rel_syntax begin syntax "_Fun_Rel_rel" :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("(_) \ (_)" [41, 40] 40) "_Dep_Fun_Rel_rel" :: "idt \ idt \ ('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _] \ (_)" [41, 41, 41, 40] 40) "_Dep_Fun_Rel_rel_if" :: "idt \ idt \ ('a \ 'b \ bool) \ bool \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _/ |/ _] \ (_)" [41, 41, 41, 41, 40] 40) "_Fun_Rel_pred" :: "('a \ bool) \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_] \ (_)" [41, 40] 40) "_Dep_Fun_Rel_pred" :: "idt \ ('a \ bool) \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_/ \/ _] \ (_)" [41, 41, 40] 40) "_Dep_Fun_Rel_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_/ \/ _/ |/ _] \ (_)" [41, 41, 41, 40] 40) end bundle no_Dep_Fun_Rel_syntax begin no_syntax "_Fun_Rel_rel" :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("(_) \ (_)" [41, 40] 40) "_Dep_Fun_Rel_rel" :: "idt \ idt \ ('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _] \ (_)" [41, 41, 41, 40] 40) "_Dep_Fun_Rel_rel_if" :: "idt \ idt \ ('a \ 'b \ bool) \ bool \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _/ |/ _] \ (_)" [41, 41, 41, 41, 40] 40) "_Fun_Rel_pred" :: "('a \ bool) \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_] \ (_)" [41, 40] 40) "_Dep_Fun_Rel_pred" :: "idt \ ('a \ bool) \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_/ \/ _] \ (_)" [41, 41, 40] 40) "_Dep_Fun_Rel_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_/ \/ _/ |/ _] \ (_)" [41, 41, 41, 40] 40) end unbundle Dep_Fun_Rel_syntax translations "R \ S" \ "CONST Fun_Rel_rel R S" "[x y \ R] \ S" \ "CONST Dep_Fun_Rel_rel R (\x y. S)" "[x y \ R | B] \ S" \ "CONST Dep_Fun_Rel_rel R (\x y. CONST rel_if B S)" "[P] \ R" \ "CONST Fun_Rel_pred P R" "[x \ P] \ R" \ "CONST Dep_Fun_Rel_pred P (\x. R)" "[x \ P | B] \ R" \ "CONST Dep_Fun_Rel_pred P (\x. CONST rel_if B R)" lemma Dep_Fun_Rel_relI [intro]: assumes "\x y. R x y \ S x y (f x) (g y)" shows "([x y \ R] \ S x y) f g" unfolding Dep_Fun_Rel_rel_def using assms by blast lemma Dep_Fun_Rel_relD: assumes "([x y \ R] \ S x y) f g" and "R x y" shows "S x y (f x) (g y)" using assms unfolding Dep_Fun_Rel_rel_def by blast lemma Dep_Fun_Rel_relE [elim]: assumes "([x y \ R] \ S x y) f g" and "R x y" obtains "S x y (f x) (g y)" using assms unfolding Dep_Fun_Rel_rel_def by blast lemma Dep_Fun_Rel_predI [intro]: assumes "\x. P x \ R x (f x) (g x)" shows "([x \ P] \ R x) f g" unfolding Dep_Fun_Rel_pred_def using assms by blast lemma Dep_Fun_Rel_predD: assumes "([x \ P] \ R x) f g" and "P x" shows "R x (f x) (g x)" using assms unfolding Dep_Fun_Rel_pred_def by blast lemma Dep_Fun_Rel_predE [elim]: assumes "([x \ P] \ R x) f g" and "P x" obtains "R x (f x) (g x)" using assms unfolding Dep_Fun_Rel_pred_def by blast lemma rel_inv_Dep_Fun_Rel_rel_eq [simp]: "([x y \ R] \ S x y)\ = ([y x \ R\] \ (S x y)\)" by (intro ext) auto lemma rel_inv_Dep_Fun_Rel_pred_eq [simp]: "([x \ P] \ R x)\ = ([x \ P] \ (R x)\)" by (intro ext) auto lemma Dep_Fun_Rel_pred_eq_Dep_Fun_Rel_rel: "([x \ P] \ R x) = ([x _ \ (((\) P) \ (=))] \ R x)" by (intro ext) (auto intro!: Dep_Fun_Rel_predI Dep_Fun_Rel_relI) lemma Fun_Rel_eq_eq_eq [simp]: "((=) \ (=)) = (=)" by (intro ext) auto paragraph \Composition\ lemma Dep_Fun_Rel_rel_compI: assumes Dep_Fun_Rel1: "([x y \ R] \ S x y) f g" and Dep_Fun_Rel2: "\x y. R x y \ ([x' y' \ T x y] \ U x y x' y') f' g'" and le: "\x y. R x y \ S x y (f x) (g y) \ T x y (f x) (g y)" shows "([x y \ R] \ U x y (f x) (g y)) (f' \ f) (g' \ g)" using assms by (intro Dep_Fun_Rel_relI) (auto 6 0) corollary Dep_Fun_Rel_rel_compI': assumes "([x y \ R] \ S x y) f g" and "\x y. R x y \ ([x' y' \ S x y] \ T x y x' y') f' g'" shows "([x y \ R] \ T x y (f x) (g y)) (f' \ f) (g' \ g)" using assms by (intro Dep_Fun_Rel_rel_compI) lemma Dep_Fun_Rel_pred_comp_Dep_Fun_Rel_rel_compI: assumes Dep_Fun_Rel1: "([x \ P] \ R x) f g" and Dep_Fun_Rel2: "\x. P x \ ([x' y' \ S x] \ T x x' y') f' g'" and le: "\x. P x \ R x (f x) (g x) \ S x (f x) (g x)" shows "([x \ P] \ T x (f x) (g x)) (f' \ f) (g' \ g)" using assms by (intro Dep_Fun_Rel_predI) (auto 6 0) corollary Dep_Fun_Rel_pred_comp_Dep_Fun_Rel_rel_compI': assumes "([x \ P] \ R x) f g" and "\x. P x \ ([x' y' \ R x] \ S x x' y') f' g'" shows "([x \ P] \ S x (f x) (g x)) (f' \ f) (g' \ g)" using assms by (intro Dep_Fun_Rel_pred_comp_Dep_Fun_Rel_rel_compI) paragraph \Restrictions\ lemma restrict_left_Dep_Fun_Rel_rel_restrict_left_eq: fixes R :: "'a1 \ 'a2 \ bool" and S :: "'a1 \ 'a2 \ 'b1 \ 'b2 \ bool" and P :: "'a1 \ 'a2 \ 'b1 \ bool" assumes "\f x y. Q f \ R x y \ P x y (f x)" shows "([x y \ R] \ (S x y)\\<^bsub>P x y\<^esub>)\\<^bsub>Q\<^esub> = ([x y \ R] \ S x y)\\<^bsub>Q\<^esub>" - using assms by (intro ext iffI restrict_leftI Dep_Fun_Rel_relI) + using assms by (intro ext iffI bin_rel_restrict_leftI Dep_Fun_Rel_relI) (auto dest!: Dep_Fun_Rel_relD) lemma restrict_right_Dep_Fun_Rel_rel_restrict_right_eq: fixes R :: "'a1 \ 'a2 \ bool" and S :: "'a1 \ 'a2 \ 'b1 \ 'b2 \ bool" and P :: "'a1 \ 'a2 \ 'b2 \ bool" assumes "\f x y. Q f \ R x y \ P x y (f y)" shows "(([x y \ R] \ (S x y)\\<^bsub>P x y\<^esub>)\\<^bsub>Q\<^esub>) = (([x y \ R] \ S x y)\\<^bsub>Q\<^esub>)" - unfolding restrict_right_eq + unfolding bin_rel_restrict_right_eq using assms restrict_left_Dep_Fun_Rel_rel_restrict_left_eq[where ?R="R\" and ?S="\y x. (S x y)\"] by simp end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Bijection.thy b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Bijection.thy --- a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Bijection.thy +++ b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Bijection.thy @@ -1,150 +1,149 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Bijection\ theory Functions_Bijection imports Functions_Inverse Functions_Monotone begin consts bijection_on :: "'a \ 'b \ ('c \ 'd) \ ('d \ 'c) \ bool" overloading - bijection_on_pred \ "bijection_on :: ('a \ bool) \ ('b \ bool) \ - ('a \ 'b) \ ('b \ 'a) \ bool" + bijection_on_pred \ "bijection_on :: ('a \ bool) \ ('b \ bool) \ ('a \ 'b) \ ('b \ 'a) \ bool" begin definition "bijection_on_pred P P' f g \ ([P] \\<^sub>m P') f \ ([P'] \\<^sub>m P) g \ inverse_on P f g \ inverse_on P' g f" end lemma bijection_onI [intro]: assumes "([P] \\<^sub>m P') f" and "([P'] \\<^sub>m P) g" and "inverse_on P f g" and "inverse_on P' g f" shows "bijection_on P P' f g" using assms unfolding bijection_on_pred_def by blast lemma bijection_onE: assumes "bijection_on P P' f g" obtains "([P] \\<^sub>m P') f" "([P'] \\<^sub>m P) g" "inverse_on P f g" "inverse_on P' g f" using assms unfolding bijection_on_pred_def by blast context fixes P :: "'a \ bool" and P' :: "'b \ bool" and f :: "'a \ 'b" begin lemma mono_wrt_pred_if_bijection_on_left: assumes "bijection_on P P' f g" shows "([P] \\<^sub>m P') f" using assms by (elim bijection_onE) lemma mono_wrt_pred_if_bijection_on_right: assumes "bijection_on P P' f g" shows "([P'] \\<^sub>m P) g" using assms by (elim bijection_onE) lemma bijection_on_pred_right: assumes "bijection_on P P' f g" and "P x" shows "P' (f x)" using assms by (blast elim: bijection_onE) lemma bijection_on_pred_left: assumes "bijection_on P P' f g" and "P' y" shows "P (g y)" using assms by (blast elim: bijection_onE) lemma inverse_on_if_bijection_on_left_right: assumes "bijection_on P P' f g" shows "inverse_on P f g" using assms by (elim bijection_onE) lemma inverse_on_if_bijection_on_right_left: assumes "bijection_on P P' f g" shows "inverse_on P' g f" using assms by (elim bijection_onE) lemma bijection_on_left_right_eq_self: assumes "bijection_on P P' f g" and "P x" shows "g (f x) = x" using assms inverse_on_if_bijection_on_left_right by (intro inverse_onD) lemma bijection_on_right_left_eq_self': assumes "bijection_on P P' f g" and "P' y" shows "f (g y) = y" using assms inverse_on_if_bijection_on_right_left by (intro inverse_onD) lemma bijection_on_right_left_if_bijection_on_left_right: assumes "bijection_on P P' f g" shows "bijection_on P' P g f" using assms by (auto elim: bijection_onE) lemma injective_on_if_bijection_on_left: assumes "bijection_on P P' f g" shows "injective_on P f" using assms by (intro injective_on_if_inverse_on inverse_on_if_bijection_on_left_right) lemma injective_on_if_bijection_on_right: assumes "bijection_on P P' f g" shows "injective_on P' g" by (intro injective_on_if_inverse_on) (fact inverse_on_if_bijection_on_right_left[OF assms]) end -definition "bijection (f :: 'a \ 'b) \ bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool) f" +definition "(bijection :: ('a \ 'b) \ _) \ bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool)" lemma bijection_eq_bijection_on: - "bijection (f :: 'a \ 'b) = bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool) f" + "(bijection :: ('a \ 'b) \ _) = bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool)" unfolding bijection_def .. lemma bijectionI [intro]: assumes "inverse f g" and "inverse g f" shows "bijection f g" unfolding bijection_eq_bijection_on using assms by (intro bijection_onI inverse_on_if_inverse dep_mono_wrt_predI) simp_all lemma bijectionE [elim]: assumes "bijection f g" obtains "inverse f g" "inverse g f" using assms unfolding bijection_eq_bijection_on inverse_eq_inverse_on by (blast elim: bijection_onE) lemma inverse_if_bijection_left_right: assumes "bijection f g" shows "inverse f g" using assms by (elim bijectionE) lemma inverse_if_bijection_right_left: assumes "bijection f g" shows "inverse g f" using assms by (elim bijectionE) lemma bijection_right_left_if_bijection_left_right: assumes "bijection f g" shows "bijection g f" using assms by auto paragraph \Instantiations\ lemma bijection_on_self_id: fixes P :: "'a \ bool" shows "bijection_on P P (id :: 'a \ _) id" by (intro bijection_onI inverse_onI dep_mono_wrt_predI) simp_all end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Injective.thy b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Injective.thy --- a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Injective.thy +++ b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Injective.thy @@ -1,58 +1,57 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Injective\ theory Functions_Injective imports Functions_Base HOL_Syntax_Bundles_Lattices begin consts injective_on :: "'a \ ('b \ 'c) \ bool" overloading injective_on_pred \ "injective_on :: ('a \ bool) \ ('a \ 'b) \ bool" begin definition "injective_on_pred P f \ \x x'. P x \ P x' \ f x = f x' \ x = x'" end lemma injective_onI [intro]: assumes "\x x'. P x \ P x' \ f x = f x' \ x = x'" shows "injective_on P f" unfolding injective_on_pred_def using assms by blast lemma injective_onD: assumes "injective_on P f" and "P x" "P x'" and "f x = f x'" shows "x = x'" using assms unfolding injective_on_pred_def by blast -definition "injective (f :: 'a \ _) \ injective_on (\ :: 'a \ bool) f" +definition "(injective :: ('a \ _) \ _) \ injective_on (\ :: 'a \ bool)" -lemma injective_eq_injective_on: - "injective (f :: 'a \ _) = injective_on (\ :: 'a \ bool) f" +lemma injective_eq_injective_on: "(injective :: ('a \ _) \ _) = injective_on (\ :: 'a \ bool)" unfolding injective_def .. lemma injectiveI [intro]: assumes "\x x'. f x = f x' \ x = x'" shows "injective f" unfolding injective_eq_injective_on using assms by (intro injective_onI) lemma injectiveD: assumes "injective f" and "f x = f x'" shows "x = x'" using assms unfolding injective_eq_injective_on by (auto dest: injective_onD) lemma injective_on_if_injective: fixes P :: "'a \ bool" and f :: "'a \ _" assumes "injective f" shows "injective_on P f" using assms by (intro injective_onI) (blast dest: injectiveD) paragraph \Instantiations\ lemma injective_id: "injective id" by auto end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Inverse.thy b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Inverse.thy --- a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Inverse.thy +++ b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Inverse.thy @@ -1,62 +1,61 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Inverse\ theory Functions_Inverse imports Functions_Injective begin consts inverse_on :: "'a \ ('b \ 'c) \ ('c \ 'b) \ bool" overloading inverse_on_pred \ "inverse_on :: ('a \ bool) \ ('a \ 'b) \ ('b \ 'a) \ bool" begin definition "inverse_on_pred P f g \ \x. P x \ g (f x) = x" end lemma inverse_onI [intro]: assumes "\x. P x \ g (f x) = x" shows "inverse_on P f g" unfolding inverse_on_pred_def using assms by blast lemma inverse_onD: assumes "inverse_on P f g" and "P x" shows "g (f x) = x" using assms unfolding inverse_on_pred_def by blast lemma injective_on_if_inverse_on: assumes inv: "inverse_on (P :: 'a \ bool) (f :: 'a \ _) g" shows "injective_on P f" proof (rule injective_onI) fix x x' assume Px: "P x" and Px': "P x'" and f_x_eq_f_x': "f x = f x'" from inv have "x = g (f x)" using Px by (intro inverse_onD[symmetric]) also have "... = g (f x')" by (simp only: f_x_eq_f_x') also have "... = x'" using inv Px' by (intro inverse_onD) finally show "x = x'" . qed -definition "inverse (f :: 'a \ _) \ inverse_on (\ :: 'a \ bool) f" +definition "(inverse :: ('a \ _) \ _) \ inverse_on (\ :: 'a \ bool)" -lemma inverse_eq_inverse_on: - "inverse (f :: 'a \ _) = inverse_on (\ :: 'a \ bool) f" +lemma inverse_eq_inverse_on: "(inverse :: ('a \ _) \ _) = inverse_on (\ :: 'a \ bool)" unfolding inverse_def .. lemma inverseI [intro]: assumes "\x. g (f x) = x" shows "inverse f g" unfolding inverse_eq_inverse_on using assms by (intro inverse_onI) lemma inverseD: assumes "inverse f g" shows "g (f x) = x" using assms unfolding inverse_eq_inverse_on by (auto dest: inverse_onD) lemma inverse_on_if_inverse: fixes P :: "'a \ bool" and f :: "'a \ 'b" assumes "inverse f g" shows "inverse_on P f g" using assms by (intro inverse_onI) (blast dest: inverseD) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Monotone.thy b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Monotone.thy --- a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Monotone.thy +++ b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Monotone.thy @@ -1,345 +1,340 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Monotonicity\ theory Functions_Monotone imports Binary_Relations_Order_Base Function_Relators Predicates begin paragraph \Summary\ text \Introduces the concept of monotone functions. A function is monotone if it is related to itself - see \<^term>\Dep_Fun_Rel_rel\.\ declare le_funI[intro] declare le_funE[elim] definition "dep_mono_wrt_rel R S f \ ([x y \ R] \ S x y) f f" abbreviation "mono_wrt_rel R S \ dep_mono_wrt_rel R (\_ _. S)" definition "dep_mono_wrt_pred P Q f \ ([x \ P] \ (\_. Q x)) f f" abbreviation "mono_wrt_pred P Q \ dep_mono_wrt_pred P (\_. Q)" bundle dep_mono_wrt_syntax begin syntax "_mono_wrt_rel" :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("(_) \\<^sub>m (_)" [41, 40] 40) "_dep_mono_wrt_rel" :: "idt \ idt \ ('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ _/ \/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) "_dep_mono_wrt_rel_if" :: "idt \ idt \ ('a \ 'a \ bool) \ bool \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ _/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 41, 40] 40) "_mono_wrt_pred" :: "('a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_] \\<^sub>m (_)" [41, 40] 40) "_dep_mono_wrt_pred" :: "idt \ ('a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ \/ _] \\<^sub>m (_)" [41, 41, 40] 40) - (*TODO: this only works if we introduce a pred_if constant first*) - (* "_dep_mono_wrt_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'b \ bool) \ *) - (* ('a \ 'b) \ bool" ("[_/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) *) + "_dep_mono_wrt_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'b \ bool) \ + ('a \ 'b) \ bool" ("[_/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) end bundle no_dep_mono_wrt_syntax begin no_syntax "_mono_wrt_rel" :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("(_) \\<^sub>m (_)" [41, 40] 40) "_dep_mono_wrt_rel" :: "idt \ idt \ ('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ _/ \/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) "_dep_mono_wrt_rel_if" :: "idt \ idt \ ('a \ 'a \ bool) \ bool \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ _/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 41, 40] 40) "_mono_wrt_pred" :: "('a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_] \\<^sub>m (_)" [41, 40] 40) "_dep_mono_wrt_pred" :: "idt \ ('a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ \/ _] \\<^sub>m (_)" [41, 41, 40] 40) - (* "_dep_mono_wrt_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'b \ bool) \ *) - (* ('a \ 'b) \ bool" ("[_/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) *) + "_dep_mono_wrt_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'b \ bool) \ + ('a \ 'b) \ bool" ("[_/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) end unbundle dep_mono_wrt_syntax translations "R \\<^sub>m S" \ "CONST mono_wrt_rel R S" "[x y \ R] \\<^sub>m S" \ "CONST dep_mono_wrt_rel R (\x y. S)" "[x y \ R | B] \\<^sub>m S" \ "CONST dep_mono_wrt_rel R (\x y. CONST rel_if B S)" "[P] \\<^sub>m Q" \ "CONST mono_wrt_pred P Q" "[x \ P] \\<^sub>m Q" \ "CONST dep_mono_wrt_pred P (\x. Q)" - (* "[x \ P | B] \\<^sub>m Q" \ "CONST dep_mono_wrt_pred P (\x. CONST rel_if B Q)" *) + "[x \ P | B] \\<^sub>m Q" \ "CONST dep_mono_wrt_pred P (\x. CONST pred_if B Q)" lemma dep_mono_wrt_relI [intro]: assumes "\x y. R x y \ S x y (f x) (f y)" shows "([x y \ R] \\<^sub>m S x y) f" using assms unfolding dep_mono_wrt_rel_def by blast lemma dep_mono_wrt_relE [elim]: assumes "([x y \ R] \\<^sub>m S x y) f" - and "R x y" - obtains "S x y (f x) (f y)" + obtains "\x y. R x y \ S x y (f x) (f y)" using assms unfolding dep_mono_wrt_rel_def by blast lemma dep_mono_wrt_relD: assumes "([x y \ R] \\<^sub>m S x y) f" and "R x y" shows "S x y (f x) (f y)" using assms by blast lemma dep_mono_wrt_predI [intro]: assumes "\x. P x \ Q x (f x)" shows "([x \ P] \\<^sub>m Q x) f" using assms unfolding dep_mono_wrt_pred_def by blast lemma dep_mono_wrt_predE [elim]: assumes "([x \ P] \\<^sub>m Q x) f" - and "P x" - obtains "Q x (f x)" + obtains "\x. P x \ Q x (f x)" using assms unfolding dep_mono_wrt_pred_def by blast lemma dep_mono_wrt_predD: assumes "([x \ P] \\<^sub>m Q x) f" and "P x" shows "Q x (f x)" using assms by blast lemma dep_mono_wrt_rel_if_Dep_Fun_Rel_rel_self: assumes "([x y \ R] \ S x y) f f" shows "([x y \ R] \\<^sub>m S x y) f" using assms by blast lemma dep_mono_wrt_pred_if_Dep_Fun_Rel_pred_self: assumes "([x \ P] \ (\_. Q x)) f f" shows "([x \ P] \\<^sub>m Q x) f" using assms by blast lemma Dep_Fun_Rel_rel_self_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" shows "([x y \ R] \ S x y) f f" using assms by blast lemma Dep_Fun_Rel_pred_self_if_dep_mono_wrt_pred: assumes "([x \ P] \\<^sub>m Q x) f" shows "([x \ P] \ (\_. Q x)) f f" using assms by blast corollary Dep_Fun_Rel_rel_self_iff_dep_mono_wrt_rel: "([x y \ R] \ S x y) f f \ ([x y \ R] \\<^sub>m S x y) f" using dep_mono_wrt_rel_if_Dep_Fun_Rel_rel_self Dep_Fun_Rel_rel_self_if_dep_mono_wrt_rel by blast corollary Dep_Fun_Rel_pred_self_iff_dep_mono_wrt_pred: "([x \ P] \ (\_. Q x)) f f \ ([x \ P] \\<^sub>m Q x) f" using dep_mono_wrt_pred_if_Dep_Fun_Rel_pred_self Dep_Fun_Rel_pred_self_if_dep_mono_wrt_pred by blast lemma dep_mono_wrt_rel_inv_eq [simp]: "([y x \ R\] \\<^sub>m (S x y)\) = ([x y \ R] \\<^sub>m S x y)" - by (intro ext) auto + by (intro ext) force lemma in_dom_if_rel_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" and "R x y" shows "in_dom (S x y) (f x)" using assms by (intro in_domI) blast corollary in_dom_if_in_dom_if_mono_wrt_rel: assumes "(R \\<^sub>m S) f" shows "([in_dom R] \\<^sub>m in_dom S) f" using assms in_dom_if_rel_if_dep_mono_wrt_rel by fast lemma in_codom_if_rel_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" and "R x y" shows "in_codom (S x y) (f y)" using assms by (intro in_codomI) blast corollary in_codom_if_in_codom_if_mono_wrt_rel: assumes "(R \\<^sub>m S) f" shows "([in_codom R] \\<^sub>m in_codom S) f" using assms in_dom_if_rel_if_dep_mono_wrt_rel by fast corollary in_field_if_in_field_if_mono_wrt_rel: assumes "(R \\<^sub>m S) f" shows "([in_field R] \\<^sub>m in_field S) f" using assms by (intro dep_mono_wrt_predI) blast lemma le_rel_map_if_mono_wrt_rel: assumes "(R \\<^sub>m S) f" shows "R \ rel_map f S" using assms by (intro le_relI) auto lemma le_pred_map_if_mono_wrt_pred: assumes "([P] \\<^sub>m Q) f" shows "P \ pred_map f Q" using assms by (intro le_predI) auto lemma mono_wrt_rel_if_le_rel_map: assumes "R \ rel_map f S" shows "(R \\<^sub>m S) f" using assms by (intro dep_mono_wrt_relI) auto lemma mono_wrt_pred_if_le_pred_map: assumes "P \ pred_map f Q" shows "([P] \\<^sub>m Q) f" using assms by (intro dep_mono_wrt_predI) auto corollary mono_wrt_rel_iff_le_rel_map: "(R \\<^sub>m S) f \ R \ rel_map f S" using mono_wrt_rel_if_le_rel_map le_rel_map_if_mono_wrt_rel by auto corollary mono_wrt_pred_iff_le_pred_map: "([P] \\<^sub>m Q) f \ P \ pred_map f Q" using mono_wrt_pred_if_le_pred_map le_pred_map_if_mono_wrt_pred by auto definition "mono \ ((\) \\<^sub>m (\))" definition "antimono \ ((\) \\<^sub>m (\))" lemma monoI [intro]: assumes "\x y. x \ y \ f x \ f y" shows "mono f" unfolding mono_def using assms by blast lemma monoE [elim]: assumes "mono f" - and "x \ y" - obtains "f x \ f y" + obtains "\x y. x \ y \ f x \ f y" using assms unfolding mono_def by blast lemma monoD: assumes "mono f" and "x \ y" shows "f x \ f y" using assms by blast lemma antimonoI [intro]: assumes "\x y. x \ y \ f y \ f x" shows "antimono f" unfolding antimono_def using assms by blast lemma antimonoE [elim]: assumes "antimono f" - and "x \ y" - obtains "f y \ f x" + obtains "\x y. x \ y \ f y \ f x" using assms unfolding antimono_def by blast lemma antimonoD: assumes "antimono f" and "x \ y" shows "f y \ f x" using assms by blast lemma antimono_Dep_Fun_Rel_rel_left: "antimono (\R. [x y \ R] \ S x y)" by (intro antimonoI) auto lemma antimono_Dep_Fun_Rel_pred_left: "antimono (\P. [x \ P] \ Q x)" by (intro antimonoI) auto lemma antimono_dep_mono_wrt_rel_left: "antimono (\R. [x y \ R] \\<^sub>m S x y)" - by (intro antimonoI) auto + by (intro antimonoI) blast lemma antimono_dep_mono_wrt_pred_left: "antimono (\P. [x \ P] \\<^sub>m Q x)" - by (intro antimonoI) auto + by (intro antimonoI) blast lemma Dep_Fun_Rel_rel_if_le_left_if_Dep_Fun_Rel_rel: assumes "([x y \ R] \ S x y) f g" and "T \ R" shows "([x y \ T] \ S x y) f g" using assms by blast lemma Dep_Fun_Rel_pred_if_le_left_if_Dep_Fun_Rel_pred: assumes "([x \ P] \ Q x) f g" and "T \ P" shows "([x \ T] \ Q x) f g" using assms by blast lemma dep_mono_wrt_rel_if_le_left_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" and "T \ R" shows "([x y \ T] \\<^sub>m S x y) f" using assms by blast lemma dep_mono_wrt_pred_if_le_left_if_dep_mono_wrt_pred: assumes "([x \ P] \\<^sub>m Q x) f" and "T \ P" shows "([x \ T] \\<^sub>m Q x) f" using assms by blast lemma mono_Dep_Fun_Rel_rel_right: "mono (\S. [x y \ R] \ S x y)" by (intro monoI) blast lemma mono_Dep_Fun_Rel_pred_right: "mono (\Q. [x \ P] \ Q x)" by (intro monoI) blast lemma mono_dep_mono_wrt_rel_right: "mono (\S. [x y \ R] \\<^sub>m S x y)" by (intro monoI) blast lemma mono_dep_mono_wrt_pred_right: "mono (\Q. [x \ P] \\<^sub>m Q x)" by (intro monoI) blast lemma Dep_Fun_Rel_rel_if_le_right_if_Dep_Fun_Rel_rel: assumes "([x y \ R] \ S x y) f g" and "\x y. R x y \ S x y (f x) (g y) \ T x y (f x) (g y)" shows "([x y \ R] \ T x y) f g" using assms by (intro Dep_Fun_Rel_relI) blast lemma Dep_Fun_Rel_pred_if_le_right_if_Dep_Fun_Rel_pred: assumes "([x \ P] \ Q x) f g" and "\x. P x \ Q x (f x) (g x) \ T x (f x) (g x)" shows "([x \ P] \ T x) f g" using assms by blast lemma dep_mono_wrt_rel_if_le_right_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" and "\x y. R x y \ S x y (f x) (f y) \ T x y (f x) (f y)" shows "([x y \ R] \\<^sub>m T x y) f" using assms by (intro dep_mono_wrt_relI) blast lemma dep_mono_wrt_pred_if_le_right_if_dep_mono_wrt_pred: assumes "([x \ P] \\<^sub>m Q x) f" and "\x. P x \ Q x (f x) \ T x (f x)" shows "([x \ P] \\<^sub>m T x) f" using assms by blast paragraph \Composition\ lemma dep_mono_wrt_rel_compI: assumes "([x y \ R] \\<^sub>m S x y) f" and "\x y. R x y \ ([x' y' \ T x y] \\<^sub>m U x y x' y') f'" and "\x y. R x y \ S x y (f x) (f y) \ T x y (f x) (f y)" shows "([x y \ R] \\<^sub>m U x y (f x) (f y)) (f' \ f)" - using assms by (intro dep_mono_wrt_relI) (auto 6 0) + using assms by (intro dep_mono_wrt_relI) force corollary dep_mono_wrt_rel_compI': assumes "([x y \ R] \\<^sub>m S x y) f" and "\x y. R x y \ ([x' y' \ S x y] \\<^sub>m T x y x' y') f'" shows "([x y \ R] \\<^sub>m T x y (f x) (f y)) (f' \ f)" using assms by (intro dep_mono_wrt_rel_compI) lemma dep_mono_wrt_pred_comp_dep_mono_wrt_rel_compI: assumes "([x \ P] \\<^sub>m Q x) f" and "\x. P x \ ([x' y' \ R x] \\<^sub>m S x x' y') f'" and "\x. P x \ Q x (f x) \ R x (f x) (f x)" shows "([x \ P] \\<^sub>m (\y. S x (f x) (f x) y y)) (f' \ f)" - using assms by (intro dep_mono_wrt_predI) (auto 6 0) + using assms by (intro dep_mono_wrt_predI) force lemma dep_mono_wrt_pred_comp_dep_mono_wrt_pred_compI: assumes "([x \ P] \\<^sub>m Q x) f" and "\x. P x \ ([x' \ R x] \\<^sub>m S x x') f'" and "\x. P x \ Q x (f x) \ R x (f x)" shows "([x \ P] \\<^sub>m S x (f x)) (f' \ f)" - using assms by (intro dep_mono_wrt_predI) (auto 6 0) + using assms by (intro dep_mono_wrt_predI) force corollary dep_mono_wrt_pred_comp_dep_mono_wrt_pred_compI': assumes "([x \ P] \\<^sub>m Q x) f" and "\x. P x \ ([x' \ Q x] \\<^sub>m S x x') f'" shows "([x \ P] \\<^sub>m S x (f x)) (f' \ f)" using assms by (intro dep_mono_wrt_pred_comp_dep_mono_wrt_pred_compI) paragraph \Instantiations\ lemma mono_wrt_rel_self_id: "(R \\<^sub>m R) id" by auto lemma mono_wrt_pred_self_id: "([P] \\<^sub>m P) id" by auto lemma mono_in_dom: "mono in_dom" by (intro monoI) fast lemma mono_in_codom: "mono in_codom" by (intro monoI) fast lemma mono_in_field: "mono in_field" by (intro monoI) fast lemma mono_rel_comp1: "mono (\\)" by (intro monoI) fast lemma mono_rel_comp2: "mono ((\\) x)" by (intro monoI) fast end diff --git a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Surjective.thy b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Surjective.thy --- a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Surjective.thy +++ b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Surjective.thy @@ -1,50 +1,49 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Surjective\ theory Functions_Surjective imports HOL_Syntax_Bundles_Lattices begin consts surjective_at :: "'a \ ('b \ 'c) \ bool" overloading surjective_at_pred \ "surjective_at :: ('a \ bool) \ ('b \ 'a) \ bool" begin definition "surjective_at_pred P f \ \y. P y \ (\x. y = f x)" end lemma surjective_atI [intro]: assumes "\y. P y \ \x. y = f x" shows "surjective_at P f" unfolding surjective_at_pred_def using assms by blast lemma surjective_atE [elim]: assumes "surjective_at P f" and "P y" obtains x where "y = f x" using assms unfolding surjective_at_pred_def by blast -definition "surjective (f :: _ \ 'a) \ surjective_at (\ :: 'a \ bool) f" +definition "(surjective :: (_ \ 'a) \ _) \ surjective_at (\ :: 'a \ bool)" -lemma surjective_eq_surjective_at: - "surjective (f :: _ \ 'a) = surjective_at (\ :: 'a \ bool) f" +lemma surjective_eq_surjective_at: "(surjective :: (_ \ 'a) \ _) = surjective_at (\ :: 'a \ bool)" unfolding surjective_def .. lemma surjectiveI [intro]: assumes "\y. \x. y = f x" shows "surjective f" unfolding surjective_eq_surjective_at using assms by (intro surjective_atI) lemma surjectiveE: assumes "surjective f" obtains x where "y = f x" using assms unfolding surjective_eq_surjective_at by (blast intro: top1I) lemma surjective_at_if_surjective: fixes P :: "'a \ bool" and f :: "_ \ 'a" assumes "surjective f" shows "surjective_at P f" using assms by (intro surjective_atI) (blast elim: surjectiveE) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Galois/Galois_Relator.thy b/thys/Transport/HOL_Basics/Galois/Galois_Relator.thy --- a/thys/Transport/HOL_Basics/Galois/Galois_Relator.thy +++ b/thys/Transport/HOL_Basics/Galois/Galois_Relator.thy @@ -1,180 +1,180 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Relator For Galois Connections\ theory Galois_Relator imports Galois_Relator_Base Galois_Property begin context galois_prop begin interpretation flip_inv : galois_rel "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" l . lemma left_Galois_if_Galois_right_if_half_galois_prop_right: assumes "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" and "x \\<^bsub>R\<^esub> y" shows "x \<^bsub>L\<^esub>\ y" using assms by (intro left_GaloisI) auto lemma Galois_right_if_left_Galois_if_half_galois_prop_left: assumes "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" and "x \<^bsub>L\<^esub>\ y" shows "x \\<^bsub>R\<^esub> y" using assms by blast corollary Galois_right_iff_left_Galois_if_galois_prop [iff]: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" shows "x \\<^bsub>R\<^esub> y \ x \<^bsub>L\<^esub>\ y" using assms left_Galois_if_Galois_right_if_half_galois_prop_right Galois_right_if_left_Galois_if_half_galois_prop_left by blast lemma rel_inv_Galois_eq_flip_Galois_rel_inv_if_galois_prop [simp]: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" shows "(\\<^bsub>L\<^esub>) = (\<^bsub>R\<^esub>\)" using assms by blast corollary flip_Galois_rel_inv_iff_Galois_if_galois_prop [iff]: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" shows "y \<^bsub>R\<^esub>\ x \ x \<^bsub>L\<^esub>\ y" using assms by blast corollary inv_flip_Galois_rel_inv_eq_Galois_if_galois_prop [simp]: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" shows "(\\<^bsub>R\<^esub>) = (\<^bsub>L\<^esub>\)" \\Note that @{term "(\\<^bsub>R\<^esub>) = (galois_rel.Galois (\\<^bsub>R\<^esub>) (\\<^bsub>L\<^esub>) l)\"}\ using assms by (subst rel_inv_eq_iff_eq[symmetric]) simp end context galois begin interpretation flip_inv : galois "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" r l . context begin interpretation flip : galois R L r l . lemma left_Galois_left_if_left_relI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" and "x \\<^bsub>L\<^esub> x'" shows "x \<^bsub>L\<^esub>\ l x'" using assms - by (intro left_Galois_if_Galois_right_if_half_galois_prop_right) auto + by (intro left_Galois_if_Galois_right_if_half_galois_prop_right) (auto 5 0) corollary left_Galois_left_if_reflexive_on_if_half_galois_prop_rightI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" and "reflexive_on P (\\<^bsub>L\<^esub>)" and "P x" shows "x \<^bsub>L\<^esub>\ l x" using assms by (intro left_Galois_left_if_left_relI) auto lemma left_Galois_left_if_in_codom_if_inflationary_onI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "inflationary_on P (\\<^bsub>L\<^esub>) \" and "in_codom (\\<^bsub>L\<^esub>) x" and "P x" shows "x \<^bsub>L\<^esub>\ l x" using assms by (intro left_GaloisI) (auto elim!: in_codomE) lemma left_Galois_left_if_in_codom_if_inflationary_on_in_codomI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "inflationary_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" and "in_codom (\\<^bsub>L\<^esub>) x" shows "x \<^bsub>L\<^esub>\ l x" using assms by (auto intro!: left_Galois_left_if_in_codom_if_inflationary_onI) lemma left_Galois_left_if_left_rel_if_inflationary_on_in_fieldI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "inflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" and "x \\<^bsub>L\<^esub> x" shows "x \<^bsub>L\<^esub>\ l x" using assms by (auto intro!: left_Galois_left_if_in_codom_if_inflationary_onI) lemma right_left_Galois_if_right_relI: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "y \\<^bsub>R\<^esub> y'" shows "r y \<^bsub>L\<^esub>\ y'" using assms by (intro left_GaloisI) auto corollary right_left_Galois_if_reflexive_onI: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "reflexive_on P (\\<^bsub>R\<^esub>)" and "P y" shows "r y \<^bsub>L\<^esub>\ y" using assms by (intro right_left_Galois_if_right_relI) auto lemma left_Galois_if_right_rel_if_left_GaloisI: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "transitive (\\<^bsub>L\<^esub>)" and "x \<^bsub>L\<^esub>\ y" and "y \\<^bsub>R\<^esub> z" shows "x \<^bsub>L\<^esub>\ z" - using assms by (intro left_GaloisI) auto + using assms by (intro left_GaloisI) (auto 5 0) lemma left_Galois_if_left_Galois_if_left_relI: assumes "transitive (\\<^bsub>L\<^esub>)" and "x \\<^bsub>L\<^esub> y" and "y \<^bsub>L\<^esub>\ z" shows "x \<^bsub>L\<^esub>\ z" using assms by (intro left_GaloisI) auto lemma left_rel_if_right_Galois_if_left_GaloisI: assumes "((\\<^bsub>R\<^esub>) \<^sub>h\ (\\<^bsub>L\<^esub>)) r l" and "transitive (\\<^bsub>L\<^esub>)" and "x \<^bsub>L\<^esub>\ y" and "y \<^bsub>R\<^esub>\ z" shows "x \\<^bsub>L\<^esub> z" using assms by auto lemma Dep_Fun_Rel_left_Galois_right_Galois_if_mono_wrt_rel [intro]: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" shows "((\<^bsub>L\<^esub>\) \ (\<^bsub>R\<^esub>\)) l r" - using assms by auto + using assms by blast lemma left_ge_Galois_eq_left_Galois_if_in_codom_eq_in_dom_if_symmetric: assumes "symmetric (\\<^bsub>L\<^esub>)" and "in_codom (\\<^bsub>R\<^esub>) = in_dom (\\<^bsub>R\<^esub>)" shows "(\<^bsub>L\<^esub>\) = (\<^bsub>L\<^esub>\)" \\Note that @{term "(\<^bsub>L\<^esub>\) = galois_rel.Galois (\\<^bsub>L\<^esub>) (\\<^bsub>R\<^esub>) r"}\ using assms by (intro ext iffI) (auto elim!: galois_rel.left_GaloisE intro!: galois_rel.left_GaloisI) end interpretation flip : galois R L r l . lemma ge_Galois_right_eq_left_Galois_if_symmetric_if_in_codom_eq_in_dom_if_galois_prop: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" and "in_codom (\\<^bsub>L\<^esub>) = in_dom (\\<^bsub>L\<^esub>)" and "symmetric (\\<^bsub>R\<^esub>)" shows "(\\<^bsub>R\<^esub>) = (\<^bsub>L\<^esub>\)" \\Note that @{term "(\\<^bsub>R\<^esub>) = (galois_rel.Galois (\\<^bsub>R\<^esub>) (\\<^bsub>L\<^esub>) l)\"}\ using assms by (simp only: inv_flip_Galois_rel_inv_eq_Galois_if_galois_prop flip: flip.left_ge_Galois_eq_left_Galois_if_in_codom_eq_in_dom_if_symmetric) interpretation gp : galois_prop "(\<^bsub>L\<^esub>\)" "(\<^bsub>R\<^esub>\)" l l . lemma half_galois_prop_left_left_Galois_right_Galois_if_half_galois_prop_leftI [intro]: assumes "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" shows "((\<^bsub>L\<^esub>\) \<^sub>h\ (\<^bsub>R\<^esub>\)) l l" using assms by fast lemma half_galois_prop_right_left_Galois_right_Galois_if_half_galois_prop_rightI [intro]: assumes "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" shows "((\<^bsub>L\<^esub>\) \\<^sub>h (\<^bsub>R\<^esub>\)) l l" using assms by fast corollary galois_prop_left_Galois_right_Galois_if_galois_prop [intro]: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" shows "((\<^bsub>L\<^esub>\) \ (\<^bsub>R\<^esub>\)) l l" using assms by blast end end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Galois/Galois_Relator_Base.thy b/thys/Transport/HOL_Basics/Galois/Galois_Relator_Base.thy --- a/thys/Transport/HOL_Basics/Galois/Galois_Relator_Base.thy +++ b/thys/Transport/HOL_Basics/Galois/Galois_Relator_Base.thy @@ -1,62 +1,62 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Basics For Relator For Galois Connections\ theory Galois_Relator_Base imports Galois_Base begin locale galois_rel = orders L R for L :: "'a \ 'b \ bool" and R :: "'c \ 'd \ bool" and r :: "'d \ 'b" begin text \Morally speaking, the Galois relator characterises when two terms \<^term>\x :: 'a\ and \<^term>\y :: 'b\ are "similar".\ definition "Galois x y \ in_codom (\\<^bsub>R\<^esub>) y \ x \\<^bsub>L\<^esub> r y" abbreviation "left_Galois \ Galois" notation left_Galois (infix "\<^bsub>L\<^esub>\" 50) abbreviation (input) "ge_Galois_left \ (\<^bsub>L\<^esub>\)\" notation ge_Galois_left (infix "\\<^bsub>L\<^esub>" 50) text \Here we only introduced the (left) Galois relator @{term "(\<^bsub>L\<^esub>\)"}. All other variants can be introduced by considering suitable flipped and inversed interpretations (see @{file "Half_Galois_Property.thy"}).\ lemma left_GaloisI [intro]: assumes "in_codom (\\<^bsub>R\<^esub>) y" and "x \\<^bsub>L\<^esub> r y" shows "x \<^bsub>L\<^esub>\ y" unfolding Galois_def using assms by blast lemma left_GaloisE [elim]: assumes "x \<^bsub>L\<^esub>\ y" obtains "in_codom (\\<^bsub>R\<^esub>) y" "x \\<^bsub>L\<^esub> r y" using assms unfolding Galois_def by blast corollary in_dom_left_if_left_Galois: assumes "x \<^bsub>L\<^esub>\ y" shows "in_dom (\\<^bsub>L\<^esub>) x" using assms by blast corollary left_Galois_iff_in_codom_and_left_rel_right: "x \<^bsub>L\<^esub>\ y \ in_codom (\\<^bsub>R\<^esub>) y \ x \\<^bsub>L\<^esub> r y" by blast lemma left_Galois_restrict_left_eq_left_Galois_left_restrict_left: "(\<^bsub>L\<^esub>\)\\<^bsub>P :: 'a \ bool\<^esub> = galois_rel.Galois (\\<^bsub>L\<^esub>)\\<^bsub>P\<^esub> (\\<^bsub>R\<^esub>) r" - by (intro ext iffI galois_rel.left_GaloisI restrict_leftI) + by (intro ext iffI galois_rel.left_GaloisI bin_rel_restrict_leftI) (auto elim: galois_rel.left_GaloisE) lemma left_Galois_restrict_right_eq_left_Galois_right_restrict_right: "(\<^bsub>L\<^esub>\)\\<^bsub>P :: 'd \ bool\<^esub> = galois_rel.Galois (\\<^bsub>L\<^esub>) (\\<^bsub>R\<^esub>)\\<^bsub>P\<^esub> r" - by (intro ext iffI galois_rel.left_GaloisI restrict_rightI) - (auto elim!: galois_rel.left_GaloisE restrict_rightE) + by (intro ext iffI galois_rel.left_GaloisI bin_rel_restrict_rightI) + (auto elim!: galois_rel.left_GaloisE bin_rel_restrict_rightE) end end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Galois/Half_Galois_Property.thy b/thys/Transport/HOL_Basics/Galois/Half_Galois_Property.thy --- a/thys/Transport/HOL_Basics/Galois/Half_Galois_Property.thy +++ b/thys/Transport/HOL_Basics/Galois/Half_Galois_Property.thy @@ -1,283 +1,283 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Half Galois Property\ theory Half_Galois_Property imports Galois_Relator_Base Order_Equivalences begin text \As the definition of the Galois property also works on heterogeneous relations, we define the concepts in a locale that generalises @{locale galois}.\ locale galois_prop = orders L R for L :: "'a \ 'b \ bool" and R :: "'c \ 'd \ bool" and l :: "'a \ 'c" and r :: "'d \ 'b" begin sublocale galois_rel L R r . interpretation gr_flip_inv : galois_rel "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" l . abbreviation "right_ge_Galois \ gr_flip_inv.Galois" notation right_ge_Galois (infix "\<^bsub>R\<^esub>\" 50) abbreviation (input) "Galois_right \ gr_flip_inv.ge_Galois_left" notation Galois_right (infix "\\<^bsub>R\<^esub>" 50) lemma Galois_rightI [intro]: assumes "in_dom (\\<^bsub>L\<^esub>) x" and "l x \\<^bsub>R\<^esub> y" shows "x \\<^bsub>R\<^esub> y" using assms by blast lemma Galois_rightE [elim]: assumes "x \\<^bsub>R\<^esub> y" obtains "in_dom (\\<^bsub>L\<^esub>) x" "l x \\<^bsub>R\<^esub> y" using assms by blast corollary Galois_right_iff_in_dom_and_left_right_rel: "x \\<^bsub>R\<^esub> y \ in_dom (\\<^bsub>L\<^esub>) x \ l x \\<^bsub>R\<^esub> y" by blast text \Unlike common literature, we split the definition of the Galois property into two halves. This has its merits in modularity of proofs and preciser statement of required assumptions.\ definition "half_galois_prop_left \ \x y. x \<^bsub>L\<^esub>\ y \ l x \\<^bsub>R\<^esub> y" notation galois_prop.half_galois_prop_left (infix "\<^sub>h\" 50) lemma half_galois_prop_leftI [intro]: assumes "\x y. x \<^bsub>L\<^esub>\ y \ l x \\<^bsub>R\<^esub> y" shows "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" unfolding half_galois_prop_left_def using assms by blast lemma half_galois_prop_leftD [dest]: assumes "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" and " x \<^bsub>L\<^esub>\ y" shows "l x \\<^bsub>R\<^esub> y" using assms unfolding half_galois_prop_left_def by blast text\Observe that the second half can be obtained by creating an appropriately flipped and inverted interpretation of @{locale galois_prop}. Indeed, many concepts in our formalisation are "closed" under inversion, i.e. taking their inversion yields a statement for a related concept. Many theorems can thus be derived for free by inverting (and flipping) the concepts at hand. In such cases, we only state those theorems that require some non-trivial setup. All other theorems can simply be obtained by creating a suitable locale interpretation.\ interpretation flip_inv : galois_prop "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" r l . definition "half_galois_prop_right \ flip_inv.half_galois_prop_left" notation galois_prop.half_galois_prop_right (infix "\\<^sub>h" 50) lemma half_galois_prop_rightI [intro]: assumes "\x y. x \\<^bsub>R\<^esub> y \ x \\<^bsub>L\<^esub> r y" shows "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" unfolding half_galois_prop_right_def using assms by blast lemma half_galois_prop_rightD [dest]: assumes "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" and "x \\<^bsub>R\<^esub> y" shows "x \\<^bsub>L\<^esub> r y" using assms unfolding half_galois_prop_right_def by blast interpretation g : galois_prop S T f g for S T f g . lemma rel_inv_half_galois_prop_right_eq_half_galois_prop_left_rel_inv [simp]: "((\\<^bsub>R\<^esub>) \\<^sub>h (\\<^bsub>L\<^esub>))\ = ((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>))" by (intro ext) blast corollary half_galois_prop_left_rel_inv_iff_half_galois_prop_right [iff]: "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) f g \ ((\\<^bsub>R\<^esub>) \\<^sub>h (\\<^bsub>L\<^esub>)) g f" by (simp flip: rel_inv_half_galois_prop_right_eq_half_galois_prop_left_rel_inv) lemma rel_inv_half_galois_prop_left_eq_half_galois_prop_right_rel_inv [simp]: "((\\<^bsub>R\<^esub>) \<^sub>h\ (\\<^bsub>L\<^esub>))\ = ((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>))" by (intro ext) blast corollary half_galois_prop_right_rel_inv_iff_half_galois_prop_left [iff]: "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) f g \ ((\\<^bsub>R\<^esub>) \<^sub>h\ (\\<^bsub>L\<^esub>)) g f" by (simp flip: rel_inv_half_galois_prop_left_eq_half_galois_prop_right_rel_inv) end context galois begin sublocale galois_prop L R l r . interpretation flip : galois R L r l . abbreviation "right_Galois \ flip.Galois" notation right_Galois (infix "\<^bsub>R\<^esub>\" 50) abbreviation (input) "ge_Galois_right \ flip.ge_Galois_left" notation ge_Galois_right (infix "\\<^bsub>R\<^esub>" 50) abbreviation "left_ge_Galois \ flip.right_ge_Galois" notation left_ge_Galois (infix "\<^bsub>L\<^esub>\" 50) abbreviation (input) "Galois_left \ flip.Galois_right" notation Galois_left (infix "\\<^bsub>L\<^esub>" 50) context begin interpretation flip_inv : galois "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" r l . lemma rel_unit_if_left_rel_if_mono_wrt_relI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "x \\<^bsub>R\<^esub> l x' \ x \\<^bsub>L\<^esub> \ x'" and "x \\<^bsub>L\<^esub> x'" shows "x \\<^bsub>L\<^esub> \ x'" - using assms by auto + using assms by blast corollary rel_unit_if_left_rel_if_half_galois_prop_right_if_mono_wrt_rel: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" and "x \\<^bsub>L\<^esub> x'" shows "x \\<^bsub>L\<^esub> \ x'" - using assms by (auto intro: rel_unit_if_left_rel_if_mono_wrt_relI) + using assms by (fastforce intro: rel_unit_if_left_rel_if_mono_wrt_relI) corollary rel_unit_if_reflexive_on_if_half_galois_prop_right_if_mono_wrt_rel: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" and "reflexive_on P (\\<^bsub>L\<^esub>)" and "P x" shows "x \\<^bsub>L\<^esub> \ x" using assms by (blast intro: rel_unit_if_left_rel_if_half_galois_prop_right_if_mono_wrt_rel) corollary inflationary_on_unit_if_reflexive_on_if_half_galois_prop_rightI: fixes P :: "'a \ bool" assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" and "reflexive_on P (\\<^bsub>L\<^esub>)" shows "inflationary_on P (\\<^bsub>L\<^esub>) \" using assms by (intro inflationary_onI) (fastforce intro: rel_unit_if_reflexive_on_if_half_galois_prop_right_if_mono_wrt_rel) interpretation flip : galois_prop R L r l . lemma right_rel_if_Galois_left_right_if_deflationary_onI: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "((\\<^bsub>R\<^esub>) \\<^sub>h (\\<^bsub>L\<^esub>)) r l" and "deflationary_on P (\\<^bsub>R\<^esub>) \" and "transitive (\\<^bsub>R\<^esub>)" and "y \\<^bsub>L\<^esub> r y'" and "P y'" shows "y \\<^bsub>R\<^esub> y'" using assms by force lemma half_galois_prop_left_left_right_if_transitive_if_deflationary_on_if_mono_wrt_rel: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "deflationary_on (in_codom (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>) \" and "transitive (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" using assms by (intro half_galois_prop_leftI) fastforce end interpretation flip_inv : galois "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" r l rewrites "flip_inv.unit \ \" and "flip_inv.counit \ \" and "\R S. (R\ \\<^sub>m S\) \ (R \\<^sub>m S)" and "\R S f g. (R\ \\<^sub>h S\) f g \ (S \<^sub>h\ R) g f" and "((\\<^bsub>R\<^esub>) \<^sub>h\ (\\<^bsub>L\<^esub>)) r l \ ((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" and "\R. R\\ \ R" and "\P R. inflationary_on P R\ \ deflationary_on P R" and "\P R. deflationary_on P R\ \ inflationary_on P R" and "\(P :: 'b \ bool). reflexive_on P (\\<^bsub>R\<^esub>) \ reflexive_on P (\\<^bsub>R\<^esub>)" and "\R. transitive R\ \ transitive R" and "\R. in_codom R\ \ in_dom R" by (simp_all add: flip_unit_eq_counit flip_counit_eq_unit galois_prop.half_galois_prop_left_rel_inv_iff_half_galois_prop_right galois_prop.half_galois_prop_right_rel_inv_iff_half_galois_prop_left) corollary counit_rel_if_right_rel_if_mono_wrt_relI: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "r y \<^bsub>L\<^esub>\ y' \ \ y \\<^bsub>R\<^esub> y'" and "y \\<^bsub>R\<^esub> y'" shows "\ y \\<^bsub>R\<^esub> y'" using assms by (fact flip_inv.rel_unit_if_left_rel_if_mono_wrt_relI [simplified rel_inv_iff_rel]) corollary counit_rel_if_right_rel_if_half_galois_prop_left_if_mono_wrt_rel: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" and "y \\<^bsub>R\<^esub> y'" shows "\ y \\<^bsub>R\<^esub> y'" using assms by (fact flip_inv.rel_unit_if_left_rel_if_half_galois_prop_right_if_mono_wrt_rel [simplified rel_inv_iff_rel]) corollary counit_rel_if_reflexive_on_if_half_galois_prop_left_if_mono_wrt_rel: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" and "reflexive_on P (\\<^bsub>R\<^esub>)" and "P y" shows "\ y \\<^bsub>R\<^esub> y" using assms by (fact flip_inv.rel_unit_if_reflexive_on_if_half_galois_prop_right_if_mono_wrt_rel [simplified rel_inv_iff_rel]) corollary deflationary_on_counit_if_reflexive_on_if_half_galois_prop_leftI: fixes P :: "'b \ bool" assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" and "reflexive_on P (\\<^bsub>R\<^esub>)" shows "deflationary_on P (\\<^bsub>R\<^esub>) \" using assms by (fact flip_inv.inflationary_on_unit_if_reflexive_on_if_half_galois_prop_rightI) corollary left_rel_if_left_right_Galois_if_inflationary_onI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \<^sub>h\ (\\<^bsub>L\<^esub>)) r l" and "inflationary_on P (\\<^bsub>L\<^esub>) \" and "transitive (\\<^bsub>L\<^esub>)" and "l x \<^bsub>R\<^esub>\ x'" and "P x" shows "x \\<^bsub>L\<^esub> x'" using assms by (intro flip_inv.right_rel_if_Galois_left_right_if_deflationary_onI [simplified rel_inv_iff_rel]) corollary half_galois_prop_right_left_right_if_transitive_if_inflationary_on_if_mono_wrt_rel: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "inflationary_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" and "transitive (\\<^bsub>L\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" using assms by (fact flip_inv.half_galois_prop_left_left_right_if_transitive_if_deflationary_on_if_mono_wrt_rel) end context order_functors begin interpretation g : galois L R l r . interpretation flip_g : galois R L r l rewrites "flip_g.unit \ \" and "flip_g.counit \ \" by (simp_all only: flip_unit_eq_counit flip_counit_eq_unit) lemma left_rel_if_left_right_rel_left_if_order_equivalenceI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" and "transitive (\\<^bsub>L\<^esub>)" and "l x \\<^bsub>R\<^esub> l x'" and "in_dom (\\<^bsub>L\<^esub>) x" and "in_codom (\\<^bsub>L\<^esub>) x'" shows "x \\<^bsub>L\<^esub> x'" using assms by (auto intro!: flip_g.right_rel_if_Galois_left_right_if_deflationary_onI g.half_galois_prop_right_left_right_if_transitive_if_inflationary_on_if_mono_wrt_rel elim!: rel_equivalence_onE intro: inflationary_on_if_le_pred_if_inflationary_on in_field_if_in_dom in_field_if_in_codom) end end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Algebra_Alignment_Galois.thy b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Algebra_Alignment_Galois.thy --- a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Algebra_Alignment_Galois.thy +++ b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Algebra_Alignment_Galois.thy @@ -1,80 +1,81 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Alignment With Definitions from HOL-Algebra\ theory HOL_Algebra_Alignment_Galois imports "HOL-Algebra.Galois_Connection" HOL_Algebra_Alignment_Orders Galois begin named_theorems HOL_Algebra_galois_alignment context galois_connection begin context fixes L R l r defines "L \ (\\<^bsub>\\<^esub>)\\<^bsub>carrier \\<^esub>\\<^bsub>carrier \\<^esub>" and "R \ (\\<^bsub>\\<^esub>)\\<^bsub>carrier \\<^esub>\\<^bsub>carrier \\<^esub>" and "l \ \\<^sup>*" and "r \ \\<^sub>*" - notes defs[simp] = L_def R_def l_def r_def and restrict_right_eq[simp] - and restrict_leftI[intro!] restrict_leftE[elim!] + notes defs[simp] = L_def R_def l_def r_def and bin_rel_restrict_right_eq[simp] + and bin_rel_restrict_leftI[intro!] bin_rel_restrict_leftE[elim!] begin interpretation galois L R l r . lemma mono_wrt_rel_lower [HOL_Algebra_galois_alignment]: "(L \\<^sub>m R) l" using lower_closed upper_closed by (fastforce intro: use_iso2[OF lower_iso]) lemma mono_wrt_rel_upper [HOL_Algebra_galois_alignment]: "(R \\<^sub>m L) r" using lower_closed upper_closed by (fastforce intro: use_iso2[OF upper_iso]) lemma half_galois_prop_left [HOL_Algebra_galois_alignment]: "(L \<^sub>h\ R) l r" - using galois_property lower_closed by fastforce + using galois_property lower_closed by (intro half_galois_prop_leftI) fastforce lemma half_galois_prop_right [HOL_Algebra_galois_alignment]: "(L \\<^sub>h R) l r" - using galois_property upper_closed by fastforce + using galois_property upper_closed by (intro half_galois_prop_rightI) fastforce lemma galois_prop [HOL_Algebra_galois_alignment]: "(L \ R) l r" using half_galois_prop_left half_galois_prop_right by blast lemma galois_connection [HOL_Algebra_galois_alignment]: "(L \ R) l r" using mono_wrt_rel_lower mono_wrt_rel_upper galois_prop by blast end end context galois_bijection begin context fixes L R l r defines "L \ (\\<^bsub>\\<^esub>)\\<^bsub>carrier \\<^esub>\\<^bsub>carrier \\<^esub>" and "R \ (\\<^bsub>\\<^esub>)\\<^bsub>carrier \\<^esub>\\<^bsub>carrier \\<^esub>" and "l \ \\<^sup>*" and "r \ \\<^sub>*" - notes defs[simp] = L_def R_def l_def r_def and restrict_right_eq[simp] - and restrict_leftI[intro!] restrict_leftE[elim!] in_codom_restrict_leftE[elim!] + notes defs[simp] = L_def R_def l_def r_def and bin_rel_restrict_right_eq[simp] + and bin_rel_restrict_leftI[intro!] bin_rel_restrict_leftE[elim!] + in_codom_bin_rel_restrict_leftE[elim!] begin interpretation galois R L r l . lemma half_galois_prop_left_right_left [HOL_Algebra_galois_alignment]: "(R \<^sub>h\ L) r l" using gal_bij_conn.right lower_inv_eq upper_closed upper_inv_eq by (intro half_galois_prop_leftI; elim left_GaloisE) (auto; metis) lemma half_galois_prop_right_right_left [HOL_Algebra_galois_alignment]: "(R \\<^sub>h L) r l" using gal_bij_conn.left lower_closed lower_inv_eq upper_inv_eq by (intro half_galois_prop_rightI; elim Galois_rightE) (auto; metis) lemma prop_right_right_left [HOL_Algebra_galois_alignment]: "(R \ L) r l" using half_galois_prop_left_right_left half_galois_prop_right_right_left by blast lemma galois_equivalence [HOL_Algebra_galois_alignment]: "(L \\<^sub>G R) l r" using gal_bij_conn.galois_connection prop_right_right_left by (intro galois.galois_equivalenceI) auto end end end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Binary_Relations.thy b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Binary_Relations.thy --- a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Binary_Relations.thy +++ b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Binary_Relations.thy @@ -1,306 +1,296 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Alignment With Definitions from HOL.Main\ theory HOL_Alignment_Binary_Relations imports Main HOL_Mem_Of HOL_Syntax_Bundles_Relations LBinary_Relations begin unbundle no_HOL_relation_syntax named_theorems HOL_bin_rel_alignment paragraph \Properties\ subparagraph \Antisymmetric\ overloading antisymmetric_on_set \ "antisymmetric_on :: 'a set \ ('a \ 'a \ bool) \ bool" begin definition "antisymmetric_on_set (S :: 'a set) :: ('a \ _) \ _ \ antisymmetric_on (mem_of S)" end lemma antisymmetric_on_set_eq_antisymmetric_on_pred [simp]: "(antisymmetric_on (S :: 'a set) :: ('a \ _) \ bool) = antisymmetric_on (mem_of S)" unfolding antisymmetric_on_set_def by simp lemma antisymmetric_on_set_iff_antisymmetric_on_pred [iff]: "antisymmetric_on (S :: 'a set) (R :: 'a \ _) \ antisymmetric_on (mem_of S) R" by simp lemma antisymp_eq_antisymmetric [HOL_bin_rel_alignment]: "antisymp = antisymmetric" by (intro ext) (auto intro: antisympI dest: antisymmetricD antisympD) subparagraph \Injective\ overloading rel_injective_on_set \ "rel_injective_on :: 'a set \ ('a \ 'b \ bool) \ bool" rel_injective_at_set \ "rel_injective_at :: 'a set \ ('b \ 'a \ bool) \ bool" begin definition "rel_injective_on_set (S :: 'a set) :: ('a \ _) \ _ \ rel_injective_on (mem_of S)" definition "rel_injective_at_set (S :: 'a set) :: ('b \ 'a \ _) \ _ \ rel_injective_at (mem_of S)" end lemma rel_injective_on_set_eq_rel_injective_on_pred [simp]: "(rel_injective_on (S :: 'a set) :: ('a \ _) \ bool) = rel_injective_on (mem_of S)" unfolding rel_injective_on_set_def by simp lemma rel_injective_on_set_iff_rel_injective_on_pred [iff]: "rel_injective_on (S :: 'a set) (R :: 'a \ _) \ rel_injective_on (mem_of S) R" by simp lemma rel_injective_at_set_eq_rel_injective_at_pred [simp]: "(rel_injective_at (S :: 'a set) :: ('b \ 'a \ bool) \ bool) = rel_injective_at (mem_of S)" unfolding rel_injective_at_set_def by simp lemma rel_injective_at_set_iff_rel_injective_at_pred [iff]: "rel_injective_at (S :: 'a set) (R :: 'b \ 'a \ bool) \ rel_injective_at (mem_of S) R" by simp lemma left_unique_eq_rel_injective [HOL_bin_rel_alignment]: "left_unique = rel_injective" by (intro ext) (blast intro: left_uniqueI dest: rel_injectiveD left_uniqueD) subparagraph \Irreflexive\ overloading irreflexive_on_set \ "irreflexive_on :: 'a set \ ('a \ 'a \ bool) \ bool" begin definition "irreflexive_on_set (S :: 'a set) :: ('a \ _) \ _ \ irreflexive_on (mem_of S)" end lemma irreflexive_on_set_eq_irreflexive_on_pred [simp]: "(irreflexive_on (S :: 'a set) :: ('a \ _) \ bool) = irreflexive_on (mem_of S)" unfolding irreflexive_on_set_def by simp lemma irreflexive_on_set_iff_irreflexive_on_pred [iff]: "irreflexive_on (S :: 'a set) (R :: 'a \ _) \ irreflexive_on (mem_of S) R" by simp lemma irreflp_on_eq_irreflexive_on [HOL_bin_rel_alignment]: "irreflp_on = irreflexive_on" by (intro ext) (blast intro: irreflp_onI dest: irreflp_onD) lemma irreflp_eq_irreflexive [HOL_bin_rel_alignment]: "irreflp = irreflexive" by (intro ext) (blast intro: irreflpI dest: irreflexiveD irreflpD) subparagraph \Left-Total\ overloading left_total_on_set \ "left_total_on :: 'a set \ ('a \ 'b \ bool) \ bool" begin definition "left_total_on_set (S :: 'a set) :: ('a \ _) \ _ \ left_total_on (mem_of S)" end lemma left_total_on_set_eq_left_total_on_pred [simp]: "(left_total_on (S :: 'a set) :: ('a \ _) \ bool) = left_total_on (mem_of S)" unfolding left_total_on_set_def by simp lemma left_total_on_set_iff_left_total_on_pred [iff]: "left_total_on (S :: 'a set) (R :: 'a \ _) \ left_total_on (mem_of S) R" by simp lemma Transfer_left_total_eq_left_total [HOL_bin_rel_alignment]: "Transfer.left_total = Binary_Relations_Left_Total.left_total" by (intro ext) (fast intro: Transfer.left_totalI elim: Transfer.left_totalE Binary_Relations_Left_Total.left_totalE) subparagraph \Reflexive\ overloading reflexive_on_set \ "reflexive_on :: 'a set \ ('a \ 'a \ bool) \ bool" begin definition "reflexive_on_set (S :: 'a set) :: ('a \ _) \ _ \ reflexive_on (mem_of S)" end lemma reflexive_on_set_eq_reflexive_on_pred [simp]: "(reflexive_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool) = reflexive_on (mem_of S)" unfolding reflexive_on_set_def by simp lemma reflexive_on_set_iff_reflexive_on_pred [iff]: "reflexive_on (S :: 'a set) (R :: 'a \ 'a \ bool) \ reflexive_on (mem_of S) R" by simp lemma reflp_on_eq_reflexive_on [HOL_bin_rel_alignment]: "reflp_on = reflexive_on" by (intro ext) (blast intro: reflp_onI dest: reflp_onD) lemma reflp_eq_reflexive [HOL_bin_rel_alignment]: "reflp = reflexive" by (intro ext) (blast intro: reflpI dest: reflexiveD reflpD) subparagraph \Right-Unique\ overloading right_unique_on_set \ "right_unique_on :: 'a set \ ('a \ 'b \ bool) \ bool" right_unique_at_set \ "right_unique_at :: 'a set \ ('b \ 'a \ bool) \ bool" begin definition "right_unique_on_set (S :: 'a set) :: ('a \ _) \ _ \ right_unique_on (mem_of S)" definition "right_unique_at_set (S :: 'a set) :: ('b \ 'a \ _) \ _ \ right_unique_at (mem_of S)" end lemma right_unique_on_set_eq_right_unique_on_pred [simp]: "(right_unique_on (S :: 'a set) :: ('a \ _) \ bool) = right_unique_on (mem_of S)" unfolding right_unique_on_set_def by simp lemma right_unique_on_set_iff_right_unique_on_pred [iff]: "right_unique_on (S :: 'a set) (R :: 'a \ _) \ right_unique_on (mem_of S) R" by simp lemma right_unique_at_set_eq_right_unique_at_pred [simp]: "(right_unique_at (S :: 'a set) :: ('b \ 'a \ bool) \ bool) = right_unique_at (mem_of S)" unfolding right_unique_at_set_def by simp lemma right_unique_at_set_iff_right_unique_at_pred [iff]: "right_unique_at (S :: 'a set) (R :: 'b \ 'a \ bool) \ right_unique_at (mem_of S) R" by simp lemma Transfer_right_unique_eq_right_unique [HOL_bin_rel_alignment]: "Transfer.right_unique = Binary_Relations_Right_Unique.right_unique" by (intro ext) (blast intro: Transfer.right_uniqueI dest: Transfer.right_uniqueD Binary_Relations_Right_Unique.right_uniqueD) subparagraph \Surjective\ overloading rel_surjective_at_set \ "rel_surjective_at :: 'a set \ ('b \ 'a \ bool) \ bool" begin definition "rel_surjective_at_set (S :: 'a set) :: ('b \ 'a \ _) \ _ \ rel_surjective_at (mem_of S)" end lemma rel_surjective_at_set_eq_rel_surjective_at_pred [simp]: "(rel_surjective_at (S :: 'a set) :: ('b \ 'a \ _) \ bool) = rel_surjective_at (mem_of S)" unfolding rel_surjective_at_set_def by simp lemma rel_surjective_at_set_iff_rel_surjective_at_pred [iff]: "rel_surjective_at (S :: 'a set) (R :: 'b \ 'a \ _) \ rel_surjective_at (mem_of S) R" by simp lemma Transfer_right_total_eq_rel_surjective [HOL_bin_rel_alignment]: "Transfer.right_total = rel_surjective" by (intro ext) (fast intro: Transfer.right_totalI rel_surjectiveI elim: Transfer.right_totalE rel_surjectiveE) subparagraph \Symmetric\ overloading symmetric_on_set \ "symmetric_on :: 'a set \ ('a \ 'a \ bool) \ bool" begin definition "symmetric_on_set (S :: 'a set) :: ('a \ _) \ _ \ symmetric_on (mem_of S)" end lemma symmetric_on_set_eq_symmetric_on_pred [simp]: "(symmetric_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool) = symmetric_on (mem_of S)" unfolding symmetric_on_set_def by simp lemma symmetric_on_set_iff_symmetric_on_pred [iff]: "symmetric_on (S :: 'a set) (R :: 'a \ 'a \ bool) \ symmetric_on (mem_of S) R" by simp lemma symp_eq_symmetric [HOL_bin_rel_alignment]: "symp = symmetric" by (intro ext) (blast intro: sympI dest: symmetricD sympD) subparagraph \Transitive\ overloading transitive_on_set \ "transitive_on :: 'a set \ ('a \ 'a \ bool) \ bool" begin definition "transitive_on_set (S :: 'a set) :: ('a \ _) \ _ \ transitive_on (mem_of S)" end lemma transitive_on_set_eq_transitive_on_pred [simp]: "(transitive_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool) = transitive_on (mem_of S)" unfolding transitive_on_set_def by simp lemma transitive_on_set_iff_transitive_on_pred [iff]: "transitive_on (S :: 'a set) (R :: 'a \ 'a \ bool) \ transitive_on (mem_of S) R" by simp lemma transp_eq_transitive [HOL_bin_rel_alignment]: "transp = transitive" by (intro ext) (blast intro: transpI dest: transpD) paragraph \Functions\ lemma relcompp_eq_rel_comp [HOL_bin_rel_alignment]: "relcompp = rel_comp" by (intro ext) auto lemma conversep_eq_rel_inv [HOL_bin_rel_alignment]: "conversep = rel_inv" by (intro ext) auto lemma Domainp_eq_in_dom [HOL_bin_rel_alignment]: "Domainp = in_dom" by (intro ext) auto lemma Rangep_eq_in_codom [HOL_bin_rel_alignment]: "Rangep = in_codom" by (intro ext) auto -overloading - restrict_left_set \ "restrict_left :: ('a \ 'b \ bool) \ ('a set) \ 'a \ 'b \ bool" -begin - definition "restrict_left_set (R :: 'a \ _) (S :: 'a set) \ R\\<^bsub>mem_of S\<^esub>" -end - -lemma restrict_left_set_eq_restrict_left_pred [simp]: - "(R\\<^bsub>S :: 'a set\<^esub> :: 'a \ _) = R\\<^bsub>mem_of S\<^esub>" - unfolding restrict_left_set_def by simp - -lemma restrict_left_set_iff_restrict_left_pred [iff]: - "(R\\<^bsub>S :: 'a set\<^esub> :: 'a \ _) x y \ R\\<^bsub>mem_of S\<^esub> x y" - by simp - - - -paragraph \Restricted Equality\ - -lemma eq_onp_eq_eq_restrict [HOL_bin_rel_alignment]: "eq_onp = eq_restrict" +lemma eq_onp_eq_eq_restrict [HOL_bin_rel_alignment]: "eq_onp = restrict_left (=)" unfolding eq_onp_def by (intro ext) auto overloading - eq_restrict_set \ "eq_restrict :: 'a set \ 'a \ 'a \ bool" + bin_rel_restrict_left_set \ "restrict_left :: ('a \ 'b \ bool) \ 'a set \ 'a \ 'b \ bool" + bin_rel_restrict_right_set \ "restrict_right :: ('a \ 'b \ bool) \ 'b set \ 'a \ 'b \ bool" begin - definition "eq_restrict_set (S :: 'a set) \ ((=\<^bsub>mem_of S\<^esub>) :: 'a \ _)" + definition "bin_rel_restrict_left_set (R :: 'a \ 'b \ bool) (S :: 'a set) \ R\\<^bsub>mem_of S\<^esub>" + definition "bin_rel_restrict_right_set (R :: 'a \ 'b \ bool) (S :: 'b set) \ R\\<^bsub>mem_of S\<^esub>" end -lemma eq_restrict_set_eq_eq_restrict_pred [simp]: - "((=\<^bsub>S :: 'a set\<^esub>) :: 'a \ _) = (=\<^bsub>mem_of S\<^esub>)" - unfolding eq_restrict_set_def by simp +lemma bin_rel_restrict_left_set_eq_restrict_left_pred [simp]: + "(R\\<^bsub>S :: 'a set\<^esub> :: 'a \ 'b \ bool) = R\\<^bsub>mem_of S\<^esub>" + unfolding bin_rel_restrict_left_set_def by simp -lemma eq_restrict_set_iff_eq_restrict_pred [iff]: - "(x :: 'a) =\<^bsub>(S :: 'a set)\<^esub> y \ x =\<^bsub>mem_of S\<^esub> y" +lemma bin_rel_restrict_right_set_eq_restrict_right_pred [simp]: + "(R\\<^bsub>S :: 'b set\<^esub> :: 'a \ 'b \ bool) = R\\<^bsub>mem_of S\<^esub>" + unfolding bin_rel_restrict_right_set_def by simp + +lemma restrict_left_set_iff_restrict_left_pred [iff]: "(R\\<^bsub>S :: 'a set\<^esub> :: 'a \ _) x y \ R\\<^bsub>mem_of S\<^esub> x y" by simp +lemma restrict_right_set_iff_restrict_right_pred [iff]: + "(R\\<^bsub>S :: 'b set\<^esub> :: _ \ 'b \ _) x y \ R\\<^bsub>mem_of S\<^esub> x y" + by simp end diff --git a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Functions.thy b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Functions.thy --- a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Functions.thy +++ b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Functions.thy @@ -1,161 +1,160 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Alignment With Definitions from HOL.Main\ theory HOL_Alignment_Functions imports HOL_Alignment_Binary_Relations HOL_Syntax_Bundles_Functions LFunctions begin unbundle no_HOL_function_syntax named_theorems HOL_fun_alignment paragraph \Functions\ subparagraph \Bijection\ overloading bijection_on_set \ "bijection_on :: 'a set \ 'b set \ ('a \ 'b) \ ('b \ 'a) \ bool" begin definition "bijection_on_set (S :: 'a set) (S' :: 'b set) :: ('a \ 'b) \ ('b \ 'a) \ bool \ bijection_on (mem_of S) (mem_of S')" end lemma bijection_on_set_eq_bijection_on_pred [simp]: "(bijection_on (S :: 'a set) (S' :: 'b set) :: ('a \ 'b) \ _) = bijection_on (mem_of S) (mem_of S')" unfolding bijection_on_set_def by simp lemma bijection_on_set_iff_bijection_on_pred [iff]: "bijection_on (S :: 'a set) (S' :: 'b set) (f :: 'a \ 'b) g \ bijection_on (mem_of S) (mem_of S') f g" by simp lemma bij_betw_bijection_onE: assumes "bij_betw f S S'" obtains g where "bijection_on S S' f g" proof let ?g = "the_inv_into S f" from assms bij_betw_the_inv_into have "bij_betw ?g S' S" by blast with assms show "bijection_on S S' f ?g" by (auto intro!: bijection_onI dest: bij_betw_apply bij_betw_imp_inj_on the_inv_into_f_f simp: f_the_inv_into_f_bij_betw) qed lemma bij_betw_if_bijection_on: assumes "bijection_on S S' f g" shows "bij_betw f S S'" using assms by (intro bij_betw_byWitness[where ?f'=g]) (auto elim: bijection_onE dest: inverse_onD) corollary bij_betw_iff_ex_bijection_on [HOL_fun_alignment]: "bij_betw f S S' \ (\g. bijection_on S S' f g)" by (intro iffI) (auto elim!: bij_betw_bijection_onE intro: bij_betw_if_bijection_on) subparagraph \Injective\ overloading injective_on_set \ "injective_on :: 'a set \ ('a \ 'b) \ bool" begin definition "injective_on_set (S :: 'a set) :: ('a \ 'b) \ bool \ injective_on (mem_of S)" end lemma injective_on_set_eq_injective_on_pred [simp]: "(injective_on (S :: 'a set) :: ('a \ 'b) \ _) = injective_on (mem_of S)" unfolding injective_on_set_def by simp lemma injective_on_set_iff_injective_on_pred [iff]: "injective_on (S :: 'a set) (f :: 'a \ 'b) \ injective_on (mem_of S) f" by simp lemma inj_on_iff_injective_on [HOL_fun_alignment]: "inj_on f P \ injective_on P f" by (auto intro: inj_onI dest: inj_onD injective_onD) lemma inj_eq_injective [HOL_fun_alignment]: "inj = injective" by (auto intro: injI dest: injD injectiveD) subparagraph \Inverse\ overloading inverse_on_set \ "inverse_on :: 'a set \ ('a \ 'b) \ ('b \ 'a) \ bool" begin definition "inverse_on_set (S :: 'a set) :: ('a \ 'b) \ _ \ inverse_on (mem_of S)" end lemma inverse_on_set_eq_inverse_on_pred [simp]: "(inverse_on (S :: 'a set) :: ('a \ 'b) \ _) = inverse_on (mem_of S)" unfolding inverse_on_set_def by simp lemma inverse_on_set_iff_inverse_on_pred [iff]: "inverse_on (S :: 'a set) (f :: 'a \ 'b) g \ inverse_on (mem_of S) f g" by simp subparagraph \Monotone\ lemma monotone_on_eq_mono_wrt_rel_restrict_left_right [HOL_fun_alignment]: - "monotone_on S R = mono_wrt_rel (R\\<^bsub>S\<^esub>\\<^bsub>S\<^esub>)" - unfolding restrict_right_eq - by (intro ext) (blast intro: monotone_onI dest: monotone_onD) + "monotone_on S R = mono_wrt_rel R\\<^bsub>S\<^esub>\\<^bsub>S\<^esub>" + by (intro ext) (auto intro!: monotone_onI dest: monotone_onD) lemma monotone_eq_mono_wrt_rel [HOL_fun_alignment]: "monotone = mono_wrt_rel" by (intro ext) (auto intro: monotoneI dest: monotoneD) lemma pred_fun_eq_mono_wrt_pred [HOL_fun_alignment]: "pred_fun = mono_wrt_pred" by (intro ext) auto lemma Fun_mono_eq_mono [HOL_fun_alignment]: "Fun.mono = mono" by (intro ext) (auto intro: Fun.mono_onI dest: Fun.monoD) lemma Fun_antimono_eq_antimono [HOL_fun_alignment]: "Fun.antimono = antimono" by (intro ext) (auto intro: monotoneI dest: monotoneD) subparagraph \Surjective\ overloading surjective_at_set \ "surjective_at :: 'a set \ ('b \ 'a) \ bool" begin definition "surjective_at_set (S :: 'a set) :: ('b \ 'a) \ bool \ surjective_at (mem_of S)" end lemma surjective_at_set_eq_surjective_at_pred [simp]: "(surjective_at (S :: 'a set) :: ('b \ 'a) \ _) = surjective_at (mem_of S)" unfolding surjective_at_set_def by simp lemma surjective_at_set_iff_surjective_at_pred [iff]: "surjective_at (S :: 'a set) (f :: 'b \ 'a) \ surjective_at (mem_of S) f" by simp lemma surj_eq_surjective [HOL_fun_alignment]: "surj = surjective" by (intro ext) (fast intro: surjI dest: surjD elim: surjectiveE) paragraph \Functions\ lemma Fun_id_eq_id [HOL_fun_alignment]: "Fun.id = Functions_Base.id" by (intro ext) simp lemma Fun_comp_eq_comp [HOL_fun_alignment]: "Fun.comp = Functions_Base.comp" by (intro ext) simp lemma map_fun_eq_fun_map [HOL_fun_alignment]: "map_fun = fun_map" by (intro ext) simp paragraph \Relators\ lemma rel_fun_eq_Fun_Rel_rel [HOL_fun_alignment]: "rel_fun = Fun_Rel_rel" by (intro ext) (auto dest: rel_funD) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Orders/Equivalence_Relations.thy b/thys/Transport/HOL_Basics/Orders/Equivalence_Relations.thy --- a/thys/Transport/HOL_Basics/Orders/Equivalence_Relations.thy +++ b/thys/Transport/HOL_Basics/Orders/Equivalence_Relations.thy @@ -1,72 +1,71 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Equivalences\ theory Equivalence_Relations imports Partial_Equivalence_Relations begin -definition "equivalence_rel_on P R \ - partial_equivalence_rel_on P R \ reflexive_on P R" +definition "equivalence_rel_on \ partial_equivalence_rel_on \ reflexive_on" lemma equivalence_rel_onI [intro]: assumes "partial_equivalence_rel_on P R" and "reflexive_on P R" shows "equivalence_rel_on P R" - unfolding equivalence_rel_on_def using assms by blast + unfolding equivalence_rel_on_def using assms by auto lemma equivalence_rel_onE [elim]: assumes "equivalence_rel_on P R" obtains "partial_equivalence_rel_on P R" "reflexive_on P R" - using assms unfolding equivalence_rel_on_def by blast + using assms unfolding equivalence_rel_on_def by auto lemma equivalence_rel_on_in_field_if_partial_equivalence_rel: assumes "partial_equivalence_rel R" shows "equivalence_rel_on (in_field R) R" using assms by (intro equivalence_rel_onI reflexive_on_in_field_if_partial_equivalence_rel) auto corollary partial_equivalence_rel_iff_equivalence_rel_on_in_field: "partial_equivalence_rel R \ equivalence_rel_on (in_field R) R" using equivalence_rel_on_in_field_if_partial_equivalence_rel by auto -definition "equivalence_rel (R :: 'a \ _) \ equivalence_rel_on (\ :: 'a \ bool) R" +definition "(equivalence_rel :: ('a \ _) \ bool) \ equivalence_rel_on (\ :: 'a \ bool)" lemma equivalence_rel_eq_equivalence_rel_on: - "equivalence_rel (R :: 'a \ _) = equivalence_rel_on (\ :: 'a \ bool) R" + "(equivalence_rel :: ('a \ _) \ bool) = equivalence_rel_on (\ :: 'a \ bool)" unfolding equivalence_rel_def .. lemma equivalence_relI [intro]: assumes "partial_equivalence_rel R" and "reflexive R" shows "equivalence_rel R" unfolding equivalence_rel_eq_equivalence_rel_on using assms by (intro equivalence_rel_onI partial_equivalence_rel_on_if_partial_equivalence_rel reflexive_on_if_reflexive) lemma equivalence_relE [elim]: assumes "equivalence_rel R" obtains "partial_equivalence_rel R" "reflexive R" using assms unfolding equivalence_rel_eq_equivalence_rel_on by (elim equivalence_rel_onE) (simp only: partial_equivalence_rel_eq_partial_equivalence_rel_on reflexive_eq_reflexive_on) lemma equivalence_rel_on_if_equivalence: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "equivalence_rel R" shows "equivalence_rel_on P R" using assms by (elim equivalence_relE) (intro equivalence_rel_onI partial_equivalence_rel_on_if_partial_equivalence_rel reflexive_on_if_reflexive) paragraph \Instantiations\ lemma equivalence_eq: "equivalence_rel (=)" using partial_equivalence_rel_eq reflexive_eq by (rule equivalence_relI) lemma equivalence_top: "equivalence_rel \" using partial_equivalence_rel_top reflexive_top by (rule equivalence_relI) end diff --git a/thys/Transport/HOL_Basics/Orders/Functions/Closure_Operators.thy b/thys/Transport/HOL_Basics/Orders/Functions/Closure_Operators.thy --- a/thys/Transport/HOL_Basics/Orders/Functions/Closure_Operators.thy +++ b/thys/Transport/HOL_Basics/Orders/Functions/Closure_Operators.thy @@ -1,86 +1,85 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Closure Operators\ theory Closure_Operators imports Order_Functions_Base begin definition "idempotent_on P R f \ rel_equivalence_on P (rel_map f R) f" lemma idempotent_onI [intro]: assumes "\x. P x \ f x \\<^bsub>R\<^esub> f (f x)" shows "idempotent_on P R f" unfolding idempotent_on_def using assms by fastforce lemma idempotent_onE [elim]: assumes "idempotent_on P R f" and "P x" obtains "R (f (f x)) (f x)" "R (f x) (f (f x))" using assms unfolding idempotent_on_def by fastforce lemma rel_equivalence_on_rel_map_iff_idempotent_on [iff]: "rel_equivalence_on P (rel_map f R) f \ idempotent_on P R f" unfolding idempotent_on_def by simp lemma bi_related_if_idempotent_onD: assumes "idempotent_on P R f" and "P x" shows "f x \\<^bsub>R\<^esub> f (f x)" using assms by blast -definition "idempotent (R :: 'a \ _) f \ idempotent_on (\ :: 'a \ bool) R f" +definition "(idempotent :: ('a \ _) \ _) \ idempotent_on (\ :: 'a \ bool)" -lemma idempotent_eq_idempotent_on: - "idempotent (R :: 'a \ _) f = idempotent_on (\ :: 'a \ bool) R f" +lemma idempotent_eq_idempotent_on: "(idempotent :: ('a \ _) \ _) = idempotent_on (\ :: 'a \ bool)" unfolding idempotent_def .. lemma idempotentI [intro]: assumes "\x. R (f (f x)) (f x)" and "\x. R (f x) (f (f x))" shows "idempotent R f" unfolding idempotent_eq_idempotent_on using assms by blast lemma idempotentE [elim]: assumes "idempotent R f" obtains "R (f (f x)) (f x)" "R (f x) (f (f x))" using assms unfolding idempotent_eq_idempotent_on by (blast intro: top1I) lemma idempotent_on_if_idempotent: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "idempotent R f" shows "idempotent_on P R f" using assms by (intro idempotent_onI) auto -definition "closure_operator R f \ - (R \\<^sub>m R) f \ inflationary_on (in_field R) R f \ idempotent_on (in_field R) R f" +definition "closure_operator R \ + (R \\<^sub>m R) \ inflationary_on (in_field R) R \ idempotent_on (in_field R) R" lemma closure_operatorI [intro]: assumes "(R \\<^sub>m R) f" and "inflationary_on (in_field R) R f" and "idempotent_on (in_field R) R f" shows "closure_operator R f" unfolding closure_operator_def using assms by blast lemma closure_operatorE [elim]: assumes "closure_operator R f" obtains "(R \\<^sub>m R) f" "inflationary_on (in_field R) R f" "idempotent_on (in_field R) R f" using assms unfolding closure_operator_def by blast lemma mono_wrt_rel_if_closure_operator: assumes "closure_operator R f" shows "(R \\<^sub>m R) f" using assms by (elim closure_operatorE) lemma inflationary_on_in_field_if_closure_operator: assumes "closure_operator R f" shows "inflationary_on (in_field R) R f" using assms by (elim closure_operatorE) lemma idempotent_on_in_field_if_closure_operator: assumes "closure_operator R f" shows "idempotent_on (in_field R) R f" using assms by (elim closure_operatorE) end diff --git a/thys/Transport/HOL_Basics/Orders/Functions/Order_Functions_Base.thy b/thys/Transport/HOL_Basics/Orders/Functions/Order_Functions_Base.thy --- a/thys/Transport/HOL_Basics/Orders/Functions/Order_Functions_Base.thy +++ b/thys/Transport/HOL_Basics/Orders/Functions/Order_Functions_Base.thy @@ -1,445 +1,447 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Functions On Orders\ subsubsection \Basics\ theory Order_Functions_Base imports Functions_Monotone - Restricted_Equality + Binary_Relations_Antisymmetric + Binary_Relations_Symmetric + Preorders begin subparagraph \Bi-Relation\ definition "bi_related R x y \ R x y \ R y x" (*Note: we are not using (\\) as infix here because it would produce an ambiguous grammar whenever using a of the form "definition c \ t"*) bundle bi_related_syntax begin syntax "_bi_related" :: "'a \ ('a \ 'a \ bool) \ 'a \ bool" ("(_) \\<^bsub>(_)\<^esub> (_)" [51,51,51] 50) notation bi_related ("'(\(\<^bsub>_\<^esub>)')") end bundle no_bi_related_syntax begin no_syntax "_bi_related" :: "'a \ ('a \ 'a \ bool) \ 'a \ bool" ("(_) \\<^bsub>(_)\<^esub> (_)" [51,51,51] 50) no_notation bi_related ("'(\(\<^bsub>_\<^esub>)')") end unbundle bi_related_syntax translations "x \\<^bsub>R\<^esub> y" \ "CONST bi_related R x y" lemma bi_relatedI [intro]: assumes "R x y" and "R y x" shows "x \\<^bsub>R\<^esub> y" unfolding bi_related_def using assms by blast lemma bi_relatedE [elim]: assumes "x \\<^bsub>R\<^esub> y" obtains "R x y" "R y x" using assms unfolding bi_related_def by blast lemma symmetric_bi_related [iff]: "symmetric (\\<^bsub>R\<^esub>)" by (intro symmetricI) blast lemma reflexive_bi_related_if_reflexive [intro]: assumes "reflexive R" shows "reflexive (\\<^bsub>R\<^esub>)" using assms by (intro reflexiveI) (blast dest: reflexiveD) lemma transitive_bi_related_if_transitive [intro]: assumes "transitive R" shows "transitive (\\<^bsub>R\<^esub>)" using assms by (intro transitiveI bi_relatedI) auto lemma mono_bi_related [iff]: "mono bi_related" by (intro monoI) blast lemma bi_related_if_le_rel_if_bi_related: assumes "x \\<^bsub>R\<^esub> y" and "R \ S" shows "x \\<^bsub>S\<^esub> y" using assms by blast lemma eq_if_bi_related_if_antisymmetric_on: assumes "antisymmetric_on P R" and "x \\<^bsub>R\<^esub> y" and "P x" "P y" shows "x = y" using assms by (blast dest: antisymmetric_onD) lemma eq_if_bi_related_if_in_field_le_if_antisymmetric_on: assumes "antisymmetric_on P R" and "in_field R \ P" and "x \\<^bsub>R\<^esub> y" shows "x = y" using assms by (intro eq_if_bi_related_if_antisymmetric_on) blast+ lemma bi_related_le_eq_if_antisymmetric_on_in_field: assumes "antisymmetric_on (in_field R) R" shows "(\\<^bsub>R\<^esub>) \ (=)" using assms by (intro le_relI eq_if_bi_related_if_in_field_le_if_antisymmetric_on) blast+ lemma bi_related_if_all_rel_iff_if_reflexive_on: assumes "reflexive_on P R" and "\z. P z \ R x z \ R y z" and "P x" "P y" shows "x \\<^bsub>R\<^esub> y" using assms by blast lemma bi_related_if_all_rel_iff_if_reflexive_on': assumes "reflexive_on P R" and "\z. P z \ R z x \ R z y" and "P x" "P y" shows "x \\<^bsub>R\<^esub> y" using assms by blast corollary eq_if_all_rel_iff_if_antisymmetric_on_if_reflexive_on: assumes "reflexive_on P R" and "antisymmetric_on P R" and "\z. P z \ R x z \ R y z" and "P x" "P y" shows "x = y" using assms by (blast intro: eq_if_bi_related_if_antisymmetric_on bi_related_if_all_rel_iff_if_reflexive_on) corollary eq_if_all_rel_iff_if_antisymmetric_on_if_reflexive_on': assumes "reflexive_on P R" and "antisymmetric_on P R" and "\z. P z \ R z x \ R z y" and "P x" "P y" shows "x = y" using assms by (blast intro: eq_if_bi_related_if_antisymmetric_on bi_related_if_all_rel_iff_if_reflexive_on') subparagraph \Inflationary\ consts inflationary_on :: "'a \ ('b \ 'b \ bool) \ ('b \ 'b) \ bool" overloading inflationary_on_pred \ "inflationary_on :: ('a \ bool) \ ('a \ 'a \ bool) \ ('a \ 'a) \ bool" begin text \Often also called "extensive".\ definition "inflationary_on_pred P (R :: 'a \ 'a \ _) f \ \x. P x \ R x (f x)" end lemma inflationary_onI [intro]: assumes "\x. P x \ R x (f x)" shows "inflationary_on P R f" unfolding inflationary_on_pred_def using assms by blast lemma inflationary_onD [dest]: assumes "inflationary_on P R f" and "P x" shows "R x (f x)" using assms unfolding inflationary_on_pred_def by blast lemma inflationary_on_eq_dep_mono_wrt_pred: "inflationary_on = dep_mono_wrt_pred" by blast lemma antimono_inflationary_on_pred [iff]: "antimono (\(P :: 'a \ bool). inflationary_on P (R :: 'a \ _))" by (intro antimonoI) auto lemma inflationary_on_if_le_pred_if_inflationary_on: fixes P P' :: "'a \ bool" and R :: "'a \ _" assumes "inflationary_on P R f" and "P' \ P" shows "inflationary_on P' R f" using assms by blast lemma mono_inflationary_on_rel [iff]: "mono (\(R :: 'a \ _). inflationary_on (P :: 'a \ bool) R)" by (intro monoI) auto lemma inflationary_on_if_le_rel_if_inflationary_on: assumes "inflationary_on P R f" and "\x. P x \ R x (f x) \ R' x (f x)" shows "inflationary_on P R' f" using assms by blast lemma le_in_dom_if_inflationary_on: assumes "inflationary_on P R f" shows "P \ in_dom R" using assms by blast lemma inflationary_on_sup_eq [simp]: "(inflationary_on :: ('a \ bool) \ ('a \ _) \ _) ((P :: 'a \ bool) \ Q) = inflationary_on P \ inflationary_on Q" by (intro ext iffI inflationary_onI) (auto intro: inflationary_on_if_le_pred_if_inflationary_on) -definition "inflationary (R :: 'a \ _) f \ inflationary_on (\ :: 'a \ bool) R f" +definition "(inflationary :: ('a \ _) \ _) \ inflationary_on (\ :: 'a \ bool)" lemma inflationary_eq_inflationary_on: - "inflationary (R :: 'a \ _) f = inflationary_on (\ :: 'a \ bool) R f" + "(inflationary :: ('a \ _) \ _) = inflationary_on (\ :: 'a \ bool)" unfolding inflationary_def .. lemma inflationaryI [intro]: assumes "\x. R x (f x)" shows "inflationary R f" unfolding inflationary_eq_inflationary_on using assms by (intro inflationary_onI) lemma inflationaryD: assumes "inflationary R f" shows "R x (f x)" using assms unfolding inflationary_eq_inflationary_on by auto lemma inflationary_on_if_inflationary: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "inflationary R f" shows "inflationary_on P R f" using assms by (intro inflationary_onI) (blast dest: inflationaryD) lemma inflationary_eq_dep_mono_wrt_pred: "inflationary = dep_mono_wrt_pred \" by (intro ext) (fastforce dest: inflationaryD) subparagraph \Deflationary\ definition "deflationary_on P R \ inflationary_on P R\" lemma deflationary_on_eq_inflationary_on_rel_inv: "deflationary_on P R = inflationary_on P R\" unfolding deflationary_on_def .. declare deflationary_on_eq_inflationary_on_rel_inv[symmetric, simp] corollary deflationary_on_rel_inv_eq_inflationary_on [simp]: "deflationary_on P R\ = inflationary_on P R" unfolding deflationary_on_eq_inflationary_on_rel_inv by simp lemma deflationary_onI [intro]: assumes "\x. P x \ R (f x) x" shows "deflationary_on P R f" unfolding deflationary_on_eq_inflationary_on_rel_inv using assms by (intro inflationary_onI rel_invI) lemma deflationary_onD [dest]: assumes "deflationary_on P R f" and "P x" shows "R (f x) x" using assms unfolding deflationary_on_eq_inflationary_on_rel_inv by blast lemma deflationary_on_eq_dep_mono_wrt_pred_rel_inv: "deflationary_on P R = ([x \ P] \\<^sub>m R\ x)" by blast lemma antimono_deflationary_on_pred [iff]: "antimono (\(P :: 'a \ bool). deflationary_on P (R :: 'a \ _))" by (intro antimonoI) auto lemma deflationary_on_if_le_pred_if_deflationary_on: fixes P P' :: "'a \ bool" and R :: "'a \ _" assumes "deflationary_on P R f" and "P' \ P" shows "deflationary_on P' R f" using assms by blast lemma mono_deflationary_on_rel [iff]: "mono (\(R :: 'a \ _). deflationary_on (P :: 'a \ bool) R)" by (intro monoI) auto lemma deflationary_on_if_le_rel_if_deflationary_on: assumes "deflationary_on P R f" and "\x. P x \ R (f x) x \ R' (f x) x" shows "deflationary_on P R' f" using assms by auto lemma le_in_dom_if_deflationary_on: assumes "deflationary_on P R f" shows "P \ in_codom R" using assms by blast lemma deflationary_on_sup_eq [simp]: "(deflationary_on :: ('a \ bool) \ ('a \ _) \ _) ((P :: 'a \ bool) \ Q) = deflationary_on P \ deflationary_on Q" unfolding deflationary_on_eq_inflationary_on_rel_inv by auto -definition "deflationary R (f :: 'a \ _) \ deflationary_on (\ :: 'a \ bool) R f" +definition "(deflationary :: ('a \ _) \ _) \ deflationary_on (\ :: 'a \ bool)" lemma deflationary_eq_deflationary_on: - "deflationary R (f :: 'a \ _) = deflationary_on (\ :: 'a \ bool) R f" + "(deflationary :: ('a \ _) \ _) = deflationary_on (\ :: 'a \ bool)" unfolding deflationary_def .. lemma deflationaryI [intro]: assumes "\x. R (f x) x" shows "deflationary R f" unfolding deflationary_eq_deflationary_on using assms by (intro deflationary_onI) lemma deflationaryD: assumes "deflationary R f" shows "R (f x) x" using assms unfolding deflationary_eq_deflationary_on by auto lemma deflationary_on_if_deflationary: fixes P :: "'a \ bool" and f :: "'a \ _" assumes "deflationary R f" shows "deflationary_on P R f" using assms by (intro deflationary_onI) (blast dest: deflationaryD) lemma deflationary_eq_dep_mono_wrt_pred_rel_inv: "deflationary R = dep_mono_wrt_pred \ R\" by (intro ext) (fastforce dest: deflationaryD) subparagraph \Relational Equivalence\ definition "rel_equivalence_on \ inflationary_on \ deflationary_on" lemma rel_equivalence_on_eq: "rel_equivalence_on = inflationary_on \ deflationary_on" unfolding rel_equivalence_on_def .. lemma rel_equivalence_onI [intro]: assumes "inflationary_on P R f" and "deflationary_on P R f" shows "rel_equivalence_on P R f" unfolding rel_equivalence_on_eq using assms by auto lemma rel_equivalence_onE [elim]: assumes "rel_equivalence_on P R f" obtains "inflationary_on P R f" "deflationary_on P R f" using assms unfolding rel_equivalence_on_eq by auto lemma rel_equivalence_on_eq_dep_mono_wrt_pred_inf: "rel_equivalence_on P R = dep_mono_wrt_pred P (R \ R\)" by (intro ext) fastforce lemma bi_related_if_rel_equivalence_on: assumes "rel_equivalence_on P R f" and "P x" shows "x \\<^bsub>R\<^esub> f x" using assms by (intro bi_relatedI) auto lemma rel_equivalence_on_if_all_bi_related: assumes "\x. P x \ x \\<^bsub>R\<^esub> f x" shows "rel_equivalence_on P R f" using assms by auto corollary rel_equivalence_on_iff_all_bi_related: "rel_equivalence_on P R f \ (\x. P x \ x \\<^bsub>R\<^esub> f x)" using rel_equivalence_on_if_all_bi_related bi_related_if_rel_equivalence_on by blast lemma rel_equivalence_onD [dest]: assumes "rel_equivalence_on P R f" and "P x" shows "R x (f x)" "R (f x) x" using assms by (auto dest: bi_related_if_rel_equivalence_on) lemma rel_equivalence_on_rel_inv_eq_rel_equivalence_on [simp]: "rel_equivalence_on P R\ = rel_equivalence_on P R" by (intro ext) fastforce lemma antimono_rel_equivalence_on_pred [iff]: "antimono (\(P :: 'a \ bool). rel_equivalence_on P (R :: 'a \ _))" by (intro antimonoI) blast lemma rel_equivalence_on_if_le_pred_if_rel_equivalence_on: fixes P P' :: "'a \ bool" and R :: "'a \ _" assumes "rel_equivalence_on P R f" and "P' \ P" shows "rel_equivalence_on P' R f" using assms by blast lemma rel_equivalence_on_sup_eq [simp]: "(rel_equivalence_on :: ('a \ bool) \ ('a \ _) \ _) ((P :: 'a \ bool) \ Q) = rel_equivalence_on P \ rel_equivalence_on Q" unfolding rel_equivalence_on_eq by (simp add: inf_aci) lemma in_codom_eq_in_dom_if_rel_equivalence_on_in_field: assumes "rel_equivalence_on (in_field R) R f" shows "in_codom R = in_dom R" using assms by (intro ext) blast lemma reflexive_on_if_transitive_on_if_mon_wrt_pred_if_rel_equivalence_on: assumes "rel_equivalence_on P R f" and "([P] \\<^sub>m P) f" and "transitive_on P R" shows "reflexive_on P R" using assms by (blast dest: transitive_onD) lemma inflationary_on_eq_rel_equivalence_on_if_symmetric: assumes "symmetric R" shows "inflationary_on P R = rel_equivalence_on P R" using assms by (simp add: rel_equivalence_on_eq deflationary_on_eq_inflationary_on_rel_inv) lemma deflationary_on_eq_rel_equivalence_on_if_symmetric: assumes "symmetric R" shows "deflationary_on P R = rel_equivalence_on P R" using assms by (simp add: deflationary_on_eq_inflationary_on_rel_inv rel_equivalence_on_eq) -definition "rel_equivalence (R :: 'a \ _) f \ rel_equivalence_on (\ :: 'a \ bool) R f" +definition "(rel_equivalence :: ('a \ _) \ _ ) \ rel_equivalence_on (\ :: 'a \ bool)" lemma rel_equivalence_eq_rel_equivalence_on: - "rel_equivalence (R :: 'a \ _) f = rel_equivalence_on (\ :: 'a \ bool) R f" + "(rel_equivalence :: ('a \ _) \ _ ) = rel_equivalence_on (\ :: 'a \ bool)" unfolding rel_equivalence_def .. lemma rel_equivalenceI [intro]: assumes "inflationary R f" and "deflationary R f" shows "rel_equivalence R f" unfolding rel_equivalence_eq_rel_equivalence_on using assms by (intro rel_equivalence_onI) (auto dest: inflationary_on_if_inflationary deflationary_on_if_deflationary) lemma rel_equivalenceE [elim]: assumes "rel_equivalence R f" obtains "inflationary R f" "deflationary R f" using assms unfolding rel_equivalence_eq_rel_equivalence_on by (elim rel_equivalence_onE) (simp only: inflationary_eq_inflationary_on deflationary_eq_deflationary_on) lemma inflationary_if_rel_equivalence: assumes "rel_equivalence R f" shows "inflationary R f" using assms by (elim rel_equivalenceE) lemma deflationary_if_rel_equivalence: assumes "rel_equivalence R f" shows "deflationary R f" using assms by (elim rel_equivalenceE) lemma rel_equivalence_on_if_rel_equivalence: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "rel_equivalence R f" shows "rel_equivalence_on P R f" using assms by (intro rel_equivalence_onI) (auto dest: inflationary_on_if_inflationary deflationary_on_if_deflationary) lemma bi_related_if_rel_equivalence: assumes "rel_equivalence R f" shows "x \\<^bsub>R\<^esub> f x" using assms by (intro bi_relatedI) (auto dest: inflationaryD deflationaryD) lemma rel_equivalence_if_all_bi_related: assumes "\x. x \\<^bsub>R\<^esub> f x" shows "rel_equivalence R f" using assms by auto lemma rel_equivalenceD: assumes "rel_equivalence R f" shows "R x (f x)" "R (f x) x" using assms by (auto dest: bi_related_if_rel_equivalence) lemma reflexive_on_in_field_if_transitive_if_rel_equivalence_on: assumes "rel_equivalence_on (in_field R) R f" and "transitive R" shows "reflexive_on (in_field R) R" using assms by (intro reflexive_onI) blast corollary preorder_on_in_field_if_transitive_if_rel_equivalence_on: assumes "rel_equivalence_on (in_field R) R f" and "transitive R" shows "preorder_on (in_field R) R" using assms reflexive_on_in_field_if_transitive_if_rel_equivalence_on - using assms by blast + by blast end diff --git a/thys/Transport/HOL_Basics/Orders/Functors/Order_Functors_Base.thy b/thys/Transport/HOL_Basics/Orders/Functors/Order_Functors_Base.thy --- a/thys/Transport/HOL_Basics/Orders/Functors/Order_Functors_Base.thy +++ b/thys/Transport/HOL_Basics/Orders/Functors/Order_Functors_Base.thy @@ -1,196 +1,196 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Order Functors\ subsubsection \Basic Setup and Results\ theory Order_Functors_Base imports Functions_Inverse Order_Functions_Base begin text \In the following, we do not add any assumptions to our locales but rather add them as needed to the theorem statements. This allows consumers to state preciser results; particularly, the development of Transport depends on this setup.\ locale orders = fixes L :: "'a \ 'b \ bool" and R :: "'c \ 'd \ bool" begin notation L (infix "\\<^bsub>L\<^esub>" 50) notation R (infix "\\<^bsub>R\<^esub>" 50) text\We call @{term "(\\<^bsub>L\<^esub>)"} the \<^emph>\left relation\ and @{term "(\\<^bsub>R\<^esub>)"} the \<^emph>\right relation\.\ abbreviation (input) "ge_left \ (\\<^bsub>L\<^esub>)\" notation ge_left (infix "\\<^bsub>L\<^esub>" 50) abbreviation (input) "ge_right \ (\\<^bsub>R\<^esub>)\" notation ge_right (infix "\\<^bsub>R\<^esub>" 50) end text \Homogeneous orders\ locale hom_orders = orders L R for L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" locale order_functor = hom_orders L R for L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" begin lemma left_right_rel_left_self_if_reflexive_on_left_if_mono_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on P (\\<^bsub>L\<^esub>)" and "P x" shows "l x \\<^bsub>R\<^esub> l x" using assms by blast lemma left_right_rel_left_self_if_reflexive_on_in_dom_right_if_mono_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on (in_dom (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) x" shows "l x \\<^bsub>R\<^esub> l x" using assms by blast lemma left_right_rel_left_self_if_reflexive_on_in_codom_right_if_mono_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on (in_codom (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" and "in_codom (\\<^bsub>L\<^esub>) x" shows "l x \\<^bsub>R\<^esub> l x" using assms by blast lemma left_right_rel_left_self_if_reflexive_on_in_field_right_if_mono_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" and "in_field (\\<^bsub>L\<^esub>) x" shows "l x \\<^bsub>R\<^esub> l x" using assms by blast lemma mono_wrt_rel_left_if_reflexive_on_if_le_eq_if_mono_wrt_in_field: assumes "([in_field (\\<^bsub>L\<^esub>)] \\<^sub>m P) l" and "(\\<^bsub>L\<^esub>) \ (=)" and "reflexive_on P (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" using assms by (intro dep_mono_wrt_relI) auto end locale order_functors = order_functor L R l + flip_of : order_functor R L r for L R l r begin text \We call the composition \<^term>\r \ l\ the \<^emph>\unit\ and the term \<^term>\l \ r\ the \<^emph>\counit\ of the order functors pair. This is terminology is borrowed from category theory - the functors are an \<^emph>\adjoint\.\ definition "unit \ r \ l" notation unit ("\") lemma unit_eq_comp: "\ = r \ l" unfolding unit_def by simp lemma unit_eq [simp]: "\ x = r (l x)" by (simp add: unit_eq_comp) context begin text \Note that by flipping the roles of the left and rights functors, we obtain a flipped interpretation of @{locale order_functors}. In many cases, this allows us to obtain symmetric definitions and theorems for free. As such, in many cases, we do we do not explicitly state those free results but users can obtain them as needed by creating said flipped interpretation.\ interpretation flip : order_functors R L r l . definition "counit \ flip.unit" notation counit ("\") lemma counit_eq_comp: "\ = l \ r" unfolding counit_def flip.unit_def by simp lemma counit_eq [simp]: "\ x = l (r x)" by (simp add: counit_eq_comp) end context begin interpretation flip : order_functors R L r l . lemma flip_counit_eq_unit: "flip.counit = \" by (intro ext) simp lemma flip_unit_eq_counit: "flip.unit = \" by (intro ext) simp lemma inflationary_on_unit_if_left_rel_right_if_left_right_relI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on P (\\<^bsub>L\<^esub>)" and "\x y. P x \ l x \\<^bsub>R\<^esub> y \ x \\<^bsub>L\<^esub> r y" shows "inflationary_on P (\\<^bsub>L\<^esub>) \" - using assms by (intro inflationary_onI) auto + using assms by (intro inflationary_onI) fastforce lemma deflationary_on_unit_if_right_left_rel_if_right_rel_leftI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on P (\\<^bsub>L\<^esub>)" and "\x y. P x \ y \\<^bsub>R\<^esub> l x \ r y \\<^bsub>L\<^esub> x" shows "deflationary_on P (\\<^bsub>L\<^esub>) \" - using assms by (intro deflationary_onI) auto + using assms by (intro deflationary_onI) fastforce context fixes P :: "'a \ bool" begin lemma rel_equivalence_on_unit_iff_inflationary_on_if_inverse_on: assumes "inverse_on P l r" shows "rel_equivalence_on P (\\<^bsub>L\<^esub>) \ \ inflationary_on P (\\<^bsub>L\<^esub>) \" using assms by (intro iffI rel_equivalence_onI inflationary_onI deflationary_onI) (auto dest!: inverse_onD) lemma reflexive_on_left_if_inflationary_on_unit_if_inverse_on: assumes "inverse_on P l r" and "inflationary_on P (\\<^bsub>L\<^esub>) \" shows "reflexive_on P (\\<^bsub>L\<^esub>)" using assms by (intro reflexive_onI) (auto dest!: inverse_onD) lemma rel_equivalence_on_unit_if_reflexive_on_if_inverse_on: assumes "inverse_on P l r" and "reflexive_on P (\\<^bsub>L\<^esub>)" shows "rel_equivalence_on P (\\<^bsub>L\<^esub>) \" using assms by (intro rel_equivalence_onI inflationary_onI deflationary_onI) (auto dest!: inverse_onD) end corollary rel_equivalence_on_unit_iff_reflexive_on_if_inverse_on: fixes P :: "'a \ bool" assumes "inverse_on P l r" shows "rel_equivalence_on P (\\<^bsub>L\<^esub>) \ \ reflexive_on P (\\<^bsub>L\<^esub>)" using assms reflexive_on_left_if_inflationary_on_unit_if_inverse_on rel_equivalence_on_unit_if_reflexive_on_if_inverse_on by (intro iffI) auto end text \Here is an example of a free theorem.\ notepad begin interpret flip : order_functors R L r l rewrites "flip.unit \ \" by (simp only: flip_unit_eq_counit) have "\((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r; reflexive_on P (\\<^bsub>R\<^esub>); \x y. \P x; r x \\<^bsub>L\<^esub> y\ \ x \\<^bsub>R\<^esub> l y\ \ inflationary_on P (\\<^bsub>R\<^esub>) \" for P by (fact flip.inflationary_on_unit_if_left_rel_right_if_left_right_relI) end end end diff --git a/thys/Transport/HOL_Basics/Orders/Partial_Equivalence_Relations.thy b/thys/Transport/HOL_Basics/Orders/Partial_Equivalence_Relations.thy --- a/thys/Transport/HOL_Basics/Orders/Partial_Equivalence_Relations.thy +++ b/thys/Transport/HOL_Basics/Orders/Partial_Equivalence_Relations.thy @@ -1,115 +1,116 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Partial Equivalence Relations\ theory Partial_Equivalence_Relations imports Binary_Relations_Symmetric Preorders begin -definition "partial_equivalence_rel_on P R \ transitive_on P R \ symmetric_on P R" +definition "partial_equivalence_rel_on \ transitive_on \ symmetric_on" lemma partial_equivalence_rel_onI [intro]: assumes "transitive_on P R" and "symmetric_on P R" shows "partial_equivalence_rel_on P R" - unfolding partial_equivalence_rel_on_def using assms by blast + unfolding partial_equivalence_rel_on_def using assms by auto lemma partial_equivalence_rel_onE [elim]: assumes "partial_equivalence_rel_on P R" obtains "transitive_on P R" "symmetric_on P R" - using assms unfolding partial_equivalence_rel_on_def by blast + using assms unfolding partial_equivalence_rel_on_def by auto lemma partial_equivalence_rel_on_rel_self_if_rel_dom: assumes "partial_equivalence_rel_on (P :: 'a \ bool) (R :: 'a \ 'a \ bool)" and "P x" "P y" and "R x y" shows "R x x" using assms by (blast dest: symmetric_onD transitive_onD) lemma partial_equivalence_rel_on_rel_self_if_rel_codom: assumes "partial_equivalence_rel_on (P :: 'a \ bool) (R :: 'a \ 'a \ bool)" and "P x" "P y" and "R x y" shows "R y y" using assms by (blast dest: symmetric_onD transitive_onD) lemma partial_equivalence_rel_on_rel_inv_iff_partial_equivalence_rel_on [iff]: "partial_equivalence_rel_on P R\ \ partial_equivalence_rel_on (P :: 'a \ bool) (R :: 'a \ _)" by blast -definition "partial_equivalence_rel (R :: 'a \ _) \ partial_equivalence_rel_on (\ :: 'a \ bool) R" +definition "(partial_equivalence_rel :: ('a \ _) \ bool) \ + partial_equivalence_rel_on (\ :: 'a \ bool)" lemma partial_equivalence_rel_eq_partial_equivalence_rel_on: - "partial_equivalence_rel (R :: 'a \ _) = partial_equivalence_rel_on (\ :: 'a \ bool) R" + "(partial_equivalence_rel :: ('a \ _) \ bool) = partial_equivalence_rel_on (\ :: 'a \ bool)" unfolding partial_equivalence_rel_def .. lemma partial_equivalence_relI [intro]: assumes "transitive R" and "symmetric R" shows "partial_equivalence_rel R" unfolding partial_equivalence_rel_eq_partial_equivalence_rel_on using assms by (intro partial_equivalence_rel_onI transitive_on_if_transitive symmetric_on_if_symmetric) lemma reflexive_on_in_field_if_partial_equivalence_rel: assumes "partial_equivalence_rel R" shows "reflexive_on (in_field R) R" using assms unfolding partial_equivalence_rel_eq_partial_equivalence_rel_on by (intro reflexive_onI) (blast intro: top1I partial_equivalence_rel_on_rel_self_if_rel_dom partial_equivalence_rel_on_rel_self_if_rel_codom) lemma partial_equivalence_relE [elim]: assumes "partial_equivalence_rel R" obtains "preorder_on (in_field R) R" "symmetric R" using assms unfolding partial_equivalence_rel_eq_partial_equivalence_rel_on by (elim partial_equivalence_rel_onE) (auto intro: reflexive_on_in_field_if_partial_equivalence_rel simp flip: transitive_eq_transitive_on symmetric_eq_symmetric_on) lemma partial_equivalence_rel_on_if_partial_equivalence_rel: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "partial_equivalence_rel R" shows "partial_equivalence_rel_on P R" using assms by (elim partial_equivalence_relE preorder_on_in_fieldE) (intro partial_equivalence_rel_onI transitive_on_if_transitive symmetric_on_if_symmetric) lemma partial_equivalence_rel_rel_inv_iff_partial_equivalence_rel [iff]: "partial_equivalence_rel R\ \ partial_equivalence_rel R" unfolding partial_equivalence_rel_eq_partial_equivalence_rel_on by blast corollary in_codom_eq_in_dom_if_partial_equivalence_rel: assumes "partial_equivalence_rel R" shows "in_codom R = in_dom R" using assms reflexive_on_in_field_if_partial_equivalence_rel in_codom_eq_in_dom_if_reflexive_on_in_field by auto lemma partial_equivalence_rel_rel_comp_self_eq_self: assumes "partial_equivalence_rel R" shows "(R \\ R) = R" using assms by (intro ext) (blast dest: symmetricD) lemma partial_equivalence_rel_if_partial_equivalence_rel_on_in_field: assumes "partial_equivalence_rel_on (in_field R) R" shows "partial_equivalence_rel R" using assms by (intro partial_equivalence_relI) (auto intro: transitive_if_transitive_on_in_field symmetric_if_symmetric_on_in_field) corollary partial_equivalence_rel_on_in_field_iff_partial_equivalence_rel [iff]: "partial_equivalence_rel_on (in_field R) R \ partial_equivalence_rel R" using partial_equivalence_rel_if_partial_equivalence_rel_on_in_field partial_equivalence_rel_on_if_partial_equivalence_rel by blast paragraph \Instantiations\ lemma partial_equivalence_rel_eq: "partial_equivalence_rel (=)" using transitive_eq symmetric_eq by (rule partial_equivalence_relI) lemma partial_equivalence_rel_top: "partial_equivalence_rel \" using transitive_top symmetric_top by (rule partial_equivalence_relI) end diff --git a/thys/Transport/HOL_Basics/Orders/Partial_Orders.thy b/thys/Transport/HOL_Basics/Orders/Partial_Orders.thy --- a/thys/Transport/HOL_Basics/Orders/Partial_Orders.thy +++ b/thys/Transport/HOL_Basics/Orders/Partial_Orders.thy @@ -1,61 +1,61 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Partial Orders\ theory Partial_Orders imports Binary_Relations_Antisymmetric Preorders begin -definition "partial_order_on P R \ preorder_on P R \ antisymmetric_on P R" +definition "partial_order_on \ preorder_on \ antisymmetric_on" lemma partial_order_onI [intro]: assumes "preorder_on P R" and "antisymmetric_on P R" shows "partial_order_on P R" - unfolding partial_order_on_def using assms by blast + unfolding partial_order_on_def using assms by auto lemma partial_order_onE [elim]: assumes "partial_order_on P R" obtains "preorder_on P R" "antisymmetric_on P R" - using assms unfolding partial_order_on_def by blast + using assms unfolding partial_order_on_def by auto lemma transitive_if_partial_order_on_in_field: assumes "partial_order_on (in_field R) R" shows "transitive R" using assms by (elim partial_order_onE) (rule transitive_if_preorder_on_in_field) lemma antisymmetric_if_partial_order_on_in_field: assumes "partial_order_on (in_field R) R" shows "antisymmetric R" using assms by (elim partial_order_onE) - (rule antisymmetric_if_antisymmetric_on_in_field) + (rule antisymmetric_if_antisymmetric_on_in_field) -definition "partial_order (R :: 'a \ _) \ partial_order_on (\ :: 'a \ bool) R" +definition "(partial_order :: ('a \ _) \ bool) \ partial_order_on (\ :: 'a \ bool)" lemma partial_order_eq_partial_order_on: - "partial_order (R :: 'a \ _) = partial_order_on (\ :: 'a \ bool) R" + "(partial_order :: ('a \ _) \ bool) = partial_order_on (\ :: 'a \ bool)" unfolding partial_order_def .. lemma partial_orderI [intro]: assumes "preorder R" and "antisymmetric R" shows "partial_order R" unfolding partial_order_eq_partial_order_on using assms by (intro partial_order_onI preorder_on_if_preorder antisymmetric_on_if_antisymmetric) lemma partial_orderE [elim]: assumes "partial_order R" obtains "preorder R" "antisymmetric R" using assms unfolding partial_order_eq_partial_order_on by (elim partial_order_onE) (simp only: preorder_eq_preorder_on antisymmetric_eq_antisymmetric_on) lemma partial_order_on_if_partial_order: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "partial_order R" shows "partial_order_on P R" using assms by (elim partial_orderE) (intro partial_order_onI preorder_on_if_preorder antisymmetric_on_if_antisymmetric) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Orders/Preorders.thy b/thys/Transport/HOL_Basics/Orders/Preorders.thy --- a/thys/Transport/HOL_Basics/Orders/Preorders.thy +++ b/thys/Transport/HOL_Basics/Orders/Preorders.thy @@ -1,94 +1,94 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Preorders\ theory Preorders imports Binary_Relations_Reflexive Binary_Relations_Transitive begin -definition "preorder_on P R \ reflexive_on P R \ transitive_on P R" +definition "preorder_on \ reflexive_on \ transitive_on" lemma preorder_onI [intro]: assumes "reflexive_on P R" and "transitive_on P R" shows "preorder_on P R" - unfolding preorder_on_def using assms by blast + unfolding preorder_on_def using assms by auto lemma preorder_onE [elim]: assumes "preorder_on P R" obtains "reflexive_on P R" "transitive_on P R" - using assms unfolding preorder_on_def by blast + using assms unfolding preorder_on_def by auto lemma reflexive_on_if_preorder_on: assumes "preorder_on P R" shows "reflexive_on P R" using assms by (elim preorder_onE) lemma transitive_on_if_preorder_on: assumes "preorder_on P R" shows "transitive_on P R" using assms by (elim preorder_onE) lemma transitive_if_preorder_on_in_field: assumes "preorder_on (in_field R) R" shows "transitive R" using assms by (elim preorder_onE) (rule transitive_if_transitive_on_in_field) corollary preorder_on_in_fieldE [elim]: assumes "preorder_on (in_field R) R" obtains "reflexive_on (in_field R) R" "transitive R" using assms by (blast dest: reflexive_on_if_preorder_on transitive_if_preorder_on_in_field) lemma preorder_on_rel_inv_if_preorder_on [iff]: "preorder_on P R\ \ preorder_on (P :: 'a \ bool) (R :: 'a \ _)" by auto lemma rel_if_all_rel_if_rel_if_reflexive_on: assumes "reflexive_on P R" and "\z. P z \ R x z \ R y z" and "P x" shows "R y x" using assms by blast lemma rel_if_all_rel_if_rel_if_reflexive_on': assumes "reflexive_on P R" and "\z. P z \ R z x \ R z y" and "P x" shows "R x y" using assms by blast -definition "preorder (R :: 'a \ _) \ preorder_on (\ :: 'a \ bool) R" +definition "(preorder :: ('a \ _) \ bool) \ preorder_on (\ :: 'a \ bool)" lemma preorder_eq_preorder_on: - "preorder (R :: 'a \ _) = preorder_on (\ :: 'a \ bool) R" + "(preorder :: ('a \ _) \ bool) = preorder_on (\ :: 'a \ bool)" unfolding preorder_def .. lemma preorderI [intro]: assumes "reflexive R" and "transitive R" shows "preorder R" unfolding preorder_eq_preorder_on using assms by (intro preorder_onI reflexive_on_if_reflexive transitive_on_if_transitive) lemma preorderE [elim]: assumes "preorder R" obtains "reflexive R" "transitive R" using assms unfolding preorder_eq_preorder_on by (elim preorder_onE) (simp only: reflexive_eq_reflexive_on transitive_eq_transitive_on) lemma preorder_on_if_preorder: fixes P :: "'a \ bool" and R :: "'a \ _" assumes "preorder R" shows "preorder_on P R" using assms by (elim preorderE) (intro preorder_onI reflexive_on_if_reflexive transitive_on_if_transitive) paragraph \Instantiations\ lemma preorder_eq: "preorder (=)" using reflexive_eq transitive_eq by (rule preorderI) end diff --git a/thys/Transport/HOL_Basics/Predicates/Predicates.thy b/thys/Transport/HOL_Basics/Predicates/Predicates.thy --- a/thys/Transport/HOL_Basics/Predicates/Predicates.thy +++ b/thys/Transport/HOL_Basics/Predicates/Predicates.thy @@ -1,23 +1,55 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Predicates\ theory Predicates imports Functions_Base Predicates_Order Predicates_Lattice begin paragraph \Summary\ text \Basic concepts on predicates.\ definition "pred_map f (P :: 'a \ bool) x \ P (f x)" lemma pred_map_eq [simp]: "pred_map f P x = P (f x)" unfolding pred_map_def by simp lemma comp_eq_pred_map [simp]: "P \ f = pred_map f P" by (intro ext) simp +definition "pred_if B P x \ B \ P x" + +bundle pred_if_syntax begin notation (output) pred_if (infixl "\" 50) end +bundle no_pred_if_syntax begin no_notation (output) pred_if (infixl "\" 50) end +unbundle pred_if_syntax + +lemma pred_if_eq_pred_if_pred [simp]: + assumes "B" + shows "(pred_if B P) = P" + unfolding pred_if_def using assms by blast + +lemma pred_if_eq_top_if_not_pred [simp]: + assumes "\B" + shows "(pred_if B P) = (\_. True)" + unfolding pred_if_def using assms by blast + +lemma pred_if_if_impI [intro]: + assumes "B \ P x" + shows "(pred_if B P) x" + unfolding pred_if_def using assms by blast + +lemma pred_ifE [elim]: + assumes "(pred_if B P) x" + obtains "\B" | "B" "P x" + using assms unfolding pred_if_def by blast + +lemma pred_ifD: + assumes "(pred_if B P) x" + and "B" + shows "P x" + using assms by blast + end \ No newline at end of file diff --git a/thys/Transport/ROOT b/thys/Transport/ROOT --- a/thys/Transport/ROOT +++ b/thys/Transport/ROOT @@ -1,50 +1,51 @@ chapter AFP -session Transport = "HOL-Algebra" + +session Transport = "HOL" + description "Transport via partial Galois connections and equivalences and basic libraries on top of HOL axioms." options [timeout = 600] sessions + "HOL-Algebra" "HOL-Library" "ML_Unification" directories "HOL_Basics" "HOL_Basics/Binary_Relations" "HOL_Basics/Binary_Relations/Order" "HOL_Basics/Binary_Relations/Properties" "HOL_Basics/Functions" "HOL_Basics/Functions/Properties" "HOL_Basics/HOL_Alignments" "HOL_Basics/HOL_Syntax_Bundles" "HOL_Basics/Orders" "HOL_Basics/Orders/Functions" "HOL_Basics/Orders/Functors" "HOL_Basics/Predicates" "HOL_Basics/Galois" "Transport" "Transport/Compositions" "Transport/Compositions/Agree" "Transport/Compositions/Generic" "Transport/Examples" "Transport/Examples/Prototype" "Transport/Examples/Typedef" "Transport/Functions" "Transport/Natural_Functors" theories HOL_Basics HOL_Alignments HOL_Algebra_Alignments HOL_Syntax_Bundles Transport Transport_Natural_Functors (*Examples*) Transport_Dep_Fun_Rel_Examples Transport_Lists_Sets_Examples Transport_Partial_Quotient_Types Transport_Typedef (*Paper*) Transport_Via_Partial_Galois_Connections_Equivalences_Paper document_files "root.tex" "root.bib" diff --git a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Relator.thy b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Relator.thy --- a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Relator.thy +++ b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Relator.thy @@ -1,91 +1,91 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Relator\ theory Transport_Compositions_Agree_Galois_Relator imports Transport_Compositions_Agree_Base begin context transport_comp_agree begin lemma left_Galois_le_comp_left_GaloisI: assumes in_codom_mono_r2: "([in_codom (\\<^bsub>R2\<^esub>)] \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" and r2_L2_self_if_in_codom: "\z. in_codom (\\<^bsub>R2\<^esub>) z \ r2 z \\<^bsub>L2\<^esub> r2 z" shows "(\<^bsub>L\<^esub>\) \ ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" proof (rule le_relI) fix x z assume "x \<^bsub>L\<^esub>\ z" then have "x \\<^bsub>L1\<^esub> r z" "in_codom (\\<^bsub>R\<^esub>) z" by auto - with \x \\<^bsub>L1\<^esub> r z\ in_codom_mono_r2 have "x \<^bsub>L1\<^esub>\ r2 z" by auto + with \x \\<^bsub>L1\<^esub> r z\ in_codom_mono_r2 have "x \<^bsub>L1\<^esub>\ r2 z" by fastforce moreover from \in_codom (\\<^bsub>R2\<^esub>) z\ r2_L2_self_if_in_codom have "r2 z \<^bsub>L2\<^esub>\ z" by (intro g2.left_GaloisI) auto ultimately show "((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\)) x z" by blast qed lemma comp_left_Galois_le_left_GaloisI: assumes mono_r1: "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and trans_L1: "transitive (\\<^bsub>L1\<^esub>)" and R1_r2_if_in_codom: "\y z. in_codom (\\<^bsub>R2\<^esub>) z \ y \\<^bsub>L2\<^esub> r2 z \ y \\<^bsub>R1\<^esub> r2 z" shows "((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\)) \ (\<^bsub>L\<^esub>\)" proof (rule le_relI) fix x z assume "((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\)) x z" then obtain y where "x \<^bsub>L1\<^esub>\ y" "y \<^bsub>L2\<^esub>\ z" by blast then have "x \\<^bsub>L1\<^esub> r1 y" "y \\<^bsub>L2\<^esub> r2 z" "in_codom (\\<^bsub>R\<^esub>) z" by auto with R1_r2_if_in_codom have "y \\<^bsub>R1\<^esub> r2 z" by blast with mono_r1 have "r1 y \\<^bsub>L1\<^esub> r z" by auto with \x \\<^bsub>L1\<^esub> r1 y\ \in_codom (\\<^bsub>R\<^esub>) z\ show "x \<^bsub>L\<^esub>\ z" using trans_L1 by blast qed corollary left_Galois_eq_comp_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "transitive (\\<^bsub>L1\<^esub>)" and "\z. in_codom (\\<^bsub>R2\<^esub>) z \ r2 z \\<^bsub>L2\<^esub> r2 z" and "\y z. in_codom (\\<^bsub>R2\<^esub>) z \ y \\<^bsub>L2\<^esub> r2 z \ y \\<^bsub>R1\<^esub> r2 z" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" using assms by (intro antisym left_Galois_le_comp_left_GaloisI comp_left_Galois_le_left_GaloisI dep_mono_wrt_predI) fastforce corollary left_Galois_eq_comp_left_GaloisI': assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "reflexive_on (in_codom (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>)" and "\y z. in_codom (\\<^bsub>R2\<^esub>) z \ y \\<^bsub>L2\<^esub> r2 z \ y \\<^bsub>R1\<^esub> r2 z" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" - using assms by (intro left_Galois_eq_comp_left_GaloisI) auto + using assms by (intro left_Galois_eq_comp_left_GaloisI) (auto 5 0) corollary left_Galois_eq_comp_left_GaloisI'': assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "reflexive_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "\y z. in_codom (\\<^bsub>R2\<^esub>) z \ y \\<^bsub>L2\<^esub> r2 z \ y \\<^bsub>R1\<^esub> r2 z" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" - using assms by (intro left_Galois_eq_comp_left_GaloisI) (auto 0 4) + using assms by (intro left_Galois_eq_comp_left_GaloisI) (auto 0 6) end context transport_comp_same begin lemma left_Galois_eq_comp_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) r2" and "reflexive_on (in_codom (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" using assms by (intro left_Galois_eq_comp_left_GaloisI') auto lemma left_Galois_eq_comp_left_GaloisI': assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "transitive (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) r2" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" using assms by (intro left_Galois_eq_comp_left_GaloisI'') auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Order_Equivalence.thy b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Order_Equivalence.thy --- a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Order_Equivalence.thy +++ b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Order_Equivalence.thy @@ -1,131 +1,131 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Order Equivalence\ theory Transport_Compositions_Agree_Order_Equivalence imports Transport_Compositions_Agree_Monotone begin context transport_comp_agree begin subsubsection \Unit\ paragraph \Inflationary\ lemma inflationary_on_unitI: assumes mono_l1: "([P] \\<^sub>m P') l1" and mono_r1: "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and inflationary_unit1: "inflationary_on P (\\<^bsub>L1\<^esub>) \\<^sub>1" and trans_L1: "transitive (\\<^bsub>L1\<^esub>)" and inflationary_unit2: "inflationary_on P' (\\<^bsub>L2\<^esub>) \\<^sub>2" and L2_le_R1: "\x. P x \ l1 x \\<^bsub>L2\<^esub> r2 (l x) \ l1 x \\<^bsub>R1\<^esub> r2 (l x)" shows "inflationary_on P (\\<^bsub>L\<^esub>) \" proof (rule inflationary_onI) fix x assume "P x" with mono_l1 have "P' (l1 x)" by blast with inflationary_unit2 have "l1 x \\<^bsub>L2\<^esub> r2 (l x)" by auto with L2_le_R1 \P x\ have "l1 x \\<^bsub>R1\<^esub> r2 (l x)" by blast with mono_r1 have "\\<^sub>1 x \\<^bsub>L1\<^esub> \ x" by auto moreover from inflationary_unit1 \P x\ have "x \\<^bsub>L1\<^esub> \\<^sub>1 x" by auto ultimately show "x \\<^bsub>L\<^esub> \ x" using trans_L1 by blast qed corollary inflationary_on_in_field_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "inflationary_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "inflationary_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ l1 x \\<^bsub>L2\<^esub> r2 (l x) \ l1 x \\<^bsub>R1\<^esub> r2 (l x)" shows "inflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" - using assms by (intro inflationary_on_unitI dep_mono_wrt_predI) auto + using assms by (intro inflationary_on_unitI dep_mono_wrt_predI) (auto 5 0) paragraph \Deflationary\ context begin interpretation inv : transport_comp_agree "(\\<^bsub>L1\<^esub>)" "(\\<^bsub>R1\<^esub>)" l1 r1 "(\\<^bsub>L2\<^esub>)" "(\\<^bsub>R2\<^esub>)" l2 r2 rewrites "\R S. (R\ \\<^sub>m S\) \ (R \\<^sub>m S)" and "\R. inflationary_on P R\ \ deflationary_on P R" and "\R. transitive R\ \ transitive R" and "\R. in_field R\ \ in_field R" by simp_all lemma deflationary_on_in_field_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "deflationary_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "deflationary_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ r2 (l x) \\<^bsub>L2\<^esub> l1 x \ r2 (l x) \\<^bsub>R1\<^esub> l1 x" shows "deflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms by (intro inv.inflationary_on_in_field_unitI[simplified rel_inv_iff_rel]) auto end text \Relational Equivalence\ corollary rel_equivalence_on_in_field_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "rel_equivalence_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ l1 x \\<^bsub>L2\<^esub> r2 (l x) \ l1 x \\<^bsub>R1\<^esub> r2 (l x)" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ r2 (l x) \\<^bsub>L2\<^esub> l1 x \ r2 (l x) \\<^bsub>R1\<^esub> l1 x" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms by (intro rel_equivalence_onI inflationary_on_in_field_unitI deflationary_on_in_field_unitI) auto subsubsection \Counit\ text \Corresponding lemmas for the counit can be obtained by flipping the interpretation of the locale.\ subsubsection \Order Equivalence\ interpretation flip : transport_comp_agree R2 L2 r2 l2 R1 L1 r1 l1 rewrites "flip.g1.unit \ \\<^sub>2" and "flip.g2.unit \ \\<^sub>1" and "flip.unit \ \" by (simp_all only: g1.flip_unit_eq_counit g2.flip_unit_eq_counit flip_unit_eq_counit) lemma order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>o (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>R2\<^esub>)" and "\x y. x \\<^bsub>L1\<^esub> y \ l1 x \\<^bsub>R1\<^esub> l1 y \ l1 x \\<^bsub>L2\<^esub> l1 y" and "\x y. x \\<^bsub>R2\<^esub> y \ r2 x \\<^bsub>L2\<^esub> r2 y \ r2 x \\<^bsub>R1\<^esub> r2 y" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ l1 x \\<^bsub>L2\<^esub> r2 (l x) \ l1 x \\<^bsub>R1\<^esub> r2 (l x)" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ r2 (l x) \\<^bsub>L2\<^esub> l1 x \ r2 (l x) \\<^bsub>R1\<^esub> l1 x" and "\x. in_field (\\<^bsub>R2\<^esub>) x \ r2 x \\<^bsub>R1\<^esub> l1 (r x) \ r2 x \\<^bsub>L2\<^esub> l1 (r x)" and "\x. in_field (\\<^bsub>R2\<^esub>) x \ l1 (r x) \\<^bsub>R1\<^esub> r2 x \ l1 (r x) \\<^bsub>L2\<^esub> r2 x" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalenceI rel_equivalence_on_in_field_unitI flip.rel_equivalence_on_in_field_unitI mono_wrt_rel_leftI flip.mono_wrt_rel_leftI dep_mono_wrt_relI) (auto elim!: g1.order_equivalenceE g2.order_equivalenceE) end context transport_comp_same begin lemma order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\<^sub>o (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>R2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (rule order_equivalenceI) auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Base.thy b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Base.thy --- a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Base.thy +++ b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Base.thy @@ -1,588 +1,589 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Generic Compositions\ subsection \Basic Setup\ theory Transport_Compositions_Generic_Base imports + Equivalence_Relations Transport_Base begin locale transport_comp = t1 : transport L1 R1 l1 r1 + t2 : transport L2 R2 l2 r2 for L1 :: "'a \ 'a \ bool" and R1 :: "'b \ 'b \ bool" and l1 :: "'a \ 'b" and r1 :: "'b \ 'a" and L2 :: "'b \ 'b \ bool" and R2 :: "'c \ 'c \ bool" and l2 :: "'b \ 'c" and r2 :: "'c \ 'b" begin text \This locale collects results about the composition of transportable components under some generic compatibility conditions on @{term "R1"} and @{term "L2"} (cf. below). The composition is rather subtle, but in return can cover quite general cases. Explanations and intuition about the construction can be found in \<^cite>\"transport"\.\ notation L1 (infix "\\<^bsub>L1\<^esub>" 50) notation R1 (infix "\\<^bsub>R1\<^esub>" 50) notation L2 (infix "\\<^bsub>L2\<^esub>" 50) notation R2 (infix "\\<^bsub>R2\<^esub>" 50) notation t1.ge_left (infix "\\<^bsub>L1\<^esub>" 50) notation t1.ge_right (infix "\\<^bsub>R1\<^esub>" 50) notation t2.ge_left (infix "\\<^bsub>L2\<^esub>" 50) notation t2.ge_right (infix "\\<^bsub>R2\<^esub>" 50) notation t1.left_Galois (infix "\<^bsub>L1\<^esub>\" 50) notation t1.right_Galois (infix "\<^bsub>R1\<^esub>\" 50) notation t2.left_Galois (infix "\<^bsub>L2\<^esub>\" 50) notation t2.right_Galois (infix "\<^bsub>R2\<^esub>\" 50) notation t1.ge_Galois_left (infix "\\<^bsub>L1\<^esub>" 50) notation t1.ge_Galois_right (infix "\\<^bsub>R1\<^esub>" 50) notation t2.ge_Galois_left (infix "\\<^bsub>L2\<^esub>" 50) notation t2.ge_Galois_right (infix "\\<^bsub>R2\<^esub>" 50) notation t1.right_ge_Galois (infix "\<^bsub>R1\<^esub>\" 50) notation t1.Galois_right (infix "\\<^bsub>R1\<^esub>" 50) notation t2.right_ge_Galois (infix "\<^bsub>R2\<^esub>\" 50) notation t2.Galois_right (infix "\\<^bsub>R2\<^esub>" 50) notation t1.left_ge_Galois (infix "\<^bsub>L1\<^esub>\" 50) notation t1.Galois_left (infix "\\<^bsub>L1\<^esub>" 50) notation t2.left_ge_Galois (infix "\<^bsub>L2\<^esub>\" 50) notation t2.Galois_left (infix "\\<^bsub>L2\<^esub>" 50) notation t1.unit ("\\<^sub>1") notation t1.counit ("\\<^sub>1") notation t2.unit ("\\<^sub>2") notation t2.counit ("\\<^sub>2") definition "L \ (\<^bsub>L1\<^esub>\) \\ (\\<^bsub>L2\<^esub>) \\ (\<^bsub>R1\<^esub>\)" lemma left_rel_eq_comp: "L = (\<^bsub>L1\<^esub>\) \\ (\\<^bsub>L2\<^esub>) \\ (\<^bsub>R1\<^esub>\)" unfolding L_def .. definition "l \ l2 \ l1" lemma left_eq_comp: "l = l2 \ l1" unfolding l_def .. lemma left_eq [simp]: "l x = l2 (l1 x)" unfolding left_eq_comp by simp context begin interpretation flip : transport_comp R2 L2 r2 l2 R1 L1 r1 l1 . abbreviation "R \ flip.L" abbreviation "r \ flip.l" lemma right_rel_eq_comp: "R = (\<^bsub>R2\<^esub>\) \\ (\\<^bsub>R1\<^esub>) \\ (\<^bsub>L2\<^esub>\)" unfolding flip.L_def .. lemma right_eq_comp: "r = r1 \ r2" unfolding flip.l_def .. lemma right_eq [simp]: "r z = r1 (r2 z)" unfolding right_eq_comp by simp lemmas transport_defs = left_rel_eq_comp left_eq_comp right_rel_eq_comp right_eq_comp end sublocale transport L R l r . (*FIXME: somehow the notation for the fixed parameters L and R, defined in Order_Functions_Base.thy, is lost. We hence re-declare it here.*) notation L (infix "\\<^bsub>L\<^esub>" 50) notation R (infix "\\<^bsub>R\<^esub>" 50) lemma left_relI [intro]: assumes "x \<^bsub>L1\<^esub>\ y" and "y \\<^bsub>L2\<^esub> y'" and "y' \<^bsub>R1\<^esub>\ x'" shows "x \\<^bsub>L\<^esub> x'" unfolding left_rel_eq_comp using assms by blast lemma left_relE [elim]: assumes "x \\<^bsub>L\<^esub> x'" obtains y y' where "x \<^bsub>L1\<^esub>\ y" "y \\<^bsub>L2\<^esub> y'" "y' \<^bsub>R1\<^esub>\ x'" using assms unfolding left_rel_eq_comp by blast context begin interpretation flip : transport_comp R2 L2 r2 l2 R1 L1 r1 l1 . interpretation inv : transport_comp "(\\<^bsub>L1\<^esub>)" "(\\<^bsub>R1\<^esub>)" l1 r1 "(\\<^bsub>L2\<^esub>)" "(\\<^bsub>R2\<^esub>)" l2 r2 . lemma ge_left_rel_eq_left_rel_inv_if_galois_prop [simp]: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" "((\\<^bsub>R1\<^esub>) \ (\\<^bsub>L1\<^esub>)) r1 l1" shows "(\\<^bsub>L\<^esub>) = transport_comp.L (\\<^bsub>L1\<^esub>) (\\<^bsub>R1\<^esub>) l1 r1 (\\<^bsub>L2\<^esub>)" using assms unfolding left_rel_eq_comp inv.left_rel_eq_comp by (simp add: rel_comp_assoc) corollary left_rel_inv_iff_left_rel_if_galois_prop [iff]: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" "((\\<^bsub>R1\<^esub>) \ (\\<^bsub>L1\<^esub>)) r1 l1" shows "(transport_comp.L (\\<^bsub>L1\<^esub>) (\\<^bsub>R1\<^esub>) l1 r1 (\\<^bsub>L2\<^esub>)) x x' \ x' \\<^bsub>L\<^esub> x" using assms by (simp flip: ge_left_rel_eq_left_rel_inv_if_galois_prop) subsubsection \Simplification of Relations\ lemma left_rel_le_left_rel1I: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>R1\<^esub>) \<^sub>h\ (\\<^bsub>L1\<^esub>)) r1 l1" and trans_L1: "transitive (\\<^bsub>L1\<^esub>)" and mono_l1: "((\\<^bsub>L\<^esub>) \\<^sub>m ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>R1\<^esub>))) l1" shows "(\\<^bsub>L\<^esub>) \ (\\<^bsub>L1\<^esub>)" proof (rule le_relI) fix x x' assume "x \\<^bsub>L\<^esub> x'" with mono_l1 obtain y where "l1 x \\<^bsub>R1\<^esub> y" "y \\<^bsub>R1\<^esub> l1 x'" by blast with \((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1\ \x \\<^bsub>L\<^esub> x'\ have "x \\<^bsub>L1\<^esub> r1 y" by blast moreover from \((\\<^bsub>R1\<^esub>) \<^sub>h\ (\\<^bsub>L1\<^esub>)) r1 l1\ \y \\<^bsub>R1\<^esub> l1 x'\ \x \\<^bsub>L\<^esub> x'\ have "... \\<^bsub>L1\<^esub> x'" by blast ultimately show "x \\<^bsub>L1\<^esub> x'" using trans_L1 by blast qed lemma left_rel1_le_left_relI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and mono_l1: "((\\<^bsub>L1\<^esub>) \\<^sub>m ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))) l1" shows "(\\<^bsub>L1\<^esub>) \ (\\<^bsub>L\<^esub>)" proof (rule le_relI) fix x x' assume "x \\<^bsub>L1\<^esub> x'" with mono_l1 obtain y y' where "l1 x \\<^bsub>R1\<^esub> y" "y \\<^bsub>L2\<^esub> y'" "y' \\<^bsub>R1\<^esub> l1 x'" by blast with \((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1\ \x \\<^bsub>L1\<^esub> x'\ have "x \<^bsub>L1\<^esub>\ y" by blast moreover note \y \\<^bsub>L2\<^esub> y'\ moreover from \y' \\<^bsub>R1\<^esub> l1 x'\ \x \\<^bsub>L1\<^esub> x'\ have "y' \<^bsub>R1\<^esub>\ x'" by blast ultimately show "x \\<^bsub>L\<^esub> x'" by blast qed corollary left_rel_eq_left_rel1I: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>R1\<^esub>) \<^sub>h\ (\\<^bsub>L1\<^esub>)) r1 l1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L\<^esub>) \\<^sub>m ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>R1\<^esub>))) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>m ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))) l1" shows "(\\<^bsub>L\<^esub>) = (\\<^bsub>L1\<^esub>)" using assms by (intro antisym left_rel_le_left_rel1I left_rel1_le_left_relI) text \Note that we may not necessarily have @{term "(\\<^bsub>L\<^esub>) = (\\<^bsub>L1\<^esub>)"}, even in case of equivalence relations. Depending on the use case, one thus may wish to use an alternative composition operation.\ lemma ex_order_equiv_left_rel_neq_left_rel1: "\(L1 :: bool \ _) (R1 :: bool \ _) l1 r1 (L2 :: bool \ _) (R2 :: bool \ _) l2 r2. (L1 \\<^sub>o R1) l1 r1 \ equivalence_rel L1 \ equivalence_rel R1 \ (L2 \\<^sub>o R2) l2 r2 \ equivalence_rel L2 \ equivalence_rel R2 \ transport_comp.L L1 R1 l1 r1 L2 \ L1" proof (intro exI conjI) let ?L1 = "(=) :: bool \ _" let ?R1 = ?L1 let ?l1 = id let ?r1 = ?l1 let ?L2 = "\ :: bool \ bool \ bool" let ?R2 = ?L2 let ?l2 = id let ?r2 = ?l2 interpret tc : transport_comp ?L1 ?R1 ?l1 ?r1 ?L2 ?R2 ?l2 ?r2 . show "(?L1 \\<^sub>o ?R1) ?l1 ?r1" by fastforce show "equivalence_rel ?L1" "equivalence_rel ?R1" by (fact equivalence_eq)+ show "(?L2 \\<^sub>o ?R2) ?l2 ?r2" by fastforce show "equivalence_rel ?L2" "equivalence_rel ?R2" by (fact equivalence_top)+ show "tc.L \ ?L1" proof - have "\(?L1 False True)" by blast moreover have "tc.L False True" by (intro tc.left_relI) auto ultimately show ?thesis by auto qed qed end subsubsection \Generic Left to Right Introduction Rules\ text \The following lemmas generalise the proof outline used, for example, when proving monotonicity and the Galois property (cf. the paper \<^cite>\"transport"\).\ interpretation flip : transport_comp R2 L2 r2 l2 R1 L1 r1 l1 . lemma right_rel_if_left_relI: assumes "x \\<^bsub>L\<^esub> x'" and l1_R1_if_L1_r1: "\y. in_codom (\\<^bsub>R1\<^esub>) y \ x \\<^bsub>L1\<^esub> r1 y \ l1 x \\<^bsub>R1\<^esub> y" and t_R1_if_l1_R1: "\y. l1 x \\<^bsub>R1\<^esub> y \ t y \\<^bsub>R1\<^esub> y" and R2_l2_if_t_L2_if_l1_R1: "\y y'. l1 x \\<^bsub>R1\<^esub> y \ t y \\<^bsub>L2\<^esub> y' \ z \\<^bsub>R2\<^esub> l2 y'" and R1_b_if_R1_l1_if_R1_l1: "\y y'. y \\<^bsub>R1\<^esub> l1 x' \ y' \\<^bsub>R1\<^esub> l1 x' \ y' \\<^bsub>R1\<^esub> b y" and b_L2_r2_if_in_codom_L2_b_if_R1_l1: "\y. y \\<^bsub>R1\<^esub> l1 x' \ in_codom (\\<^bsub>L2\<^esub>) (b y) \ b y \\<^bsub>L2\<^esub> r2 z'" and in_codom_R2_if_in_codom_L2_b_if_R1_l1: "\y. y \\<^bsub>R1\<^esub> l1 x' \ in_codom (\\<^bsub>L2\<^esub>) (b y) \ in_codom (\\<^bsub>R2\<^esub>) z'" and rel_comp_le: "(\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)" and in_codom_rel_comp_le: "in_codom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom (\\<^bsub>L2\<^esub>)" shows "z \\<^bsub>R\<^esub> z'" proof - from \x \\<^bsub>L\<^esub> x'\ obtain yl yl' where "l1 x \\<^bsub>R1\<^esub> yl" "yl \\<^bsub>L2\<^esub> yl'" "yl' \\<^bsub>R1\<^esub> l1 x'" using l1_R1_if_L1_r1 by blast moreover then have "t yl \\<^bsub>R1\<^esub> yl" by (intro t_R1_if_l1_R1) ultimately have "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) (t yl) (l1 x')" using rel_comp_le by blast then obtain y where "t yl \\<^bsub>L2\<^esub> y" "y \\<^bsub>R1\<^esub> l1 x'" by blast show "z \\<^bsub>R\<^esub> z'" proof (rule flip.left_relI) from \t yl \\<^bsub>L2\<^esub> y\ \l1 x \\<^bsub>R1\<^esub> yl\ show "z \<^bsub>R2\<^esub>\ y" by (auto intro: R2_l2_if_t_L2_if_l1_R1) from \yl' \\<^bsub>R1\<^esub> l1 x'\ \y \\<^bsub>R1\<^esub> l1 x'\ show "y \\<^bsub>R1\<^esub> b yl'" by (rule R1_b_if_R1_l1_if_R1_l1) show "b yl' \<^bsub>L2\<^esub>\ z'" proof (rule t2.left_GaloisI) from \yl' \\<^bsub>R1\<^esub> l1 x'\ have "yl' \\<^bsub>R1\<^esub> b yl'" by (intro R1_b_if_R1_l1_if_R1_l1) with \l1 x \\<^bsub>R1\<^esub> yl\ \yl \\<^bsub>L2\<^esub> yl'\ in_codom_rel_comp_le have "in_codom (\\<^bsub>L2\<^esub>) (b yl')" by blast with \yl' \\<^bsub>R1\<^esub> l1 x'\ show "b yl' \\<^bsub>L2\<^esub> r2 z'" "in_codom (\\<^bsub>R2\<^esub>) z'" by (auto intro: b_L2_r2_if_in_codom_L2_b_if_R1_l1 in_codom_R2_if_in_codom_L2_b_if_R1_l1) qed qed qed lemma right_rel_if_left_relI': assumes "x \\<^bsub>L\<^esub> x'" and l1_R1_if_L1_r1: "\y. in_codom (\\<^bsub>R1\<^esub>) y \ x \\<^bsub>L1\<^esub> r1 y \ l1 x \\<^bsub>R1\<^esub> y" and R1_b_if_R1_l1: "\y. y \\<^bsub>R1\<^esub> l1 x' \ y \\<^bsub>R1\<^esub> b y" and L2_r2_if_L2_b_if_R1_l1: "\y y'. y \\<^bsub>R1\<^esub> l1 x' \ y' \\<^bsub>L2\<^esub> b y \ y' \\<^bsub>L2\<^esub> r2 z'" and in_codom_R2_if_L2_b_if_R1_l1: "\y y'. y \\<^bsub>R1\<^esub> l1 x' \ y' \\<^bsub>L2\<^esub> b y \ in_codom (\\<^bsub>R2\<^esub>) z'" and t_R1_if_R1_l1_if_l1_R1: "\y y' y''. l1 x \\<^bsub>R1\<^esub> y \ l1 x \\<^bsub>R1\<^esub> y' \ t y \\<^bsub>R1\<^esub> y'" and R2_l2_t_if_in_dom_L2_t_if_l1_R1: "\y y'. l1 x \\<^bsub>R1\<^esub> y \ in_dom (\\<^bsub>L2\<^esub>) (t y) \ z \\<^bsub>R2\<^esub> l2 (t y)" and in_codom_L2_t_if_in_dom_L2_t_if_l1_R1: "\y y'. l1 x \\<^bsub>R1\<^esub> y \ in_dom (\\<^bsub>L2\<^esub>) (t y) \ in_codom (\\<^bsub>L2\<^esub>) (t y)" and rel_comp_le: "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" and in_dom_rel_comp_le: "in_dom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_dom (\\<^bsub>L2\<^esub>)" shows "z \\<^bsub>R\<^esub> z'" proof - from \x \\<^bsub>L\<^esub> x'\ obtain yl yl' where "l1 x \\<^bsub>R1\<^esub> yl" "yl \\<^bsub>L2\<^esub> yl'" "yl' \\<^bsub>R1\<^esub> l1 x'" using l1_R1_if_L1_r1 by blast moreover then have "yl' \\<^bsub>R1\<^esub> b yl'" by (intro R1_b_if_R1_l1) ultimately have "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) (l1 x) (b yl')" using rel_comp_le by blast then obtain y where "l1 x \\<^bsub>R1\<^esub> y" "y \\<^bsub>L2\<^esub> b yl'" by blast show "z \\<^bsub>R\<^esub> z'" proof (rule flip.left_relI) from \yl' \\<^bsub>R1\<^esub> l1 x'\ \y \\<^bsub>L2\<^esub> b yl'\ have "in_codom (\\<^bsub>R2\<^esub>) z'" "y \\<^bsub>L2\<^esub> r2 z'" by (auto intro: in_codom_R2_if_L2_b_if_R1_l1 L2_r2_if_L2_b_if_R1_l1) then show "y \<^bsub>L2\<^esub>\ z'" by blast from \l1 x \\<^bsub>R1\<^esub> yl\ \l1 x \\<^bsub>R1\<^esub> y\ show "t yl \\<^bsub>R1\<^esub> y" by (rule t_R1_if_R1_l1_if_l1_R1) show "z \<^bsub>R2\<^esub>\ t yl" proof (rule flip.t1.left_GaloisI) from \l1 x \\<^bsub>R1\<^esub> yl\ have "t yl \\<^bsub>R1\<^esub> yl" by (intro t_R1_if_R1_l1_if_l1_R1) with \yl \\<^bsub>L2\<^esub> yl'\ \yl' \\<^bsub>R1\<^esub> l1 x'\ in_dom_rel_comp_le have "in_dom (\\<^bsub>L2\<^esub>) (t yl)" by blast with \l1 x \\<^bsub>R1\<^esub> yl\ show "z \\<^bsub>R2\<^esub> l2 (t yl)" "in_codom (\\<^bsub>L2\<^esub>) (t yl)" by (auto intro: R2_l2_t_if_in_dom_L2_t_if_l1_R1 in_codom_L2_t_if_in_dom_L2_t_if_l1_R1) qed qed qed subsubsection \Simplification of Monotonicity Assumptions\ text \Some sufficient conditions for monotonicity assumptions that repeatedly arise in various places.\ lemma mono_in_dom_left_rel_left1_if_in_dom_rel_comp_le: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "in_dom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_dom (\\<^bsub>L2\<^esub>)" shows "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" using assms by (intro dep_mono_wrt_predI) blast lemma mono_in_codom_left_rel_left1_if_in_codom_rel_comp_le: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "in_codom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom (\\<^bsub>L2\<^esub>)" shows "([in_codom (\\<^bsub>L\<^esub>)] \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" using assms by (intro dep_mono_wrt_predI) blast subsubsection \Simplification of Compatibility Conditions\ text \Most results will depend on certain compatibility conditions between @{term "(\\<^bsub>R1\<^esub>)"} and @{term "(\\<^bsub>L2\<^esub>)"}. We next derive some sufficient assumptions for these conditions.\ end lemma rel_comp_comp_le_rel_comp_if_rel_comp_comp_if_in_dom_leI: assumes trans_R: "transitive R" and refl_S: "reflexive_on P S" and in_dom_le: "in_dom (R \\ S \\ R) \ P" and rel_comp_le: "(S \\ R \\ S) \ (S \\ R)" shows "(R \\ S \\ R) \ (S \\ R)" proof (intro le_relI) fix x y assume "(R \\ S \\ R) x y" moreover with in_dom_le refl_S have "S x x" by blast ultimately have "((S \\ R \\ S) \\ R) x y" by blast with rel_comp_le have "(S \\ R \\ R) x y" by blast with trans_R show "(S \\ R) x y" by blast qed lemma rel_comp_comp_le_rel_comp_if_rel_comp_comp_if_in_codom_leI: assumes trans_R: "transitive R" and refl_S: "reflexive_on P S" and in_codom_le: "in_codom (R \\ S \\ R) \ P" and rel_comp_le: "(S \\ R \\ S) \ (R \\ S)" shows "(R \\ S \\ R) \ (R \\ S)" proof (intro le_relI) fix x y assume "(R \\ S \\ R) x y" moreover with in_codom_le refl_S have "S y y" by blast ultimately have "(R \\ (S \\ R \\ S)) x y" by blast with rel_comp_le have "(R \\ R \\ S) x y" by blast with trans_R show "(R \\ S) x y" by blast qed lemma rel_comp_comp_le_rel_comp_if_rel_comp_le_if_transitive: assumes trans_R: "transitive R" and R_S_le: "(R \\ S) \ (S \\ R)" shows "(R \\ S \\ R) \ (S \\ R)" proof - from trans_R have R_R_le: "(R \\ R) \ R" by (intro rel_comp_le_self_if_transitive) have "(R \\ S \\ R) \ (S \\ R \\ R)" using monoD[OF mono_rel_comp1 R_S_le] by blast also have "... \ (S \\ R)" using monoD[OF mono_rel_comp2 R_R_le] by (auto simp flip: rel_comp_assoc) finally show ?thesis . qed lemma rel_comp_comp_le_rel_comp_if_rel_comp_le_if_transitive': assumes trans_R: "transitive R" and S_R_le: "(S \\ R) \ (R \\ S)" shows "(R \\ S \\ R) \ (R \\ S)" proof - from trans_R have R_R_le: "(R \\ R) \ R" by (intro rel_comp_le_self_if_transitive) have "(R \\ S \\ R) \ (R \\ R \\ S)" using monoD[OF mono_rel_comp2 S_R_le] by (auto simp flip: rel_comp_assoc) also have "... \ (R \\ S)" using monoD[OF mono_rel_comp1 R_R_le] by blast finally show ?thesis . qed lemma rel_comp_eq_rel_comp_if_le_if_transitive_if_reflexive: assumes refl_R: "reflexive_on (in_field S) R" and trans_S: "transitive S" and R_le: "R \ S \ (=)" shows "(R \\ S) = (S \\ R)" proof (intro ext iffI) fix x y assume "(R \\ S) x y" then obtain z where "R x z" "S z y" by blast with R_le have "(S \ (=)) x z" by blast with \S z y\ trans_S have "S x y" by auto moreover from \S z y\ refl_R have "R y y" by blast ultimately show "(S \\ R) x y" by blast next fix x y assume "(S \\ R) x y" then obtain z where "S x z" "R z y" by blast with R_le have "(S \ (=)) z y" by blast with \S x z\ trans_S have "S x y" by auto moreover from \S x y\ refl_R have "R x x" by blast ultimately show "(R \\ S) x y" by blast qed lemma rel_comp_eq_rel_comp_if_in_field_le_if_le_eq: assumes le_eq: "R \ (=)" and in_field_le: "in_field S \ in_field R" shows "(R \\ S) = (S \\ R)" proof (intro ext iffI) fix x y assume "(R \\ S) x y" then obtain z where "R x z" "S z y" by blast with le_eq have "S x y" by blast with assms show "(S \\ R) x y" by blast next fix x y assume "(S \\ R) x y" then obtain z where "S x z" "R z y" by blast with le_eq have "S x y" by blast with assms show "(R \\ S) x y" by blast qed context transport_comp begin lemma left2_right1_left2_le_left2_right1_if_right1_left2_right1_le_left2_right1: assumes "reflexive_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "transitive (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" shows "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" using assms by (intro rel_comp_comp_le_rel_comp_if_rel_comp_comp_if_in_codom_leI) auto lemma left2_right1_left2_le_right1_left2_if_right1_left2_right1_le_right1_left2I: assumes "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "transitive (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" and "in_dom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_dom (\\<^bsub>R1\<^esub>)" shows "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" using assms by (intro rel_comp_comp_le_rel_comp_if_rel_comp_comp_if_in_dom_leI) auto lemma in_dom_right1_left2_right1_le_if_right1_left2_right1_le: assumes "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" shows "in_dom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_dom (\\<^bsub>L2\<^esub>)" using monoD[OF mono_in_dom assms] by (auto intro: in_dom_if_in_dom_rel_comp) lemma in_codom_right1_left2_right1_le_if_right1_left2_right1_le: assumes "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" shows "in_codom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom (\\<^bsub>L2\<^esub>)" using monoD[OF mono_in_codom assms] by (auto intro: in_codom_if_in_codom_rel_comp) text \Our main results will be derivable for two different sets of compatibility conditions. The next two lemmas show the equivalence between those two sets under certain assumptions. In cases where these assumptions are met, we will only state the result for one of the two compatibility conditions. The other one will then be derivable using one of the following lemmas.\ definition "middle_compatible_dom \ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \ in_dom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_dom (\\<^bsub>L2\<^esub>) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_dom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_dom (\\<^bsub>R1\<^esub>)" lemma middle_compatible_domI [intro]: assumes "(\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)" and "in_dom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_dom (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" and "in_dom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_dom (\\<^bsub>R1\<^esub>)" shows "middle_compatible_dom" unfolding middle_compatible_dom_def using assms by blast lemma middle_compatible_domE [elim]: assumes "middle_compatible_dom" obtains "(\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)" and "in_dom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_dom (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" and "in_dom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_dom (\\<^bsub>R1\<^esub>)" using assms unfolding middle_compatible_dom_def by blast definition "middle_compatible_codom \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom (\\<^bsub>L2\<^esub>) \ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \ in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" lemma middle_compatible_codomI [intro]: assumes "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" and "in_codom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom (\\<^bsub>L2\<^esub>)" and "(\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" shows "middle_compatible_codom" unfolding middle_compatible_codom_def using assms by blast lemma middle_compatible_codomE [elim]: assumes "middle_compatible_codom" obtains "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" and "in_codom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom (\\<^bsub>L2\<^esub>)" and "(\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" using assms unfolding middle_compatible_codom_def by blast context begin interpretation flip : transport_comp R2 L2 r2 l2 R1 L1 r1 l1 . lemma rel_comp_comp_le_assms_if_in_codom_rel_comp_comp_leI: assumes "preorder_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "preorder_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "middle_compatible_codom" shows "middle_compatible_dom" using assms by (intro middle_compatible_domI) (auto intro!: left2_right1_left2_le_left2_right1_if_right1_left2_right1_le_left2_right1 flip.left2_right1_left2_le_left2_right1_if_right1_left2_right1_le_left2_right1 in_dom_right1_left2_right1_le_if_right1_left2_right1_le flip.in_dom_right1_left2_right1_le_if_right1_left2_right1_le intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom) lemma rel_comp_comp_le_assms_if_in_dom_rel_comp_comp_leI: assumes "preorder_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "preorder_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "middle_compatible_dom" shows "middle_compatible_codom" using assms by (intro middle_compatible_codomI) (auto intro!: left2_right1_left2_le_right1_left2_if_right1_left2_right1_le_right1_left2I flip.left2_right1_left2_le_right1_left2_if_right1_left2_right1_le_right1_left2I in_codom_right1_left2_right1_le_if_right1_left2_right1_le flip.in_codom_right1_left2_right1_le_if_right1_left2_right1_le intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) lemma middle_compatible_dom_iff_middle_compatible_codom_if_preorder_on: assumes "preorder_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "preorder_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" shows "middle_compatible_dom \ middle_compatible_codom" using assms by (intro iffI rel_comp_comp_le_assms_if_in_codom_rel_comp_comp_leI) (auto intro!: rel_comp_comp_le_assms_if_in_dom_rel_comp_comp_leI) end text \Finally we derive some sufficient assumptions for the compatibility conditions.\ lemma right1_left2_right1_le_assms_if_right1_left2_eqI: assumes "transitive (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) = ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" shows "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" using assms rel_comp_comp_le_rel_comp_if_rel_comp_le_if_transitive[of R1 L2] by auto interpretation flip : transport_comp R2 L2 r2 l2 R1 L1 r1 l1 rewrites "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) = ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) = ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" by (simp only: eq_commute) lemma middle_compatible_codom_if_rel_comp_eq_if_transitive: assumes "transitive (\\<^bsub>R1\<^esub>)" "transitive (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) = ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" shows "middle_compatible_codom" using assms by (intro middle_compatible_codomI in_codom_right1_left2_right1_le_if_right1_left2_right1_le flip.in_codom_right1_left2_right1_le_if_right1_left2_right1_le right1_left2_right1_le_assms_if_right1_left2_eqI flip.right1_left2_right1_le_assms_if_right1_left2_eqI) auto lemma middle_compatible_codom_if_right1_le_left2_eqI: assumes "preorder_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" "transitive (\\<^bsub>L2\<^esub>)" and "(\\<^bsub>R1\<^esub>) \ (\\<^bsub>L2\<^esub>) \ (=)" and "in_field (\\<^bsub>L2\<^esub>) \ in_field (\\<^bsub>R1\<^esub>)" shows "middle_compatible_codom" using assms by (intro middle_compatible_codomI in_codom_right1_left2_right1_le_if_right1_left2_right1_le flip.in_codom_right1_left2_right1_le_if_right1_left2_right1_le right1_left2_right1_le_assms_if_right1_left2_eqI flip.right1_left2_right1_le_assms_if_right1_left2_eqI rel_comp_eq_rel_comp_if_le_if_transitive_if_reflexive) (auto intro: reflexive_on_if_le_pred_if_reflexive_on) lemma middle_compatible_codom_if_right1_le_eqI: assumes "(\\<^bsub>R1\<^esub>) \ (=)" and "transitive (\\<^bsub>L2\<^esub>)" and "in_field (\\<^bsub>L2\<^esub>) \ in_field (\\<^bsub>R1\<^esub>)" shows "middle_compatible_codom" using assms by (intro middle_compatible_codomI in_codom_right1_left2_right1_le_if_right1_left2_right1_le flip.in_codom_right1_left2_right1_le_if_right1_left2_right1_le right1_left2_right1_le_assms_if_right1_left2_eqI flip.right1_left2_right1_le_assms_if_right1_left2_eqI rel_comp_eq_rel_comp_if_in_field_le_if_le_eq) auto end end diff --git a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Relator.thy b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Relator.thy --- a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Relator.thy +++ b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Relator.thy @@ -1,151 +1,151 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Relator\ theory Transport_Compositions_Generic_Galois_Relator imports Transport_Compositions_Generic_Base begin context transport_comp begin interpretation flip : transport_comp R2 L2 r2 l2 R1 L1 r1 l1 rewrites "flip.t2.unit \ \\<^sub>1" by (simp only: t1.flip_unit_eq_counit) lemma left_Galois_le_comp_left_GaloisI: assumes mono_r1: "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and galois_prop1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and preorder_R1: "preorder_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and rel_comp_le: "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" and mono_in_codom_r2: "([in_codom (\\<^bsub>R\<^esub>)] \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" shows "(\<^bsub>L\<^esub>\) \ ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" proof (rule le_relI) fix x z assume "x \<^bsub>L\<^esub>\ z" then have "in_codom (\\<^bsub>R\<^esub>) z" "x \\<^bsub>L\<^esub> r z" by auto with galois_prop1 obtain y y' where "in_dom (\\<^bsub>L1\<^esub>) x" "l1 x \\<^bsub>R1\<^esub> y" "y \\<^bsub>L2\<^esub> y'" "y' \\<^bsub>R1\<^esub> \\<^sub>1 (r2 z)" by (auto elim!: left_relE) moreover have "\\<^sub>1 (r2 z) \\<^bsub>R1\<^esub> r2 z" proof - from mono_in_codom_r2 \in_codom (\\<^bsub>R\<^esub>) z\ have "in_codom (\\<^bsub>R1\<^esub>) (r2 z)" by blast with mono_r1 galois_prop1 preorder_R1 show ?thesis by (blast intro!: t1.counit_rel_if_reflexive_on_if_half_galois_prop_left_if_mono_wrt_rel) qed ultimately have "y' \\<^bsub>R1\<^esub> r2 z" using preorder_R1 by blast with \l1 x \\<^bsub>R1\<^esub> y\ \y \\<^bsub>L2\<^esub> y'\ have "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) (l1 x) (r2 z)" by blast with rel_comp_le obtain y'' where "l1 x \\<^bsub>R1\<^esub> y''" "y'' \\<^bsub>L2\<^esub> r2 z" by blast with galois_prop1 \in_dom (\\<^bsub>L1\<^esub>) x\ have "x \<^bsub>L1\<^esub>\ y''" by (intro t1.left_Galois_if_Galois_right_if_half_galois_prop_right t1.left_GaloisI) auto moreover from \in_codom (\\<^bsub>R\<^esub>) z\ \y'' \\<^bsub>L2\<^esub> r2 z\ have "y'' \<^bsub>L2\<^esub>\ z" by (intro t2.left_GaloisI) auto ultimately show "((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\)) x z" by blast qed lemma comp_left_Galois_le_left_GaloisI: assumes mono_r1: "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and half_galois_prop_left1: "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and half_galois_prop_right1: "((\\<^bsub>R1\<^esub>) \\<^sub>h (\\<^bsub>L1\<^esub>)) r1 l1" and refl_R1: "reflexive_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and mono_l2: "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and refl_L2: "reflexive_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and in_codom_rel_comp_le: "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" shows "((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\)) \ (\<^bsub>L\<^esub>\)" proof (intro le_relI left_GaloisI) fix x z assume "((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\)) x z" from \((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\)) x z\ obtain y where "x \<^bsub>L1\<^esub>\ y" "y \<^bsub>L2\<^esub>\ z" by blast with half_galois_prop_left1 have "l1 x \\<^bsub>R1\<^esub> y" "y \\<^bsub>L2\<^esub> r2 z" by auto with refl_R1 refl_L2 have "y \\<^bsub>R1\<^esub> y" "y \\<^bsub>L2\<^esub> y" by auto show "in_codom (\\<^bsub>R\<^esub>) z" proof (intro in_codomI flip.left_relI) from mono_l2 \y \\<^bsub>L2\<^esub> y\ show "l2 y \<^bsub>R2\<^esub>\ y" by blast show "y \\<^bsub>R1\<^esub> y" "y \<^bsub>L2\<^esub>\ z" by fact+ qed show "x \\<^bsub>L\<^esub> r z" proof (intro left_relI) show "x \<^bsub>L1\<^esub>\ y" "y \\<^bsub>L2\<^esub> r2 z" by fact+ show "r2 z \<^bsub>R1\<^esub>\ r z" proof (intro flip.t2.left_GaloisI) from \y \\<^bsub>L2\<^esub> y\ \y \\<^bsub>R1\<^esub> y\ \y \\<^bsub>L2\<^esub> r2 z\ have "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) y (r2 z)" by blast with in_codom_rel_comp_le have "in_codom (\\<^bsub>R1\<^esub>) (r2 z)" by blast with refl_R1 have "r2 z \\<^bsub>R1\<^esub> r2 z" by blast with mono_r1 show "in_codom (\\<^bsub>L1\<^esub>) (r z)" by auto with \r2 z \\<^bsub>R1\<^esub> r2 z\ half_galois_prop_right1 mono_r1 - show "r2 z \\<^bsub>R1\<^esub> l1 (r z)" by (auto intro: + show "r2 z \\<^bsub>R1\<^esub> l1 (r z)" by (fastforce intro: flip.t2.rel_unit_if_left_rel_if_half_galois_prop_right_if_mono_wrt_rel) qed qed qed corollary left_Galois_eq_comp_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>R1\<^esub>) \\<^sub>h (\\<^bsub>L1\<^esub>)) r1 l1" and "preorder_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "reflexive_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" and "([in_codom (\\<^bsub>R\<^esub>)] \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" using assms by (intro antisym left_Galois_le_comp_left_GaloisI comp_left_Galois_le_left_GaloisI) (auto elim!: preorder_on_in_fieldE intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom) corollary left_Galois_eq_comp_left_GaloisI': assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>R1\<^esub>) \\<^sub>h (\\<^bsub>L1\<^esub>)) r1 l1" and "preorder_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \<^sub>h\ (\\<^bsub>L2\<^esub>)) r2 l2" and "reflexive_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" using assms by (intro left_Galois_eq_comp_left_GaloisI flip.mono_in_codom_left_rel_left1_if_in_codom_rel_comp_le) auto theorem left_Galois_eq_comp_left_Galois_if_galois_connection_if_galois_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "preorder_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \ (\\<^bsub>L2\<^esub>)) r2 l2" and "reflexive_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" using assms by (intro left_Galois_eq_comp_left_GaloisI') (auto elim!: t1.galois_equivalenceE) corollary left_Galois_eq_comp_left_Galois_if_galois_connection_if_galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "preorder_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \ (\\<^bsub>L2\<^esub>)) r2 l2" and "reflexive_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "in_codom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" using assms by (intro left_Galois_eq_comp_left_Galois_if_galois_connection_if_galois_equivalenceI' flip.left2_right1_left2_le_left2_right1_if_right1_left2_right1_le_left2_right1) auto corollary left_Galois_eq_comp_left_Galois_if_preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>R2\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>L2\<^esub>)) r2 l2" and "in_codom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom (\\<^bsub>L2\<^esub>)" and "(\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom (\\<^bsub>R1\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \\ (\<^bsub>L2\<^esub>\))" using assms by (intro left_Galois_eq_comp_left_Galois_if_galois_connection_if_galois_equivalenceI) auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Monotone.thy b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Monotone.thy --- a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Monotone.thy +++ b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Monotone.thy @@ -1,57 +1,57 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Monotonicity\ theory Transport_Compositions_Generic_Monotone imports Transport_Compositions_Generic_Base begin context transport_comp begin lemma mono_wrt_rel_leftI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and inflationary_unit2: "inflationary_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" and "in_codom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_codom (\\<^bsub>L2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" proof (rule dep_mono_wrt_relI) fix x x' assume "x \\<^bsub>L\<^esub> x'" then show "l x \\<^bsub>R\<^esub> l x'" proof (rule right_rel_if_left_relI) fix y' assume "l1 x \\<^bsub>L2\<^esub> y'" with \((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2\ show "l x \\<^bsub>R2\<^esub> l2 y'" by auto next assume "in_codom (\\<^bsub>L2\<^esub>) (l1 x')" with inflationary_unit2 show "l1 x' \\<^bsub>L2\<^esub> r2 (l x')" by auto from \in_codom (\\<^bsub>L2\<^esub>) (l1 x')\ \((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2\ - show "in_codom (\\<^bsub>R2\<^esub>) (l x')" by auto + show "in_codom (\\<^bsub>R2\<^esub>) (l x')" by fastforce qed (insert assms, auto) qed lemma mono_wrt_rel_leftI': assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>L2\<^esub>) \\<^sub>h (\\<^bsub>R2\<^esub>)) l2 r2" and refl_L2: "reflexive_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" and "in_dom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_dom (\\<^bsub>L2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" proof (rule dep_mono_wrt_relI) fix x x' assume "x \\<^bsub>L\<^esub> x'" then show "l x \\<^bsub>R\<^esub> l x'" proof (rule right_rel_if_left_relI') fix y' assume "y' \\<^bsub>L2\<^esub> l1 x'" moreover with \((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2\ have "l2 y' \\<^bsub>R2\<^esub> l x'" by auto ultimately show "in_codom (\\<^bsub>R2\<^esub>) (l x')" "y' \\<^bsub>L2\<^esub> r2 (l x')" using \((\\<^bsub>L2\<^esub>) \\<^sub>h (\\<^bsub>R2\<^esub>)) l2 r2\ by auto next assume "in_dom (\\<^bsub>L2\<^esub>) (l1 x)" with refl_L2 \((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2\ show "l x \\<^bsub>R2\<^esub> l2 (l1 x)" by auto qed (insert assms, auto) qed end end \ No newline at end of file diff --git a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Order_Base.thy b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Order_Base.thy --- a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Order_Base.thy +++ b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Order_Base.thy @@ -1,237 +1,237 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Basic Order Properties\ theory Transport_Compositions_Generic_Order_Base imports Transport_Compositions_Generic_Base begin context transport_comp begin interpretation flip1 : galois R1 L1 r1 l1 . subsubsection \Reflexivity\ lemma reflexive_on_in_dom_leftI: assumes galois_prop: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and in_dom_L1_le: "in_dom (\\<^bsub>L1\<^esub>) \ in_codom (\\<^bsub>L1\<^esub>)" and refl_R1: "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and refl_L2: "reflexive_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and mono_in_dom_l1: "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" shows "reflexive_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" proof (rule reflexive_onI) fix x assume "in_dom (\\<^bsub>L\<^esub>) x" then obtain x' where "x \\<^bsub>L\<^esub> x'" "in_dom (\\<^bsub>L1\<^esub>) x" by blast show "x \\<^bsub>L\<^esub> x" proof (rule left_relI) from refl_R1 have "l1 x \\<^bsub>R1\<^esub> l1 x" proof (rule reflexive_onD) from \x \\<^bsub>L\<^esub> x'\ galois_prop show "in_dom (\\<^bsub>R1\<^esub>) (l1 x)" by blast qed then show "x \<^bsub>L1\<^esub>\ l1 x" proof (intro t1.left_GaloisI) from galois_prop \in_dom (\\<^bsub>L1\<^esub>) x\ \l1 x \\<^bsub>R1\<^esub> l1 x\ show "x \\<^bsub>L1\<^esub> r1 (l1 x)" by blast qed blast from refl_L2 show "l1 x \\<^bsub>L2\<^esub> l1 x" proof (rule reflexive_onD) from mono_in_dom_l1 \x \\<^bsub>L\<^esub> x'\ show "in_dom (\\<^bsub>L2\<^esub>) (l1 x)" by blast qed from \l1 x \\<^bsub>R1\<^esub> l1 x\ show "l1 x \<^bsub>R1\<^esub>\ x" proof (intro flip1.left_GaloisI) from \in_dom (\\<^bsub>L1\<^esub>) x\ in_dom_L1_le show "in_codom (\\<^bsub>L1\<^esub>) x" by blast qed qed qed lemma reflexive_on_in_codom_leftI: assumes L1_r1_l1I: "\x. in_dom (\\<^bsub>L1\<^esub>) x \ l1 x \\<^bsub>R1\<^esub> l1 x \ x \\<^bsub>L1\<^esub> r1 (l1 x)" and in_codom_L1_le: "in_codom (\\<^bsub>L1\<^esub>) \ in_dom (\\<^bsub>L1\<^esub>)" and refl_R1: "reflexive_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and refl_L2: "reflexive_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and mono_in_codom_l1: "([in_codom (\\<^bsub>L\<^esub>)] \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" shows "reflexive_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" proof (rule reflexive_onI) fix x assume "in_codom (\\<^bsub>L\<^esub>) x" then obtain x' where "x' \\<^bsub>L\<^esub> x" "in_codom (\\<^bsub>L1\<^esub>) x" "in_codom (\\<^bsub>R1\<^esub>) (l1 x)" by blast show "x \\<^bsub>L\<^esub> x" proof (rule left_relI) from refl_R1 \in_codom (\\<^bsub>R1\<^esub>) (l1 x)\ have "l1 x \\<^bsub>R1\<^esub> l1 x" by blast show "x \<^bsub>L1\<^esub>\ l1 x" proof (rule t1.left_GaloisI) from in_codom_L1_le \in_codom (\\<^bsub>L1\<^esub>) x\ have "in_dom (\\<^bsub>L1\<^esub>) x" by blast with \l1 x \\<^bsub>R1\<^esub> l1 x\ show "x \\<^bsub>L1\<^esub> r1 (l1 x)" by (intro L1_r1_l1I) qed fact from refl_L2 show "l1 x \\<^bsub>L2\<^esub> l1 x" proof (rule reflexive_onD) from mono_in_codom_l1 \x' \\<^bsub>L\<^esub> x\ show "in_codom (\\<^bsub>L2\<^esub>) (l1 x)" by blast qed show "l1 x \<^bsub>R1\<^esub>\ x" by (rule flip1.left_GaloisI) fact+ qed qed corollary reflexive_on_in_field_leftI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "in_codom (\\<^bsub>L1\<^esub>) = in_dom (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "([in_field (\\<^bsub>L\<^esub>)] \\<^sub>m in_field (\\<^bsub>L2\<^esub>)) l1" shows "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" proof - from assms have "reflexive_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" by (intro reflexive_on_in_dom_leftI) - (auto 0 4 intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) + (auto 0 5 intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) moreover from assms have "reflexive_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" by (intro reflexive_on_in_codom_leftI) - (auto 0 4 intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom) + (auto 0 5 intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom) ultimately show ?thesis by (auto iff: in_field_iff_in_dom_or_in_codom) qed subsubsection \Transitivity\ text\There are many similar proofs for transitivity. They slightly differ in their assumptions, particularly which of @{term "(\\<^bsub>R1\<^esub>)"} and @{term "(\\<^bsub>L2\<^esub>)"} has to be transitive and the order of commutativity for the relations. In the following, we just give two of them that suffice for many purposes.\ lemma transitive_leftI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and trans_L2: "transitive (\\<^bsub>L2\<^esub>)" and R1_L2_R1_le: "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" shows "transitive (\\<^bsub>L\<^esub>)" proof (rule transitiveI) fix x1 x2 x3 assume "x1 \\<^bsub>L\<^esub> x2" "x2 \\<^bsub>L\<^esub> x3" from \x1 \\<^bsub>L\<^esub> x2\ obtain y1 y2 where "x1 \<^bsub>L1\<^esub>\ y1" "y1 \\<^bsub>L2\<^esub> y2" "y2 \\<^bsub>R1\<^esub> l1 x2" by blast from \x2 \\<^bsub>L\<^esub> x3\ \((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1\ obtain y3 y4 where "l1 x2 \\<^bsub>R1\<^esub> y3" "y3 \\<^bsub>L2\<^esub> y4" "y4 \\<^bsub>R1\<^esub> l1 x3" "in_codom (\\<^bsub>L1\<^esub>) x3" by blast with R1_L2_R1_le have "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) (l1 x2) (l1 x3)" by blast then obtain y where "l1 x2 \\<^bsub>L2\<^esub> y" "y \\<^bsub>R1\<^esub> l1 x3" by blast with \y2 \\<^bsub>R1\<^esub> l1 x2\ R1_L2_R1_le have "((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) y2 (l1 x3)" by blast then obtain y' where "y2 \\<^bsub>L2\<^esub> y'" "y' \\<^bsub>R1\<^esub> l1 x3" by blast with \y1 \\<^bsub>L2\<^esub> y2\ have "y1 \\<^bsub>L2\<^esub> y'" using trans_L2 by blast show "x1 \\<^bsub>L\<^esub> x3" proof (rule left_relI) show "x1 \<^bsub>L1\<^esub>\ y1" "y1 \\<^bsub>L2\<^esub> y'" by fact+ show "y' \<^bsub>R1\<^esub>\ x3" by (rule flip1.left_GaloisI) fact+ qed qed lemma transitive_leftI': assumes galois_prop: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and trans_L2: "transitive (\\<^bsub>L2\<^esub>)" and R1_L2_R1_le: "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" shows "transitive (\\<^bsub>L\<^esub>)" proof (rule transitiveI) fix x1 x2 x3 assume "x1 \\<^bsub>L\<^esub> x2" "x2 \\<^bsub>L\<^esub> x3" from \x1 \\<^bsub>L\<^esub> x2\ galois_prop obtain y1 y2 where "in_dom (\\<^bsub>L1\<^esub>) x1" "l1 x1 \\<^bsub>R1\<^esub> y1" "y1 \\<^bsub>L2\<^esub> y2" "y2 \\<^bsub>R1\<^esub> l1 x2" by blast with R1_L2_R1_le have "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) (l1 x1) (l1 x2)" by blast then obtain y where "l1 x1 \\<^bsub>R1\<^esub> y" "y \\<^bsub>L2\<^esub> l1 x2" by blast moreover from \x2 \\<^bsub>L\<^esub> x3\ galois_prop obtain y3 y4 where "l1 x2 \\<^bsub>R1\<^esub> y3" "y3 \\<^bsub>L2\<^esub> y4" "y4 \<^bsub>R1\<^esub>\ x3" by blast moreover note R1_L2_R1_le ultimately have "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) (l1 x1) y3" by blast then obtain y' where "l1 x1 \\<^bsub>R1\<^esub> y'" "y' \\<^bsub>L2\<^esub> y3" by blast with \y3 \\<^bsub>L2\<^esub> y4\ have "y' \\<^bsub>L2\<^esub> y4" using trans_L2 by blast show "x1 \\<^bsub>L\<^esub> x3" proof (rule left_relI) from \in_dom (\\<^bsub>L1\<^esub>) x1\ \l1 x1 \\<^bsub>R1\<^esub> y'\ galois_prop show "x1 \<^bsub>L1\<^esub>\ y'" by (intro t1.left_Galois_if_Galois_right_if_half_galois_prop_right t1.left_GaloisI) auto show "y' \\<^bsub>L2\<^esub> y4" by fact from \y' \\<^bsub>L2\<^esub> y4\ \y4 \<^bsub>R1\<^esub>\ x3\ show "y4 \<^bsub>R1\<^esub>\ x3" by blast qed qed subsubsection \Preorders\ lemma preorder_on_in_field_leftI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "in_codom (\\<^bsub>L1\<^esub>) = in_dom (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "preorder_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and mono_in_codom_l1: "([in_codom (\\<^bsub>L\<^esub>)] \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" and R1_L2_R1_le: "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" shows "preorder_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" proof - have "([in_field (\\<^bsub>L\<^esub>)] \\<^sub>m in_field (\\<^bsub>L2\<^esub>)) l1" proof - from \((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1\ R1_L2_R1_le have "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" by (intro mono_in_dom_left_rel_left1_if_in_dom_rel_comp_le in_dom_right1_left2_right1_le_if_right1_left2_right1_le) auto with mono_in_codom_l1 show ?thesis by (intro dep_mono_wrt_predI) blast qed with assms show ?thesis by (intro preorder_onI) (auto intro: reflexive_on_in_field_leftI transitive_leftI) qed lemma preorder_on_in_field_leftI': assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "in_codom (\\<^bsub>L1\<^esub>) = in_dom (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "preorder_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and mono_in_dom_l1: "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" and R1_L2_R1_le: "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" shows "preorder_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" proof - have "([in_field (\\<^bsub>L\<^esub>)] \\<^sub>m in_field (\\<^bsub>L2\<^esub>)) l1" proof - from \((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1\ R1_L2_R1_le have "([in_codom (\\<^bsub>L\<^esub>)] \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" by (intro mono_in_codom_left_rel_left1_if_in_codom_rel_comp_le in_codom_right1_left2_right1_le_if_right1_left2_right1_le) auto with mono_in_dom_l1 show ?thesis by (intro dep_mono_wrt_predI) blast qed with assms show ?thesis by (intro preorder_onI) (auto intro: reflexive_on_in_field_leftI transitive_leftI') qed subsubsection \Symmetry\ lemma symmetric_leftI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "in_codom (\\<^bsub>L1\<^esub>) = in_dom (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>R1\<^esub>)" and "symmetric (\\<^bsub>L2\<^esub>)" shows "symmetric (\\<^bsub>L\<^esub>)" proof - from assms have "(\\<^bsub>R1\<^esub>) = (\<^bsub>L1\<^esub>\)" by (intro t1.ge_Galois_right_eq_left_Galois_if_symmetric_if_in_codom_eq_in_dom_if_galois_prop) moreover then have "(\<^bsub>R1\<^esub>\) = (\\<^bsub>L1\<^esub>)" by (subst rel_inv_eq_iff_eq[symmetric]) simp ultimately show ?thesis using assms unfolding left_rel_eq_comp by (subst symmetric_iff_rel_inv_eq_self) (simp add: rel_comp_assoc) qed lemma partial_equivalence_rel_leftI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "in_codom (\\<^bsub>L1\<^esub>) = in_dom (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>R1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>))" shows "partial_equivalence_rel (\\<^bsub>L\<^esub>)" using assms by (intro partial_equivalence_relI transitive_leftI symmetric_leftI) auto lemma partial_equivalence_rel_leftI': assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "in_codom (\\<^bsub>L1\<^esub>) = in_dom (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>R1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" shows "partial_equivalence_rel (\\<^bsub>L\<^esub>)" using assms by (intro partial_equivalence_relI transitive_leftI' symmetric_leftI) auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Examples/Prototype/Transport_Prototype.thy b/thys/Transport/Transport/Examples/Prototype/Transport_Prototype.thy --- a/thys/Transport/Transport/Examples/Prototype/Transport_Prototype.thy +++ b/thys/Transport/Transport/Examples/Prototype/Transport_Prototype.thy @@ -1,168 +1,168 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Transport via Equivalences on PERs (Prototype)\ theory Transport_Prototype imports Transport_Rel_If ML_Unification.ML_Unification_HOL_Setup ML_Unification.Unify_Resolve_Tactics keywords "trp_term" :: thy_goal_defn begin paragraph \Summary\ text \We implement a simple Transport prototype. The prototype is restricted to work with equivalences on partial equivalence relations. It is also not forming the compositions of equivalences so far. The support for dependent function relators is restricted to the form described in @{thm transport_Dep_Fun_Rel_no_dep_fun.partial_equivalence_rel_equivalenceI}: The relations can be dependent, but the functions must be simple. This is not production ready, but a proof of concept. The package provides a command @{command trp_term}, which sets up the required goals to prove a given term. See the examples in this directory for some use cases and refer to \<^cite>\"transport"\ for more details.\ paragraph \Theorem Setups\ context transport begin lemma left_Galois_left_if_left_rel_if_partial_equivalence_rel_equivalence: assumes "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" and "x \\<^bsub>L\<^esub> x'" shows "x \<^bsub>L\<^esub>\ l x" using assms by (intro left_Galois_left_if_left_rel_if_inflationary_on_in_fieldI) (blast elim: preorder_equivalence_order_equivalenceE)+ definition "transport_per x y \ ((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r \ x \<^bsub>L\<^esub>\ y" text \The choice of @{term "x'"} is arbitrary. All we need is @{term "in_dom (\\<^bsub>L\<^esub>) x"}.\ lemma transport_per_start: assumes "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" and "x \\<^bsub>L\<^esub> x'" shows "transport_per x (l x)" using assms unfolding transport_per_def by (blast intro: left_Galois_left_if_left_rel_if_partial_equivalence_rel_equivalence) lemma left_Galois_if_transport_per: assumes "transport_per x y" shows "x \<^bsub>L\<^esub>\ y" using assms unfolding transport_per_def by blast end context transport_Fun_Rel begin text \Simplification of Galois relator for simple function relator.\ corollary left_Galois_eq_Fun_Rel_left_Galois: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2\<^esub>)) l2 r2" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))" proof (intro ext) fix f g show "f \<^bsub>L\<^esub>\ g \ ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" proof assume "f \<^bsub>L\<^esub>\ g" moreover have "g \\<^bsub>R\<^esub> g" proof - from assms have per: "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" by (intro partial_equivalence_rel_equivalenceI) auto with \f \<^bsub>L\<^esub>\ g\ show ?thesis by blast qed ultimately show "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" using assms by (intro Fun_Rel_left_Galois_if_left_GaloisI) (auto elim!: tdfrs.t1.partial_equivalence_rel_equivalenceE tdfrs.t1.preorder_equivalence_galois_equivalenceE tdfrs.t1.galois_equivalenceE intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) next assume "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" with assms have "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> f g" by (subst Fun_Rel_left_Galois_restrict_left_right_eq_Fun_Rel_left_GaloisI) blast+ with assms show "f \<^bsub>L\<^esub>\ g" by (intro left_Galois_if_Fun_Rel_left_GaloisI) blast+ qed qed end lemmas related_Fun_Rel_combI = Dep_Fun_Rel_relD[where ?S="\_ _. S" for S, rotated] lemma related_Fun_Rel_lambdaI: assumes "\x y. R x y \ S (f x) (g y)" and "T = (R \ S)" shows "T f g" using assms by blast paragraph \General ML setups\ ML_file\transport_util.ML\ paragraph \Unification Setup\ ML\ @{functor_instance struct_name = Transport_Unification_Combine and functor_name = Unification_Combine and id = Transport_Util.transport_id} \ local_setup \Transport_Unification_Combine.setup_attribute NONE\ ML\ @{functor_instance struct_name = Transport_Mixed_Unification and functor_name = Mixed_Unification and id = Transport_Util.transport_id and more_args = \structure UC = Transport_Unification_Combine\} \ ML\ @{functor_instance struct_name = Transport_Unification_Hints and functor_name = Term_Index_Unification_Hints and id = Transport_Util.transport_id and more_args = \ structure TI = Discrimination_Tree val init_args = { concl_unifier = SOME Higher_Order_Pattern_Unification.unify, - normalisers = SOME Transport_Mixed_Unification.norms_first_higherp_first_comb_higher_unify, prems_unifier = SOME (Transport_Mixed_Unification.first_higherp_first_comb_higher_unify |> Unification_Combinator.norm_unifier Envir_Normalisation.beta_norm_term_unif), + normalisers = SOME Transport_Mixed_Unification.norms_first_higherp_first_comb_higher_unify, retrieval = SOME (Term_Index_Unification_Hints_Args.mk_sym_retrieval TI.norm_term TI.unifiables), hint_preprocessor = SOME (K I) }\} \ local_setup \Transport_Unification_Hints.setup_attribute NONE\ declare [[trp_uhint where hint_preprocessor = \Unification_Hints_Base.obj_logic_hint_preprocessor @{thm atomize_eq[symmetric]} (Conv.rewr_conv @{thm eq_eq_True})\]] declare [[trp_ucombine add = \Transport_Unification_Combine.eunif_data (Transport_Unification_Hints.try_hints |> Unification_Combinator.norm_unifier (#norm_term Transport_Mixed_Unification.norms_first_higherp_first_comb_higher_unify) |> K) (Transport_Unification_Combine.default_metadata Transport_Unification_Hints.binding)\]] paragraph \Prototype\ ML_file\transport.ML\ declare transport_Dep_Fun_Rel.transport_defs[trp_def] transport_Fun_Rel.transport_defs[trp_def] declare (*dependent case currently disabled by default since they easily make the unifier enumerate many undesired instantiations*) (* transport_Dep_Fun_Rel.partial_equivalence_rel_equivalenceI[per_intro] *) (* transport.rel_if_partial_equivalence_rel_equivalence_if_iff_if_partial_equivalence_rel_equivalenceI[rotated, per_intro] transport_Dep_Fun_Rel_no_dep_fun.partial_equivalence_rel_equivalenceI [ML_Krattr \Conversion_Util.move_prems_to_front_conv [1] |> Conversion_Util.thm_conv\, ML_Krattr \Conversion_Util.move_prems_to_front_conv [2,3] |> Conversion_Util.thm_conv\, per_intro] *) transport_Fun_Rel.partial_equivalence_rel_equivalenceI[rotated, per_intro] transport_eq_id.partial_equivalence_rel_equivalenceI[per_intro] transport_eq_restrict_id.partial_equivalence_rel_equivalence[per_intro] declare transport_id.left_Galois_eq_left[trp_relator_rewrite] transport_Fun_Rel.left_Galois_eq_Fun_Rel_left_Galois[trp_relator_rewrite] end diff --git a/thys/Transport/Transport/Examples/Prototype/Transport_Rel_If.thy b/thys/Transport/Transport/Examples/Prototype/Transport_Rel_If.thy --- a/thys/Transport/Transport/Examples/Prototype/Transport_Rel_If.thy +++ b/thys/Transport/Transport/Examples/Prototype/Transport_Rel_If.thy @@ -1,218 +1,218 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Transport for Dependent Function Relator with Non-Dependent Functions\ theory Transport_Rel_If imports Transport begin paragraph \Summary\ text \We introduce a special case of @{locale transport_Dep_Fun_Rel}. The derived theorem is easier to apply and supported by the current prototype.\ context fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" begin lemma reflexive_on_rel_if_if_reflexive_onI [intro]: assumes "B \ reflexive_on P R" shows "reflexive_on P (rel_if B R)" using assms by (intro reflexive_onI) blast lemma transitive_on_rel_if_if_transitive_onI [intro]: assumes "B \ transitive_on P R" shows "transitive_on P (rel_if B R)" using assms by (intro transitive_onI) (blast dest: transitive_onD) lemma preorder_on_rel_if_if_preorder_onI [intro]: assumes "B \ preorder_on P R" shows "preorder_on P (rel_if B R)" using assms by (intro preorder_onI) auto lemma symmetric_on_rel_if_if_symmetric_onI [intro]: assumes "B \ symmetric_on P R" shows "symmetric_on P (rel_if B R)" using assms by (intro symmetric_onI) (blast dest: symmetric_onD) lemma partial_equivalence_rel_on_rel_if_if_partial_equivalence_rel_onI [intro]: assumes "B \ partial_equivalence_rel_on P R" shows "partial_equivalence_rel_on P (rel_if B R)" using assms by (intro partial_equivalence_rel_onI) auto lemma rel_if_dep_mono_wrt_rel_if_iff_if_dep_mono_wrt_relI: assumes "B \ B' \ ([x y \ R] \\<^sub>m S x y) f" and "B \ B'" shows "([x y \ (rel_if B R)] \\<^sub>m (rel_if B' (S x y))) f" using assms by (intro dep_mono_wrt_relI) auto end corollary reflexive_rel_if_if_reflexiveI [intro]: assumes "B \ reflexive R" shows "reflexive (rel_if B R)" using assms unfolding reflexive_eq_reflexive_on by blast corollary transitive_rel_if_if_transitiveI [intro]: assumes "B \ transitive R" shows "transitive (rel_if B R)" using assms unfolding transitive_eq_transitive_on by blast corollary preorder_rel_if_if_preorderI [intro]: assumes "B \ preorder R" shows "preorder (rel_if B R)" using assms unfolding preorder_eq_preorder_on by blast corollary symmetric_rel_if_if_symmetricI [intro]: assumes "B \ symmetric R" shows "symmetric (rel_if B R)" using assms unfolding symmetric_eq_symmetric_on by blast corollary partial_equivalence_rel_rel_if_if_partial_equivalence_relI [intro]: assumes "B \ partial_equivalence_rel R" shows "partial_equivalence_rel (rel_if B R)" using assms unfolding partial_equivalence_rel_eq_partial_equivalence_rel_on by blast context galois_prop begin interpretation rel_if : galois_prop "rel_if B (\\<^bsub>L\<^esub>)" "rel_if B' (\\<^bsub>R\<^esub>)" l r . interpretation flip_inv : galois_prop "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" r l . lemma rel_if_half_galois_prop_left_if_iff_if_half_galois_prop_leftI: assumes "B \ B' \ ((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" and "B \ B'" shows "((rel_if B (\\<^bsub>L\<^esub>)) \<^sub>h\ (rel_if B' (\\<^bsub>R\<^esub>))) l r" using assms by (intro rel_if.half_galois_prop_leftI) auto lemma rel_if_half_galois_prop_right_if_iff_if_half_galois_prop_rightI: assumes "B \ B' \ ((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" and "B \ B'" shows "((rel_if B (\\<^bsub>L\<^esub>)) \\<^sub>h (rel_if B' (\\<^bsub>R\<^esub>))) l r" using assms by (intro rel_if.half_galois_prop_rightI) fastforce lemma rel_if_galois_prop_if_iff_if_galois_propI: assumes "B \ B' \ ((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" and "B \ B'" shows "((rel_if B (\\<^bsub>L\<^esub>)) \ (rel_if B' (\\<^bsub>R\<^esub>))) l r" using assms by (intro rel_if.galois_propI rel_if_half_galois_prop_left_if_iff_if_half_galois_prop_leftI rel_if_half_galois_prop_right_if_iff_if_half_galois_prop_rightI) auto end context galois begin interpretation rel_if : galois "rel_if B (\\<^bsub>L\<^esub>)" "rel_if B' (\\<^bsub>R\<^esub>)" l r . lemma rel_if_galois_connection_if_iff_if_galois_connectionI: assumes "B \ B' \ ((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" and "B \ B'" shows "((rel_if B (\\<^bsub>L\<^esub>)) \ (rel_if B' (\\<^bsub>R\<^esub>))) l r" using assms by (intro rel_if.galois_connectionI rel_if_dep_mono_wrt_rel_if_iff_if_dep_mono_wrt_relI rel_if_galois_prop_if_iff_if_galois_propI) auto lemma rel_if_galois_equivalence_if_iff_if_galois_equivalenceI: assumes "B \ B' \ ((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" and "B \ B'" shows "((rel_if B (\\<^bsub>L\<^esub>)) \\<^sub>G (rel_if B' (\\<^bsub>R\<^esub>))) l r" using assms by (intro rel_if.galois_equivalenceI rel_if_galois_connection_if_iff_if_galois_connectionI galois_prop.rel_if_galois_prop_if_iff_if_galois_propI) (auto elim: galois.galois_connectionE) end context transport begin interpretation rel_if : transport "rel_if B (\\<^bsub>L\<^esub>)" "rel_if B' (\\<^bsub>R\<^esub>)" l r . lemma rel_if_preorder_equivalence_if_iff_if_preorder_equivalenceI: assumes "B \ B' \ ((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R\<^esub>)) l r" and "B \ B'" shows "((rel_if B (\\<^bsub>L\<^esub>)) \\<^bsub>pre\<^esub> (rel_if B' (\\<^bsub>R\<^esub>))) l r" using assms by (intro rel_if.preorder_equivalence_if_galois_equivalenceI rel_if_galois_equivalence_if_iff_if_galois_equivalenceI preorder_on_rel_if_if_preorder_onI) blast+ lemma rel_if_partial_equivalence_rel_equivalence_if_iff_if_partial_equivalence_rel_equivalenceI: assumes "B \ B' \ ((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" and "B \ B'" shows "((rel_if B (\\<^bsub>L\<^esub>)) \\<^bsub>PER\<^esub> (rel_if B' (\\<^bsub>R\<^esub>))) l r" using assms by (intro rel_if.partial_equivalence_rel_equivalence_if_galois_equivalenceI rel_if_galois_equivalence_if_iff_if_galois_equivalenceI) blast+ end locale transport_Dep_Fun_Rel_no_dep_fun = transport_Dep_Fun_Rel_syntax L1 R1 l1 r1 L2 R2 "\_ _. l2" "\_ _. r2" + tdfr : transport_Dep_Fun_Rel L1 R1 l1 r1 L2 R2 "\_ _. l2" "\_ _. r2" for L1 :: "'a1 \ 'a1 \ bool" and R1 :: "'a2 \ 'a2 \ bool" and l1 :: "'a1 \ 'a2" and r1 :: "'a2 \ 'a1" and L2 :: "'a1 \ 'a1 \ 'b1 \ 'b1 \ bool" and R2 :: "'a2 \ 'a2 \ 'b2 \ 'b2 \ bool" and l2 :: "'b1 \ 'b2" and r2 :: "'b2 \ 'b1" begin notation t2.unit ("\\<^sub>2") notation t2.counit ("\\<^sub>2") abbreviation "L \ tdfr.L" abbreviation "R \ tdfr.R" abbreviation "l \ tdfr.l" abbreviation "r \ tdfr.r" notation tdfr.L (infix "\\<^bsub>L\<^esub>" 50) notation tdfr.R (infix "\\<^bsub>R\<^esub>" 50) notation tdfr.ge_left (infix "\\<^bsub>L\<^esub>" 50) notation tdfr.ge_right (infix "\\<^bsub>R\<^esub>" 50) notation tdfr.unit ("\") notation tdfr.counit ("\") theorem partial_equivalence_rel_equivalenceI: assumes per_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and per_equiv2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) l2 r2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" shows "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" proof - have per2I: "((\\<^bsub>L2 x1 (r1 x2')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2 r2" if hyps: "x1 \\<^bsub>L1\<^esub> x2" "x2 \<^bsub>L1\<^esub>\ x1'" "x1' \\<^bsub>R1\<^esub> x2'" for x1 x2 x1' x2' proof - from hyps have "x1 \<^bsub>L1\<^esub>\ x2'" using per_equiv1 t1.left_Galois_if_left_Galois_if_left_relI t1.left_Galois_if_right_rel_if_left_GaloisI by fast with per_equiv2 show ?thesis by blast qed have "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) (\_ _. l2)" by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI Dep_Fun_Rel_predI rel_if_if_impI) - (auto 8 0 dest!: per2I) + (auto 10 0 dest!: per2I) moreover have "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) (\_ _. r2)" by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI Dep_Fun_Rel_predI rel_if_if_impI) - (auto 8 0 dest!: per2I) + (auto 10 0 dest!: per2I) ultimately show ?thesis using assms by (intro tdfr.partial_equivalence_rel_equivalenceI) auto qed end end \ No newline at end of file diff --git a/thys/Transport/Transport/Examples/Prototype/transport.ML b/thys/Transport/Transport/Examples/Prototype/transport.ML --- a/thys/Transport/Transport/Examples/Prototype/transport.ML +++ b/thys/Transport/Transport/Examples/Prototype/transport.ML @@ -1,409 +1,409 @@ (* Title: Transport/transport.ML Author: Kevin Kappelmann, Paul Bachmann Prototype for Transport. See README.md for future work. *) -(*TODO: signature*) +(*TODO: signature for final implementation*) structure Transport = struct structure Util = Transport_Util (*definitions used by Transport that need to be folded before a PER proof and unfolded after success.*) structure Transport_Defs = Named_Thms( val name = @{binding "trp_def"} val description = "Definitions used by Transport" ) val _ = Theory.setup Transport_Defs.setup (* simplifying definitions *) val simp_rhs = Simplifier.rewrite #> Conversion_Util.rhs_conv #> Conversion_Util.thm_conv (*simplifies the generated definition of a transported term*) fun simp_transported_def ctxt simps y_def = let val ctxt = ctxt addsimps simps val y_def_eta_expanded = Util.equality_eta_expand ctxt "x" y_def in apply2 (simp_rhs ctxt) (y_def, y_def_eta_expanded) end (* resolution setup *) val any_unify_trp_hints_resolve_tac = Unify_Resolve_Base.unify_resolve_any_tac Transport_Mixed_Unification.norms_first_higherp_first_comb_higher_unify Transport_Mixed_Unification.first_higherp_first_comb_higher_unify fun get_theorems_tac f get_theorems ctxt = f (get_theorems ctxt) ctxt val get_theorems_resolve_tac = get_theorems_tac any_unify_trp_hints_resolve_tac val _ = Theory.setup ( Method.setup @{binding trp_hints_resolve} (Attrib.thms >> (SIMPLE_METHOD' oo any_unify_trp_hints_resolve_tac)) "Resolution with unification hints for Transport" ) (* PER equivalence prover *) (*introduction rules*) structure PER_Intros = Named_Thms( val name = @{binding "per_intro"} val description = "Introduction rules for per_prover" ) val _ = Theory.setup PER_Intros.setup fun per_prover_tac ctxt = REPEAT_ALL_NEW (get_theorems_resolve_tac PER_Intros.get ctxt) val _ = Theory.setup ( Method.setup @{binding per_prover} (Scan.succeed (SIMPLE_METHOD' o per_prover_tac)) "PER prover for Transport" ) (* domain prover *) structure Transport_in_dom = Named_Thms( val name = @{binding "trp_in_dom"} val description = "In domain theorems for Transport" ) val _ = Theory.setup Transport_in_dom.setup (*discharges the @{term "L x x'"} goals by registered lemmas*) fun transport_in_dom_prover_tac ctxt = get_theorems_resolve_tac Transport_in_dom.get ctxt val _ = Theory.setup ( Method.setup @{binding trp_in_dom_prover} (Scan.succeed (SIMPLE_METHOD' o transport_in_dom_prover_tac)) "in_dom prover for Transport" ) (* blackbox prover *) (*first derives the PER equivalence, then looks for registered domain lemmas.*) fun unfold_tac thms ctxt = simp_tac (clear_simpset ctxt addsimps thms) val unfold_transport_defs_tac = get_theorems_tac unfold_tac Transport_Defs.get fun transport_prover ctxt i = per_prover_tac ctxt i THEN TRY (SOMEGOAL (TRY o unfold_transport_defs_tac ctxt THEN' transport_in_dom_prover_tac ctxt) ) val _ = Theory.setup ( Method.setup @{binding trp_prover} (Scan.succeed (SIMPLE_METHOD' o transport_prover)) "Blackbox prover for Transport" ) (* whitebox prover intro rules *) structure Transport_Related_Intros = Named_Thms( val name = @{binding "trp_related_intro"} val description = "Introduction rules for Transport whitebox proofs" ) val _ = Theory.setup Transport_Related_Intros.setup (* relator rewriter *) (*rewrite rules to simplify the derived Galois relator*) structure Transport_Relator_Rewrites = Named_Thms( val name = @{binding "trp_relator_rewrite"} val description = "Rewrite rules for relators used by Transport" ) val _ = Theory.setup Transport_Relator_Rewrites.setup (*simple rewrite tactic for Galois relators*) fun per_simp_prover ctxt thm = let val prems = Thm.cprems_of thm val per_prover_tac = per_prover_tac ctxt fun prove_prem prem = Goal.prove_internal ctxt [] prem (fn _ => HEADGOAL per_prover_tac) in try (map prove_prem) prems |> Option.map (curry (op OF) thm) end fun transport_relator_rewrite ctxt thm = let val transport_defs = Transport_Defs.get ctxt val transport_relator_rewrites = Transport_Relator_Rewrites.get ctxt val ctxt = (clear_simpset ctxt) addsimps transport_relator_rewrites in Local_Defs.fold ctxt transport_defs thm |> Raw_Simplifier.rewrite_thm (false, false, false) per_simp_prover ctxt end fun transport_relator_rewrite_tac ctxt = EqSubst.eqsubst_tac ctxt [0] (Transport_Relator_Rewrites.get ctxt) THEN_ALL_NEW TRY o SOLVED' (per_prover_tac ctxt) val _ = Theory.setup ( Method.setup @{binding trp_relator_rewrite} (Scan.succeed (SIMPLE_METHOD' o transport_relator_rewrite_tac)) "Rewrite Transport relator" ) (* term transport command *) (*parsing*) @{parse_entries (struct) PA [L, R, x, y]} val parse_cmd_entries = let val parse_value = PA.parse_entry Parse.term Parse.term Parse.term Parse.term val parse_entry = Parse_Key_Value.parse_entry PA.parse_key Parse_Util.eq parse_value in PA.parse_entries_required Parse.and_list1 [PA.key PA.x] parse_entry (PA.empty_entries ()) end (*some utilities to destruct terms*) val transport_per_start_thm = @{thm "transport.transport_per_start"} val related_if_transport_per_thm = @{thm "transport.left_Galois_if_transport_per"} fun dest_transport_per \<^Const_>\transport.transport_per S T for L R l r x y\ = ((S, T), (L, R, l, r, x, y)) val dest_transport_per_y = dest_transport_per #> (fn (_, (_, _, _, _, _, y)) => y) fun mk_hom_Galois Ta Tb L R r x y = \<^Const>\galois_rel.Galois Ta Ta Tb Tb for L R r x y\ fun dest_hom_Galois \<^Const_>\galois_rel.Galois Ta _ Tb _ for L R r x y\ = ((Ta, Tb), (L, R, r, x, y)) val dest_hom_Galois_y = dest_hom_Galois #> (fn (_, (_, _, _, _, y)) => y) (*bindings for generated theorems and definitions*) val binding_transport_per = \<^binding>\transport_per\ val binding_per = \<^binding>\per\ val binding_in_dom = \<^binding>\in_dom\ val binding_related = \<^binding>\related\ val binding_related_rewritten = \<^binding>\related'\ val binding_def_simplified = \<^binding>\eq\ val binding_def_eta_expanded_simplified = \<^binding>\app_eq\ fun note_facts (binding, mixfix) ctxt related_thm y binding_thms_attribs = let val ((_, (_, y_def)), ctxt) = Util.create_local_theory_def (binding, mixfix) [] y ctxt (*create simplified definition theorems*) val transport_defs = Transport_Defs.get ctxt val (y_def_simplified, y_def_eta_expanded_simplified) = simp_transported_def ctxt transport_defs y_def (*create relatedness theorems*) val related_thm_rewritten = transport_relator_rewrite ctxt related_thm fun prepare_fact (suffix, thm, attribs) = let val binding = (Util.add_suffix binding suffix, []) val ctxt = (clear_simpset ctxt) addsimps transport_defs val folded_thm = (*fold definition of transported term*) Local_Defs.fold ctxt [y_def] thm (*simplify other transport definitions in theorem*) |> (Simplifier.rewrite ctxt |> Conversion_Util.thm_conv) val thm_attribs = ([folded_thm], attribs) in (binding, [thm_attribs]) end val facts = map prepare_fact ([ (binding_related, related_thm, []), (binding_related_rewritten, related_thm_rewritten, [Util.attrib_to_src (Binding.pos_of binding) Transport_Related_Intros.add]), (binding_def_simplified, y_def_simplified, []), (binding_def_eta_expanded_simplified, y_def_eta_expanded_simplified, []) ] @ binding_thms_attribs) in Local_Theory.notes facts ctxt |> snd end (*black-box transport as described in the Transport paper*) fun after_qed_blackbox (binding, mixfix) [thms as [per_thm, in_dom_thm]] ctxt = let val transport_per_thm = List.foldl (op INCR_COMP) transport_per_start_thm thms (*fix possibly occurring meta type variables*) val ((_, [transport_per_thm]), ctxt) = Variable.importT [transport_per_thm] ctxt val y = Util.real_thm_concl transport_per_thm |> dest_transport_per_y val related_thm = transport_per_thm INCR_COMP related_if_transport_per_thm val binding_thms = [ (binding_transport_per, transport_per_thm, []), (binding_per, per_thm, []), (binding_in_dom, in_dom_thm, [Util.attrib_to_src (Binding.pos_of binding) Transport_in_dom.add]) ] in note_facts (binding, mixfix) ctxt related_thm y binding_thms end (*experimental white-box transport support*) fun after_qed_whitebox (binding, mixfix) [[related_thm]] ctxt = let (*fix possibly occurring meta type variables*) val ((_, [related_thm]), ctxt) = Variable.importT [related_thm] ctxt val y = Util.real_thm_concl related_thm |> dest_hom_Galois_y in note_facts (binding, mixfix) ctxt related_thm y [] end fun setup_goals_blackbox ctxt (L, R, cx) maxidx = let (*check*) val [cL, cR] = Syntax.check_terms ctxt [L, R] |> map (Thm.cterm_of ctxt) (*instantiate theorem*) val transport_per_start_thm = Thm.incr_indexes (maxidx + 1) transport_per_start_thm val args = [SOME cL, SOME cR, NONE, NONE, SOME cx] val transport_per_start_thm = Drule.infer_instantiate' ctxt args transport_per_start_thm val transport_defs = Transport_Defs.get ctxt val goals = Local_Defs.fold ctxt transport_defs transport_per_start_thm |> Thm.prems_of |> map (rpair []) in goals end fun setup_goals_whitebox ctxt (yT, L, R, cx, y) maxidx = let val (r, _) = Term_Util.fresh_var "r" dummyT maxidx (*check*) val Galois = mk_hom_Galois (Thm.typ_of_cterm cx) yT L R r (Thm.term_of cx) y |> Syntax.check_term ctxt val goal = Util.mk_judgement Galois |> rpair [] in [goal] end fun setup_proof ((((binding, opt_yT, mixfix), params), unfolds), whitebox) lthy = let val ctxt = Util.set_proof_mode_schematic lthy (*type of transported term*) val yT = Option.map (Syntax.read_typ ctxt) opt_yT |> the_default dummyT (*theorems to unfold*) val unfolds = map (Proof_Context.get_fact ctxt o fst) unfolds |> flat (*term to transport*) val cx = (**read term**) Syntax.read_term ctxt (PA.get_x params) |> Thm.cterm_of ctxt (**unfold passed theorems**) |> Drule.cterm_rule (Local_Defs.unfold ctxt unfolds) (*transport relations and transport term goal*) val ([L, R, y], maxidx) = let (**configuration**) val opts = [PA.get_L_safe params, PA.get_R_safe params, PA.get_y_safe params] val opts_default_names = ["L", "R", "y"] val opts_constraints = [Util.mk_hom_rel_type (Thm.typ_of_cterm cx), Util.mk_hom_rel_type yT, yT] |> map Type.constraint (**parse**) val opts = map (Syntax.parse_term ctxt |> Option.map) opts val params_maxidx = Util.list_max (the_default ~1 o Option.map Term.maxidx_of_term) (Thm.maxidx_of_cterm cx) opts fun create_var (NONE, n) maxidx = Term_Util.fresh_var n dummyT params_maxidx ||> Integer.max maxidx | create_var (SOME t, _) created = (t, created) val (ts, maxidx) = fold_map create_var (opts ~~ opts_default_names) params_maxidx |>> map2 I opts_constraints in (ts, maxidx) end (*initialise goals and callback*) val (goals, after_qed) = if whitebox then (setup_goals_whitebox ctxt (yT, L, R, cx, y) maxidx, after_qed_whitebox) (*TODO: consider y in blackbox proofs*) else (setup_goals_blackbox ctxt (L, R, cx) maxidx, after_qed_blackbox) in Proof.theorem NONE (after_qed (binding, mixfix)) [goals] ctxt |> Proof.refine_singleton Util.split_conjunctions end val parse_strings = (*binding for transported term*) Parse_Spec.constdecl (*other params*) -- parse_cmd_entries (*optionally pass unfold theorems in case of white-box transports*) -- Scan.optional (Parse.reserved "unfold" |-- Parse.thms1) [] (*use a bang "!" to start white-box transport mode (experimental)*) -- Parse.opt_bang val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\trp_term\ "Transport term" (parse_strings >> setup_proof) (* experimental white-box prover *) val any_match_resolve_related_tac = let fun unif binders = Higher_Ordern_Pattern_First_Decomp_Unification.e_match Unification_Util.match_types unif Unification_Combinator.fail_unify binders in Unify_Resolve_Base.unify_resolve_any_tac Higher_Ordern_Pattern_First_Decomp_Unification.norms_match unif end val related_comb_tac = any_match_resolve_related_tac @{thms related_Fun_Rel_combI} val related_lambda_tac = any_match_resolve_related_tac @{thms related_Fun_Rel_lambdaI} val related_tac = any_unify_trp_hints_resolve_tac val related_assume_tac = assume_tac fun mk_transport_related_tac cc_comb cc_lambda ctxt = let val transport_related_intros = Transport_Related_Intros.get ctxt val related_tac = related_tac transport_related_intros ctxt val comb_tac = related_comb_tac ctxt val lambda_tac = related_lambda_tac ctxt val assume_tac = related_assume_tac ctxt in Tactic_Util.CONCAT' [ related_tac, cc_comb comb_tac, cc_lambda lambda_tac, assume_tac ] end val transport_related_step_tac = let fun cc_comb tac i = tac i THEN prefer_tac i THEN prefer_tac (i + 1) in mk_transport_related_tac cc_comb I end fun transport_related_tac ctxt = let fun transport_related_tac cc = let fun cc_comb tac = tac THEN_ALL_NEW TRY o cc fun cc_lambda tac = tac THEN' TRY o cc in mk_transport_related_tac cc_comb cc_lambda ctxt end fun fix tac i thm = tac (fix tac) i thm in fix transport_related_tac end val _ = Theory.setup ( Method.setup @{binding trp_related_prover} (Scan.succeed (SIMPLE_METHOD' o transport_related_tac)) "Relatedness prover for Transport" ) fun instantiate_tac name ct ctxt = PRIMITIVE (Drule.infer_instantiate_types ctxt [((name, Thm.typ_of_cterm ct), ct)]) |> CHANGED val map_dummyT = Term.map_types (K dummyT) fun mk_term_skeleton maxidx t = let val consts = Term.add_consts t [] val (vars, _) = fold_map (uncurry Term_Util.fresh_var o apfst Long_Name.base_name) consts maxidx val t' = Term.subst_atomic (map2 (pair o Const) consts vars) t in t' end fun instantiate_skeleton_tac ctxt = let fun tac ct = let val (x, y) = Transport_Util.cdest_judgement ct |> Thm.dest_binop val default_sort = Proof_Context.default_sort ctxt val skeleton = mk_term_skeleton (Thm.maxidx_of_cterm ct) (Thm.term_of x) |> map_dummyT |> Type.constraint (Thm.typ_of_cterm y) |> Syntax.check_term (Util.set_proof_mode_pattern ctxt) (*add sort constraints for type variables*) |> Term.map_types (Term.map_atyps (map_type_tvar (fn (n, _) => TVar (n, default_sort n)))) |> Thm.cterm_of ctxt in instantiate_tac (Thm.term_of y |> dest_Var |> fst) skeleton ctxt end in Tactic_Util.CSUBGOAL_DATA I (K o tac) end fun transport_whitebox_tac ctxt = instantiate_skeleton_tac ctxt THEN' transport_related_tac ctxt THEN_ALL_NEW ( TRY o REPEAT1 o transport_relator_rewrite_tac ctxt THEN' TRY o any_unify_trp_hints_resolve_tac @{thms refl} ctxt ) val _ = Theory.setup ( Method.setup @{binding trp_whitebox_prover} (Scan.succeed (SIMPLE_METHOD' o transport_whitebox_tac)) "Whitebox prover for Transport" ) end \ No newline at end of file diff --git a/thys/Transport/Transport/Examples/Transport_Partial_Quotient_Types.thy b/thys/Transport/Transport/Examples/Transport_Partial_Quotient_Types.thy --- a/thys/Transport/Transport/Examples/Transport_Partial_Quotient_Types.thy +++ b/thys/Transport/Transport/Examples/Transport_Partial_Quotient_Types.thy @@ -1,98 +1,98 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Transport for Partial Quotient Types\ theory Transport_Partial_Quotient_Types imports HOL.Lifting Transport begin paragraph \Summary\ text \Every partial quotient type @{term Quotient}, as used by the Lifting package, is transportable.\ context transport begin interpretation t : transport L "(=)" l r . lemma Quotient_T_eq_Galois: assumes "Quotient (\\<^bsub>L\<^esub>) l r T" shows "T = t.Galois" proof (intro ext iffI) fix x y assume "T x y" with assms have "x \\<^bsub>L\<^esub> x" "l x = y" using Quotient_cr_rel by auto with assms have "r (l x) \\<^bsub>L\<^esub> x" "r (l x) \\<^bsub>L\<^esub> r y" using Quotient_rep_abs Quotient_rep_reflp by auto with assms have "x \\<^bsub>L\<^esub> r y" using Quotient_part_equivp by (blast elim: part_equivpE dest: transpD sympD) then show "t.Galois x y" by blast next fix x y assume "t.Galois x y" with assms show "T x y" using Quotient_cr_rel Quotient_refl1 Quotient_symp by (fastforce intro: Quotient_rel_abs2[symmetric] dest: sympD) qed lemma Quotient_if_preorder_equivalence: assumes "((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (=)) l r" shows "Quotient (\\<^bsub>L\<^esub>) l r t.Galois" proof (rule QuotientI) from assms show g2: "l (r y) = y" for y by fastforce from assms show "r y \\<^bsub>L\<^esub> r y" for y by blast show g1: "x \\<^bsub>L\<^esub> x' \ x \\<^bsub>L\<^esub> x \ x' \\<^bsub>L\<^esub> x' \ l x = l x'" (is "?lhs \ ?rhs") for x x' proof (rule iffI) assume ?rhs - with assms have "\ x \\<^bsub>L\<^esub> \ x'" by fastforce + with assms have "\ x \\<^bsub>L\<^esub> \ x'" by force moreover from \?rhs\ assms have "x \\<^bsub>L\<^esub> \ x" "\ x' \\<^bsub>L\<^esub> x'" by (blast elim: t.preorder_equivalence_order_equivalenceE)+ moreover from assms have "transitive (\\<^bsub>L\<^esub>)" by blast ultimately show "x \\<^bsub>L\<^esub> x'" by blast next assume ?lhs with assms show ?rhs by blast qed from assms show "t.Galois = (\x y. x \\<^bsub>L\<^esub> x \ l x = y)" by (intro ext iffI) (metis g1 g2 t.left_GaloisE, auto intro!: t.left_Galois_left_if_left_rel_if_inflationary_on_in_fieldI elim!: t.preorder_equivalence_order_equivalenceE) qed lemma partial_equivalence_rel_equivalence_if_Quotient: assumes "Quotient (\\<^bsub>L\<^esub>) l r T" shows "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (=)) l r" proof (rule t.partial_equivalence_rel_equivalence_if_order_equivalenceI) from Quotient_part_equivp[OF assms] show "partial_equivalence_rel (\\<^bsub>L\<^esub>)" by (blast elim: part_equivpE dest: transpD sympD) have "x \\<^bsub>L\<^esub> r (l x)" if "in_field (\\<^bsub>L\<^esub>) x" for x proof - from assms \in_field (\\<^bsub>L\<^esub>) x\ have "x \\<^bsub>L\<^esub> x" using Quotient_refl1 Quotient_refl2 by fastforce with assms Quotient_rep_abs Quotient_symp show ?thesis by (fastforce dest: sympD) qed with assms show "((\\<^bsub>L\<^esub>) \\<^sub>o (=)) l r" using Quotient_abs_rep Quotient_rel_abs Quotient_rep_reflp Quotient_abs_rep[symmetric] by (intro t.order_equivalenceI dep_mono_wrt_relI rel_equivalence_onI inflationary_onI deflationary_onI) auto qed auto corollary Quotient_iff_partial_equivalence_rel_equivalence: "Quotient (\\<^bsub>L\<^esub>) l r t.Galois \ ((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (=)) l r" using Quotient_if_preorder_equivalence partial_equivalence_rel_equivalence_if_Quotient by blast corollary Quotient_T_eq_ge_Galois_right: assumes "Quotient (\\<^bsub>L\<^esub>) l r T" shows "T = t.ge_Galois_right" using assms by (subst t.ge_Galois_right_eq_left_Galois_if_symmetric_if_in_codom_eq_in_dom_if_galois_prop) (blast dest: partial_equivalence_rel_equivalence_if_Quotient intro: in_codom_eq_in_dom_if_reflexive_on_in_field Quotient_T_eq_Galois)+ end end \ No newline at end of file diff --git a/thys/Transport/Transport/Examples/Typedef/Transport_Typedef.thy b/thys/Transport/Transport/Examples/Typedef/Transport_Typedef.thy --- a/thys/Transport/Transport/Examples/Typedef/Transport_Typedef.thy +++ b/thys/Transport/Transport/Examples/Typedef/Transport_Typedef.thy @@ -1,210 +1,210 @@ \<^marker>\creator "Kevin Kappelmann"\ theory Transport_Typedef imports "HOL-Library.FSet" Transport_Typedef_Base Transport_Prototype Transport_Syntax begin context includes galois_rel_syntax transport_syntax begin typedef pint = "{i :: int. 0 \ i}" by auto interpretation typedef_pint : type_definition Rep_pint Abs_pint "{i :: int. 0 \ i}" by (fact type_definition_pint) lemma [trp_relator_rewrite, trp_uhint]: "(\<^bsub>(=\<^bsub>Collect ((\) (0 :: int))\<^esub>)\<^esub>\\<^bsub>(=) Rep_pint\<^esub>) \ typedef_pint.AR" using typedef_pint.left_Galois_eq_AR by (intro eq_reflection) simp typedef 'a fset = "{s :: 'a set. finite s}" by auto interpretation typedef_fset : type_definition Rep_fset Abs_fset "{s :: 'a set. finite s}" by (fact type_definition_fset) lemma [trp_relator_rewrite, trp_uhint]: "(\<^bsub>(=\<^bsub>{s :: 'a set. finite s}\<^esub>) :: 'a set \ _\<^esub>\\<^bsub>(=) Rep_fset\<^esub>) \ typedef_fset.AR" using typedef_fset.left_Galois_eq_AR by (intro eq_reflection) simp lemma eq_restrict_set_eq_eq_uhint [trp_uhint]: - "P \ \x. x \ A \ ((=\<^bsub>A :: 'a set\<^esub>) :: 'a \ _) \ (=\<^bsub>P\<^esub>)" - by simp + "(\x. P x \ x \ A) \ ((=\<^bsub>A :: 'a set\<^esub>) :: 'a \ _) \ (=\<^bsub>P\<^esub>)" + by (intro eq_reflection) (auto simp: fun_eq_iff) -(*could also automatically tagged for every instance in type_definition*) +(*could also automatically be tagged for every instance of type_definition*) declare typedef_pint.partial_equivalence_rel_equivalence[per_intro] typedef_fset.partial_equivalence_rel_equivalence[per_intro] lemma one_parametric [trp_in_dom]: "typedef_pint.L 1 1" by auto trp_term pint_one :: "pint" where x = "1 :: int" by trp_prover lemma add_parametric [trp_in_dom]: "(typedef_pint.L \ typedef_pint.L \ typedef_pint.L) (+) (+)" - by (intro Dep_Fun_Rel_relI) (auto intro!: eq_restrictI elim!: eq_restrictE) + by (intro Dep_Fun_Rel_relI) fastforce trp_term pint_add :: "pint \ pint \ pint" where x = "(+) :: int \ _" by trp_prover lemma empty_parametric [trp_in_dom]: "typedef_fset.L {} {}" by auto trp_term fempty :: "'a fset" where x = "{} :: 'a set" by trp_prover lemma insert_parametric [trp_in_dom]: "((=) \ typedef_fset.L \ typedef_fset.L) insert insert" by auto trp_term finsert :: "'a \ 'a fset \ 'a fset" where x = insert and L = "((=) \ typedef_fset.L \ typedef_fset.L)" and R = "((=) \ typedef_fset.R \ typedef_fset.R)" by trp_prover (*experiments with white-box transports*) context notes refl[trp_related_intro] begin trp_term insert_add_int_fset_whitebox :: "int fset" where x = "insert (1 + (1 :: int)) {}" ! by trp_whitebox_prover lemma empty_parametric' [trp_related_intro]: "(rel_set R) {} {}" by (intro Dep_Fun_Rel_relI rel_setI) (auto dest: rel_setD1 rel_setD2) lemma insert_parametric' [trp_related_intro]: "(R \ rel_set R \ rel_set R) insert insert" by (intro Dep_Fun_Rel_relI rel_setI) (auto dest: rel_setD1 rel_setD2) context assumes [trp_uhint]: (*proven for all natural functors*) "L \ rel_set (L1 :: int \ int \ bool) \ R \ rel_set (R1 :: pint \ pint \ bool) \ r \ image r1 \ S \ (\<^bsub>L1\<^esub>\\<^bsub>R1 r1\<^esub>) \ (\<^bsub>L\<^esub>\\<^bsub>R r\<^esub>) \ rel_set S" begin trp_term insert_add_pint_set_whitebox :: "pint set" where x = "insert (1 + (1 :: int)) {}" ! by trp_whitebox_prover print_statement insert_add_int_fset_whitebox_def insert_add_pint_set_whitebox_def end end lemma image_parametric [trp_in_dom]: "(((=) \ (=)) \ typedef_fset.L \ typedef_fset.L) image image" by (intro Dep_Fun_Rel_relI) auto trp_term fimage :: "('a \ 'b) \ 'a fset \ 'b fset" where x = image by trp_prover (*experiments with compositions*) lemma rel_fun_eq_Fun_Rel_rel: "rel_fun = Fun_Rel_rel" by (intro ext iffI Dep_Fun_Rel_relI) (auto elim: rel_funE) lemma image_parametric' [trp_related_intro]: "((R \ S) \ rel_set R \ rel_set S) image image" using transfer_raw[simplified rel_fun_eq_Fun_Rel_rel Transfer.Rel_def] by simp lemma Galois_id_hint [trp_uhint]: "(L :: 'a \ 'a \ bool) \ R \ r \ id \ E \ L \ (\<^bsub>L\<^esub>\\<^bsub>R r\<^esub>) \ E" by (simp only: eq_reflection[OF transport_id.left_Galois_eq_left]) lemma Freq [trp_uhint]: "L \ (=) \ (=) \ L \ (=)" by auto context fixes L1 R1 l1 r1 L R l r assumes per1: "(L1 \\<^bsub>PER\<^esub> R1) l1 r1" defines "L \ rel_set L1" and "R \ rel_set R1" and "l \ image l1" and "r \ image r1" begin interpretation transport L R l r . context (*proven for all natural functors*) assumes transport_per_set: "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" and compat: "transport_comp.middle_compatible_codom R typedef_fset.L" begin trp_term fempty_param :: "'b fset" where x = "{} :: 'a set" and L = "transport_comp.L ?L1 ?R1 (?l1 :: 'a set \ 'b set) ?r1 typedef_fset.L" and R = "transport_comp.L typedef_fset.R typedef_fset.L ?r2 ?l2 ?R1" apply (rule transport_comp.partial_equivalence_rel_equivalenceI) apply (rule transport_per_set) apply per_prover apply (fact compat) apply (rule transport_comp.left_relI[where ?y="{}" and ?y'="{}"]) apply (unfold L_def R_def l_def r_def) apply (auto intro!: galois_rel.left_GaloisI in_codomI empty_transfer) done definition "set_succ \ image ((+) (1 :: int))" lemma set_succ_parametric [trp_in_dom]: "(typedef_fset.L \ typedef_fset.L) set_succ set_succ" unfolding set_succ_def by auto trp_term fset_succ :: "int fset \ int fset" where x = set_succ and L = "typedef_fset.L \ typedef_fset.L" and R = "typedef_fset.R \ typedef_fset.R" by trp_prover trp_term fset_succ' :: "int fset \ int fset" where x = set_succ and L = "typedef_fset.L \ typedef_fset.L" and R = "typedef_fset.R \ typedef_fset.R" unfold set_succ_def ! (*current prototype gets lost in this example*) using refl[trp_related_intro] apply (tactic \Transport.instantiate_skeleton_tac @{context} 1\) apply (tactic \Transport.transport_related_step_tac @{context} 1\) apply (tactic \Transport.transport_related_step_tac @{context} 1\) apply (tactic \Transport.transport_related_step_tac @{context} 1\) apply (tactic \Transport.transport_related_step_tac @{context} 1\) apply (tactic \Transport.transport_related_step_tac @{context} 1\) apply (tactic \Transport.transport_related_step_tac @{context} 1\) apply assumption apply assumption prefer 3 apply (tactic \Transport.transport_related_step_tac @{context} 1\) apply (tactic \Transport.transport_related_step_tac @{context} 1\) apply (fold trp_def) apply (trp_relator_rewrite)+ apply (unfold trp_def) apply (trp_hints_resolve refl) done lemma pint_middle_compat: "transport_comp.middle_compatible_codom (rel_set ((=) :: pint \ _)) (=\<^bsub>Collect (finite :: pint set \ _)\<^esub>)" by (intro transport_comp.middle_compatible_codom_if_right1_le_eqI) (auto simp: rel_set_eq intro!: transitiveI) trp_term pint_fset_succ :: "pint fset \ pint fset" where x = "set_succ :: int set \ int set" (*automation for composition not supported as of now*) oops end end end end \ No newline at end of file diff --git a/thys/Transport/Transport/Examples/Typedef/Transport_Typedef_Base.thy b/thys/Transport/Transport/Examples/Typedef/Transport_Typedef_Base.thy --- a/thys/Transport/Transport/Examples/Typedef/Transport_Typedef_Base.thy +++ b/thys/Transport/Transport/Examples/Typedef/Transport_Typedef_Base.thy @@ -1,42 +1,40 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Transport for HOL Type Definitions\ theory Transport_Typedef_Base imports HOL_Alignment_Binary_Relations Transport_Bijections HOL.Typedef begin context type_definition begin abbreviation (input) "L :: 'a \ 'a \ bool \ (=\<^bsub>A\<^esub>)" abbreviation (input) "R :: 'b \ 'b \ bool \ (=)" sublocale transport? : transport_eq_restrict_bijection "mem_of A" "\ :: 'b \ bool" Abs Rep rewrites "(=\<^bsub>mem_of A\<^esub>) \ L" and "(=\<^bsub>\ :: 'b \ bool\<^esub>) \ R" and "(galois_rel.Galois (=) (=) Rep)\\<^bsub>mem_of A\<^esub>\\<^bsub>\ :: 'b \ bool\<^esub> \ (galois_rel.Galois (=) (=) Rep)" using Abs_inverse Rep_inverse by (intro transport_eq_restrict_bijection.intro bijection_onI) - (auto simp: restrict_right_eq - intro!: eq_reflection galois_rel.left_GaloisI Rep - elim: galois_rel.left_GaloisE) + (auto intro!: eq_reflection galois_rel.left_GaloisI Rep elim: galois_rel.left_GaloisE) interpretation galois L R Abs Rep . lemma Rep_left_Galois_self: "Rep y \<^bsub>L\<^esub>\ y" using Rep by (intro left_GaloisI) auto definition "AR x y \ x = Rep y" lemma left_Galois_eq_AR: "left_Galois = AR" unfolding AR_def by (auto intro!: galois_rel.left_GaloisI Rep elim: galois_rel.left_GaloisE) end end diff --git a/thys/Transport/Transport/Functions/Monotone_Function_Relator.thy b/thys/Transport/Transport/Functions/Monotone_Function_Relator.thy --- a/thys/Transport/Transport/Functions/Monotone_Function_Relator.thy +++ b/thys/Transport/Transport/Functions/Monotone_Function_Relator.thy @@ -1,105 +1,105 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Monotone Function Relator\ theory Monotone_Function_Relator imports Reflexive_Relator begin abbreviation "Mono_Dep_Fun_Rel R S \ ([x y \ R] \ S x y)\<^sup>\" abbreviation "Mono_Fun_Rel R S \ Mono_Dep_Fun_Rel R (\_ _. S)" bundle Mono_Dep_Fun_Rel_syntax begin syntax "_Mono_Fun_Rel_rel" :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("(_) \\ (_)" [41, 40] 40) "_Mono_Dep_Fun_Rel_rel" :: "idt \ idt \ ('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _] \\ (_)" [41, 41, 41, 40] 40) "_Mono_Dep_Fun_Rel_rel_if" :: "idt \ idt \ ('a \ 'b \ bool) \ bool \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _/ |/ _] \\ (_)" [41, 41, 41, 41, 40] 40) end bundle no_Mono_Dep_Fun_Rel_syntax begin no_syntax "_Mono_Fun_Rel_rel" :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("(_) \\ (_)" [41, 40] 40) "_Mono_Dep_Fun_Rel_rel" :: "idt \ idt \ ('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _] \\ (_)" [41, 41, 41, 40] 40) "_Mono_Dep_Fun_Rel_rel_if" :: "idt \ idt \ ('a \ 'b \ bool) \ bool \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _/ |/ _] \\ (_)" [41, 41, 41, 41, 40] 40) end unbundle Mono_Dep_Fun_Rel_syntax translations "R \\ S" \ "CONST Mono_Fun_Rel R S" "[x y \ R] \\ S" \ "CONST Mono_Dep_Fun_Rel R (\x y. S)" "[x y \ R | B] \\ S" \ "CONST Mono_Dep_Fun_Rel R (\x y. CONST rel_if B S)" locale Dep_Fun_Rel_orders = fixes L :: "'a \ 'b \ bool" and R :: "'a \ 'b \ 'c \ 'd \ bool" begin sublocale o : orders L "R a b" for a b . notation L (infix "\\<^bsub>L\<^esub>" 50) notation o.ge_left (infix "\\<^bsub>L\<^esub>" 50) notation R ("(\\<^bsub>R (_) (_)\<^esub>)" 50) abbreviation "right_infix c a b d \ (\\<^bsub>R a b\<^esub>) c d" notation right_infix ("(_) \\<^bsub>R (_) (_)\<^esub> (_)" [51,51,51,51] 50) notation o.ge_right ("(\\<^bsub>R (_) (_)\<^esub>)" 50) abbreviation (input) "ge_right_infix d a b c \ (\\<^bsub>R a b\<^esub>) d c" notation ge_right_infix ("(_) \\<^bsub>R (_) (_)\<^esub> (_)" [51,51,51,51] 50) abbreviation (input) "DFR \ ([a b \ L] \ R a b)" end locale hom_Dep_Fun_Rel_orders = Dep_Fun_Rel_orders L R for L :: "'a \ 'a \ bool" and R :: "'a \ 'a \ 'b \ 'b \ bool" begin sublocale ho : hom_orders L "R a b" for a b . lemma Mono_Dep_Fun_Refl_Rel_right_eq_Mono_Dep_Fun_if_le_if_reflexive_onI: assumes refl_L: "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x2 x2\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x1\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" shows "([x y \ (\\<^bsub>L\<^esub>)] \\ (\\<^bsub>R x y\<^esub>)\<^sup>\) = ([x y \ (\\<^bsub>L\<^esub>)] \\ (\\<^bsub>R x y\<^esub>))" proof - { fix f g x1 x2 assume "([x y \ (\\<^bsub>L\<^esub>)] \ (\\<^bsub>R x y\<^esub>)) f g" "x1 \\<^bsub>L\<^esub> x1" "x1 \\<^bsub>L\<^esub> x2" with assms have "f x1 \\<^bsub>R x1 x2\<^esub> g x1" "f x2 \\<^bsub>R x1 x2\<^esub> g x2" by blast+ } with refl_L show ?thesis by (intro ext iffI Refl_RelI Dep_Fun_Rel_relI) (auto elim!: Refl_RelE) qed lemma Mono_Dep_Fun_Refl_Rel_right_eq_Mono_Dep_Fun_if_mono_if_reflexive_onI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "([x1 x2 \ (\\<^bsub>L\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L\<^esub>) | x1 \\<^bsub>L\<^esub> x3] \ (\)) R" shows "([x y \ (\\<^bsub>L\<^esub>)] \\ (\\<^bsub>R x y\<^esub>)\<^sup>\) = ([x y \ (\\<^bsub>L\<^esub>)] \\ (\\<^bsub>R x y\<^esub>))" using assms by (intro Mono_Dep_Fun_Refl_Rel_right_eq_Mono_Dep_Fun_if_le_if_reflexive_onI) - auto + (auto 6 0) end context hom_orders begin sublocale fro : hom_Dep_Fun_Rel_orders L "\_ _. R" . corollary Mono_Fun_Rel_Refl_Rel_right_eq_Mono_Fun_RelI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\ (\\<^bsub>R\<^esub>)\<^sup>\) = ((\\<^bsub>L\<^esub>) \\ (\\<^bsub>R\<^esub>))" using assms by (intro fro.Mono_Dep_Fun_Refl_Rel_right_eq_Mono_Dep_Fun_if_le_if_reflexive_onI) simp_all end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Reflexive_Relator.thy b/thys/Transport/Transport/Functions/Reflexive_Relator.thy --- a/thys/Transport/Transport/Functions/Reflexive_Relator.thy +++ b/thys/Transport/Transport/Functions/Reflexive_Relator.thy @@ -1,275 +1,275 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Reflexive Relator\ theory Reflexive_Relator imports Galois_Equivalences Galois_Relator begin definition "Refl_Rel R x y \ R x x \ R y y \ R x y" bundle Refl_Rel_syntax begin notation Refl_Rel ("(_\<^sup>\)" [1000]) end bundle no_Refl_Rel_syntax begin no_notation Refl_Rel ("(_\<^sup>\)" [1000]) end unbundle Refl_Rel_syntax lemma Refl_RelI [intro]: assumes "R x x" and "R y y" and "R x y" shows "R\<^sup>\ x y" using assms unfolding Refl_Rel_def by blast lemma Refl_Rel_selfI [intro]: assumes "R x x" shows "R\<^sup>\ x x" using assms by blast lemma Refl_RelE [elim]: assumes "R\<^sup>\ x y" obtains "R x x" "R y y" "R x y" using assms unfolding Refl_Rel_def by blast lemma Refl_Rel_reflexive_on_in_field [iff]: "reflexive_on (in_field R\<^sup>\) R\<^sup>\" by (rule reflexive_onI) auto lemma Refl_Rel_le_self [iff]: "R\<^sup>\ \ R" by blast lemma Refl_Rel_eq_self_if_reflexive_on [simp]: assumes "reflexive_on (in_field R) R" shows "R\<^sup>\ = R" using assms by blast lemma reflexive_on_in_field_if_Refl_Rel_eq_self: assumes "R\<^sup>\ = R" shows "reflexive_on (in_field R) R" by (fact Refl_Rel_reflexive_on_in_field[of R, simplified assms]) corollary Refl_Rel_eq_self_iff_reflexive_on: "R\<^sup>\ = R \ reflexive_on (in_field R) R" using Refl_Rel_eq_self_if_reflexive_on reflexive_on_in_field_if_Refl_Rel_eq_self by blast lemma Refl_Rel_Refl_Rel_eq [simp]: "(R\<^sup>\)\<^sup>\ = R\<^sup>\" by (intro ext) auto lemma rel_inv_Refl_Rel_eq [simp]: "(R\<^sup>\)\ = (R\)\<^sup>\" by (intro ext iffI Refl_RelI rel_invI) auto lemma Refl_Rel_transitive_onI [intro]: assumes "transitive_on (P :: 'a \ bool) (R :: 'a \ _)" shows "transitive_on P R\<^sup>\" using assms by (intro transitive_onI) (blast dest: transitive_onD) corollary Refl_Rel_transitiveI [intro]: assumes "transitive R" shows "transitive R\<^sup>\" using assms by blast corollary Refl_Rel_preorder_onI: assumes "transitive_on P R" and "P \ in_field R\<^sup>\" shows "preorder_on P R\<^sup>\" using assms by (intro preorder_onI reflexive_on_if_le_pred_if_reflexive_on[where ?P="in_field R\<^sup>\" and ?P'=P]) auto corollary Refl_Rel_preorder_on_in_fieldI [intro]: assumes "transitive R" shows "preorder_on (in_field R\<^sup>\) R\<^sup>\" using assms by (intro Refl_Rel_preorder_onI) auto lemma Refl_Rel_symmetric_onI [intro]: assumes "symmetric_on (P :: 'a \ bool) (R :: 'a \ _)" shows "symmetric_on P R\<^sup>\" using assms by (intro symmetric_onI) (auto dest: symmetric_onD) lemma Refl_Rel_symmetricI [intro]: assumes "symmetric R" shows "symmetric R\<^sup>\" using assms by (fold symmetric_on_in_field_iff_symmetric) (blast intro: symmetric_on_if_le_pred_if_symmetric_on) lemma Refl_Rel_partial_equivalence_rel_onI [intro]: assumes "partial_equivalence_rel_on (P :: 'a \ bool) (R :: 'a \ _)" shows "partial_equivalence_rel_on P R\<^sup>\" using assms by (intro partial_equivalence_rel_onI Refl_Rel_transitive_onI Refl_Rel_symmetric_onI) auto lemma Refl_Rel_partial_equivalence_relI [intro]: assumes "partial_equivalence_rel R" shows "partial_equivalence_rel R\<^sup>\" using assms by (intro partial_equivalence_relI Refl_Rel_transitiveI Refl_Rel_symmetricI) auto lemma Refl_Rel_app_leftI: assumes "R (f x) y" and "in_field S\<^sup>\ x" and "in_field R\<^sup>\ y" and "(S \\<^sub>m R) f" shows "R\<^sup>\ (f x) y" proof (rule Refl_RelI) from \in_field R\<^sup>\ y\ show "R y y" by blast from \in_field S\<^sup>\ x\ have "S x x" by blast with \(S \\<^sub>m R) f\ show "R (f x) (f x)" by blast qed fact corollary Refl_Rel_app_rightI: assumes "R x (f y)" and "in_field S\<^sup>\ y" and "in_field R\<^sup>\ x" and "(S \\<^sub>m R) f" shows "R\<^sup>\ x (f y)" proof - from assms have "(R\)\<^sup>\ (f y) x" by (intro Refl_Rel_app_leftI[where ?S="S\"]) - (auto simp flip: rel_inv_Refl_Rel_eq) + (auto 5 0 simp flip: rel_inv_Refl_Rel_eq) then show ?thesis by blast qed lemma mono_wrt_rel_Refl_Rel_Refl_Rel_if_mono_wrt_rel [intro]: assumes "(R \\<^sub>m S) f" shows "(R\<^sup>\ \\<^sub>m S\<^sup>\) f" - using assms by (intro dep_mono_wrt_relI) auto + using assms by (intro dep_mono_wrt_relI) blast context galois begin interpretation gR : galois "(\\<^bsub>L\<^esub>)\<^sup>\" "(\\<^bsub>R\<^esub>)\<^sup>\" l r . lemma Galois_Refl_RelI: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "in_field (\\<^bsub>L\<^esub>)\<^sup>\ x" and "in_field (\\<^bsub>R\<^esub>)\<^sup>\ y" and "in_codom (\\<^bsub>R\<^esub>) y \ x \<^bsub>L\<^esub>\ y" shows "(galois_rel.Galois ((\\<^bsub>L\<^esub>)\<^sup>\) ((\\<^bsub>R\<^esub>)\<^sup>\) r) x y" using assms by (intro gR.left_GaloisI in_codomI Refl_Rel_app_rightI[where ?f=r]) auto lemma half_galois_prop_left_Refl_Rel_left_rightI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \<^sub>h\ (\\<^bsub>R\<^esub>)\<^sup>\) l r" using assms by (intro gR.half_galois_prop_leftI Refl_RelI) (auto elim!: in_codomE gR.left_GaloisE Refl_RelE) interpretation flip_inv : galois "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" r l rewrites "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) \ ((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>))" and "\R. (R\)\<^sup>\ \ (R\<^sup>\)\" and "\R S f g. (R\ \<^sub>h\ S\) f g \ (S \\<^sub>h R) g f" by (simp_all add: galois_prop.half_galois_prop_left_rel_inv_iff_half_galois_prop_right) lemma half_galois_prop_right_Refl_Rel_right_leftI: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \\<^sub>h (\\<^bsub>R\<^esub>)\<^sup>\) l r" using assms by (fact flip_inv.half_galois_prop_left_Refl_Rel_left_rightI) corollary galois_prop_Refl_Rel_left_rightI: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \ (\\<^bsub>R\<^esub>)\<^sup>\) l r" using assms by (intro gR.galois_propI half_galois_prop_left_Refl_Rel_left_rightI half_galois_prop_right_Refl_Rel_right_leftI) auto lemma galois_connection_Refl_Rel_left_rightI: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \ (\\<^bsub>R\<^esub>)\<^sup>\) l r" using assms by (intro gR.galois_connectionI galois_prop_Refl_Rel_left_rightI) auto lemma galois_equivalence_Refl_RelI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \\<^sub>G (\\<^bsub>R\<^esub>)\<^sup>\) l r" proof - interpret flip : galois R L r l . show ?thesis using assms by (intro gR.galois_equivalenceI galois_connection_Refl_Rel_left_rightI flip.galois_prop_Refl_Rel_left_rightI) auto qed end context order_functors begin lemma inflationary_on_in_field_Refl_Rel_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "inflationary_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "inflationary_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms - by (intro inflationary_onI Refl_RelI) (auto elim!: in_fieldE Refl_RelE) + by (intro inflationary_onI Refl_RelI) (auto 0 3 elim!: in_fieldE Refl_RelE) lemma inflationary_on_in_field_Refl_Rel_left': assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "inflationary_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "inflationary_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms - by (intro inflationary_onI Refl_RelI) (auto elim!: in_fieldE Refl_RelE) + by (intro inflationary_onI Refl_RelI) (auto 0 3 elim!: in_fieldE Refl_RelE) interpretation inv : galois "(\\<^bsub>L\<^esub>)" "(\\<^bsub>R\<^esub>)" l r rewrites "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) \ ((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>))" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) \ ((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>))" and "\R. (R\)\<^sup>\ \ (R\<^sup>\)\" and "\R. in_dom R\ \ in_codom R" and "\R. in_codom R\ \ in_dom R" and "\R. in_field R\ \ in_field R" and "\P R. inflationary_on P R\ \ deflationary_on P R" by simp_all lemma deflationary_on_in_field_Refl_Rel_leftI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "deflationary_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "deflationary_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms by (fact inv.inflationary_on_in_field_Refl_Rel_left') lemma deflationary_on_in_field_Refl_RelI_left': assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "deflationary_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "deflationary_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms by (fact inv.inflationary_on_in_field_Refl_Rel_left) lemma rel_equivalence_on_in_field_Refl_Rel_leftI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "rel_equivalence_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms by (intro rel_equivalence_onI inflationary_on_in_field_Refl_Rel_left deflationary_on_in_field_Refl_Rel_leftI) auto lemma rel_equivalence_on_in_field_Refl_Rel_leftI': assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "rel_equivalence_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms by (intro rel_equivalence_onI inflationary_on_in_field_Refl_Rel_left' deflationary_on_in_field_Refl_RelI_left') auto interpretation oR : order_functors "(\\<^bsub>L\<^esub>)\<^sup>\" "(\\<^bsub>R\<^esub>)\<^sup>\" l r . lemma order_equivalence_Refl_RelI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \\<^sub>o (\\<^bsub>R\<^esub>)\<^sup>\) l r" proof - interpret flip : galois R L r l rewrites "flip.unit \ \" by (simp only: flip_unit_eq_counit) show ?thesis using assms by (intro oR.order_equivalenceI mono_wrt_rel_Refl_Rel_Refl_Rel_if_mono_wrt_rel rel_equivalence_on_in_field_Refl_Rel_leftI flip.rel_equivalence_on_in_field_Refl_Rel_leftI) (auto intro: rel_equivalence_on_if_le_pred_if_rel_equivalence_on in_field_if_in_dom) qed end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions.thy b/thys/Transport/Transport/Functions/Transport_Functions.thy --- a/thys/Transport/Transport/Functions/Transport_Functions.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions.thy @@ -1,248 +1,248 @@ \<^marker>\creator "Kevin Kappelmann"\ theory Transport_Functions imports Transport_Functions_Galois_Equivalence Transport_Functions_Galois_Relator Transport_Functions_Order_Base Transport_Functions_Order_Equivalence Transport_Functions_Relation_Simplifications begin paragraph \Summary\ text \Composition under (dependent) (monotone) function relators. Refer to \<^cite>\"transport"\ for more details.\ subsection \Summary of Main Results\ text \More precise results can be found in the corresponding subtheories.\ paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin interpretation flip : transport_Mono_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.counit \ \\<^sub>1" and "flip.t1.unit \ \\<^sub>1" by (simp_all only: t1.flip_counit_eq_unit t1.flip_unit_eq_counit) subparagraph \Closure of Order and Galois Concepts\ theorem preorder_galois_connection_if_galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([_ x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>) | \\<^sub>1 x2' \\<^bsub>R1\<^esub> x1'] \\<^sub>m [x3' _ \ (\\<^bsub>R1\<^esub>) | x2' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro preorder_galois_connectionI galois_connection_left_right_if_mono_if_galois_connectionI' preorder_on_in_field_leftI flip.preorder_on_in_field_leftI tdfr.transitive_leftI' flip.tdfr.transitive_leftI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) theorem preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 _ \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' _ \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' _ \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" shows "((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro preorder_equivalence_if_galois_equivalenceI galois_equivalence_if_mono_if_preorder_equivalenceI' preorder_on_in_field_leftI flip.preorder_on_in_field_leftI tdfr.transitive_leftI' flip.tdfr.transitive_leftI tdfr.transitive_left2_if_preorder_equivalenceI tdfr.transitive_right2_if_preorder_equivalenceI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI flip.tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) theorem partial_equivalence_rel_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 _ \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' _ \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' _ \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" shows "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" - using assms by (intro partial_equivalence_rel_equivalence_if_galois_equivalenceI + using assms + by (intro partial_equivalence_rel_equivalence_if_galois_equivalenceI galois_equivalence_if_mono_if_preorder_equivalenceI' tdfr.transitive_left2_if_preorder_equivalenceI tdfr.transitive_right2_if_preorder_equivalenceI partial_equivalence_rel_leftI flip.partial_equivalence_rel_leftI tdfr.partial_equivalence_rel_left2_if_partial_equivalence_rel_equivalenceI tdfr.partial_equivalence_rel_right2_if_partial_equivalence_rel_equivalenceI) auto - subparagraph \Simplification of Left and Right Relations\ text \See @{thm "left_rel_eq_tdfr_leftI_if_equivalencesI"}.\ subparagraph \Simplification of Galois relator\ text \See @{thm "left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_mono_if_galois_connectionI" "left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_preorder_equivalenceI" "left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_preorder_equivalenceI'" "Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI" "Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq"}\ end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin interpretation flip : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . subparagraph \Closure of Order and Galois Concepts\ lemma preorder_galois_connection_if_galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \ (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>L2\<^esub>)" "transitive (\\<^bsub>R2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro tpdfr.preorder_galois_connectionI galois_connection_left_rightI tpdfr.preorder_on_in_field_leftI flip.tpdfr.preorder_on_in_field_leftI tfr.transitive_leftI' flip.tfr.transitive_leftI) auto theorem preorder_galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro preorder_galois_connection_if_galois_connectionI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) theorem preorder_equivalence_if_galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>G (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>L2\<^esub>)" "transitive (\\<^bsub>R2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro tpdfr.preorder_equivalence_if_galois_equivalenceI galois_equivalenceI tpdfr.preorder_on_in_field_leftI flip.tpdfr.preorder_on_in_field_leftI tfr.transitive_leftI flip.tfr.transitive_leftI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) theorem preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro preorder_equivalence_if_galois_equivalenceI) auto theorem partial_equivalence_rel_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro tpdfr.partial_equivalence_rel_equivalence_if_galois_equivalenceI galois_equivalenceI partial_equivalence_rel_leftI flip.partial_equivalence_rel_leftI) auto subparagraph \Simplification of Left and Right Relations\ text \See @{thm "left_rel_eq_tfr_leftI"}.\ subparagraph \Simplification of Galois relator\ text \See @{thm "left_Galois_eq_Fun_Rel_left_Galois_restrictI" "Fun_Rel_left_Galois_restrict_left_right_eq_Fun_Rel_left_GaloisI" "Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq"}.\ end paragraph \Dependent Function Relator\ text \While a general transport of functions is only possible for the monotone function relator (see above), the locales @{locale "transport_Dep_Fun_Rel"} and @{locale "transport_Fun_Rel"} contain special cases to transport functions that are proven to be monotone using the standard function space. Moreover, in the special case of equivalences on partial equivalence relations, the standard function space is monotone - see @{thm "transport_Mono_Dep_Fun_Rel.left_rel_eq_tdfr_leftI_if_equivalencesI"} As such, we can derive general transport theorems from the monotone cases above.\ context transport_Dep_Fun_Rel begin interpretation tpdfr : transport_Mono_Dep_Fun_Rel L1 R1 l1 r1 L2 R2 l2 r2 . interpretation flip : transport_Mono_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . theorem partial_equivalence_rel_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" shows "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" proof - from assms have "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) = (tpdfr.L \\<^bsub>PER\<^esub> tpdfr.R)" by (subst tpdfr.left_rel_eq_tdfr_leftI_if_equivalencesI flip.left_rel_eq_tdfr_leftI_if_equivalencesI, auto intro!: partial_equivalence_rel_left2_if_partial_equivalence_rel_equivalenceI partial_equivalence_rel_right2_if_partial_equivalence_rel_equivalenceI iff: t1.galois_equivalence_right_left_iff_galois_equivalence_left_right)+ with assms show ?thesis by (auto intro!: tpdfr.partial_equivalence_rel_equivalenceI) qed end paragraph \Function Relator\ context transport_Fun_Rel begin interpretation tpfr : transport_Mono_Fun_Rel L1 R1 l1 r1 L2 R2 l2 r2 . interpretation flip_tpfr : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . theorem partial_equivalence_rel_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" proof - from assms have "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) = (tpfr.tpdfr.L \\<^bsub>PER\<^esub> tpfr.tpdfr.R)" by (subst tpfr.left_rel_eq_tfr_leftI flip_tpfr.left_rel_eq_tfr_leftI; auto)+ with assms show ?thesis by (auto intro!: tpfr.partial_equivalence_rel_equivalenceI) qed end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions_Galois_Connection.thy b/thys/Transport/Transport/Functions/Transport_Functions_Galois_Connection.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Galois_Connection.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Galois_Connection.thy @@ -1,307 +1,307 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Connection\ theory Transport_Functions_Galois_Connection imports Transport_Functions_Galois_Property Transport_Functions_Monotone begin paragraph \Dependent Function Relator\ context transport_Dep_Fun_Rel begin subparagraph \Lemmas for Monotone Function Relator\ lemma galois_connection_left_right_if_galois_connection_mono_2_assms_leftI: assumes galois_conn1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_R1: "reflexive_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and R2_le1: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and mono_l2_2: "([x' \ in_codom (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x'] \\<^sub>m [in_field (\\<^bsub>L2 x1 (r1 x')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x'\<^esub>)) l2" shows "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ([in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>)" proof - show "([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" if "x1' \\<^bsub>R1\<^esub> x2'" for x1' x2' proof - from galois_conn1 \x1' \\<^bsub>R1\<^esub> x2'\ have "r1 x1' \\<^bsub>L1\<^esub> r1 x2'" "r1 x2' \<^bsub>L1\<^esub>\ x2'" using refl_R1 by (auto intro: t1.right_left_Galois_if_reflexive_onI) with mono_l2_2 show ?thesis using R2_le1 \x1' \\<^bsub>R1\<^esub> x2'\ by fastforce qed show "([in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>)" if "x \\<^bsub>L1\<^esub> x" for x proof - from galois_conn1 \x \\<^bsub>L1\<^esub> x\ have "x \\<^bsub>L1\<^esub> \\<^sub>1 x" "\\<^sub>1 x \<^bsub>L1\<^esub>\ l1 x" by (auto intro!: t1.right_left_Galois_if_right_relI t1.rel_unit_if_left_rel_if_half_galois_prop_right_if_mono_wrt_rel [unfolded t1.unit_eq]) with mono_l2_2 show ?thesis by fastforce qed qed lemma galois_connection_left_right_if_galois_connection_mono_assms_leftI: assumes galois_conn1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_R1: "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and R2_le1: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and mono_l2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" shows "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" and "([x' \ in_codom (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x'] \\<^sub>m [in_field (\\<^bsub>L2 x1 (r1 x')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x'\<^esub>)) l2" proof - show "([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" if "x1' \\<^bsub>R1\<^esub> x2'" for x1' x2' proof - from galois_conn1 \x1' \\<^bsub>R1\<^esub> x2'\ have "r1 x1' \\<^bsub>L1\<^esub> r1 x1'" "r1 x1' \<^bsub>L1\<^esub>\ x1'" - using refl_R1 by blast+ - with mono_l2 show ?thesis using \x1' \\<^bsub>R1\<^esub> x2'\ R2_le1 by (auto 9 0) + using refl_R1 by force+ + with mono_l2 show ?thesis using \x1' \\<^bsub>R1\<^esub> x2'\ R2_le1 by (auto 11 0) qed from mono_l2 show "([x' \ in_codom (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x'] \\<^sub>m [in_field (\\<^bsub>L2 x1 (r1 x')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x'\<^esub>)) l2" using refl_R1 by blast qed text \In theory, the following lemmas can be obtained by taking the flipped, inverse interpretation of the locale; however, rewriting the assumptions is more involved than simply copying and adapting above proofs.\ lemma galois_connection_left_right_if_galois_connection_mono_2_assms_rightI: assumes galois_conn1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and L2_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and mono_r2_2: "([x \ in_dom (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x \<^bsub>L1\<^esub>\ x1'] \\<^sub>m [in_field (\\<^bsub>R2 (l1 x) x2'\<^esub>)] \ (\\<^bsub>L2 x (r1 x2')\<^esub>)) r2" shows "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" proof - show "([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" if "x1 \\<^bsub>L1\<^esub> x2" for x1 x2 proof - from galois_conn1 \x1 \\<^bsub>L1\<^esub> x2\ have "x1 \<^bsub>L1\<^esub>\ l1 x1" "l1 x1 \\<^bsub>R1\<^esub> l1 x2" using refl_L1 by (auto intro!: t1.left_Galois_left_if_reflexive_on_if_half_galois_prop_rightI) - with mono_r2_2 show ?thesis using L2_le2 \x1 \\<^bsub>L1\<^esub> x2\ by (auto 9 0) + with mono_r2_2 show ?thesis using L2_le2 \x1 \\<^bsub>L1\<^esub> x2\ by (auto 12 0) qed show "([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" if "x' \\<^bsub>R1\<^esub> x'" for x' proof - from galois_conn1 \x' \\<^bsub>R1\<^esub> x'\ have "r1 x' \<^bsub>L1\<^esub>\ \\<^sub>1 x'" "\\<^sub>1 x' \\<^bsub>R1\<^esub> x'" by (auto intro!: t1.left_Galois_left_if_left_relI t1.counit_rel_if_right_rel_if_half_galois_prop_left_if_mono_wrt_rel [unfolded t1.counit_eq]) with mono_r2_2 show ?thesis by fastforce qed qed lemma galois_connection_left_right_if_galois_connection_mono_assms_rightI: assumes galois_conn1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and L2_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" shows "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "([x \ in_dom (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x \<^bsub>L1\<^esub>\ x1'] \\<^sub>m [in_field (\\<^bsub>R2 (l1 x) x2'\<^esub>)] \ (\\<^bsub>L2 x (r1 x2')\<^esub>)) r2" proof - show "([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" if "x1 \\<^bsub>L1\<^esub> x2" for x1 x2 proof - from galois_conn1 \x1 \\<^bsub>L1\<^esub> x2\ have "x2 \<^bsub>L1\<^esub>\ l1 x2" "l1 x2 \\<^bsub>R1\<^esub> l1 x2" using refl_L1 by (blast intro: t1.left_Galois_left_if_reflexive_on_if_half_galois_prop_rightI)+ with mono_r2 show ?thesis using \x1 \\<^bsub>L1\<^esub> x2\ L2_le2 by fastforce qed from mono_r2 show "([x \ in_dom (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x \<^bsub>L1\<^esub>\ x1'] \\<^sub>m [in_field (\\<^bsub>R2 (l1 x) x2'\<^esub>)] \ (\\<^bsub>L2 x (r1 x2')\<^esub>)) r2" using refl_L1 by blast qed end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin interpretation flip : transport_Mono_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma galois_connection_left_rightI: assumes "(tdfr.L \\<^sub>m tdfr.R) l" and "(tdfr.R \\<^sub>m tdfr.L) r" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) \ (\\<^bsub>L2 x x\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) \ (\\<^bsub>R2 x' x'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ([in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_connectionI galois_prop_left_rightI' mono_wrt_rel_leftI flip.mono_wrt_rel_leftI) auto lemma galois_connection_left_rightI': assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \\<^sub>m (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ((\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) \\<^sub>m (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ([in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_connection_left_rightI tdfr.mono_wrt_rel_left_if_transitiveI tdfr.mono_wrt_rel_right_if_transitiveI) auto lemma galois_connection_left_right_if_galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ([in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_connection_left_rightI' tdfr.mono_wrt_rel_left2_if_mono_wrt_rel_left2_if_left_GaloisI tdfr.mono_wrt_rel_right2_if_mono_wrt_rel_right2_if_left_GaloisI) (auto 7 0) corollary galois_connection_left_right_if_galois_connectionI': assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" and "([x' \ in_codom (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x'] \\<^sub>m [in_field (\\<^bsub>L2 x1 (r1 x')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x'\<^esub>)) l2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "([x \ in_dom (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x \<^bsub>L1\<^esub>\ x1'] \\<^sub>m [in_field (\\<^bsub>R2 (l1 x) x2'\<^esub>)] \ (\\<^bsub>L2 x (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_connection_left_right_if_galois_connectionI tdfr.galois_connection_left_right_if_galois_connection_mono_2_assms_leftI tdfr.galois_connection_left_right_if_galois_connection_mono_2_assms_rightI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) corollary galois_connection_left_right_if_mono_if_galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_connection_left_right_if_galois_connectionI' tdfr.galois_connection_left_right_if_galois_connection_mono_assms_leftI tdfr.galois_connection_left_right_if_galois_connection_mono_assms_rightI) auto corollary galois_connection_left_right_if_mono_if_galois_connectionI': assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([_ x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>) | \\<^sub>1 x2' \\<^bsub>R1\<^esub> x1'] \\<^sub>m [x3' _ \ (\\<^bsub>R1\<^esub>) | x2' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_connection_left_right_if_mono_if_galois_connectionI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI) auto end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin interpretation flip : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma galois_connection_left_rightI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \ (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>L2\<^esub>)" and "transitive (\\<^bsub>R2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro tpdfr.galois_connectionI galois_prop_left_rightI mono_wrt_rel_leftI flip.mono_wrt_rel_leftI) auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions_Galois_Equivalence.thy b/thys/Transport/Transport/Functions/Transport_Functions_Galois_Equivalence.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Galois_Equivalence.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Galois_Equivalence.thy @@ -1,351 +1,352 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Equivalence\ theory Transport_Functions_Galois_Equivalence imports Transport_Functions_Galois_Connection Transport_Functions_Order_Base begin paragraph \Dependent Function Relator\ context transport_Dep_Fun_Rel begin subparagraph \Lemmas for Monotone Function Relator\ lemma flip_half_galois_prop_left2_if_half_galois_prop_left2_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and half_galois_prop_left2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \<^sub>h\ (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>) (l2\<^bsub>x' x\<^esub>) " and "(\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) = (\\<^bsub>L2 x x\<^esub>)" and "(\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) = (\\<^bsub>L2 x x\<^esub>)" and "x \\<^bsub>L1\<^esub> x" shows "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \<^sub>h\ (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub>)" proof - from assms have "x \<^bsub>L1\<^esub>\ l1 x" by (intro t1.left_Galois_left_if_left_relI) auto with half_galois_prop_left2 have "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \<^sub>h\ (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub>)" by auto with assms show ?thesis by simp qed lemma flip_half_galois_prop_right2_if_half_galois_prop_right2_if_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and half_galois_prop_right2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>h (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>) (l2\<^bsub>x' x\<^esub>)" and "(\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) = (\\<^bsub>R2 x' x'\<^esub>)" and "(\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) = (\\<^bsub>R2 x' x'\<^esub>)" and "x' \\<^bsub>R1\<^esub> x'" shows "((\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) \\<^sub>h (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>) (l2\<^bsub>x' (r1 x')\<^esub>)" proof - from assms have "r1 x' \<^bsub>L1\<^esub>\ x'" by (intro t1.right_left_Galois_if_right_relI) auto with half_galois_prop_right2 have "((\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) \\<^sub>h (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>) (l2\<^bsub>x' (r1 x')\<^esub>)" by auto with assms show ?thesis by simp qed interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.counit \ \\<^sub>1" and "flip.t1.unit \ \\<^sub>1" by (simp_all only: t1.flip_counit_eq_unit t1.flip_unit_eq_counit) lemma galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI: assumes galois_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and preorder_L1: "preorder_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and mono_L2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" shows "([x1 x2 \ (\\<^bsub>L1\<^esub>) | \\<^sub>1 x2 \\<^bsub>L1\<^esub> x1] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x2 \\<^bsub>L1\<^esub> x3] \ (\)) L2" (is ?goal1) and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" (is ?goal2) proof - show ?goal1 proof (intro dep_mono_wrt_relI rel_if_if_impI Dep_Fun_Rel_relI) fix x1 x2 x3 x4 assume "x1 \\<^bsub>L1\<^esub> x2" moreover with galois_equiv1 preorder_L1 have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast intro: t1.rel_unit_if_reflexive_on_if_galois_connection) moreover assume "\\<^sub>1 x2 \\<^bsub>L1\<^esub> x1" ultimately have "x2 \\<^bsub>L1\<^esub> x1" using preorder_L1 by blast moreover assume "x3 \\<^bsub>L1\<^esub> x4" "x2 \\<^bsub>L1\<^esub> x3" - ultimately show "(\\<^bsub>L2 x1 x3\<^esub>) \ (\\<^bsub>L2 x2 x4\<^esub>)" using preorder_L1 mono_L2 by blast + ultimately show "(\\<^bsub>L2 x1 x3\<^esub>) \ (\\<^bsub>L2 x2 x4\<^esub>)" using preorder_L1 mono_L2 + by (intro le_relI) (blast dest!: rel_ifD elim!: dep_mono_wrt_relE) qed show ?goal2 proof (intro dep_mono_wrt_relI rel_if_if_impI Dep_Fun_Rel_relI) fix x1 x2 x3 x4 presume "x3 \\<^bsub>L1\<^esub> x4" "x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3" moreover with galois_equiv1 preorder_L1 have "\\<^sub>1 x3 \\<^bsub>L1\<^esub> x3" by (blast intro: flip.t1.counit_rel_if_reflexive_on_if_galois_connection) ultimately have "x3 \\<^bsub>L1\<^esub> x4" using preorder_L1 by blast moreover presume "x1 \\<^bsub>L1\<^esub> x2" "x2 \\<^bsub>L1\<^esub> x3" ultimately show "(\\<^bsub>L2 x2 x4\<^esub>) \ (\\<^bsub>L2 x1 x3\<^esub>)" using preorder_L1 mono_L2 by fast qed auto qed lemma galois_equivalence_if_mono_if_galois_equivalence_Dep_Fun_Rel_pred_assm_leftI: assumes galois_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and refl_R1: "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and mono_L2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and mono_R2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and mono_l2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "x \\<^bsub>L1\<^esub> x" shows "([in_codom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub>)" proof (intro Dep_Fun_Rel_predI) fix y assume "in_codom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) y" moreover from \x \\<^bsub>L1\<^esub> x\ galois_equiv1 refl_L1 have "x \\<^bsub>L1\<^esub> \\<^sub>1 x" by (blast intro: bi_related_if_rel_equivalence_on t1.rel_equivalence_on_unit_if_reflexive_on_if_galois_equivalence) moreover with refl_L1 have "\\<^sub>1 x \\<^bsub>L1\<^esub> \\<^sub>1 x" by blast ultimately have "in_codom (\\<^bsub>L2 (\\<^sub>1 x) (\\<^sub>1 x)\<^esub>) y" using mono_L2 by blast moreover from \x \\<^bsub>L1\<^esub> x\ galois_equiv1 have "l1 x \\<^bsub>R1\<^esub> l1 x" "\\<^sub>1 x \\<^bsub>L1\<^esub> x" "x \<^bsub>L1\<^esub>\ l1 x" by (blast intro: t1.left_Galois_left_if_left_relI flip.t1.counit_rel_if_right_rel_if_galois_connection)+ moreover note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_l2 \l1 x \\<^bsub>R1\<^esub> l1 x\] \\\<^sub>1 x \\<^bsub>L1\<^esub> x\] ultimately have "l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (\\<^sub>1 (l1 x)) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" by auto moreover note \l1 x \\<^bsub>R1\<^esub> l1 x\ moreover with galois_equiv1 refl_R1 have "l1 x \\<^bsub>R1\<^esub> \\<^sub>1 (l1 x)" by (blast intro: bi_related_if_rel_equivalence_on flip.t1.rel_equivalence_on_unit_if_reflexive_on_if_galois_equivalence) ultimately show "l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" using mono_R2 by blast qed lemma galois_equivalence_if_mono_if_galois_equivalence_Dep_Fun_Rel_pred_assm_right: assumes galois_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and refl_R1: "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and mono_L2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and mono_R2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "x' \\<^bsub>R1\<^esub> x'" shows "([in_dom (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>)" proof (intro Dep_Fun_Rel_predI) fix y assume "in_dom (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) y" moreover from \x' \\<^bsub>R1\<^esub> x'\ galois_equiv1 refl_R1 have "x' \\<^bsub>R1\<^esub> \\<^sub>1 x'" by (blast intro: bi_related_if_rel_equivalence_on flip.t1.rel_equivalence_on_unit_if_reflexive_on_if_galois_equivalence) moreover with refl_R1 have "\\<^sub>1 x' \\<^bsub>R1\<^esub> \\<^sub>1 x'" by blast ultimately have "in_dom (\\<^bsub>R2 (\\<^sub>1 x') (\\<^sub>1 x')\<^esub>) y" using mono_R2 by blast moreover from \x' \\<^bsub>R1\<^esub> x'\ galois_equiv1 have "r1 x' \\<^bsub>L1\<^esub> r1 x'" "x' \\<^bsub>R1\<^esub> \\<^sub>1 x'" "r1 x' \<^bsub>L1\<^esub>\ x'" by (blast intro: t1.right_left_Galois_if_right_relI flip.t1.rel_unit_if_left_rel_if_galois_connection)+ moreover note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \r1 x' \\<^bsub>L1\<^esub> r1 x'\] \x' \\<^bsub>R1\<^esub> \\<^sub>1 x'\] ultimately have "r2\<^bsub>(r1 x') x'\<^esub> y \\<^bsub>L2 (r1 x') (\\<^sub>1 (r1 x'))\<^esub> r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y" by auto moreover note \r1 x' \\<^bsub>L1\<^esub> r1 x'\ moreover with galois_equiv1 refl_R1 have "r1 x' \\<^bsub>L1\<^esub> \\<^sub>1 (r1 x')" by (blast intro: bi_related_if_rel_equivalence_on t1.rel_equivalence_on_unit_if_reflexive_on_if_galois_equivalence) ultimately show "r2\<^bsub>(r1 x') x'\<^esub> y \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y" using mono_L2 by blast qed end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin context begin interpretation flip : transport_Mono_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.counit \ \\<^sub>1" and "flip.t1.unit \ \\<^sub>1" by (simp_all only: t1.flip_counit_eq_unit t1.flip_unit_eq_counit) lemma galois_equivalence_if_galois_equivalenceI: assumes galois_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and galois_equiv2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>G (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) \ (\\<^bsub>L2 x x\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x2' x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) \ (\\<^bsub>R2 x' x'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ([in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ([in_codom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_dom (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" proof - from galois_equiv2 have "((\\<^bsub>L2 x (r1 x')\<^esub>) \ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" "((\\<^bsub>R2 (l1 x) x'\<^esub>) \<^sub>h\ (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>) (l2\<^bsub>x' x\<^esub>)" "((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>h (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>) (l2\<^bsub>x' x\<^esub>)" if "x \<^bsub>L1\<^esub>\ x'" for x x' using \x \<^bsub>L1\<^esub>\ x'\ by (blast elim: galois.galois_connectionE galois_prop.galois_propE)+ moreover from galois_equiv1 galois_equiv2 have "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \\<^sub>m (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>)" by (intro tdfr.mono_wrt_rel_left2_if_mono_wrt_rel_left2_if_left_GaloisI) auto moreover from galois_equiv1 galois_equiv2 have "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ((\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) \\<^sub>m (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>)" by (intro tdfr.mono_wrt_rel_right2_if_mono_wrt_rel_right2_if_left_GaloisI) (auto elim!: t1.galois_equivalenceE) moreover from galois_equiv1 refl_L1 have "\x. x \\<^bsub>L1\<^esub> x \ x \\<^bsub>L1\<^esub> \\<^sub>1 x" "\x'. x' \\<^bsub>R1\<^esub> x' \ x' \\<^bsub>R1\<^esub> \\<^sub>1 x'" by (blast intro!: bi_related_if_rel_equivalence_on t1.rel_equivalence_on_unit_if_reflexive_on_if_galois_equivalence flip.t1.rel_equivalence_on_unit_if_reflexive_on_if_galois_equivalence)+ ultimately show ?thesis using assms by (intro galois_equivalenceI galois_connection_left_right_if_galois_connectionI flip.galois_prop_left_rightI tdfr.flip_half_galois_prop_left2_if_half_galois_prop_left2_if_left_GaloisI tdfr.flip_half_galois_prop_right2_if_half_galois_prop_right2_if_GaloisI tdfr.mono_wrt_rel_left_if_transitiveI tdfr.mono_wrt_rel_right_if_transitiveI flip.tdfr.left_rel_right_if_left_right_rel_le_right2_assmI flip.tdfr.left_right_rel_if_left_rel_right_ge_left2_assmI tdfr.left_rel2_unit_eqs_left_rel2I flip.tdfr.left_rel2_unit_eqs_left_rel2I) (auto elim!: t1.galois_equivalenceE intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) qed corollary galois_equivalence_if_galois_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>G (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) \ (\\<^bsub>L2 x x\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x2' x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) \ (\\<^bsub>R2 x' x'\<^esub>)" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "\x. x \\<^bsub>L1\<^esub> x \ ([in_codom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_dom (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_equivalence_if_galois_equivalenceI tdfr.galois_connection_left_right_if_galois_connection_mono_assms_leftI tdfr.galois_connection_left_right_if_galois_connection_mono_assms_rightI tdfr.galois_connection_left_right_if_galois_connection_mono_2_assms_leftI tdfr.galois_connection_left_right_if_galois_connection_mono_2_assms_rightI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) corollary galois_equivalence_if_mono_if_galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>G (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>) | \\<^sub>1 x2 \\<^bsub>L1\<^esub> x1] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x2 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>) | \\<^sub>1 x2' \\<^bsub>R1\<^esub> x1'] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x2' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | (x2' \\<^bsub>R1\<^esub> x3' \ x4' \\<^bsub>R1\<^esub> \\<^sub>1 x3')] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "\x. x \\<^bsub>L1\<^esub> x \ ([in_codom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_dom (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_equivalence_if_galois_equivalenceI' tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI flip.tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI flip.tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI) auto end interpretation flip : transport_Mono_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.counit \ \\<^sub>1" and "flip.t1.unit \ \\<^sub>1" by (simp_all only: t1.flip_counit_eq_unit t1.flip_unit_eq_counit) lemma galois_equivalence_if_mono_if_preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>G (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_equivalence_if_mono_if_galois_equivalenceI tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI flip.tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI tdfr.galois_equivalence_if_mono_if_galois_equivalence_Dep_Fun_Rel_pred_assm_leftI tdfr.galois_equivalence_if_mono_if_galois_equivalence_Dep_Fun_Rel_pred_assm_right) auto theorem galois_equivalence_if_mono_if_preorder_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" shows "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_equivalence_if_mono_if_preorder_equivalenceI tdfr.transitive_left2_if_preorder_equivalenceI tdfr.transitive_right2_if_preorder_equivalenceI) auto end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin interpretation flip : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>G (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>L2\<^esub>)" and "transitive (\\<^bsub>R2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" using assms by (intro tpdfr.galois_equivalenceI galois_connection_left_rightI flip.galois_prop_left_rightI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions_Galois_Relator.thy b/thys/Transport/Transport/Functions/Transport_Functions_Galois_Relator.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Galois_Relator.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Galois_Relator.thy @@ -1,865 +1,865 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Relator\ theory Transport_Functions_Galois_Relator imports Transport_Functions_Relation_Simplifications begin paragraph \Dependent Function Relator\ context transport_Dep_Fun_Rel begin interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.counit \ \\<^sub>1" by (simp only: t1.flip_counit_eq_unit) lemma Dep_Fun_Rel_left_Galois_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and mono_r2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and L2_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and ge_L2_r2_le2: "\x x' y'. x \<^bsub>L1\<^esub>\ x' \ in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>) y' \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x (l1 x)\<^esub> y') \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x x'\<^esub> y')" and trans_L2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" and "f \<^bsub>L\<^esub>\ g" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" proof (intro Dep_Fun_Rel_relI) fix x x' assume "x \<^bsub>L1\<^esub>\ x'" show "f x \<^bsub>L2 x x'\<^esub>\ g x'" proof (intro t2.left_GaloisI) from \x \<^bsub>L1\<^esub>\ x'\ \((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1\ have "x \\<^bsub>L1\<^esub> r1 x'" "l1 x \\<^bsub>R1\<^esub> x'" by auto with \g \\<^bsub>R\<^esub> g\ have "g (l1 x) \\<^bsub>R2 (l1 x) x'\<^esub> g x'" by blast then show "in_codom (\\<^bsub>R2 (l1 x) x'\<^esub>) (g x')" by blast from \f \<^bsub>L\<^esub>\ g\ have "f \\<^bsub>L\<^esub> r g" by blast moreover from refl_L1 \x \<^bsub>L1\<^esub>\ x'\ have "x \\<^bsub>L1\<^esub> x" by blast ultimately have "f x \\<^bsub>L2 x x\<^esub> (r g) x" by blast with L2_le2 \x \\<^bsub>L1\<^esub> r1 x'\ have "f x \\<^bsub>L2 x (r1 x')\<^esub> (r g) x" by blast then have "f x \\<^bsub>L2 x (r1 x')\<^esub> r2\<^bsub>x (l1 x)\<^esub> (g (l1 x))" by simp with ge_L2_r2_le2 have "f x \\<^bsub>L2 x (r1 x')\<^esub> r2\<^bsub>x x'\<^esub> (g (l1 x))" using \x \<^bsub>L1\<^esub>\ x'\ \g (l1 x) \\<^bsub>R2 (l1 x) x'\<^esub> _\ by blast moreover have "... \\<^bsub>L2 x (r1 x')\<^esub> r2\<^bsub>x x'\<^esub> (g x')" using mono_r2 \x \<^bsub>L1\<^esub>\ x'\ \g (l1 x) \\<^bsub>R2 (l1 x) x'\<^esub> g x'\ by blast ultimately show "f x \\<^bsub>L2 x (r1 x')\<^esub> r2\<^bsub>x x'\<^esub> (g x')" using trans_L2 \x \<^bsub>L1\<^esub>\ x'\ by blast qed qed lemma left_rel_right_if_Dep_Fun_Rel_left_GaloisI: assumes mono_l1: "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and half_galois_prop_right1: "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and L2_unit_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and ge_L2_r2_le1: "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and rel_f_g: "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "f \\<^bsub>L\<^esub> r g" proof (intro left_relI) fix x1 x2 assume "x1 \\<^bsub>L1\<^esub> x2" with mono_l1 half_galois_prop_right1 have "x1 \<^bsub>L1\<^esub>\ l1 x2" by (intro t1.left_Galois_left_if_left_relI) auto with rel_f_g have "f x1 \<^bsub>L2 x1 (l1 x2)\<^esub>\ g (l1 x2)" by blast then have "in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) (g (l1 x2))" "f x1 \\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub> r2\<^bsub>x1 (l1 x2)\<^esub> (g (l1 x2))" by auto with L2_unit_le2 \x1 \\<^bsub>L1\<^esub> x2\ have "f x1 \\<^bsub>L2 x1 x2\<^esub> r2\<^bsub>x1 (l1 x2)\<^esub> (g (l1 x2))" by blast with ge_L2_r2_le1 \x1 \\<^bsub>L1\<^esub> x2\ \in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) (g (l1 x2))\ have "f x1 \\<^bsub>L2 x1 x2\<^esub> r2\<^bsub>x2 (l1 x2)\<^esub> (g (l1 x2))" by blast then show "f x1 \\<^bsub>L2 x1 x2\<^esub> r g x2" by simp qed lemma left_Galois_if_Dep_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and "in_codom (\\<^bsub>R\<^esub>) g" and "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro left_GaloisI left_rel_right_if_Dep_Fun_Rel_left_GaloisI) auto lemma left_right_rel_if_Dep_Fun_Rel_left_GaloisI: assumes mono_r1: "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and half_galois_prop_left2: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and R2_le1: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and R2_l2_le1: "\x1' x2' y. x1' \\<^bsub>R1\<^esub> x2' \ in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) y \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub> y) \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x1' (r1 x1')\<^esub> y)" and rel_f_g: "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "l f \\<^bsub>R\<^esub> g" proof (rule flip.left_relI) fix x1' x2' assume "x1' \\<^bsub>R1\<^esub> x2'" with mono_r1 have "r1 x1' \<^bsub>L1\<^esub>\ x2'" by blast with rel_f_g have "f (r1 x1') \<^bsub>L2 (r1 x1') x2'\<^esub>\ g x2'" by blast with half_galois_prop_left2[OF \x1' \\<^bsub>R1\<^esub> x2'\] have "l2\<^bsub>x2' (r1 x1')\<^esub> (f (r1 x1')) \\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub> g x2'" by auto with R2_le1 \x1' \\<^bsub>R1\<^esub> x2'\ have "l2\<^bsub>x2' (r1 x1')\<^esub> (f (r1 x1')) \\<^bsub>R2 x1' x2'\<^esub> g x2'" by blast with R2_l2_le1 \x1' \\<^bsub>R1\<^esub> x2'\ \f (r1 x1') \<^bsub>L2 (r1 x1') x2'\<^esub>\ _\ have "l2\<^bsub>x1' (r1 x1')\<^esub> (f (r1 x1')) \\<^bsub>R2 x1' x2'\<^esub> g x2'" by blast then show "l f x1' \\<^bsub>R2 x1' x2'\<^esub> g x2'" by simp qed lemma left_Galois_if_Dep_Fun_Rel_left_GaloisI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and "\x1' x2' y. x1' \\<^bsub>R1\<^esub> x2' \ in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) y \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub> y) \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x1' (r1 x1')\<^esub> y)" and "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro left_Galois_if_Dep_Fun_Rel_left_GaloisI in_codomI[where ?x="l f"]) (auto intro!: left_right_rel_if_Dep_Fun_Rel_left_GaloisI) lemma left_Galois_iff_Dep_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and "\x x' y'. x \<^bsub>L1\<^esub>\ x' \ in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>) y' \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x (l1 x)\<^esub> y') \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x x'\<^esub> y')" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro iffI) (auto intro!: Dep_Fun_Rel_left_Galois_if_left_GaloisI left_Galois_if_Dep_Fun_Rel_left_GaloisI) lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI: assumes "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" using assms by blast lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI': assumes "\x x'. x \<^bsub>L1\<^esub>\ x' \ ([in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>)] \ (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "\x x' y'. x \<^bsub>L1\<^esub>\ x' \ in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>) y' \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x (l1 x)\<^esub> y') \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x x'\<^esub> y')" using assms by blast lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_codom_rightI: assumes mono_l1: "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and half_galois_prop_right1: "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and L2_le_unit2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "x1 \\<^bsub>L1\<^esub> x2" shows "([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" proof (intro Dep_Fun_Rel_predI) from mono_l1 half_galois_prop_right1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "l1 x2 \\<^bsub>R1\<^esub> l1 x2" "x2 \<^bsub>L1\<^esub>\ l1 x2" by (blast intro: t1.left_Galois_left_if_left_relI)+ fix y' assume "in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y'" with Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x1 \\<^bsub>L1\<^esub> x2\] \l1 x2 \\<^bsub>R1\<^esub> l1 x2\] have "r2\<^bsub>x1 (l1 x2)\<^esub> y' \\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub> r2\<^bsub>x2 (l1 x2)\<^esub> y'" using \x2 \<^bsub>L1\<^esub>\ l1 x2\ by (auto dest: in_field_if_in_codom) with L2_le_unit2 \x1 \\<^bsub>L1\<^esub> x2\ show "r2\<^bsub>x1 (l1 x2)\<^esub> y' \\<^bsub>L2 x1 x2\<^esub> r2\<^bsub>x2 (l1 x2)\<^esub> y'" by blast qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_dom_rightI: assumes mono_l1: "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and half_galois_prop_right1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "x \<^bsub>L1\<^esub>\ x'" shows "([in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>)] \ (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>) (r2\<^bsub>x x'\<^esub>)" proof - from mono_l1 half_galois_prop_right1 refl_L1 \x \<^bsub>L1\<^esub>\ x'\ have "x \\<^bsub>L1\<^esub> x" "l1 x \\<^bsub>R1\<^esub> x'" "x \<^bsub>L1\<^esub>\ l1 x" by (auto intro!: t1.half_galois_prop_leftD t1.left_Galois_left_if_left_relI) with Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x \\<^bsub>L1\<^esub> x\] \l1 x \\<^bsub>R1\<^esub> x'\] show ?thesis by blast qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_GaloisI left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI' left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_dom_rightI left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_codom_rightI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI: assumes refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and mono_L2: "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "x1 \\<^bsub>L1\<^esub> x2" shows "(\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" proof - from refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "x1 \\<^bsub>L1\<^esub> x1" by blast with dep_mono_wrt_relD[OF dep_mono_wrt_predD[OF mono_L2] \x1 \\<^bsub>L1\<^esub> x2\] show "(\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" by auto qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assmI: assumes mono_l1: "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and half_galois_prop_right1: "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and antimono_L2: "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "x1 \\<^bsub>L1\<^esub> x2" shows "(\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" proof - from mono_l1 half_galois_prop_right1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast intro: t1.rel_unit_if_reflexive_on_if_half_galois_prop_right_if_mono_wrt_rel) with refl_L1 have "\\<^sub>1 x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by blast with dep_mono_wrt_relD[OF dep_mono_wrt_predD[OF antimono_L2] \x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2\] show "(\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" using \x1 \\<^bsub>L1\<^esub> x2\ by auto qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assmI left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom in_field_if_in_dom) corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI') auto interpretation flip_inv : galois "(\\<^bsub>R1\<^esub>)" "(\\<^bsub>L1\<^esub>)" r1 l1 . lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assm_if_galois_equivI: assumes galois_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and mono_L2: "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "x1 \\<^bsub>L1\<^esub> x2" shows "(\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" proof - from refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "x1 \\<^bsub>L1\<^esub> x1" by blast from galois_equiv1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "\\<^sub>1 x2 \\<^bsub>L1\<^esub> x2" by (intro flip.t1.counit_rel_if_reflexive_on_if_half_galois_prop_left_if_mono_wrt_rel) blast+ have "x1 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (rule t1.rel_unit_if_left_rel_if_mono_wrt_relI) (insert galois_equiv1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\, auto) with dep_mono_wrt_relD[OF dep_mono_wrt_predD[OF mono_L2] \\\<^sub>1 x2 \\<^bsub>L1\<^esub> x2\] show "(\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" by auto qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assm_if_galois_equivI) auto corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI reflexive_on_in_field_mono_assm_left2I) auto corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI') auto corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI transitive_left2_if_preorder_equivalenceI) (auto 5 0) subparagraph \Simplification of Restricted Function Relator\ lemma Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2' y. x1' \\<^bsub>R1\<^esub> x2' \ in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) y \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub> y) \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x1' (r1 x1')\<^esub> y)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))" - using assms by (intro ext iffI restrict_leftI restrict_rightI + using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI in_domI[where ?y="r _"] left_rel_right_if_Dep_Fun_Rel_left_GaloisI in_codomI[where ?x="l _"] left_right_rel_if_Dep_Fun_Rel_left_GaloisI) auto lemma Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI': assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>) | \\<^sub>1 x2' \\<^bsub>R1\<^esub> x1'] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x2' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))" using assms by (intro Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI reflexive_on_in_field_mono_assm_left2I left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI mono_wrt_rel_left_in_dom_mono_left_assm galois_connection_left_right_if_galois_connection_mono_assms_leftI galois_connection_left_right_if_galois_connection_mono_assms_rightI left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI) auto text \Simplification of Restricted Function Relator for Nested Transports\ lemma Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq: fixes S :: "'a1 \ 'a2 \ 'b1 \ 'b2 \ bool" assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (S x x')\\<^bsub>in_dom (\\<^bsub>L2 x (r1 x')\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R2 (l1 x) x'\<^esub>)\<^esub>) \\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ S x x')\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" (is "?lhs = ?rhs") proof - have "?lhs = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (S x x')\\<^bsub>in_codom (\\<^bsub>R2 (l1 x) x'\<^esub>)\<^esub>) \\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - by (subst restrict_left_right_eq_restrict_right_left, + by (subst bin_rel_restrict_left_right_eq_restrict_right_left, subst restrict_left_Dep_Fun_Rel_rel_restrict_left_eq) auto also have "... = ?rhs" - using assms by (subst restrict_left_right_eq_restrict_right_left, + using assms by (subst bin_rel_restrict_left_right_eq_restrict_right_left, subst restrict_right_Dep_Fun_Rel_rel_restrict_right_eq) (auto elim!: in_codomE t1.left_GaloisE - simp only: restrict_left_right_eq_restrict_right_left) + simp only: bin_rel_restrict_left_right_eq_restrict_right_left) finally show ?thesis . qed end paragraph \Function Relator\ context transport_Fun_Rel begin corollary Fun_Rel_left_Galois_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "transitive (\\<^bsub>L2\<^esub>)" and "g \\<^bsub>R\<^esub> g" and "f \<^bsub>L\<^esub>\ g" shows "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" using assms by (intro tdfr.Dep_Fun_Rel_left_Galois_if_left_GaloisI) simp_all corollary left_Galois_if_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "in_codom (\\<^bsub>R\<^esub>) g" and "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro tdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI) simp_all lemma left_Galois_if_Fun_Rel_left_GaloisI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \<^sub>h\ (\\<^bsub>R2\<^esub>)) l2 r2" and "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro tdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI') simp_all corollary left_Galois_iff_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "transitive (\\<^bsub>L2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" using assms by (intro tdfr.left_Galois_iff_Dep_Fun_Rel_left_GaloisI) simp_all subparagraph \Simplification of Restricted Function Relator\ lemma Fun_Rel_left_Galois_restrict_left_right_eq_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \<^sub>h\ (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))" using assms by (intro tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI) simp_all text \Simplification of Restricted Function Relator for Nested Transports\ lemma Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq: fixes S :: "'b1 \ 'b2 \ bool" assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" shows "((\<^bsub>L1\<^esub>\) \ S\\<^bsub>in_dom (\\<^bsub>L2\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R2\<^esub>)\<^esub>)\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ((\<^bsub>L1\<^esub>\) \ S)\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" using assms by (intro tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq) simp_all end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin lemma Dep_Fun_Rel_left_Galois_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x x' y'. x \<^bsub>L1\<^esub>\ x' \ in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>) y' \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x (l1 x)\<^esub> y') \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x x'\<^esub> y')" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "f \<^bsub>L\<^esub>\ g" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel by (intro tdfr.Dep_Fun_Rel_left_Galois_if_left_GaloisI tdfr.left_GaloisI) (auto elim!: galois_rel.left_GaloisE in_codomE) lemma left_Galois_if_Dep_Fun_Rel_left_GaloisI: assumes "(tdfr.R \\<^sub>m tdfr.L) r" and "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" and "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel by (intro tdfr.Galois_Refl_RelI tdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI) (auto simp: in_codom_eq_in_dom_if_reflexive_on_in_field) lemma left_Galois_iff_Dep_Fun_Rel_left_GaloisI: assumes "(tdfr.R \\<^sub>m tdfr.L) r" and "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro iffI Dep_Fun_Rel_left_Galois_if_left_GaloisI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI' tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_dom_rightI) (auto intro!: left_Galois_if_Dep_Fun_Rel_left_GaloisI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_codom_rightI intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI: assumes galois_conn1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and L2_le_unit2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and trans_L2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" (is "?lhs \ ?rhs") proof - have "(\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x1)\<^esub> y')" if hyps: "x1 \\<^bsub>L1\<^esub> x2" "in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y'" for x1 x2 y' proof - have "([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" proof (intro Dep_Fun_Rel_predI) from galois_conn1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "x1 \\<^bsub>L1\<^esub> x1" "l1 x1 \\<^bsub>R1\<^esub> l1 x2" "x1 \<^bsub>L1\<^esub>\ l1 x1" by (blast intro: t1.left_Galois_left_if_left_relI)+ fix y' assume "in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y'" with Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x1 \\<^bsub>L1\<^esub> x1\] \l1 x1 \\<^bsub>R1\<^esub> l1 x2\] have "r2\<^bsub>x1 (l1 x1)\<^esub> y' \\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub> r2\<^bsub>x1 (l1 x2)\<^esub> y'" using \x1 \<^bsub>L1\<^esub>\ l1 x1\ by (auto dest: in_field_if_in_dom) with L2_le_unit2 \x1 \\<^bsub>L1\<^esub> x2\ show "r2\<^bsub>x1 (l1 x1)\<^esub> y' \\<^bsub>L2 x1 x2\<^esub> r2\<^bsub>x1 (l1 x2)\<^esub> y'" by blast qed with hyps show ?thesis using trans_L2 by blast qed then show ?thesis using assms using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_GaloisI tdfr.mono_wrt_rel_rightI tdfr.mono_wrt_rel_right2_if_mono_wrt_rel_right2_if_left_GaloisI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_codom_rightI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom) qed corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI': assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" (is "?lhs \ ?rhs") using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assmI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI) auto corollary left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_mono_if_galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI restrict_leftI restrict_rightI + using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI iffD1[OF left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI']) (auto intro!: iffD2[OF left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI']) lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI tdfr.reflexive_on_in_field_mono_assm_left2I tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assm_if_galois_equivI) auto theorem left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI restrict_leftI restrict_rightI + using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI iffD1[OF left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI]) (auto intro!: iffD2[OF left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI]) corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI) auto corollary left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI restrict_leftI restrict_rightI + using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI iffD1[OF left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI]) (auto intro!: iffD2[OF left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI]) corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI tdfr.transitive_left2_if_preorder_equivalenceI) (auto 5 0) corollary left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_preorder_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" shows "(\<^bsub>L\<^esub>\) = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI restrict_leftI restrict_rightI + using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI iffD1[OF left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI']) (auto intro!: iffD2[OF left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI']) subparagraph \Simplification of Restricted Function Relator\ lemma Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_Galois_if_reflexive_onI: assumes "reflexive_on (in_field tdfr.L) tdfr.L" and "reflexive_on (in_field tdfr.R) tdfr.R" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>) | \\<^sub>1 x2' \\<^bsub>R1\<^esub> x1'] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x2' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))" using assms by (auto simp only: left_rel_eq_tdfr_left_rel_if_reflexive_on right_rel_eq_tdfr_right_rel_if_reflexive_on intro!: tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI') interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.unit \ \\<^sub>1" by (simp only: t1.flip_unit_eq_counit) lemma Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and PERS: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>L2 x1 x2\<^esub>)" "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ partial_equivalence_rel (\\<^bsub>R2 x1' x2'\<^esub>)" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))" using assms by (intro Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_Galois_if_reflexive_onI tdfr.reflexive_on_in_field_left_if_equivalencesI flip.reflexive_on_in_field_left_if_equivalencesI tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI flip.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI) (auto dest!: PERS) text \Simplification of Restricted Function Relator for Nested Transports\ lemma Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq: fixes S :: "'a1 \ 'a2 \ 'b1 \ 'b2 \ bool" assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (S x x')\\<^bsub>in_dom (\\<^bsub>L2 x (r1 x')\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R2 (l1 x) x'\<^esub>)\<^esub>) \\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ S x x')\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" (is "?lhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> = ?rhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub>") proof (intro ext) fix f g have "?lhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> f g \ ?lhs f g \ ?DL f \ ?CR g" by blast also have "... \ ?lhs\\<^bsub>in_dom tdfr.L\<^esub>\\<^bsub>in_codom tdfr.R\<^esub> f g \ ?DL f \ ?CR g" unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel by blast also with assms have "... \ ?rhs\\<^bsub>in_dom tdfr.L\<^esub>\\<^bsub>in_codom tdfr.R\<^esub> f g \ ?DL f \ ?CR g" by (simp only: tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq) also have "... \ ?rhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> f g" unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel by blast finally show "?lhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> f g \ ?rhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> f g" . qed end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin corollary Fun_Rel_left_Galois_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) (r2)" and "transitive (\\<^bsub>L2\<^esub>)" and "f \<^bsub>L\<^esub>\ g" shows "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" using assms by (intro tpdfr.Dep_Fun_Rel_left_Galois_if_left_GaloisI) simp_all interpretation flip : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma left_Galois_if_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" and "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro tpdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI flip.tfr.mono_wrt_rel_leftI) simp_all corollary left_Galois_iff_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) (r2)" and "transitive (\\<^bsub>L2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" using assms by (intro iffI Fun_Rel_left_Galois_if_left_GaloisI) (auto intro!: left_Galois_if_Fun_Rel_left_GaloisI) theorem left_Galois_eq_Fun_Rel_left_Galois_restrictI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "transitive (\\<^bsub>L2\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI restrict_leftI restrict_rightI + using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI iffD1[OF left_Galois_iff_Fun_Rel_left_GaloisI]) (auto elim!: tpdfr.left_GaloisE intro!: iffD2[OF left_Galois_iff_Fun_Rel_left_GaloisI]) subparagraph \Simplification of Restricted Function Relator\ lemma Fun_Rel_left_Galois_restrict_left_right_eq_Fun_Rel_left_Galois_if_reflexive_onI: assumes "reflexive_on (in_field tfr.tdfr.L) tfr.tdfr.L" and "reflexive_on (in_field tfr.tdfr.R) tfr.tdfr.R" and "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \<^sub>h\ (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))" using assms by (auto simp only: tpdfr.left_rel_eq_tdfr_left_rel_if_reflexive_on tpdfr.right_rel_eq_tdfr_right_rel_if_reflexive_on intro!: tfr.Fun_Rel_left_Galois_restrict_left_right_eq_Fun_Rel_left_GaloisI) lemma Fun_Rel_left_Galois_restrict_left_right_eq_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \<^sub>h\ (\\<^bsub>R2\<^esub>)) l2 r2" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" and "partial_equivalence_rel (\\<^bsub>R2\<^esub>)" shows "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))" using assms by (intro Fun_Rel_left_Galois_restrict_left_right_eq_Fun_Rel_left_Galois_if_reflexive_onI tfr.reflexive_on_in_field_leftI flip.tfr.reflexive_on_in_field_leftI) auto text \Simplification of Restricted Function Relator for Nested Transports\ lemma Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq: fixes S :: "'b1 \ 'b2 \ bool" assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" shows "((\<^bsub>L1\<^esub>\) \ S\\<^bsub>in_dom (\\<^bsub>L2\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R2\<^esub>)\<^esub>)\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ((\<^bsub>L1\<^esub>\) \ S)\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" using assms by (intro tpdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq) simp_all end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions_Order_Base.thy b/thys/Transport/Transport/Functions/Transport_Functions_Order_Base.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Order_Base.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Order_Base.thy @@ -1,428 +1,428 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Basic Order Properties\ theory Transport_Functions_Order_Base imports Transport_Functions_Base begin paragraph \Dependent Function Relator\ context hom_Dep_Fun_Rel_orders begin lemma reflexive_on_in_domI: assumes refl_L: "reflexive_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x2 x2\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and pequiv_R: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "reflexive_on (in_dom DFR) DFR" proof (intro reflexive_onI Dep_Fun_Rel_relI) fix f x1 x2 assume "in_dom DFR f" then obtain g where "DFR f g" by auto moreover assume "x1 \\<^bsub>L\<^esub> x2" moreover with refl_L have "x2 \\<^bsub>L\<^esub> x2" by blast ultimately have "f x1 \\<^bsub>R x1 x2\<^esub> g x2" "f x2 \\<^bsub>R x1 x2\<^esub> g x2" using R_le_R_if_L by auto moreover with pequiv_R \x1 \\<^bsub>L\<^esub> x2\ have "g x2 \\<^bsub>R x1 x2\<^esub> f x2" by (blast dest: symmetricD) ultimately show "f x1 \\<^bsub>R x1 x2\<^esub> f x2" using pequiv_R \x1 \\<^bsub>L\<^esub> x2\ by blast qed lemma reflexive_on_in_codomI: assumes refl_L: "reflexive_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x1\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and pequiv_R: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "reflexive_on (in_codom DFR) DFR" proof (intro reflexive_onI Dep_Fun_Rel_relI) fix f x1 x2 assume "in_codom DFR f" then obtain g where "DFR g f" by auto moreover assume "x1 \\<^bsub>L\<^esub> x2" moreover with refl_L have "x1 \\<^bsub>L\<^esub> x1" by blast ultimately have "g x1 \\<^bsub>R x1 x2\<^esub> f x2" "g x1 \\<^bsub>R x1 x2\<^esub> f x1" using R_le_R_if_L by auto moreover with pequiv_R \x1 \\<^bsub>L\<^esub> x2\ have "f x1 \\<^bsub>R x1 x2\<^esub> g x1" by (blast dest: symmetricD) ultimately show "f x1 \\<^bsub>R x1 x2\<^esub> f x2" using pequiv_R \x1 \\<^bsub>L\<^esub> x2\ by blast qed corollary reflexive_on_in_fieldI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x2 x2\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x1\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "reflexive_on (in_field DFR) DFR" proof - from assms have "reflexive_on (in_dom DFR) DFR" by (intro reflexive_on_in_domI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom) moreover from assms have "reflexive_on (in_codom DFR) DFR" by (intro reflexive_on_in_codomI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) ultimately show ?thesis by (auto iff: in_field_iff_in_dom_or_in_codom) qed lemma transitiveI: assumes refl_L: "reflexive_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x1\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and trans: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ transitive (\\<^bsub>R x1 x2\<^esub>)" shows "transitive DFR" proof (intro transitiveI Dep_Fun_Rel_relI) fix f1 f2 f3 x1 x2 assume "x1 \\<^bsub>L\<^esub> x2" with refl_L have "x1 \\<^bsub>L\<^esub> x1" by blast moreover assume "DFR f1 f2" ultimately have "f1 x1 \\<^bsub>R x1 x1\<^esub> f2 x1" by blast with R_le_R_if_L have "f1 x1 \\<^bsub>R x1 x2\<^esub> f2 x1" using \x1 \\<^bsub>L\<^esub> x2\ by blast assume "DFR f2 f3" with \x1 \\<^bsub>L\<^esub> x2\ have "f2 x1 \\<^bsub>R x1 x2\<^esub> f3 x2" by blast with \f1 x1 \\<^bsub>R x1 x2\<^esub> f2 x1\ show "f1 x1 \\<^bsub>R x1 x2\<^esub> f3 x2" using trans \x1 \\<^bsub>L\<^esub> x2\ by blast qed lemma transitiveI': assumes refl_L: "reflexive_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x2 x2\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and trans: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ transitive (\\<^bsub>R x1 x2\<^esub>)" shows "transitive DFR" proof (intro Binary_Relations_Transitive.transitiveI Dep_Fun_Rel_relI) fix f1 f2 f3 x1 x2 assume "DFR f1 f2" "x1 \\<^bsub>L\<^esub> x2" then have "f1 x1 \\<^bsub>R x1 x2\<^esub> f2 x2" by blast from \x1 \\<^bsub>L\<^esub> x2\ refl_L have "x2 \\<^bsub>L\<^esub> x2" by blast moreover assume "DFR f2 f3" ultimately have "f2 x2 \\<^bsub>R x2 x2\<^esub> f3 x2" by blast with R_le_R_if_L have "f2 x2 \\<^bsub>R x1 x2\<^esub> f3 x2" using \x1 \\<^bsub>L\<^esub> x2\ by blast with \f1 x1 \\<^bsub>R x1 x2\<^esub> f2 x2\ show "f1 x1 \\<^bsub>R x1 x2\<^esub> f3 x2" using trans \x1 \\<^bsub>L\<^esub> x2\ by blast qed lemma preorder_on_in_fieldI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x2 x2\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x1\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and pequiv_R: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "preorder_on (in_field DFR) DFR" using assms by (intro preorder_onI reflexive_on_in_fieldI) (auto intro!: transitiveI dest: pequiv_R elim!: partial_equivalence_relE) lemma symmetricI: assumes sym_L: "symmetric (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x2\<^esub>) \ (\\<^bsub>R x2 x1\<^esub>)" and sym_R: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ symmetric (\\<^bsub>R x1 x2\<^esub>)" shows "symmetric DFR" proof (intro symmetricI Dep_Fun_Rel_relI) fix f g x y assume "x \\<^bsub>L\<^esub> y" with sym_L have "y \\<^bsub>L\<^esub> x" by (rule symmetricD) moreover assume "DFR f g" ultimately have "f y \\<^bsub>R y x\<^esub> g x" by blast with sym_R \y \\<^bsub>L\<^esub> x\ have "g x \\<^bsub>R y x\<^esub> f y" by (blast dest: symmetricD) with R_le_R_if_L \y \\<^bsub>L\<^esub> x\ show "g x \\<^bsub>R x y\<^esub> f y" by blast qed corollary partial_equivalence_relI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and sym_L: "symmetric (\\<^bsub>L\<^esub>)" and mono_R: "([x1 x2 \ (\\<^bsub>L\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L\<^esub>) | x1 \\<^bsub>L\<^esub> x3] \ (\)) R" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "partial_equivalence_rel DFR" proof - have "(\\<^bsub>R x1 x2\<^esub>) \ (\\<^bsub>R x2 x1\<^esub>)" if "x1 \\<^bsub>L\<^esub> x2" for x1 x2 proof - from sym_L \x1 \\<^bsub>L\<^esub> x2\ have "x2 \\<^bsub>L\<^esub> x1" by (rule symmetricD) with mono_R \x1 \\<^bsub>L\<^esub> x2\ show ?thesis by blast qed with assms show ?thesis by (intro partial_equivalence_relI transitiveI symmetricI) - (auto elim: partial_equivalence_relE[OF assms(4)]) + (blast elim: partial_equivalence_relE[OF assms(4)])+ qed end context transport_Dep_Fun_Rel begin lemmas reflexive_on_in_field_leftI = dfro1.reflexive_on_in_fieldI [folded left_rel_eq_Dep_Fun_Rel] lemmas transitive_leftI = dfro1.transitiveI[folded left_rel_eq_Dep_Fun_Rel] lemmas transitive_leftI' = dfro1.transitiveI'[folded left_rel_eq_Dep_Fun_Rel] lemmas preorder_on_in_field_leftI = dfro1.preorder_on_in_fieldI [folded left_rel_eq_Dep_Fun_Rel] lemmas symmetric_leftI = dfro1.symmetricI[folded left_rel_eq_Dep_Fun_Rel] lemmas partial_equivalence_rel_leftI = dfro1.partial_equivalence_relI [folded left_rel_eq_Dep_Fun_Rel] subparagraph \Introduction Rules for Assumptions\ lemma transitive_left2_if_transitive_left2_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and L2_eq: "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ transitive (\\<^bsub>L2 x (r1 x')\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "transitive (\\<^bsub>L2 x1 x2\<^esub>)" by (subst L2_eq) (auto intro!: assms t1.left_Galois_left_if_left_relI) lemma symmetric_left2_if_symmetric_left2_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and L2_eq: "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ symmetric (\\<^bsub>L2 x (r1 x')\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "symmetric (\\<^bsub>L2 x1 x2\<^esub>)" by (subst L2_eq) (auto intro!: assms t1.left_Galois_left_if_left_relI) lemma partial_equivalence_rel_left2_if_partial_equivalence_rel_left2_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and L2_eq: "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ partial_equivalence_rel (\\<^bsub>L2 x (r1 x')\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "partial_equivalence_rel (\\<^bsub>L2 x1 x2\<^esub>)" by (subst L2_eq) (auto intro!: assms t1.left_Galois_left_if_left_relI) context assumes galois_prop: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" begin interpretation flip_inv : transport_Dep_Fun_Rel "(\\<^bsub>R1\<^esub>)" "(\\<^bsub>L1\<^esub>)" r1 l1 "flip2 R2" "flip2 L2" r2 l2 rewrites "flip_inv.t1.unit \ \\<^sub>1" and "\R x y. (flip2 R x y) \ (R y x)\" and "\R S. R\ = S\ \ R = S" and "\R S. (R\ \\<^sub>m S\) \ (R \\<^sub>m S)" and "\x x'. x' \<^bsub>R1\<^esub>\ x \ x \<^bsub>L1\<^esub>\ x'" and "((\\<^bsub>R1\<^esub>) \\<^sub>h (\\<^bsub>L1\<^esub>)) r1 l1 \ True" and "\R. transitive R\ \ transitive R" and "\R. symmetric R\ \ symmetric R" and "\R. partial_equivalence_rel R\ \ partial_equivalence_rel R" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" using galois_prop by (auto intro!: Eq_TrueI simp add: t1.flip_unit_eq_counit galois_prop.half_galois_prop_right_rel_inv_iff_half_galois_prop_left simp del: rel_inv_iff_rel) lemma transitive_right2_if_transitive_right2_if_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "(\\<^bsub>R2 x1 x2\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1) x2\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ transitive (\\<^bsub>R2 (l1 x) x'\<^esub>)" and "x1 \\<^bsub>R1\<^esub> x2" shows "transitive (\\<^bsub>R2 x1 x2\<^esub>)" using galois_prop assms by (intro flip_inv.transitive_left2_if_transitive_left2_if_left_GaloisI [simplified rel_inv_iff_rel, of x1]) auto lemma symmetric_right2_if_symmetric_right2_if_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "(\\<^bsub>R2 x1 x2\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1) x2\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ symmetric (\\<^bsub>R2 (l1 x) x'\<^esub>)" and "x1 \\<^bsub>R1\<^esub> x2" shows "symmetric (\\<^bsub>R2 x1 x2\<^esub>)" using galois_prop assms by (intro flip_inv.symmetric_left2_if_symmetric_left2_if_left_GaloisI [simplified rel_inv_iff_rel, of x1]) auto lemma partial_equivalence_rel_right2_if_partial_equivalence_rel_right2_if_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "(\\<^bsub>R2 x1 x2\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1) x2\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ partial_equivalence_rel (\\<^bsub>R2 (l1 x) x'\<^esub>)" and "x1 \\<^bsub>R1\<^esub> x2" shows "partial_equivalence_rel (\\<^bsub>R2 x1 x2\<^esub>)" using galois_prop assms by (intro flip_inv.partial_equivalence_rel_left2_if_partial_equivalence_rel_left2_if_left_GaloisI [simplified rel_inv_iff_rel, of x1]) auto end lemma transitive_left2_if_preorder_equivalenceI: assumes pre_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "transitive (\\<^bsub>L2 x1 x2\<^esub>)" proof - from \x1 \\<^bsub>L1\<^esub> x2\ pre_equiv1 have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" by (intro left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro transitive_left2_if_transitive_left2_if_left_GaloisI[of x1]) blast+ qed lemma symmetric_left2_if_partial_equivalence_rel_equivalenceI: assumes PER_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "symmetric (\\<^bsub>L2 x1 x2\<^esub>)" proof - from \x1 \\<^bsub>L1\<^esub> x2\ PER_equiv1 have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" by (intro left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro symmetric_left2_if_symmetric_left2_if_left_GaloisI[of x1]) blast+ qed lemma partial_equivalence_rel_left2_if_partial_equivalence_rel_equivalenceI: assumes PER_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "partial_equivalence_rel (\\<^bsub>L2 x1 x2\<^esub>)" proof - from \x1 \\<^bsub>L1\<^esub> x2\ PER_equiv1 have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" by (intro left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro partial_equivalence_rel_left2_if_partial_equivalence_rel_left2_if_left_GaloisI[of x1]) blast+ qed interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.counit \ \\<^sub>1" and "flip.t1.unit \ \\<^sub>1" by (simp_all only: t1.flip_counit_eq_unit t1.flip_unit_eq_counit) lemma transitive_right2_if_preorder_equivalenceI: assumes pre_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1' \\<^bsub>R1\<^esub> x2'" shows "transitive (\\<^bsub>R2 x1' x2'\<^esub>)" proof - from \x1' \\<^bsub>R1\<^esub> x2'\ pre_equiv1 have "x1' \\<^bsub>R1\<^esub> \\<^sub>1 x1'" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>R2 x1' x2'\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)" by (intro flip.left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro transitive_right2_if_transitive_right2_if_left_GaloisI[of x1']) blast+ qed lemma symmetric_right2_if_partial_equivalence_rel_equivalenceI: assumes PER_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1' \\<^bsub>R1\<^esub> x2'" shows "symmetric (\\<^bsub>R2 x1' x2'\<^esub>)" proof - from \x1' \\<^bsub>R1\<^esub> x2'\ PER_equiv1 have "x1' \\<^bsub>R1\<^esub> \\<^sub>1 x1'" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>R2 x1' x2'\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)" by (intro flip.left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro symmetric_right2_if_symmetric_right2_if_left_GaloisI[of x1']) blast+ qed lemma partial_equivalence_rel_right2_if_partial_equivalence_rel_equivalenceI: assumes PER_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1' \\<^bsub>R1\<^esub> x2'" shows "partial_equivalence_rel (\\<^bsub>R2 x1' x2'\<^esub>)" proof - from \x1' \\<^bsub>R1\<^esub> x2'\ PER_equiv1 have "x1' \\<^bsub>R1\<^esub> \\<^sub>1 x1'" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>R2 x1' x2'\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)" by (intro flip.left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro partial_equivalence_rel_right2_if_partial_equivalence_rel_right2_if_left_GaloisI[of x1']) blast+ qed end paragraph \Function Relator\ context transport_Fun_Rel begin lemma reflexive_on_in_field_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" shows "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.reflexive_on_in_field_leftI) simp_all lemma transitive_leftI: assumes "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L2\<^esub>)" shows "transitive (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.transitive_leftI) simp_all lemma transitive_leftI': assumes "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L2\<^esub>)" shows "transitive (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.transitive_leftI') simp_all lemma preorder_on_in_field_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" shows "preorder_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.preorder_on_in_field_leftI) simp_all lemma symmetric_leftI: assumes "symmetric (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>L2\<^esub>)" shows "symmetric (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.symmetric_leftI) simp_all corollary partial_equivalence_rel_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>L1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" shows "partial_equivalence_rel (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.partial_equivalence_rel_leftI) auto end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin lemmas reflexive_on_in_field_leftI = Refl_Rel_reflexive_on_in_field[of tdfr.L, folded left_rel_eq_tdfr_left_Refl_Rel] lemmas transitive_leftI = Refl_Rel_transitiveI [of tdfr.L, folded left_rel_eq_tdfr_left_Refl_Rel] lemmas preorder_on_in_field_leftI = Refl_Rel_preorder_on_in_fieldI[of tdfr.L, folded left_rel_eq_tdfr_left_Refl_Rel] lemmas symmetric_leftI = Refl_Rel_symmetricI[of tdfr.L, OF tdfr.symmetric_leftI, folded left_rel_eq_tdfr_left_Refl_Rel] lemmas partial_equivalence_rel_leftI = Refl_Rel_partial_equivalence_relI[of tdfr.L, OF tdfr.partial_equivalence_rel_leftI, folded left_rel_eq_tdfr_left_Refl_Rel] end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin lemma symmetric_leftI: assumes "symmetric (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>L2\<^esub>)" shows "symmetric (\\<^bsub>L\<^esub>)" using assms by (intro tpdfr.symmetric_leftI) auto lemma partial_equivalence_rel_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>L1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" shows "partial_equivalence_rel (\\<^bsub>L\<^esub>)" using assms by (intro tpdfr.partial_equivalence_rel_leftI) auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions_Order_Equivalence.thy b/thys/Transport/Transport/Functions/Transport_Functions_Order_Equivalence.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Order_Equivalence.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Order_Equivalence.thy @@ -1,725 +1,726 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Order Equivalence\ theory Transport_Functions_Order_Equivalence imports Transport_Functions_Monotone Transport_Functions_Galois_Equivalence begin paragraph \Dependent Function Relator\ context transport_Dep_Fun_Rel begin subparagraph \Inflationary\ lemma rel_unit_self_if_rel_selfI: assumes inflationary_unit1: "inflationary_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and refl_L1: "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and trans_L1: "transitive (\\<^bsub>L1\<^esub>)" and mono_l2: "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" and mono_r2: "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and inflationary_unit2: "\x. x \\<^bsub>L1\<^esub> x \ inflationary_on (in_codom (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and L2_le1: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and L2_unit_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and ge_R2_l2_le2: "\x y. x \\<^bsub>L1\<^esub> x \ in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and trans_L2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "f \\<^bsub>L\<^esub> f" shows "f \\<^bsub>L\<^esub> \ f" proof (intro left_relI) fix x1 x2 assume [iff]: "x1 \\<^bsub>L1\<^esub> x2" moreover with inflationary_unit1 have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by blast ultimately have "x1 \\<^bsub>L1\<^esub> \\<^sub>1 x2" using trans_L1 by blast with \f \\<^bsub>L\<^esub> f\ have "f x1 \\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub> f (\\<^sub>1 x2)" by blast with L2_unit_le2 have "f x1 \\<^bsub>L2 x1 x2\<^esub> f (\\<^sub>1 x2)" by blast moreover have "... \\<^bsub>L2 x1 x2\<^esub> \ f x2" proof - from refl_L1 \x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2\ have "\\<^sub>1 x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by blast with \f \\<^bsub>L\<^esub> f\ have "f (\\<^sub>1 x2) \\<^bsub>L2 (\\<^sub>1 x2) (\\<^sub>1 x2)\<^esub> f (\\<^sub>1 x2)" by blast with L2_le1 have "f (\\<^sub>1 x2) \\<^bsub>L2 x2 (\\<^sub>1 x2)\<^esub> f (\\<^sub>1 x2)" using \x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2\ by blast moreover from refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have [iff]: "x2 \\<^bsub>L1\<^esub> x2" by blast ultimately have "f (\\<^sub>1 x2) \\<^bsub>L2 x2 x2\<^esub> f (\\<^sub>1 x2)" using L2_unit_le2 by blast with inflationary_unit2 have "f (\\<^sub>1 x2) \\<^bsub>L2 x2 x2\<^esub> \\<^bsub>2 x2 (l1 x2)\<^esub> (f (\\<^sub>1 x2))" by blast moreover have "... \\<^bsub>L2 x2 x2\<^esub> \ f x2" proof - from \f (\\<^sub>1 x2) \\<^bsub>L2 x2 x2\<^esub> f (\\<^sub>1 x2)\ mono_l2 have "l2\<^bsub>(l1 x2) x2\<^esub> (f (\\<^sub>1 x2)) \\<^bsub>R2 (l1 x2) (l1 x2)\<^esub> l2\<^bsub>(l1 x2) x2\<^esub> (f (\\<^sub>1 x2))" by blast with ge_R2_l2_le2 have "l2\<^bsub>(l1 x2) x2\<^esub> (f (\\<^sub>1 x2)) \\<^bsub>R2 (l1 x2) (l1 x2)\<^esub> l2\<^bsub>(l1 x2) (\\<^sub>1 x2)\<^esub> (f (\\<^sub>1 x2))" using \f (\\<^sub>1 x2) \\<^bsub>L2 x2 (\\<^sub>1 x2)\<^esub> f (\\<^sub>1 x2)\ by blast with mono_r2 have "\\<^bsub>2 x2 (l1 x2)\<^esub> (f (\\<^sub>1 x2)) \\<^bsub>L2 x2 (\\<^sub>1 x2)\<^esub> \ f x2" by auto with L2_unit_le2 show ?thesis by blast qed ultimately have "f (\\<^sub>1 x2) \\<^bsub>L2 x2 x2\<^esub> \ f x2" using trans_L2 by blast with L2_le1 show ?thesis by blast qed ultimately show "f x1 \\<^bsub>L2 x1 x2\<^esub> \ f x2" using trans_L2 by blast qed subparagraph \Deflationary\ interpretation flip_inv : transport_Dep_Fun_Rel "(\\<^bsub>R1\<^esub>)" "(\\<^bsub>L1\<^esub>)" r1 l1 "flip2 R2" "flip2 L2" r2 l2 rewrites "flip_inv.L \ (\\<^bsub>R\<^esub>)" and "flip_inv.R \ (\\<^bsub>L\<^esub>)" and "flip_inv.unit \ \" and "flip_inv.t1.unit \ \\<^sub>1" and "\x y. flip_inv.t2_unit x y \ \\<^bsub>2 y x\<^esub>" and "\R x y. (flip2 R x y)\ \ R y x" and "\R. in_codom R\ \ in_dom R" and "\R x1 x2. in_codom (flip2 R x1 x2) \ in_dom (R x2 x1)" and "\x1 x2 x1' x2'. (flip2 R2 x1' x2' \\<^sub>m flip2 L2 x1 x2) \ ((\\<^bsub>R2 x2' x1'\<^esub>) \\<^sub>m (\\<^bsub>L2 x2 x1\<^esub>))" and "\x1 x2 x1' x2'. (flip2 L2 x1 x2 \\<^sub>m flip2 R2 x1' x2') \ ((\\<^bsub>L2 x2 x1\<^esub>) \\<^sub>m (\\<^bsub>R2 x2' x1'\<^esub>))" and "\P. inflationary_on P (\\<^bsub>R1\<^esub>) \ deflationary_on P (\\<^bsub>R1\<^esub>)" and "\P x. inflationary_on P (flip2 R2 x x) \ deflationary_on P (\\<^bsub>R2 x x\<^esub>)" and "\x1 x2 x3 x4. flip2 R2 x1 x2 \ flip2 R2 x3 x4 \ (\\<^bsub>R2 x2 x1\<^esub>) \ (\\<^bsub>R2 x4 x3\<^esub>)" and "\(R :: 'z \ _) (P :: 'z \ bool). reflexive_on P R\ \ reflexive_on P R" and "\R. transitive R\ \ transitive R" and "\x1' x2'. transitive (flip2 R2 x1' x2') \ transitive (\\<^bsub>R2 x2' x1'\<^esub>)" by (simp_all add: flip_inv_left_eq_ge_right flip_inv_right_eq_ge_left flip_unit_eq_counit t1.flip_unit_eq_counit t2.flip_unit_eq_counit galois_prop.rel_inv_half_galois_prop_right_eq_half_galois_prop_left_rel_inv) lemma counit_rel_self_if_rel_selfI: assumes "deflationary_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "transitive (\\<^bsub>R1\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \\<^sub>m (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)) (l2\<^bsub> x' (r1 x')\<^esub>)" and "\x' x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>R2 x' x'\<^esub>) \\<^sub>m (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ deflationary_on (in_dom (\\<^bsub>R2 x' x'\<^esub>)) (\\<^bsub>R2 x' x'\<^esub>) (\\<^bsub>2 (r1 x') x'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_dom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub> y') \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y')" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "\ g \\<^bsub>R\<^esub> g" using assms by (intro flip_inv.rel_unit_self_if_rel_selfI[simplified rel_inv_iff_rel]) subparagraph \Relational Equivalence\ lemma bi_related_unit_self_if_rel_self_aux: assumes rel_equiv_unit1: "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and mono_r2: "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and rel_equiv_unit2: "\x. x \\<^bsub>L1\<^esub> x \ rel_equivalence_on (in_field (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and L2_le1: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and L2_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and [iff]: "x \\<^bsub>L1\<^esub> x" shows "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "deflationary_on (in_dom (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) \\<^bsub>2 x (l1 x)\<^esub>" and "inflationary_on (in_codom (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) \\<^bsub>2 x (l1 x)\<^esub>" proof - from rel_equiv_unit1 have "x \\<^bsub>L1\<^esub> \\<^sub>1 x" by blast with mono_r2 show "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" using L2_le1 L2_le2 by blast+ qed (insert rel_equiv_unit2, blast+) interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.counit \ \" and "flip.t1.counit \ \\<^sub>1" and "\x y. flip.t2_counit x y \ \\<^bsub>2 y x\<^esub>" by (simp_all add: order_functors.flip_counit_eq_unit) lemma bi_related_unit_self_if_rel_selfI: assumes rel_equiv_unit1: "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and trans_L1: "transitive (\\<^bsub>L1\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ rel_equivalence_on (in_field (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 (\\<^sub>1 x1) x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_dom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "f \\<^bsub>L\<^esub> f" shows "f \\<^bsub>L\<^esub> \ f" proof - from rel_equiv_unit1 trans_L1 have "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" by (intro reflexive_on_in_field_if_transitive_if_rel_equivalence_on) with assms show ?thesis by (intro bi_relatedI rel_unit_self_if_rel_selfI flip.counit_rel_self_if_rel_selfI bi_related_unit_self_if_rel_self_aux) (auto intro: inflationary_on_if_le_pred_if_inflationary_on deflationary_on_if_le_pred_if_deflationary_on reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) qed subparagraph \Lemmas for Monotone Function Relator\ lemma order_equivalence_if_order_equivalence_mono_assms_leftI: assumes order_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and refl_R1: "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and R2_counit_le1: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and mono_l2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and [iff]: "x1' \\<^bsub>R1\<^esub> x2'" shows "([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" and "([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" proof - from refl_R1 have "x1' \\<^bsub>R1\<^esub> x1'" "x2' \\<^bsub>R1\<^esub> x2'" by auto moreover with order_equiv1 have "r1 x1' \\<^bsub>L1\<^esub> r1 x2'" "r1 x1' \\<^bsub>L1\<^esub> r1 x1'" "r1 x2' \\<^bsub>L1\<^esub> r1 x2'" by auto ultimately have "r1 x1' \<^bsub>L1\<^esub>\ x1'" "r1 x2' \<^bsub>L1\<^esub>\ x2'" by blast+ note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_l2 \x1' \\<^bsub>R1\<^esub> x2'\] \r1 x1' \\<^bsub>L1\<^esub> r1 x1'\] with \r1 x1' \<^bsub>L1\<^esub>\ x1'\ R2_counit_le1 show "([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" by (intro Dep_Fun_Rel_predI) (auto dest!: in_field_if_in_dom) note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_l2 \x2' \\<^bsub>R1\<^esub> x2'\] \r1 x1' \\<^bsub>L1\<^esub> r1 x2'\] with \r1 x2' \<^bsub>L1\<^esub>\ x2'\ R2_counit_le1 show "([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" by (intro Dep_Fun_Rel_predI) (auto dest!: in_field_if_in_codom) qed lemma order_equivalence_if_order_equivalence_mono_assms_rightI: assumes order_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and L2_unit_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and [iff]: "x1 \\<^bsub>L1\<^esub> x2" shows "([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" proof - from refl_L1 have "x1 \\<^bsub>L1\<^esub> x1" "x2 \\<^bsub>L1\<^esub> x2" by auto moreover with order_equiv1 have "l1 x1 \\<^bsub>R1\<^esub> l1 x2" "l1 x1 \\<^bsub>R1\<^esub> l1 x1" "l1 x2 \\<^bsub>R1\<^esub> l1 x2" by auto ultimately have "x1 \<^bsub>L1\<^esub>\ l1 x1" "x2 \<^bsub>L1\<^esub>\ l1 x2" using order_equiv1 by (auto intro!: t1.left_Galois_left_if_in_codom_if_inflationary_onI) note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x1 \\<^bsub>L1\<^esub> x2\] \l1 x2 \\<^bsub>R1\<^esub> l1 x2\] with \x2 \<^bsub>L1\<^esub>\ l1 x2\ L2_unit_le2 show "([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" by (intro Dep_Fun_Rel_predI) (auto dest!: in_field_if_in_codom) note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x1 \\<^bsub>L1\<^esub> x1\] \l1 x1 \\<^bsub>R1\<^esub> l1 x2\] with \x1 \<^bsub>L1\<^esub>\ l1 x1\ L2_unit_le2 show "([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" by (intro Dep_Fun_Rel_predI) (auto dest!: in_field_if_in_dom) qed lemma l2_unit_bi_rel_selfI: assumes pre_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and mono_L2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and mono_R2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | (x2' \\<^bsub>R1\<^esub> x3' \ x4' \\<^bsub>R1\<^esub> \\<^sub>1 x3')] \ (\)) R2" and mono_l2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "x \\<^bsub>L1\<^esub> x" and "in_field (\\<^bsub>L2 x x\<^esub>) y" shows "l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" proof (rule bi_relatedI) note t1.preorder_equivalence_order_equivalenceE[elim!] from \x \\<^bsub>L1\<^esub> x\ pre_equiv1 have "l1 x \\<^bsub>R1\<^esub> l1 x" "x \\<^bsub>L1\<^esub> \\<^sub>1 x" "\\<^sub>1 x \\<^bsub>L1\<^esub> x" by blast+ with pre_equiv1 have "x \<^bsub>L1\<^esub>\ l1 x" "\\<^sub>1 x \<^bsub>L1\<^esub>\ l1 x" by (auto 4 3) from pre_equiv1 \x \\<^bsub>L1\<^esub> \\<^sub>1 x\ have "x \\<^bsub>L1\<^esub> \\<^sub>1 (\\<^sub>1 x)" by fastforce moreover note \in_field (\\<^bsub>L2 x x\<^esub>) y\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_L2 \\\<^sub>1 x \\<^bsub>L1\<^esub> x\] \\\<^sub>1 x \\<^bsub>L1\<^esub> x\] Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_L2 \x \\<^bsub>L1\<^esub> x\] \\\<^sub>1 x \\<^bsub>L1\<^esub> x\] ultimately have "in_field (\\<^bsub>L2 (\\<^sub>1 x) (\\<^sub>1 x)\<^esub>) y" "in_field (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y" using \x \\<^bsub>L1\<^esub> \\<^sub>1 x\ by blast+ moreover note \x \<^bsub>L1\<^esub>\ l1 x\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_l2 \l1 x \\<^bsub>R1\<^esub> l1 x\] \\\<^sub>1 x \\<^bsub>L1\<^esub> x\] ultimately have "l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (\\<^sub>1 (l1 x)) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" by auto moreover from pre_equiv1 \l1 x \\<^bsub>R1\<^esub> l1 x\ have "\\<^sub>1 (l1 x) \\<^bsub>R1\<^esub> l1 x" "l1 x \\<^bsub>R1\<^esub> \\<^sub>1 (l1 x)" by fastforce+ moreover note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD [OF mono_R2 \l1 x \\<^bsub>R1\<^esub> \\<^sub>1 (l1 x)\] \l1 x \\<^bsub>R1\<^esub> l1 x\] ultimately show "l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" by blast note \\\<^sub>1 x \<^bsub>L1\<^esub>\ l1 x\ \in_field (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_l2 \l1 x \\<^bsub>R1\<^esub> l1 x\] \x \\<^bsub>L1\<^esub> \\<^sub>1 x\] then show "l2\<^bsub>(l1 x) x\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y" by auto qed lemma r2_counit_bi_rel_selfI: assumes pre_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and mono_L2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and mono_R2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | (x2' \\<^bsub>R1\<^esub> x3' \ x4' \\<^bsub>R1\<^esub> \\<^sub>1 x3')] \ (\)) R2" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "x' \\<^bsub>R1\<^esub> x'" and "in_field (\\<^bsub>R2 x' x'\<^esub>) y'" shows "r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') x'\<^esub> y'" proof (rule bi_relatedI) note t1.preorder_equivalence_order_equivalenceE[elim!] from \x' \\<^bsub>R1\<^esub> x'\ pre_equiv1 have "r1 x' \\<^bsub>L1\<^esub> r1 x'" "x' \\<^bsub>R1\<^esub> \\<^sub>1 x'" "\\<^sub>1 x' \\<^bsub>R1\<^esub> x'" by blast+ - with pre_equiv1 have "r1 x' \<^bsub>L1\<^esub>\ x'" "r1 x' \<^bsub>L1\<^esub>\ \\<^sub>1 x'" by auto + with pre_equiv1 have "r1 x' \<^bsub>L1\<^esub>\ x'" "r1 x' \<^bsub>L1\<^esub>\ \\<^sub>1 x'" by force+ from pre_equiv1 \x' \\<^bsub>R1\<^esub> \\<^sub>1 x'\ have "x' \\<^bsub>R1\<^esub> \\<^sub>1 (\\<^sub>1 x')" by fastforce moreover note \in_field (\\<^bsub>R2 x' x'\<^esub>) y'\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_R2 \\\<^sub>1 x' \\<^bsub>R1\<^esub> x'\] \\\<^sub>1 x' \\<^bsub>R1\<^esub> x'\] Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_R2 \\\<^sub>1 x' \\<^bsub>R1\<^esub> x'\] \x' \\<^bsub>R1\<^esub> x'\] ultimately have "in_field (\\<^bsub>R2 (\\<^sub>1 x') (\\<^sub>1 x')\<^esub>) y'" "in_field (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y'" using \x' \\<^bsub>R1\<^esub> \\<^sub>1 x'\ \x' \\<^bsub>R1\<^esub> x'\ by blast+ moreover note \r1 x' \<^bsub>L1\<^esub>\ \\<^sub>1 x'\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \r1 x' \\<^bsub>L1\<^esub> r1 x'\] \\\<^sub>1 x' \\<^bsub>R1\<^esub> x'\] ultimately show "r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') x'\<^esub> y'" by auto note \r1 x' \<^bsub>L1\<^esub>\ x'\ \in_field (\\<^bsub>R2 (\\<^sub>1 x') (\\<^sub>1 x')\<^esub>) y'\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \r1 x' \\<^bsub>L1\<^esub> r1 x'\] \x' \\<^bsub>R1\<^esub> \\<^sub>1 x'\] then have "r2\<^bsub>(r1 x') x'\<^esub> y' \\<^bsub>L2 (r1 x') (\\<^sub>1 (r1 x'))\<^esub> r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y'" by auto moreover from pre_equiv1 \r1 x' \\<^bsub>L1\<^esub> r1 x'\ have "\\<^sub>1 (r1 x') \\<^bsub>L1\<^esub> r1 x'" "r1 x' \\<^bsub>L1\<^esub> \\<^sub>1 (r1 x')" by fastforce+ moreover note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD [OF mono_L2 \r1 x' \\<^bsub>L1\<^esub> r1 x'\] \r1 x' \\<^bsub>L1\<^esub> \\<^sub>1 (r1 x')\] ultimately show "r2\<^bsub>(r1 x') x'\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y'" using pre_equiv1 by blast qed end paragraph \Function Relator\ context transport_Fun_Rel begin corollary rel_unit_self_if_rel_selfI: assumes "inflationary_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "inflationary_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>L2\<^esub>)" and "f \\<^bsub>L\<^esub> f" shows "f \\<^bsub>L\<^esub> \ f" using assms by (intro tdfr.rel_unit_self_if_rel_selfI) simp_all corollary counit_rel_self_if_rel_selfI: assumes "deflationary_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "transitive (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "deflationary_on (in_dom (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>R2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "\ g \\<^bsub>R\<^esub> g" using assms by (intro tdfr.counit_rel_self_if_rel_selfI) simp_all lemma bi_related_unit_self_if_rel_selfI: assumes "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "rel_equivalence_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>L2\<^esub>)" and "f \\<^bsub>L\<^esub> f" shows "f \\<^bsub>L\<^esub> \ f" using assms by (intro tdfr.bi_related_unit_self_if_rel_selfI) simp_all end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin subparagraph \Inflationary\ lemma inflationary_on_unitI: assumes "(tdfr.L \\<^sub>m tdfr.R) l" and "(tdfr.R \\<^sub>m tdfr.L) r" and "inflationary_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L1\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ inflationary_on (in_codom (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "inflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" unfolding left_rel_eq_tdfr_left_Refl_Rel using assms by (intro inflationary_onI Refl_RelI) - (auto intro: tdfr.rel_unit_self_if_rel_selfI[simplified unit_eq] elim!: Refl_RelE) + (auto 6 2 intro: tdfr.rel_unit_self_if_rel_selfI[simplified unit_eq] + elim!: Refl_RelE in_fieldE) subparagraph \Deflationary\ lemma deflationary_on_counitI: assumes "(tdfr.L \\<^sub>m tdfr.R) l" and "(tdfr.R \\<^sub>m tdfr.L) r" and "deflationary_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "transitive (\\<^bsub>R1\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \\<^sub>m (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)) (l2\<^bsub> x' (r1 x')\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>R2 x' x'\<^esub>) \\<^sub>m (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ deflationary_on (in_dom (\\<^bsub>R2 x' x'\<^esub>)) (\\<^bsub>R2 x' x'\<^esub>) (\\<^bsub>2 (r1 x') x'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_dom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub> y') \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y')" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "deflationary_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>) \" unfolding right_rel_eq_tdfr_right_Refl_Rel using assms by (intro deflationary_onI Refl_RelI) - (auto intro: tdfr.counit_rel_self_if_rel_selfI[simplified counit_eq] - elim!: Refl_RelE) + (auto 6 2 intro: tdfr.counit_rel_self_if_rel_selfI[simplified counit_eq] + elim!: Refl_RelE in_fieldE) subparagraph \Relational Equivalence\ context begin interpretation flip : transport_Mono_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.counit \ \" and "flip.t1.counit \ \\<^sub>1" and "\x y. flip.t2_counit x y \ \\<^bsub>2 y x\<^esub>" by (simp_all add: order_functors.flip_counit_eq_unit) lemma rel_equivalence_on_unitI: assumes "(tdfr.L \\<^sub>m tdfr.R) l" and "(tdfr.R \\<^sub>m tdfr.L) r" and rel_equiv_unit1: "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and trans_L1: "transitive (\\<^bsub>L1\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ rel_equivalence_on (in_field (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 (\\<^sub>1 x1) x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_dom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" proof - from rel_equiv_unit1 trans_L1 have "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" by (intro reflexive_on_in_field_if_transitive_if_rel_equivalence_on) with assms show ?thesis by (intro rel_equivalence_onI inflationary_on_unitI flip.deflationary_on_counitI) (auto intro!: tdfr.bi_related_unit_self_if_rel_self_aux intro: inflationary_on_if_le_pred_if_inflationary_on deflationary_on_if_le_pred_if_deflationary_on reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom elim!: rel_equivalence_onE simp only:) qed end subparagraph \Order Equivalence\ interpretation flip : transport_Mono_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.unit \ \" and "flip.t1.unit \ \\<^sub>1" and "flip.counit \ \" and "flip.t1.counit \ \\<^sub>1" and "\x y. flip.t2_unit x y \ \\<^bsub>2 y x\<^esub>" by (simp_all add: order_functors.flip_counit_eq_unit) lemma order_equivalenceI: assumes "(tdfr.L \\<^sub>m tdfr.R) l" and "(tdfr.R \\<^sub>m tdfr.L) r" and "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "rel_equivalence_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>R1\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \\<^sub>m (\\<^bsub>R2 x' x'\<^esub>)) (l2\<^bsub>x' (r1 x')\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>R2 x' x'\<^esub>) \\<^sub>m (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ rel_equivalence_on (in_field (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ rel_equivalence_on (in_field (\\<^bsub>R2 x' x'\<^esub>)) (\\<^bsub>R2 x' x'\<^esub>) (\\<^bsub>2 (r1 x') x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 (\\<^sub>1 x1) x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x2' x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' (\\<^sub>1 x2')\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_dom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_dom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub> y') \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y')" and "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_codom (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub> y') \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y')" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>R1\<^esub> x2 \ transitive (\\<^bsub>R2 x1 x2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalenceI rel_equivalence_on_unitI flip.rel_equivalence_on_unitI mono_wrt_rel_leftI flip.mono_wrt_rel_leftI) auto lemma order_equivalence_if_preorder_equivalenceI: assumes pre_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and order_equiv2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>o (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and L2_les: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 (\\<^sub>1 x1) x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and R2_les: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x2' x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' (\\<^sub>1 x2')\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" and l2_bi_rel: "\x y. x \\<^bsub>L1\<^esub> x \ in_field (\\<^bsub>L2 x x\<^esub>) y \ l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" and r2_bi_rel: "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_field (\\<^bsub>R2 x' x'\<^esub>) y' \ r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') x'\<^esub> y'" and trans_L2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and trans_R2: "\x1 x2. x1 \\<^bsub>R1\<^esub> x2 \ transitive (\\<^bsub>R2 x1 x2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" proof - from pre_equiv1 L2_les have L2_unit_eq1: "(\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) = (\\<^bsub>L2 x x\<^esub>)" and L2_unit_eq2: "(\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) = (\\<^bsub>L2 x x\<^esub>)" if "x \\<^bsub>L1\<^esub> x" for x using \x \\<^bsub>L1\<^esub> x\ by (auto elim!: t1.preorder_equivalence_order_equivalenceE intro!: tdfr.left_rel2_unit_eqs_left_rel2I bi_related_if_rel_equivalence_on simp del: t1.unit_eq) from pre_equiv1 R2_les have R2_counit_eq1: "(\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) = (\\<^bsub>R2 x' x'\<^esub>)" and R2_counit_eq2: "(\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) = (\\<^bsub>R2 x' x'\<^esub>)" (is ?goal2) if "x' \\<^bsub>R1\<^esub> x'" for x' using \x' \\<^bsub>R1\<^esub> x'\ by (auto elim!: t1.preorder_equivalence_order_equivalenceE intro!: flip.tdfr.left_rel2_unit_eqs_left_rel2I bi_related_if_rel_equivalence_on simp del: t1.counit_eq) from order_equiv2 have mono_l2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>)" and mono_r2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" by auto moreover have "rel_equivalence_on (in_field (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" (is ?goal1) and "((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" (is ?goal2) if [iff]: "x \\<^bsub>L1\<^esub> x" for x proof - from pre_equiv1 have "x \<^bsub>L1\<^esub>\ l1 x" by (auto intro!: t1.left_GaloisI elim!: t1.preorder_equivalence_order_equivalenceE t1.order_equivalenceE) with order_equiv2 have "((\\<^bsub>L2 x x\<^esub>) \\<^sub>o (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (r2\<^bsub>x (l1 x)\<^esub>)" by (auto simp flip: L2_unit_eq2) then show ?goal1 ?goal2 by (auto elim: order_functors.order_equivalenceE) qed moreover have "rel_equivalence_on (in_field (\\<^bsub>R2 x' x'\<^esub>)) (\\<^bsub>R2 x' x'\<^esub>) (\\<^bsub>2 (r1 x') x'\<^esub>)" (is ?goal1) and "((\\<^bsub>R2 x' x'\<^esub>) \\<^sub>m (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>)" (is ?goal2) if [iff]: "x' \\<^bsub>R1\<^esub> x'" for x' proof - from pre_equiv1 have "r1 x' \<^bsub>L1\<^esub>\ x'" by blast with order_equiv2 have "((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \\<^sub>o (\\<^bsub>R2 x' x'\<^esub>)) (l2\<^bsub>x' (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" by (auto simp flip: R2_counit_eq1) then show ?goal1 ?goal2 by (auto elim: order_functors.order_equivalenceE) qed moreover from mono_l2 tdfr.mono_wrt_rel_left2_if_mono_wrt_rel_left2_if_left_GaloisI have "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \\<^sub>m (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>)" using pre_equiv1 R2_les(2) by blast moreover from pre_equiv1 have "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" by (intro t1.half_galois_prop_right_left_right_if_transitive_if_order_equivalence) (auto elim!: t1.preorder_equivalence_order_equivalenceE) moreover with mono_r2 tdfr.mono_wrt_rel_right2_if_mono_wrt_rel_right2_if_left_GaloisI have "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ((\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) \\<^sub>m (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>)" using pre_equiv1 by blast moreover with L2_les have "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ((\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) \\<^sub>m (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>)" by blast moreover have "in_dom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" (is "_ \ ?goal1") and "in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" (is "_ \ ?goal2") if [iff]: "x \\<^bsub>L1\<^esub> x" for x y proof - presume "in_dom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) y \ in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y" then have "in_field (\\<^bsub>L2 x x\<^esub>) y" using L2_unit_eq1 L2_unit_eq2 by auto with l2_bi_rel have "l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" by blast moreover from pre_equiv1 have \l1 x \\<^bsub>R1\<^esub> l1 x\ by blast ultimately show ?goal1 ?goal2 using trans_R2 by blast+ qed auto moreover have "in_dom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub> y') \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y')" (is "_ \ ?goal1") and "in_codom (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub> y') \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y')" (is "_ \ ?goal2") if [iff]: "x' \\<^bsub>R1\<^esub> x'" for x' y' proof - presume "in_dom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ in_codom (\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) y'" then have "in_field (\\<^bsub>R2 x' x'\<^esub>) y'" using R2_counit_eq1 R2_counit_eq2 by auto with r2_bi_rel have "r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') x'\<^esub> y'" by blast moreover from pre_equiv1 have \r1 x' \\<^bsub>L1\<^esub> r1 x'\ by blast ultimately show ?goal1 ?goal2 using trans_L2 by blast+ qed auto ultimately show ?thesis using assms by (intro order_equivalenceI tdfr.mono_wrt_rel_left_if_transitiveI tdfr.mono_wrt_rel_left2_if_mono_wrt_rel_left2_if_left_GaloisI tdfr.mono_wrt_rel_right_if_transitiveI tdfr.mono_wrt_rel_right2_if_mono_wrt_rel_right2_if_left_GaloisI) (auto elim!: t1.preorder_equivalence_order_equivalenceE) qed lemma order_equivalence_if_preorder_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>o (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 (\\<^sub>1 x1) x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x2' x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' (\\<^sub>1 x2')\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "\x y. x \\<^bsub>L1\<^esub> x \ in_field (\\<^bsub>L2 x x\<^esub>) y \ l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_field (\\<^bsub>R2 x' x'\<^esub>) y' \ r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') x'\<^esub> y'" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>R1\<^esub> x2 \ transitive (\\<^bsub>R2 x1 x2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalence_if_preorder_equivalenceI tdfr.order_equivalence_if_order_equivalence_mono_assms_leftI tdfr.order_equivalence_if_order_equivalence_mono_assms_rightI reflexive_on_in_field_if_transitive_if_rel_equivalence_on) (auto elim!: t1.preorder_equivalence_order_equivalenceE) lemma order_equivalence_if_mono_if_preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>o (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>) | \\<^sub>1 x2 \\<^bsub>L1\<^esub> x1] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x2 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>) | \\<^sub>1 x2' \\<^bsub>R1\<^esub> x1'] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x2' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | (x2' \\<^bsub>R1\<^esub> x3' \ x4' \\<^bsub>R1\<^esub> \\<^sub>1 x3')] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>R1\<^esub> x2 \ transitive (\\<^bsub>R2 x1 x2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalence_if_preorder_equivalenceI' tdfr.l2_unit_bi_rel_selfI tdfr.r2_counit_bi_rel_selfI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI flip.tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI flip.tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI t1.galois_connection_left_right_if_transitive_if_order_equivalence flip.t1.galois_connection_left_right_if_transitive_if_order_equivalence reflexive_on_in_field_if_transitive_if_rel_equivalence_on) (auto elim!: t1.preorder_equivalence_order_equivalenceE) theorem order_equivalence_if_mono_if_preorder_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalence_if_mono_if_preorder_equivalenceI tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI flip.tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI tdfr.transitive_left2_if_preorder_equivalenceI tdfr.transitive_right2_if_preorder_equivalenceI t1.preorder_on_in_field_left_if_transitive_if_order_equivalence flip.t1.preorder_on_in_field_left_if_transitive_if_order_equivalence t1.galois_equivalence_left_right_if_transitive_if_order_equivalence flip.t1.galois_equivalence_left_right_if_transitive_if_order_equivalence) (auto elim!: t1.preorder_equivalence_order_equivalenceE t2.preorder_equivalence_order_equivalenceE) end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin interpretation flip : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma inflationary_on_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "inflationary_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "inflationary_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>L2\<^esub>)" shows "inflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms by (intro tpdfr.inflationary_on_unitI tfr.mono_wrt_rel_leftI flip.tfr.mono_wrt_rel_leftI) simp_all lemma deflationary_on_counitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "deflationary_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "transitive (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "deflationary_on (in_dom (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>R2\<^esub>)" shows "deflationary_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>) \" using assms by (intro tpdfr.deflationary_on_counitI tfr.mono_wrt_rel_leftI flip.tfr.mono_wrt_rel_leftI) simp_all lemma rel_equivalence_on_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "rel_equivalence_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>L2\<^esub>)" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms by (intro tpdfr.rel_equivalence_on_unitI tfr.mono_wrt_rel_leftI flip.tfr.mono_wrt_rel_leftI) simp_all lemma order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro tpdfr.order_equivalenceI tfr.mono_wrt_rel_leftI flip.tfr.mono_wrt_rel_leftI) (auto elim!: tdfrs.t1.preorder_equivalence_order_equivalenceE tdfrs.t2.preorder_equivalence_order_equivalenceE) end end \ No newline at end of file diff --git a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Base.thy b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Base.thy --- a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Base.thy +++ b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Base.thy @@ -1,663 +1,663 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Transport for Natural Functors\ subsection \Basic Setup\ theory Transport_Natural_Functors_Base imports HOL.BNF_Def HOL_Alignment_Functions Transport_Base begin paragraph \Summary\ text \Basic setup for closure proofs and simple lemmas.\ text \In the following, we willingly use granular apply-style proofs since, in practice, these theorems have to be automatically generated whenever we declare a new natural functor. Note that "HOL-Library" provides a command \bnf_axiomatization\ which allows one to axiomatically declare a bounded natural functor. However, we only need a subset of these axioms - the boundedness of the functor is irrelevant for our purposes. For this reason - and the sake of completeness - we state all the required axioms explicitly below.\ lemma Grp_UNIV_eq_eq_comp: "BNF_Def.Grp UNIV f = (=) \ f" by (intro ext) (auto elim: GrpE intro: GrpI) lemma eq_comp_rel_comp_eq_comp: "(=) \ f \\ R = R \ f" by (intro ext) auto lemma Domain_Collect_case_prod_eq_Collect_in_dom: "Domain {(x, y). R x y} = {x. in_dom R x}" by blast lemma ball_in_dom_iff_ball_ex: "(\x \ S. in_dom R x) \ (\x \ S. \y. R x y)" by blast lemma pair_mem_Collect_case_prod_iff: "(x, y) \ {(x, y). R x y} \ R x y" by blast paragraph \Natural Functor Axiomatisation\ typedecl ('d, 'a, 'b, 'c) F consts Fmap :: "('a1 \ 'a2) \ ('b1 \ 'b2) \ ('c1 \ 'c2) \ ('d, 'a1, 'b1, 'c1) F \ ('d, 'a2, 'b2, 'c2) F" Fset1 :: "('d, 'a, 'b, 'c) F \ 'a set" Fset2 :: "('d, 'a, 'b, 'c) F \ 'b set" Fset3 :: "('d, 'a, 'b, 'c) F \ 'c set" axiomatization where Fmap_id: "Fmap id id id = id" and Fmap_comp: "\f1 f2 f3 g1 g2 g3. Fmap (g1 \ f1) (g2 \ f2) (g3 \ f3) = Fmap g1 g2 g3 \ Fmap f1 f2 f3" and Fmap_cong: "\f1 f2 f3 g1 g2 g3 x. (\x1. x1 \ Fset1 x \ f1 x1 = g1 x1) \ (\x2. x2 \ Fset2 x \ f2 x2 = g2 x2) \ (\x3. x3 \ Fset3 x \ f3 x3 = g3 x3) \ Fmap f1 f2 f3 x = Fmap g1 g2 g3 x" and Fset1_natural: "\f1 f2 f3. Fset1 \ Fmap f1 f2 f3 = image f1 \ Fset1" and Fset2_natural: "\f1 f2 f3. Fset2 \ Fmap f1 f2 f3 = image f2 \ Fset2" and Fset3_natural: "\f1 f2 f3. Fset3 \ Fmap f1 f2 f3 = image f3 \ Fset3" lemma Fmap_id_eq_self: "Fmap id id id x = x" by (subst Fmap_id, subst id_eq_self, rule refl) lemma Fmap_comp_eq_Fmap_Fmap: "Fmap (g1 \ f1) (g2 \ f2) (g3 \ f3) x = Fmap g1 g2 g3 (Fmap f1 f2 f3 x)" by (fact fun_cong[OF Fmap_comp, simplified comp_eq]) lemma Fset1_Fmap_eq_image_Fset1: "Fset1 (Fmap f1 f2 f3 x) = f1 ` Fset1 x" by (fact fun_cong[OF Fset1_natural, simplified comp_eq]) lemma Fset2_Fmap_eq_image_Fset2: "Fset2 (Fmap f1 f2 f3 x) = f2 ` Fset2 x" by (fact fun_cong[OF Fset2_natural, simplified comp_eq]) lemma Fset3_Fmap_eq_image_Fset3: "Fset3 (Fmap f1 f2 f3 x) = f3 ` Fset3 x" by (fact fun_cong[OF Fset3_natural, simplified comp_eq]) lemmas Fset_Fmap_eqs = Fset1_Fmap_eq_image_Fset1 Fset2_Fmap_eq_image_Fset2 Fset3_Fmap_eq_image_Fset3 paragraph \Relator\ definition Frel :: "('a1 \ 'a2 \ bool) \ ('b1 \ 'b2 \ bool) \ ('c1 \ 'c2 \ bool) \ ('d, 'a1, 'b1, 'c1) F \ ('d, 'a2, 'b2, 'c2) F \ bool" where "Frel R1 R2 R3 x y \ (\z. z \ {x. Fset1 x \ {(x, y). R1 x y} \ Fset2 x \ {(x, y). R2 x y} \ Fset3 x \ {(x, y). R3 x y}} \ Fmap fst fst fst z = x \ Fmap snd snd snd z = y)" lemma FrelI: assumes "Fset1 z \ {(x, y). R1 x y}" and "Fset2 z \ {(x, y). R2 x y}" and "Fset3 z \ {(x, y). R3 x y}" and "Fmap fst fst fst z = x" and "Fmap snd snd snd z = y" shows "Frel R1 R2 R3 x y" apply (subst Frel_def) apply (intro exI conjI CollectI) apply (fact assms)+ done lemma FrelE: assumes "Frel R1 R2 R3 x y" obtains z where "Fset1 z \ {(x, y). R1 x y}" "Fset2 z \ {(x, y). R2 x y}" "Fset3 z \ {(x, y). R3 x y}" "Fmap fst fst fst z = x" "Fmap snd snd snd z = y" apply (insert assms) apply (subst (asm) Frel_def) apply (elim exE CollectE conjE) apply assumption done lemma Grp_UNIV_Fmap_eq_Frel_Grp: "BNF_Def.Grp UNIV (Fmap f1 f2 f3) = Frel (BNF_Def.Grp UNIV f1) (BNF_Def.Grp UNIV f2) (BNF_Def.Grp UNIV f3)" apply (intro ext iffI) apply (rule FrelI[where ?z="Fmap (BNF_Def.convol id f1) (BNF_Def.convol id f2) (BNF_Def.convol id f3) _"]) apply (subst Fset_Fmap_eqs, rule image_subsetI, rule convol_mem_GrpI[simplified Fun_id_eq_id], rule UNIV_I)+ apply (unfold Fmap_comp_eq_Fmap_Fmap[symmetric] fst_convol[simplified Fun_comp_eq_comp] snd_convol[simplified Fun_comp_eq_comp] Fmap_id id_eq_self) apply (rule refl) apply (subst (asm) Grp_UNIV_eq_eq_comp) apply (subst (asm) comp_eq) apply assumption apply (erule FrelE) apply hypsubst apply (subst Grp_UNIV_eq_eq_comp) apply (subst comp_eq) apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (rule Fmap_cong; rule Collect_case_prod_Grp_eqD[simplified Fun_comp_eq_comp], drule rev_subsetD, assumption+) done lemma Frel_Grp_UNIV_Fmap: "Frel (BNF_Def.Grp UNIV f1) (BNF_Def.Grp UNIV f2) (BNF_Def.Grp UNIV f3) x (Fmap f1 f2 f3 x)" apply (subst Grp_UNIV_Fmap_eq_Frel_Grp[symmetric]) apply (subst Grp_UNIV_eq_eq_comp) apply (subst comp_eq) apply (rule refl) done lemma Frel_Grp_UNIV_iff_eq_Fmap: "Frel (BNF_Def.Grp UNIV f1) (BNF_Def.Grp UNIV f2) (BNF_Def.Grp UNIV f3) x y \ (y = Fmap f1 f2 f3 x)" by (subst eq_commute[of y]) (fact fun_cong[OF fun_cong[OF Grp_UNIV_Fmap_eq_Frel_Grp], simplified Grp_UNIV_eq_eq_comp comp_eq, folded Grp_UNIV_eq_eq_comp, symmetric]) lemma Frel_eq: "Frel (=) (=) (=) = (=)" apply (unfold BNF_Def.eq_alt[simplified Fun_id_eq_id]) apply (subst Grp_UNIV_Fmap_eq_Frel_Grp[symmetric]) apply (subst Fmap_id) apply (fold BNF_Def.eq_alt[simplified Fun_id_eq_id]) apply (rule refl) done corollary Frel_eq_self: "Frel (=) (=) (=) x x" by (fact iffD2[OF fun_cong[OF fun_cong[OF Frel_eq]] refl]) lemma Frel_mono_strong: assumes "Frel R1 R2 R3 x y" and "\x1 y1. x1 \ Fset1 x \ y1 \ Fset1 y \ R1 x1 y1 \ S1 x1 y1" and "\x2 y2. x2 \ Fset2 x \ y2 \ Fset2 y \ R2 x2 y2 \ S2 x2 y2" and "\x3 y3. x3 \ Fset3 x \ y3 \ Fset3 y \ R3 x3 y3 \ S3 x3 y3" shows "Frel S1 S2 S3 x y" apply (insert assms(1)) apply (erule FrelE) apply (rule FrelI) apply (rule subsetI, frule rev_subsetD, assumption, frule imageI[of _ "Fset1 _" fst] imageI[of _ "Fset2 _" fst] imageI[of _ "Fset3 _" fst], drule imageI[of _ "Fset1 _" snd] imageI[of _ "Fset2 _" snd] imageI[of _ "Fset3 _" snd], (subst (asm) Fset_Fmap_eqs[symmetric])+, intro CollectI case_prodI2, rule assms; hypsubst, unfold fst_conv snd_conv, (elim CollectE case_prodE Pair_inject, hypsubst)?, assumption)+ apply assumption+ done corollary Frel_mono: assumes "R1 \ S1" "R2 \ S2" "R3 \ S3" shows "Frel R1 R2 R3 \ Frel S1 S2 S3" apply (intro le_relI) apply (rule Frel_mono_strong) apply assumption apply (insert assms) apply (drule le_relD[OF assms(1)] le_relD[OF assms(2)] le_relD[OF assms(3)], assumption)+ done lemma Frel_refl_strong: assumes "\x1. x1 \ Fset1 x \ R1 x1 x1" and "\x2. x2 \ Fset2 x \ R2 x2 x2" and "\x3. x3 \ Fset3 x \ R3 x3 x3" shows "Frel R1 R2 R3 x x" by (rule Frel_mono_strong[OF Frel_eq_self[of x]]; drule assms, hypsubst, assumption) lemma Frel_cong: assumes "\x1 y1. x1 \ Fset1 x \ y1 \ Fset1 y \ R1 x1 y1 \ R1' x1 y1" and "\x2 y2. x2 \ Fset2 x \ y2 \ Fset2 y \ R2 x2 y2 \ R2' x2 y2" and "\x3 y3. x3 \ Fset3 x \ y3 \ Fset3 y \ R3 x3 y3 \ R3' x3 y3" shows "Frel R1 R2 R3 x y = Frel R1' R2' R3' x y" by (rule iffI; rule Frel_mono_strong, assumption; rule iffD1[OF assms(1)] iffD1[OF assms(2)] iffD1[OF assms(3)] iffD2[OF assms(1)] iffD2[OF assms(2)] iffD2[OF assms(3)]; assumption) lemma Frel_rel_inv_eq_rel_inv_Frel: "Frel R1\ R2\ R3\ = (Frel R1 R2 R3)\" by (intro ext iffI; unfold rel_inv_iff_rel, erule FrelE, hypsubst, rule FrelI[where ?z="Fmap prod.swap prod.swap prod.swap _"]; ((subst Fset_Fmap_eqs, rule image_subsetI, drule rev_subsetD, assumption, elim CollectE case_prodE, hypsubst, subst swap_simp, subst pair_mem_Collect_case_prod_iff, assumption) | (subst Fmap_comp_eq_Fmap_Fmap[symmetric], rule Fmap_cong; unfold comp_eq fst_swap snd_swap, rule refl))) text \Given the former axioms, the following axiom - subdistributivity of the relator - is equivalent to the (F, Fmap) functor preserving weak pullbacks.\ axiomatization where Frel_comp_le_Frel_rel_comp: "\R1 R2 R3 S1 S2 S3. Frel R1 R2 R3 \\ Frel S1 S2 S3 \ Frel (R1 \\ S1) (R2 \\ S2) (R3 \\ S3)" lemma fst_sndOp_eq_snd_fstOp: "fst \ BNF_Def.sndOp P Q = snd \ BNF_Def.fstOp P Q" unfolding fstOp_def sndOp_def by (intro ext) simp lemma Frel_rel_comp_le_Frel_comp: "Frel (R1 \\ S1) (R2 \\ S2) (R3 \\ S3) \ (Frel R1 R2 R3 \\ Frel S1 S2 S3)" apply (rule le_relI) apply (erule FrelE) apply (rule rel_compI[where ?y="Fmap (snd \ BNF_Def.fstOp R1 S1) (snd \ BNF_Def.fstOp R2 S2) (snd \ BNF_Def.fstOp R3 S3) _"]) apply (rule FrelI[where ?z="Fmap (BNF_Def.fstOp R1 S1) (BNF_Def.fstOp R2 S2) (BNF_Def.fstOp R3 S3) _"]) apply (subst Fset_Fmap_eqs, intro image_subsetI, rule fstOp_in[unfolded relcompp_eq_rel_comp], drule rev_subsetD, assumption+)+ apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (fold ext[of fst "fst \ _", OF fst_fstOp[unfolded Fun_comp_eq_comp]]) apply hypsubst apply (rule refl) apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (rule refl) apply (rule FrelI[where ?z="Fmap (BNF_Def.sndOp R1 S1) (BNF_Def.sndOp R2 S2) (BNF_Def.sndOp R3 S3) _"]) apply (subst Fset_Fmap_eqs, intro image_subsetI, rule sndOp_in[unfolded relcompp_eq_rel_comp], drule rev_subsetD, assumption+)+ apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (unfold fst_sndOp_eq_snd_fstOp) apply (rule refl) apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (fold ext[of snd "snd \ _", OF snd_sndOp[unfolded Fun_comp_eq_comp]]) apply hypsubst apply (rule refl) done corollary Frel_comp_eq_Frel_rel_comp: "Frel R1 R2 R3 \\ Frel S1 S2 S3 = Frel (R1 \\ S1) (R2 \\ S2) (R3 \\ S3)" by (rule antisym; rule Frel_comp_le_Frel_rel_comp Frel_rel_comp_le_Frel_comp) lemma Frel_Fmap_eq1: "Frel R1 R2 R3 (Fmap f1 f2 f3 x) y = Frel (\x. R1 (f1 x)) (\x. R2 (f2 x)) (\x. R3 (f3 x)) x y" apply (rule iffI) apply (fold comp_eq[of R1] comp_eq[of R2] comp_eq[of R3]) apply (drule rel_compI[where ?R="Frel _ _ _" and ?S="Frel _ _ _", OF Frel_Grp_UNIV_Fmap]) apply (unfold Grp_UNIV_eq_eq_comp) apply (drule le_relD[OF Frel_comp_le_Frel_rel_comp]) apply (unfold eq_comp_rel_comp_eq_comp) apply assumption apply (fold eq_comp_rel_comp_eq_comp[where ?R=R1] eq_comp_rel_comp_eq_comp[where ?R=R2] eq_comp_rel_comp_eq_comp[where ?R=R3] Grp_UNIV_eq_eq_comp) apply (drule le_relD[OF Frel_rel_comp_le_Frel_comp]) apply (erule rel_compE) apply (subst (asm) Frel_Grp_UNIV_iff_eq_Fmap) apply hypsubst apply assumption done lemma Frel_Fmap_eq2: "Frel R1 R2 R3 x (Fmap g1 g2 g3 y) = Frel (\x y. R1 x (g1 y)) (\x y. R2 x (g2 y)) (\x y. R3 x (g3 y)) x y" apply (subst rel_inv_iff_rel[of "Frel _ _ _", symmetric]) apply (subst Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (subst Frel_Fmap_eq1) apply (rule sym) apply (subst rel_inv_iff_rel[of "Frel _ _ _", symmetric]) apply (subst Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (unfold rel_inv_iff_rel) apply (rule refl) done lemmas Frel_Fmap_eqs = Frel_Fmap_eq1 Frel_Fmap_eq2 paragraph \Predicator\ definition Fpred :: "('a \ bool) \ ('b \ bool) \ ('c \ bool) \ ('d, 'a, 'b, 'c) F \ bool" where "Fpred P1 P2 P3 x \ Frel ((=)\\<^bsub>P1\<^esub>) ((=)\\<^bsub>P2\<^esub>) ((=)\\<^bsub>P3\<^esub>) x x" lemma Fpred_mono_strong: assumes "Fpred P1 P2 P3 x" and "\x1. x1 \ Fset1 x \ P1 x1 \ Q1 x1" and "\x2. x2 \ Fset2 x \ P2 x2 \ Q2 x2" and "\x3. x3 \ Fset3 x \ P3 x3 \ Q3 x3" shows "Fpred Q1 Q2 Q3 x" apply (insert assms(1)) apply (unfold Fpred_def) apply (rule Frel_mono_strong, assumption; - erule restrict_leftE, - rule restrict_leftI, + erule bin_rel_restrict_leftE, + rule bin_rel_restrict_leftI, assumption, rule assms, assumption+) done lemma Fpred_top: "Fpred \ \ \ x" apply (subst Fpred_def) apply (rule Frel_refl_strong; - subst restrict_left_top_eq, + subst bin_rel_restrict_left_top_eq, rule refl) done lemma FpredI: assumes "\x1. x1 \ Fset1 x \ P1 x1" and "\x2. x2 \ Fset2 x \ P2 x2" and "\x3. x3 \ Fset3 x \ P3 x3" shows "Fpred P1 P2 P3 x" using assms by (rule Fpred_mono_strong[OF Fpred_top]) lemma FpredE: assumes "Fpred P1 P2 P3 x" obtains "\x1. x1 \ Fset1 x \ P1 x1" "\x2. x2 \ Fset2 x \ P2 x2" "\x3. x3 \ Fset3 x \ P3 x3" by (elim meta_impE; (assumption | insert assms, subst (asm) Fpred_def, erule FrelE, hypsubst, subst (asm) Fset_Fmap_eqs, subst (asm) Domain_fst[symmetric], drule rev_subsetD, rule Domain_mono, assumption, - unfold Domain_Collect_case_prod_eq_Collect_in_dom in_dom_restrict_left_eq, + unfold Domain_Collect_case_prod_eq_Collect_in_dom in_dom_bin_rel_restrict_left_eq, elim CollectE inf1E, assumption)) lemma Fpred_eq_ball: "Fpred P1 P2 P3 = (\x. Ball (Fset1 x) P1 \ Ball (Fset2 x) P2 \ Ball (Fset3 x) P3)" by (intro ext iffI conjI ballI FpredI; elim FpredE conjE bspec) lemma Fpred_Fmap_eq: "Fpred P1 P2 P3 (Fmap f1 f2 f3 x) = Fpred (P1 \ f1) (P2 \ f2) (P3 \ f3) x" by (unfold Fpred_def Frel_Fmap_eqs) (rule iffI; erule FrelE, hypsubst, unfold Frel_Fmap_eqs, rule Frel_refl_strong; - rule restrict_leftI, + rule bin_rel_restrict_leftI, rule refl, drule rev_subsetD, assumption, - elim CollectE case_prodE restrict_leftE, + elim CollectE case_prodE bin_rel_restrict_leftE, hypsubst, unfold comp_eq fst_conv, assumption) lemma Fpred_in_dom_if_in_dom_Frel: assumes "in_dom (Frel R1 R2 R3) x" shows "Fpred (in_dom R1) (in_dom R2) (in_dom R3) x" apply (insert assms) apply (elim in_domE FrelE) apply hypsubst apply (subst Fpred_Fmap_eq) apply (rule FpredI; drule rev_subsetD, assumption, elim CollectE case_prodE, hypsubst, unfold comp_eq fst_conv, rule in_domI, assumption) done lemma in_dom_Frel_if_Fpred_in_dom: assumes "Fpred (in_dom R1) (in_dom R2) (in_dom R3) x" shows "in_dom (Frel R1 R2 R3) x" apply (insert assms) apply (subst (asm) Fpred_eq_ball) apply (elim conjE) apply (subst (asm) ball_in_dom_iff_ball_ex, drule bchoice, \\requires the axiom of choice.\ erule exE)+ apply (rule in_domI[where ?x=x and ?y="Fmap _ _ _ x" for x]) apply (subst Frel_Fmap_eq2) apply (rule Frel_refl_strong) apply (drule bspec[of "Fset1 _"]) apply assumption+ apply (drule bspec[of "Fset2 _"]) apply assumption+ apply (drule bspec[of "Fset3 _"]) apply assumption+ done lemma in_dom_Frel_eq_Fpred_in_dom: "in_dom (Frel R1 R2 R3) = Fpred (in_dom R1) (in_dom R2) (in_dom R3)" by (intro ext iffI; rule Fpred_in_dom_if_in_dom_Frel in_dom_Frel_if_Fpred_in_dom) lemma in_codom_Frel_eq_Fpred_in_codom: "in_codom (Frel R1 R2 R3) = Fpred (in_codom R1) (in_codom R2) (in_codom R3)" apply (subst in_dom_rel_inv_eq_in_codom[symmetric]) apply (subst Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (subst in_dom_Frel_eq_Fpred_in_dom) apply (subst in_dom_rel_inv_eq_in_codom)+ apply (rule refl) done lemma in_field_Frel_eq_Fpred_in_in_field: "in_field (Frel R1 R2 R3) = Fpred (in_dom R1) (in_dom R2) (in_dom R3) \ Fpred (in_codom R1) (in_codom R2) (in_codom R3)" apply (subst in_field_eq_in_dom_sup_in_codom) apply (subst in_dom_Frel_eq_Fpred_in_dom) apply (subst in_codom_Frel_eq_Fpred_in_codom) apply (rule refl) done lemma Frel_restrict_left_Fpred_eq_Frel_restrict_left: fixes R1 :: "'a1 \ 'a2 \ bool" and R2 :: "'b1 \ 'b2 \ bool" and R3 :: "'c1 \ 'c2 \ bool" and P1 :: "'a1 \ bool" and P2 :: "'b1 \ bool" and P3 :: "'c1 \ bool" shows "(Frel R1 R2 R3 :: ('d, 'a1, 'b1, 'c1) F \ _)\\<^bsub>Fpred P1 P2 P3 :: ('d, 'a1, 'b1, 'c1) F \ _\<^esub> = Frel (R1\\<^bsub>P1\<^esub>) (R2\\<^bsub>P2\<^esub>) (R3\\<^bsub>P3\<^esub>)" apply (intro ext) apply (rule iffI) - apply (erule restrict_leftE) + apply (erule bin_rel_restrict_leftE) apply (elim FpredE) apply (rule Frel_mono_strong, assumption; - rule restrict_leftI, + rule bin_rel_restrict_leftI, assumption+) - apply (rule restrict_leftI) + apply (rule bin_rel_restrict_leftI) apply (rule Frel_mono_strong, assumption; - erule restrict_leftE, + erule bin_rel_restrict_leftE, assumption) apply (drule in_domI[of "Frel (R1\\<^bsub>P1\<^esub>) (R2\\<^bsub>P2\<^esub>) (R3\\<^bsub>P3\<^esub>)"]) apply (drule Fpred_in_dom_if_in_dom_Frel) apply (rule Fpred_mono_strong, assumption; - unfold in_dom_restrict_left_eq inf_apply inf_bool_def; + unfold in_dom_bin_rel_restrict_left_eq inf_apply inf_bool_def; rule conjunct2, assumption) done lemma Frel_restrict_right_Fpred_eq_Frel_restrict_right: fixes R1 :: "'a1 \ 'a2 \ bool" and R2 :: "'b1 \ 'b2 \ bool" and R3 :: "'c1 \ 'c2 \ bool" and P1 :: "'a2 \ bool" and P2 :: "'b2 \ bool" and P3 :: "'c2 \ bool" shows "(Frel R1 R2 R3 :: _ \ ('d, 'a2, 'b2, 'c2) F \ _)\\<^bsub>Fpred P1 P2 P3 :: ('d, 'a2, 'b2, 'c2) F \ _\<^esub> = Frel (R1\\<^bsub>P1\<^esub>) (R2\\<^bsub>P2\<^esub>) (R3\\<^bsub>P3\<^esub>)" - apply (subst restrict_right_eq) + apply (subst bin_rel_restrict_right_eq) apply (subst Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (subst Frel_restrict_left_Fpred_eq_Frel_restrict_left) apply (subst Frel_rel_inv_eq_rel_inv_Frel[symmetric]) - apply (fold restrict_right_eq) + apply (fold bin_rel_restrict_right_eq) apply (rule refl) done locale transport_natural_functor = t1 : transport L1 R1 l1 r1 + t2 : transport L2 R2 l2 r2 + t3 : transport L3 R3 l3 r3 for L1 :: "'a1 \ 'a1 \ bool" and R1 :: "'b1 \ 'b1 \ bool" and l1 :: "'a1 \ 'b1" and r1 :: "'b1 \ 'a1" and L2 :: "'a2 \ 'a2 \ bool" and R2 :: "'b2 \ 'b2 \ bool" and l2 :: "'a2 \ 'b2" and r2 :: "'b2 \ 'a2" and L3 :: "'a3 \ 'a3 \ bool" and R3 :: "'b3 \ 'b3 \ bool" and l3 :: "'a3 \ 'b3" and r3 :: "'b3 \ 'a3" begin notation L1 (infix "\\<^bsub>L1\<^esub>" 50) notation R1 (infix "\\<^bsub>R1\<^esub>" 50) notation L2 (infix "\\<^bsub>L2\<^esub>" 50) notation R2 (infix "\\<^bsub>R2\<^esub>" 50) notation L3 (infix "\\<^bsub>L3\<^esub>" 50) notation R3 (infix "\\<^bsub>R3\<^esub>" 50) notation t1.ge_left (infix "\\<^bsub>L1\<^esub>" 50) notation t1.ge_right (infix "\\<^bsub>R1\<^esub>" 50) notation t2.ge_left (infix "\\<^bsub>L2\<^esub>" 50) notation t2.ge_right (infix "\\<^bsub>R2\<^esub>" 50) notation t3.ge_left (infix "\\<^bsub>L3\<^esub>" 50) notation t3.ge_right (infix "\\<^bsub>R3\<^esub>" 50) notation t1.left_Galois (infix "\<^bsub>L1\<^esub>\" 50) notation t1.right_Galois (infix "\<^bsub>R1\<^esub>\" 50) notation t2.left_Galois (infix "\<^bsub>L2\<^esub>\" 50) notation t2.right_Galois (infix "\<^bsub>R2\<^esub>\" 50) notation t3.left_Galois (infix "\<^bsub>L3\<^esub>\" 50) notation t3.right_Galois (infix "\<^bsub>R3\<^esub>\" 50) notation t1.ge_Galois_left (infix "\\<^bsub>L1\<^esub>" 50) notation t1.ge_Galois_right (infix "\\<^bsub>R1\<^esub>" 50) notation t2.ge_Galois_left (infix "\\<^bsub>L2\<^esub>" 50) notation t2.ge_Galois_right (infix "\\<^bsub>R2\<^esub>" 50) notation t3.ge_Galois_left (infix "\\<^bsub>L3\<^esub>" 50) notation t3.ge_Galois_right (infix "\\<^bsub>R3\<^esub>" 50) notation t1.right_ge_Galois (infix "\<^bsub>R1\<^esub>\" 50) notation t1.Galois_right (infix "\\<^bsub>R1\<^esub>" 50) notation t2.right_ge_Galois (infix "\<^bsub>R2\<^esub>\" 50) notation t2.Galois_right (infix "\\<^bsub>R2\<^esub>" 50) notation t3.right_ge_Galois (infix "\<^bsub>R3\<^esub>\" 50) notation t3.Galois_right (infix "\\<^bsub>R3\<^esub>" 50) notation t1.left_ge_Galois (infix "\<^bsub>L1\<^esub>\" 50) notation t1.Galois_left (infix "\\<^bsub>L1\<^esub>" 50) notation t2.left_ge_Galois (infix "\<^bsub>L2\<^esub>\" 50) notation t2.Galois_left (infix "\\<^bsub>L2\<^esub>" 50) notation t3.left_ge_Galois (infix "\<^bsub>L3\<^esub>\" 50) notation t3.Galois_left (infix "\\<^bsub>L3\<^esub>" 50) notation t1.unit ("\\<^sub>1") notation t1.counit ("\\<^sub>1") notation t2.unit ("\\<^sub>2") notation t2.counit ("\\<^sub>2") notation t3.unit ("\\<^sub>3") notation t3.counit ("\\<^sub>3") definition "L \ Frel (\\<^bsub>L1\<^esub>) (\\<^bsub>L2\<^esub>) (\\<^bsub>L3\<^esub>)" lemma left_rel_eq_Frel: "L = Frel (\\<^bsub>L1\<^esub>) (\\<^bsub>L2\<^esub>) (\\<^bsub>L3\<^esub>)" unfolding L_def .. definition "l \ Fmap l1 l2 l3" lemma left_eq_Fmap: "l = Fmap l1 l2 l3" unfolding l_def .. context begin interpretation flip : transport_natural_functor R1 L1 r1 l1 R2 L2 r2 l2 R3 L3 r3 l3 . abbreviation "R \ flip.L" abbreviation "r \ flip.l" lemma right_rel_eq_Frel: "R = Frel (\\<^bsub>R1\<^esub>) (\\<^bsub>R2\<^esub>) (\\<^bsub>R3\<^esub>)" unfolding flip.left_rel_eq_Frel .. lemma right_eq_Fmap: "r = Fmap r1 r2 r3" unfolding flip.left_eq_Fmap .. lemmas transport_defs = left_rel_eq_Frel left_eq_Fmap right_rel_eq_Frel right_eq_Fmap end sublocale transport L R l r . (*FIXME: somehow the notation for the fixed parameters L and R, defined in Order_Functions_Base.thy, is lost. We hence re-declare it here.*) notation L (infix "\\<^bsub>L\<^esub>" 50) notation R (infix "\\<^bsub>R\<^esub>" 50) lemma unit_eq_Fmap: "\ = Fmap \\<^sub>1 \\<^sub>2 \\<^sub>3" unfolding unit_eq_comp by (simp only: right_eq_Fmap left_eq_Fmap flip: Fmap_comp t1.unit_eq_comp t2.unit_eq_comp t3.unit_eq_comp) interpretation flip_inv : transport_natural_functor "(\\<^bsub>R1\<^esub>)" "(\\<^bsub>L1\<^esub>)" r1 l1 "(\\<^bsub>R2\<^esub>)" "(\\<^bsub>L2\<^esub>)" r2 l2 "(\\<^bsub>R3\<^esub>)" "(\\<^bsub>L3\<^esub>)" r3 l3 rewrites "flip_inv.unit \ \" and "flip_inv.t1.unit \ \\<^sub>1" and "flip_inv.t2.unit \ \\<^sub>2" and "flip_inv.t3.unit \ \\<^sub>3" by (simp_all only: order_functors.flip_counit_eq_unit) lemma counit_eq_Fmap: "\ = Fmap \\<^sub>1 \\<^sub>2 \\<^sub>3" by (fact flip_inv.unit_eq_Fmap) lemma flip_inv_right_eq_ge_left: "flip_inv.R = (\\<^bsub>L\<^esub>)" unfolding left_rel_eq_Frel flip_inv.right_rel_eq_Frel by (fact Frel_rel_inv_eq_rel_inv_Frel) interpretation flip : transport_natural_functor R1 L1 r1 l1 R2 L2 r2 l2 R3 L3 r3 l3 . lemma flip_inv_left_eq_ge_right: "flip_inv.L \ (\\<^bsub>R\<^esub>)" unfolding flip.flip_inv_right_eq_ge_left . lemma mono_wrt_rel_leftI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>L3\<^esub>) \\<^sub>m (\\<^bsub>R3\<^esub>)) l3" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" apply (unfold left_rel_eq_Frel right_rel_eq_Frel left_eq_Fmap) apply (rule dep_mono_wrt_relI) apply (unfold Frel_Fmap_eqs) apply (fold rel_map_eq) apply (rule le_relD[OF Frel_mono]) apply (subst mono_wrt_rel_iff_le_rel_map[symmetric], rule assms)+ apply assumption done end end \ No newline at end of file diff --git a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Galois_Relator.thy b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Galois_Relator.thy --- a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Galois_Relator.thy +++ b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Galois_Relator.thy @@ -1,55 +1,55 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Relator\ theory Transport_Natural_Functors_Galois_Relator imports Transport_Natural_Functors_Base begin context transport_natural_functor begin lemma left_Galois_Frel_left_Galois: "(\<^bsub>L\<^esub>\) \ Frel (\<^bsub>L1\<^esub>\) (\<^bsub>L2\<^esub>\) (\<^bsub>L3\<^esub>\)" apply (rule le_relI) apply (erule left_GaloisE) apply (unfold left_rel_eq_Frel right_rel_eq_Frel right_eq_Fmap) apply (subst (asm) in_codom_Frel_eq_Fpred_in_codom) apply (erule FpredE) apply (subst (asm) Frel_Fmap_eq2) apply (rule Frel_mono_strong, assumption; rule t1.left_GaloisI t2.left_GaloisI t3.left_GaloisI; assumption) done lemma Frel_left_Galois_le_left_Galois: "Frel (\<^bsub>L1\<^esub>\) (\<^bsub>L2\<^esub>\) (\<^bsub>L3\<^esub>\) \ (\<^bsub>L\<^esub>\)" apply (rule le_relI) apply (unfold t1.left_Galois_iff_in_codom_and_left_rel_right t2.left_Galois_iff_in_codom_and_left_rel_right t3.left_Galois_iff_in_codom_and_left_rel_right) apply (fold - restrict_right_eq[of "\x y. x \\<^bsub>L1\<^esub> r1 y" "in_codom (\\<^bsub>R1\<^esub>)", - unfolded restrict_left_pred_def rel_inv_iff_rel] - restrict_right_eq[of "\x y. x \\<^bsub>L2\<^esub> r2 y" "in_codom (\\<^bsub>R2\<^esub>)", - unfolded restrict_left_pred_def rel_inv_iff_rel] - restrict_right_eq[of "\x y. x \\<^bsub>L3\<^esub> r3 y" "in_codom (\\<^bsub>R3\<^esub>)", - unfolded restrict_left_pred_def rel_inv_iff_rel]) + bin_rel_restrict_right_eq[of "\x y. x \\<^bsub>L1\<^esub> r1 y" "in_codom (\\<^bsub>R1\<^esub>)", + unfolded bin_rel_restrict_left_pred_def rel_inv_iff_rel] + bin_rel_restrict_right_eq[of "\x y. x \\<^bsub>L2\<^esub> r2 y" "in_codom (\\<^bsub>R2\<^esub>)", + unfolded bin_rel_restrict_left_pred_def rel_inv_iff_rel] + bin_rel_restrict_right_eq[of "\x y. x \\<^bsub>L3\<^esub> r3 y" "in_codom (\\<^bsub>R3\<^esub>)", + unfolded bin_rel_restrict_left_pred_def rel_inv_iff_rel]) apply (subst (asm) Frel_restrict_right_Fpred_eq_Frel_restrict_right[symmetric]) - apply (erule restrict_rightE) + apply (erule bin_rel_restrict_rightE) apply (subst (asm) in_codom_Frel_eq_Fpred_in_codom[symmetric]) apply (erule in_codomE) apply (rule left_GaloisI) apply (rule in_codomI) apply (subst right_rel_eq_Frel) apply assumption apply (unfold left_rel_eq_Frel right_eq_Fmap Frel_Fmap_eq2) apply assumption done corollary left_Galois_eq_Frel_left_Galois: "(\<^bsub>L\<^esub>\) = Frel (\<^bsub>L1\<^esub>\) (\<^bsub>L2\<^esub>\) (\<^bsub>L3\<^esub>\)" by (intro antisym left_Galois_Frel_left_Galois Frel_left_Galois_le_left_Galois) end end \ No newline at end of file diff --git a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Base.thy b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Base.thy --- a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Base.thy +++ b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Base.thy @@ -1,99 +1,99 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Basic Order Properties\ theory Transport_Natural_Functors_Order_Base imports Transport_Natural_Functors_Base begin lemma reflexive_on_in_field_FrelI: assumes "reflexive_on (in_field R1) R1" and "reflexive_on (in_field R2) R2" and "reflexive_on (in_field R3) R3" defines "R \ Frel R1 R2 R3" shows "reflexive_on (in_field R) R" - apply (subst reflexive_on_iff_eq_restrict_left_le) + apply (subst reflexive_on_iff_eq_restrict_le) apply (subst Frel_eq[symmetric]) apply (unfold R_def) apply (subst in_field_Frel_eq_Fpred_in_in_field) - apply (subst restrict_left_sup_eq) + apply (subst bin_rel_restrict_left_sup_eq) apply (subst Frel_restrict_left_Fpred_eq_Frel_restrict_left)+ apply (rule le_supI; rule Frel_mono; - subst reflexive_on_iff_eq_restrict_left_le[symmetric], + subst reflexive_on_iff_eq_restrict_le[symmetric], rule reflexive_on_if_le_pred_if_reflexive_on, rule assms, rule le_predI[OF in_field_if_in_dom] le_predI[OF in_field_if_in_codom], assumption) done lemma transitive_FrelI: assumes "transitive R1" and "transitive R2" and "transitive R3" shows "transitive (Frel R1 R2 R3)" apply (subst transitive_iff_rel_comp_le_self) apply (subst Frel_comp_eq_Frel_rel_comp) apply (rule Frel_mono; subst transitive_iff_rel_comp_le_self[symmetric], rule assms) done lemma preorder_on_in_field_FrelI: assumes "preorder_on (in_field R1) R1" and "preorder_on (in_field R2) R2" and "preorder_on (in_field R3) R3" defines "R \ Frel R1 R2 R3" shows "preorder_on (in_field R) R" apply (unfold R_def) apply (insert assms) apply (elim preorder_on_in_fieldE) apply (rule preorder_onI) apply (rule reflexive_on_in_field_FrelI; assumption) apply (subst transitive_on_in_field_iff_transitive) apply (rule transitive_FrelI; assumption) done lemma symmetric_FrelI: assumes "symmetric R1" and "symmetric R2" and "symmetric R3" shows "symmetric (Frel R1 R2 R3)" apply (subst symmetric_iff_rel_inv_eq_self) apply (subst Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (subst rel_inv_eq_self_if_symmetric, fact)+ apply (rule refl) done lemma partial_equivalence_rel_FrelI: assumes "partial_equivalence_rel R1" and "partial_equivalence_rel R2" and "partial_equivalence_rel R3" shows "partial_equivalence_rel (Frel R1 R2 R3)" apply (insert assms) apply (elim partial_equivalence_relE preorder_on_in_fieldE) apply (rule partial_equivalence_relI; rule transitive_FrelI symmetric_FrelI; assumption) done context transport_natural_functor begin lemmas reflexive_on_in_field_leftI = reflexive_on_in_field_FrelI [of L1 L2 L3, folded transport_defs] lemmas transitive_leftI = transitive_FrelI[of L1 L2 L3, folded transport_defs] lemmas preorder_on_in_field_leftI = preorder_on_in_field_FrelI [of L1 L2 L3, folded transport_defs] lemmas symmetricI = symmetric_FrelI[of L1 L2 L3, folded transport_defs] lemmas partial_equivalence_rel_leftI = partial_equivalence_rel_FrelI [of L1 L2 L3, folded transport_defs] end end \ No newline at end of file diff --git a/thys/Transport/Transport/Transport_Bijections.thy b/thys/Transport/Transport/Transport_Bijections.thy --- a/thys/Transport/Transport/Transport_Bijections.thy +++ b/thys/Transport/Transport/Transport_Bijections.thy @@ -1,202 +1,199 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Transport using Bijections\ theory Transport_Bijections imports + Restricted_Equality Functions_Bijection Transport_Base begin paragraph \Summary\ text \Setup for Transport using bijective transport functions.\ locale transport_bijection = fixes L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes mono_wrt_rel_left: "(L \\<^sub>m R) l" and mono_wrt_rel_right: "(R \\<^sub>m L) r" and inverse_left_right: "inverse_on (in_field L) l r" and inverse_right_left: "inverse_on (in_field R) r l" begin interpretation transport L R l r . interpretation g_flip_inv : galois "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" r l . lemma bijection_on_in_field: "bijection_on (in_field (\\<^bsub>L\<^esub>)) (in_field (\\<^bsub>R\<^esub>)) l r" using mono_wrt_rel_left mono_wrt_rel_right inverse_left_right inverse_right_left by (intro bijection_onI in_field_if_in_field_if_mono_wrt_rel) auto lemma half_galois_prop_left: "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" using mono_wrt_rel_left inverse_right_left - by (intro half_galois_prop_leftI) - (auto dest!: in_field_if_in_codom inverse_onD) + by (intro half_galois_prop_leftI) (fastforce dest: inverse_onD) lemma half_galois_prop_right: "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" using mono_wrt_rel_right inverse_left_right by (intro half_galois_prop_rightI) (force dest: in_field_if_in_dom inverse_onD) lemma galois_prop: "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using half_galois_prop_left half_galois_prop_right by (intro galois_propI) lemma galois_connection: "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using mono_wrt_rel_left mono_wrt_rel_right galois_prop by (intro galois_connectionI) lemma rel_equivalence_on_unitI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms inverse_left_right by (subst rel_equivalence_on_unit_iff_reflexive_on_if_inverse_on) interpretation flip : transport_bijection R L r l rewrites "order_functors.unit r l \ \" using mono_wrt_rel_left mono_wrt_rel_right inverse_left_right inverse_right_left by unfold_locales (simp_all only: flip_unit_eq_counit) lemma galois_equivalence: "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" using galois_connection flip.galois_prop by (intro galois_equivalenceI) lemmas rel_equivalence_on_counitI = flip.rel_equivalence_on_unitI lemma order_equivalenceI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms mono_wrt_rel_left mono_wrt_rel_right rel_equivalence_on_unitI rel_equivalence_on_counitI by (intro order_equivalenceI) lemma preorder_equivalenceI: assumes "preorder_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "preorder_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro preorder_equivalence_if_galois_equivalenceI galois_equivalence) simp_all lemma partial_equivalence_rel_equivalenceI: assumes "partial_equivalence_rel (\\<^bsub>L\<^esub>)" and "partial_equivalence_rel (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro partial_equivalence_rel_equivalence_if_galois_equivalenceI galois_equivalence) simp_all end locale transport_reflexive_on_in_field_bijection = fixes L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes reflexive_on_in_field_left: "reflexive_on (in_field L) L" and reflexive_on_in_field_right: "reflexive_on (in_field R) R" and transport_bijection: "transport_bijection L R l r" begin sublocale tbij? : transport_bijection L R l r rewrites "reflexive_on (in_field L) L \ True" and "reflexive_on (in_field R) R \ True" and "\P. (True \ P) \ Trueprop P" using transport_bijection reflexive_on_in_field_left reflexive_on_in_field_right by auto lemmas rel_equivalence_on_unit = rel_equivalence_on_unitI lemmas rel_equivalence_on_counit = rel_equivalence_on_counitI lemmas order_equivalence = order_equivalenceI end locale transport_preorder_on_in_field_bijection = fixes L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes preorder_on_in_field_left: "preorder_on (in_field L) L" and preorder_on_in_field_right: "preorder_on (in_field R) R" and transport_bijection: "transport_bijection L R l r" begin sublocale trefl_bij? : transport_reflexive_on_in_field_bijection L R l r rewrites "preorder_on (in_field L) L \ True" and "preorder_on (in_field R) R \ True" and "\P. (True \ P) \ Trueprop P" using transport_bijection by (intro transport_reflexive_on_in_field_bijection.intro) (insert preorder_on_in_field_left preorder_on_in_field_right, auto) lemmas preorder_equivalence = preorder_equivalenceI end locale transport_partial_equivalence_rel_bijection = fixes L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes partial_equivalence_rel_left: "partial_equivalence_rel L" and partial_equivalence_rel_right: "partial_equivalence_rel R" and transport_bijection: "transport_bijection L R l r" begin sublocale tpre_bij? : transport_preorder_on_in_field_bijection L R l r rewrites "partial_equivalence_rel L \ True" and "partial_equivalence_rel R \ True" and "\P. (True \ P) \ Trueprop P" using transport_bijection by (intro transport_preorder_on_in_field_bijection.intro) (insert partial_equivalence_rel_left partial_equivalence_rel_right, auto) lemmas partial_equivalence_rel_equivalence = partial_equivalence_rel_equivalenceI end locale transport_eq_restrict_bijection = fixes P :: "'a \ bool" and Q :: "'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes bijection_on_in_field: "bijection_on (in_field ((=\<^bsub>P\<^esub>) :: 'a \ _)) (in_field ((=\<^bsub>Q\<^esub>) :: 'b \ _)) l r" begin interpretation transport "(=\<^bsub>P\<^esub>)" "(=\<^bsub>Q\<^esub>)" l r . sublocale tper_bij? : transport_partial_equivalence_rel_bijection "(=\<^bsub>P\<^esub>)" "(=\<^bsub>Q\<^esub>)" l r using bijection_on_in_field partial_equivalence_rel_eq_restrict - eq_restrict_le_eq by unfold_locales (auto elim: bijection_onE intro!: - mono_wrt_rel_left_if_reflexive_on_if_le_eq_if_mono_wrt_in_field - [of "in_field (=\<^bsub>Q\<^esub>)"] - flip_of.mono_wrt_rel_left_if_reflexive_on_if_le_eq_if_mono_wrt_in_field - [of "in_field (=\<^bsub>P\<^esub>)"]) + mono_wrt_rel_left_if_reflexive_on_if_le_eq_if_mono_wrt_in_field[of "in_field (=\<^bsub>Q\<^esub>)"] + flip_of.mono_wrt_rel_left_if_reflexive_on_if_le_eq_if_mono_wrt_in_field[of "in_field (=\<^bsub>P\<^esub>)"]) lemma left_Galois_eq_Galois_eq_eq_restrict: "(\<^bsub>L\<^esub>\) = (galois_rel.Galois (=) (=) r)\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub>" by (subst galois_rel.left_Galois_restrict_left_eq_left_Galois_left_restrict_left galois_rel.left_Galois_restrict_right_eq_left_Galois_right_restrict_right - restrict_right_eq rel_inv_eq_self_if_symmetric)+ - (auto simp: eq_restrict_eq_eq_restrict_left) + bin_rel_restrict_right_eq rel_inv_eq_self_if_symmetric)+ + auto end locale transport_eq_bijection = fixes l :: "'a \ 'b" and r :: "'b \ 'a" assumes bijection_on_in_field: "bijection_on (in_field ((=) :: 'a \ _)) (in_field ((=) :: 'b \ _)) l r" begin sublocale teq_restr_bij? : transport_eq_restrict_bijection \ \ l r rewrites "(=\<^bsub>\ :: 'a \ bool\<^esub>) = ((=) :: 'a \ _)" and "(=\<^bsub>\ :: 'b \ bool\<^esub>) = ((=) :: 'b \ _)" using bijection_on_in_field by unfold_locales simp_all end end \ No newline at end of file diff --git a/thys/Transport/document/root.bib b/thys/Transport/document/root.bib --- a/thys/Transport/document/root.bib +++ b/thys/Transport/document/root.bib @@ -1,20 +1,23 @@ -@misc{transport, - title={{Transport via Partial Galois Connections and Equivalences}}, - author={{Kevin Kappelmann}}, - year={2023}, - eprint={2303.05244}, - archivePrefix={arXiv}, - primaryClass={cs.PL}, - doi={10.48550/arXiv.2303.05244} +@InProceedings{transport, + author="Kappelmann, Kevin", + editor="Hur, Chung-Kil", + title="Transport via Partial Galois Connections and Equivalences", + booktitle="Programming Languages and Systems", + year="2023", + publisher="Springer Nature Singapore", + address="Singapore", + pages="225--245", + abstract="Multiple types can represent the same concept. For example, lists and trees can both represent sets. Unfortunately, this easily leads to incomplete libraries: some set-operations may only be available on lists, others only on trees. Similarly, subtypes and quotients are commonly used to construct new type abstractions in formal verification. In such cases, one often wishes to reuse operations on the representation type for the new type abstraction, but to no avail: the types are not the same. To address these problems, we present a new framework that transports programs via equivalences. Existing transport frameworks are either designed for dependently typed, constructive proof assistants, use univalence, or are restricted to partial quotient types. Our framework (1) is designed for simple type theory, (2) generalises previous approaches working on partial quotient types, and (3) is based on standard mathematical concepts, particularly Galois connections and equivalences. We introduce the notions of partial Galois connection and equivalence and prove their closure properties under (dependent) function relators, (co)datatypes, and compositions. We formalised the framework in Isabelle/HOL and provide a prototype.", + isbn="978-981-99-8311-7" } @inproceedings{lifting, title={{Lifting and Transfer: A modular design for quotients in Isabelle/HOL}}, author={{Huffman, Brian and Kun{\v{c}}ar, Ond{\v{r}}ej}}, booktitle={Certified Programs and Proofs: Third International Conference, CPP 2013, Melbourne, VIC, Australia, December 11-13, 2013, Proceedings 3}, pages={131--146}, year={2013}, organization={Springer}, doi={10.1007/978-3-319-03545-1_9} }