diff --git a/thys/Transport/HOL_Basics/Adhoc_Overloading/Adhoc_Overloading.thy b/thys/Transport/HOL_Basics/Adhoc_Overloading/Adhoc_Overloading.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Adhoc_Overloading/Adhoc_Overloading.thy @@ -0,0 +1,252 @@ +theory Adhoc_Overloading + imports Pure + keywords "adhoc_overloading" "no_adhoc_overloading" :: thy_decl +begin + +(*adapted from HOL-Library*) + +ML\ +signature ADHOC_OVERLOADING = +sig + val is_overloaded: Proof.context -> string -> bool + val generic_add_overloaded: string -> Context.generic -> Context.generic + val generic_remove_overloaded: string -> Context.generic -> Context.generic + val generic_add_variant: string -> term -> Context.generic -> Context.generic + (*If the list of variants is empty at the end of "generic_remove_variant", then + "generic_remove_overloaded" is called implicitly.*) + val generic_remove_variant: string -> term -> Context.generic -> Context.generic + val show_variants: bool Config.T +end + +structure Adhoc_Overloading: ADHOC_OVERLOADING = +struct + +val show_variants = Attrib.setup_config_bool \<^binding>\show_variants\ (K false); + + +(* errors *) + +fun err_duplicate_variant oconst = + error ("Duplicate variant of " ^ quote oconst); + +fun err_not_a_variant oconst = + error ("Not a variant of " ^ quote oconst); + +fun err_not_overloaded oconst = + error ("Constant " ^ quote oconst ^ " is not declared as overloaded"); + +fun err_unresolved_overloading ctxt0 (c, T) t instances = + let + val ctxt = Config.put show_variants true ctxt0 + val const_space = Proof_Context.const_space ctxt + val prt_const = + Pretty.block [Name_Space.pretty ctxt const_space c, Pretty.str " ::", Pretty.brk 1, + Pretty.quote (Syntax.pretty_typ ctxt T)] + in + error (Pretty.string_of (Pretty.chunks [ + Pretty.block [ + Pretty.str "Unresolved adhoc overloading of constant", Pretty.brk 1, + prt_const, Pretty.brk 1, Pretty.str "in term", Pretty.brk 1, + Pretty.block [Pretty.quote (Syntax.pretty_term ctxt t)]], + Pretty.block ( + (if null instances then [Pretty.str "no instances"] + else Pretty.fbreaks ( + Pretty.str "multiple instances:" :: + map (Pretty.mark Markup.item o Syntax.pretty_term ctxt) instances)))])) + end; + + +(* generic data *) + +fun variants_eq ((v1, T1), (v2, T2)) = + Term.aconv_untyped (v1, v2) andalso T1 = T2; + +structure Overload_Data = Generic_Data +( + type T = + {variants : (term * typ) list Symtab.table, + oconsts : string Termtab.table}; + val empty = {variants = Symtab.empty, oconsts = Termtab.empty}; + + fun merge + ({variants = vtab1, oconsts = otab1}, + {variants = vtab2, oconsts = otab2}) : T = + let + fun merge_oconsts _ (oconst1, oconst2) = + if oconst1 = oconst2 then oconst1 + else err_duplicate_variant oconst1; + in + {variants = Symtab.merge_list variants_eq (vtab1, vtab2), + oconsts = Termtab.join merge_oconsts (otab1, otab2)} + end; +); + +fun map_tables f g = + Overload_Data.map (fn {variants = vtab, oconsts = otab} => + {variants = f vtab, oconsts = g otab}); + +val is_overloaded = Symtab.defined o #variants o Overload_Data.get o Context.Proof; +val get_variants = Symtab.lookup o #variants o Overload_Data.get o Context.Proof; +val get_overloaded = Termtab.lookup o #oconsts o Overload_Data.get o Context.Proof; + +fun generic_add_overloaded oconst context = + if is_overloaded (Context.proof_of context) oconst then context + else map_tables (Symtab.update (oconst, [])) I context; + +fun generic_remove_overloaded oconst context = + let + fun remove_oconst_and_variants context oconst = + let + val remove_variants = + (case get_variants (Context.proof_of context) oconst of + NONE => I + | SOME vs => fold (Termtab.remove (op =) o rpair oconst o fst) vs); + in map_tables (Symtab.delete_safe oconst) remove_variants context end; + in + if is_overloaded (Context.proof_of context) oconst then remove_oconst_and_variants context oconst + else err_not_overloaded oconst + end; + +local + fun generic_variant add oconst t context = + let + val ctxt = Context.proof_of context; + val _ = if is_overloaded ctxt oconst then () else err_not_overloaded oconst; + val T = t |> fastype_of; + val t' = Term.map_types (K dummyT) t; + in + if add then + let + val _ = + (case get_overloaded ctxt t' of + NONE => () + | SOME oconst' => err_duplicate_variant oconst'); + in + map_tables (Symtab.cons_list (oconst, (t', T))) (Termtab.update (t', oconst)) context + end + else + let + val _ = + if member variants_eq (the (get_variants ctxt oconst)) (t', T) then () + else err_not_a_variant oconst; + in + map_tables (Symtab.map_entry oconst (remove1 variants_eq (t', T))) + (Termtab.delete_safe t') context + |> (fn context => + (case get_variants (Context.proof_of context) oconst of + SOME [] => generic_remove_overloaded oconst context + | _ => context)) + end + end; +in + val generic_add_variant = generic_variant true; + val generic_remove_variant = generic_variant false; +end; + + +(* check / uncheck *) + +fun unifiable_with thy T1 T2 = + let + val maxidx1 = Term.maxidx_of_typ T1; + val T2' = Logic.incr_tvar (maxidx1 + 1) T2; + val maxidx2 = Term.maxidx_typ T2' maxidx1; + in can (Sign.typ_unify thy (T1, T2')) (Vartab.empty, maxidx2) end; + +fun get_candidates ctxt (c, T) = + get_variants ctxt c + |> Option.map (map_filter (fn (t, T') => + if unifiable_with (Proof_Context.theory_of ctxt) T T' + (*keep the type constraint for the type-inference check phase*) + then SOME (Type.constraint T t) + else NONE)); + +fun insert_variants ctxt t (oconst as Const (c, T)) = + (case get_candidates ctxt (c, T) of + SOME [] => err_unresolved_overloading ctxt (c, T) t [] + | SOME [variant] => variant + | _ => oconst) + | insert_variants _ _ oconst = oconst; + +fun insert_overloaded ctxt = + let + fun proc t = + Term.map_types (K dummyT) t + |> get_overloaded ctxt + |> Option.map (Const o rpair (Term.type_of t)); + in + Pattern.rewrite_term_top (Proof_Context.theory_of ctxt) [] [proc] + end; + +fun check ctxt = + map (fn t => Term.map_aterms (insert_variants ctxt t) t); + +fun uncheck ctxt ts = + if Config.get ctxt show_variants orelse exists (is_none o try Term.type_of) ts then ts + else map (insert_overloaded ctxt) ts; + +fun reject_unresolved ctxt = + let + val the_candidates = the o get_candidates ctxt; + fun check_unresolved t = + (case filter (is_overloaded ctxt o fst) (Term.add_consts t []) of + [] => t + | (cT :: _) => err_unresolved_overloading ctxt cT t (the_candidates cT)); + in map check_unresolved end; + + +(* setup *) + +val _ = Context.>> + (Syntax_Phases.term_check 0 "adhoc_overloading" check + #> Syntax_Phases.term_check 1 "adhoc_overloading_unresolved_check" reject_unresolved + #> Syntax_Phases.term_uncheck 0 "adhoc_overloading" uncheck); + + +(* commands *) + +fun generic_adhoc_overloading_cmd add = + if add then + fold (fn (oconst, ts) => + generic_add_overloaded oconst + #> fold (generic_add_variant oconst) ts) + else + fold (fn (oconst, ts) => + fold (generic_remove_variant oconst) ts); + +fun adhoc_overloading_cmd' add args phi = + let val args' = args + |> map (apsnd (map_filter (fn t => + let val t' = Morphism.term phi t; + in if Term.aconv_untyped (t, t') then SOME t' else NONE end))); + in generic_adhoc_overloading_cmd add args' end; + +fun adhoc_overloading_cmd add raw_args lthy = + let + fun const_name ctxt = + fst o dest_Const o Proof_Context.read_const {proper = false, strict = false} ctxt; (* FIXME {proper = true, strict = true} (!?) *) + fun read_term ctxt = singleton (Variable.polymorphic ctxt) o Syntax.read_term ctxt; + val args = + raw_args + |> map (apfst (const_name lthy)) + |> map (apsnd (map (read_term lthy))); + in + Local_Theory.declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} + (adhoc_overloading_cmd' add args) lthy + end; + +val _ = + Outer_Syntax.local_theory \<^command_keyword>\adhoc_overloading\ + "add adhoc overloading for constants / fixed variables" + (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd true); + +val _ = + Outer_Syntax.local_theory \<^command_keyword>\no_adhoc_overloading\ + "delete adhoc overloading for constants / fixed variables" + (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd false); + +end; +\ + + +end \ No newline at end of file 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,342 +1,429 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Binary Relations\ subsection \Basic Functions\ theory Binary_Relation_Functions imports - HOL_Basics_Base - HOL_Syntax_Bundles_Lattices - ML_Unification.ML_Unification_HOL_Setup + Bounded_Quantifiers + ML_Unification.Unify_Resolve_Tactics begin paragraph \Summary\ text \Basic functions on binary relations.\ subsubsection \Domain, Codomain, and Field\ 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 -definition "in_codom R y \ \x. R x y" +lemma in_dom_bottom_eq_bottom [simp]: "in_dom \ = \" by fastforce +lemma in_dom_top_eq_top [simp]: "in_dom \ = \" by fastforce + +lemma in_dom_sup_eq_in_dom_sup_in_dom [simp]: "in_dom (R \ S) = in_dom R \ in_dom S" by fastforce + +consts in_codom_on :: "'a \ 'b \ 'c \ bool" + +definition "in_codom_on_pred (P :: 'a \ bool) R y \ \x : P. R x y" +adhoc_overloading in_codom_on in_codom_on_pred + +lemma in_codom_onI [intro]: + assumes "R x y" + and "P x" + shows "in_codom_on P R y" + using assms unfolding in_codom_on_pred_def by blast + +lemma in_codom_onE [elim]: + assumes "in_codom_on P R y" + obtains x where "P x" "R x y" + using assms unfolding in_codom_on_pred_def by blast + +lemma in_codom_on_pred_iff_bex_rel: "in_codom_on P R y \ (\x : P. R x y)" by blast + +lemma in_codom_bottom_pred_eq_bottom [simp]: "in_codom_on \ = \" by fastforce +lemma in_codom_bottom_rel_eq_bottom [simp]: "in_codom_on P \ = \" by fastforce +lemma in_codom_top_top_eq [simp]: "in_codom_on \ \ = \" by fastforce + +lemma in_codom_on_eq_pred_eq [simp]: "in_codom_on ((=) P) R = R P" + by auto + +lemma in_codom_on_sup_pred_eq_in_codom_on_sup_in_codom_on [simp]: + "in_codom_on (P \ Q) = in_codom_on P \ in_codom_on Q" + by fastforce + +lemma in_codom_on_sup_rel_eq_in_codom_on_sup_in_codom_on [simp]: + "in_codom_on P (R \ S) = in_codom_on P R \ in_codom_on P S" + by fastforce + +definition "in_codom \ in_codom_on (\ :: 'a \ bool)" + +lemma in_codom_eq_in_codom_on_top: "in_codom = in_codom_on \" unfolding in_codom_def by auto + +lemma in_codom_eq_in_codom_on_top_uhint [uhint]: + assumes "P \ (\ :: 'a \ bool)" + shows "in_codom \ in_codom_on P" + using assms by (simp add: in_codom_eq_in_codom_on_top) lemma in_codomI [intro]: assumes "R x y" shows "in_codom R y" - using assms unfolding in_codom_def by blast + using assms by (urule in_codom_onI) simp lemma in_codomE [elim]: assumes "in_codom R y" obtains x where "R x y" - using assms unfolding in_codom_def by blast + using assms by (urule (e) in_codom_onE) 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 R x \ in_dom R x \ in_codom R x" 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') +lemma in_field_bottom_eq_bottom [simp]: "in_field \ = \" by fastforce +lemma in_field_top_eq_top [simp]: "in_field \ = \" by fastforce + +lemma in_field_sup_eq_in_field_sup_in_field [simp]: "in_field (R \ S) = in_field R \ in_field S" + by fastforce subsubsection \Composition\ -definition "rel_comp R S x y \ \z. R x z \ S z y" +consts rel_comp :: "'a \ 'b \ 'c" 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 +definition "rel_comp_rel R S x y \ \z. R x z \ S z y" +adhoc_overloading rel_comp rel_comp_rel + 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 + using assms unfolding rel_comp_rel_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 + using assms unfolding rel_comp_rel_def by blast lemma rel_comp_assoc: "R \\ (S \\ T) = (R \\ S) \\ T" by (intro ext) blast lemma in_dom_if_in_dom_rel_comp: assumes "in_dom (R \\ S) x" shows "in_dom R x" using assms 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_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 subsubsection \Inverse\ -definition rel_inv :: "('a \ 'b \ bool) \ 'b \ 'a \ bool" - where "rel_inv R x y \ R y x" +consts rel_inv :: "'a \ 'b" 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 +definition rel_inv_rel :: "('a \ 'b \ bool) \ 'b \ 'a \ bool" + where "rel_inv_rel R x y \ R y x" +adhoc_overloading rel_inv rel_inv_rel + lemma rel_invI [intro]: assumes "R x y" shows "R\ y x" - using assms unfolding rel_inv_def . + using assms unfolding rel_inv_rel_def . lemma rel_invD [dest]: assumes "R\ x y" shows "R y x" - using assms unfolding rel_inv_def . + using assms unfolding rel_inv_rel_def . lemma rel_inv_iff_rel [simp]: "R\ x y \ R y x" by blast -lemma in_dom_rel_inv_eq_in_codom [simp]: "in_dom (R\) = in_codom R" +lemma in_dom_rel_inv_eq_in_codom [simp]: "in_dom R\ = in_codom R" by (intro ext) blast -lemma in_codom_rel_inv_eq_in_dom [simp]: "in_codom (R\) = in_dom R" +lemma in_codom_rel_inv_eq_in_dom [simp]: "in_codom R\ = in_dom R" by (intro ext) blast -lemma in_field_rel_inv_eq [simp]: "in_field R\ = in_field R" +lemma in_field_rel_inv_eq_in_field [simp]: "in_field R\ = in_field R" by (intro ext) auto 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) lemma rel_inv_le_rel_inv_iff [iff]: "R\ \ S\ \ R \ S" by (auto intro: predicate2I dest: predicate2D) +lemma rel_inv_top_eq [simp]: "\\ = \" by fastforce +lemma rel_inv_bottom_eq [simp]: "\\ = \" by fastforce + subsubsection \Restrictions\ -definition "rel_if B R x y \ B \ R x y" +consts rel_if :: "bool \ 'a \ 'a" 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 +definition "rel_if_rel B R x y \ B \ R x y" +adhoc_overloading rel_if rel_if_rel + lemma rel_if_eq_rel_if_pred [simp]: assumes "B" shows "(rel_if B R) = R" - unfolding rel_if_def using assms by blast + unfolding rel_if_rel_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 + unfolding rel_if_rel_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 + unfolding rel_if_rel_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 + using assms unfolding rel_if_rel_def by blast lemma rel_ifD: assumes "(rel_if B R) x y" and "B" shows "R x y" using assms by blast consts rel_restrict_left :: "'a \ 'b \ 'a" consts rel_restrict_right :: "'a \ 'b \ 'a" -overloading - rel_restrict_left_pred \ "rel_restrict_left :: ('a \ 'b \ bool) \ ('a \ bool) \ 'a \ 'b \ bool" - rel_restrict_right_pred \ "rel_restrict_right :: ('a \ 'b \ bool) \ ('b \ bool) \ 'a \ 'b \ bool" -begin - definition "rel_restrict_left_pred R P x y \ P x \ R x y" - definition "rel_restrict_right_pred R P x y \ P y \ R x y" -end - bundle rel_restrict_syntax begin notation rel_restrict_left ("(_)\(\<^bsub>_\<^esub>)" [1000]) notation rel_restrict_right ("(_)\(\<^bsub>_\<^esub>)" [1000]) end bundle no_rel_restrict_syntax begin no_notation rel_restrict_left ("(_)\(\<^bsub>_\<^esub>)" [1000]) no_notation rel_restrict_right ("(_)\(\<^bsub>_\<^esub>)" [1000]) end unbundle rel_restrict_syntax +definition "rel_restrict_left_pred R P x y \ P x \ R x y" +adhoc_overloading rel_restrict_left rel_restrict_left_pred +definition "rel_restrict_right_pred R P x y \ P y \ R x y" +adhoc_overloading rel_restrict_right rel_restrict_right_pred + lemma rel_restrict_leftI [intro]: assumes "R x y" and "P x" shows "R\\<^bsub>P\<^esub> x y" using assms unfolding rel_restrict_left_pred_def by blast lemma rel_restrict_leftE [elim]: assumes "R\\<^bsub>P\<^esub> x y" obtains "P x" "R x y" using assms unfolding rel_restrict_left_pred_def by blast lemma rel_restrict_left_cong: assumes "\x. P x \ P' x" and "\x y. P' x \ R x y \ R' x y" shows "R\\<^bsub>P\<^esub> = R'\\<^bsub>P'\<^esub>" using assms by (intro ext iffI) blast+ -lemma rel_restrict_left_top_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>\ ::'a \ bool\<^esub> = R" +lemma rel_restrict_left_top_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>\ :: 'a \ bool\<^esub> = R" by (intro ext rel_restrict_leftI) auto lemma rel_restrict_left_top_eq_uhint [uhint]: - assumes "P \ (\ ::'a \ bool)" + assumes "P \ (\ :: 'a \ bool)" shows "(R :: 'a \ 'b \ bool)\\<^bsub>P\<^esub> \ R" using assms by simp +lemma rel_restrict_left_bottom_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>\ :: 'a \ bool\<^esub> = \" + by (intro ext rel_restrict_leftI) auto + +lemma bottom_restrict_left_eq [simp]: "\\\<^bsub>P :: 'a \ bool\<^esub> = \" + by fastforce + lemma rel_restrict_left_le_self: "(R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'a \ bool)\<^esub> \ R" by (auto intro: predicate2I dest: predicate2D) +lemma le_rel_restrict_left_self_if_in_dom_le: + assumes "in_dom R \ P" + shows "R \ (R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'a \ bool)\<^esub>" + using assms by (auto intro: predicate2I dest: predicate2D predicate1D) + +corollary rel_restrict_left_eq_self_if_in_dom_le [simp]: + assumes "in_dom R \ P" + shows "R\\<^bsub>P\<^esub> = R" + using assms rel_restrict_left_le_self le_rel_restrict_left_self_if_in_dom_le by (intro antisym) + +lemma ex_rel_restrict_left_iff_in_codom_on [iff]: "(\x. R\\<^bsub>P\<^esub> x y) \ (in_codom_on P R y)" by blast + lemma rel_restrict_rightI [intro]: assumes "R x y" and "P y" shows "R\\<^bsub>P\<^esub> x y" using assms unfolding rel_restrict_right_pred_def by blast lemma rel_restrict_rightE [elim]: assumes "R\\<^bsub>P\<^esub> x y" obtains "P y" "R x y" using assms unfolding rel_restrict_right_pred_def by blast lemma rel_restrict_right_cong: assumes "\y. P y \ P' y" and "\x y. P' y \ R x y \ R' x y" shows "R\\<^bsub>P\<^esub> = R'\\<^bsub>P'\<^esub>" using assms by (intro ext iffI) blast+ lemma rel_restrict_right_top_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>\ ::'b \ bool\<^esub> = R" by (intro ext rel_restrict_rightI) auto lemma rel_restrict_right_top_eq_uhint [uhint]: assumes "P \ (\ ::'b \ bool)" shows "(R :: 'a \ 'b \ bool)\\<^bsub>P\<^esub> \ R" using assms by simp +lemma rel_restrict_right_bottom_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>\ :: 'b \ bool\<^esub> = \" + by (intro ext rel_restrict_rightI) auto + +lemma bottom_restrict_right_eq [simp]: "\\\<^bsub>P :: 'b \ bool\<^esub> = \" + by fastforce + lemma rel_restrict_right_le_self: "(R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'b \ bool)\<^esub> \ R" by (auto intro: predicate2I dest: predicate2D) +lemma le_rel_restrict_right_self_if_in_codom_le: + assumes "in_codom R \ P" + shows "R \ (R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'b \ bool)\<^esub>" + using assms by (auto intro: predicate2I dest: predicate2D predicate1D) + +corollary rel_restrict_right_eq_self_if_in_codom_le [simp]: + assumes "in_codom R \ P" + shows "R\\<^bsub>P\<^esub> = R" + using assms rel_restrict_right_le_self le_rel_restrict_right_self_if_in_codom_le + by (intro antisym) + context fixes R :: "'a \ 'b \ bool" begin lemma 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 rel_restrict_right_iff_restrict_left: "R\\<^bsub>P :: 'b \ bool\<^esub> x y \ (R\)\\<^bsub>P\<^esub> y x" unfolding rel_restrict_right_eq by simp end lemma rel_inv_rel_restrict_left_inv_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 rel_restrict_leftI rel_invI) auto lemma 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 rel_restrict_right_eq by (fact rel_inv_rel_restrict_left_inv_rel_restrict_left_eq) lemma in_dom_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_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_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_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_rel_restrict_leftE [elim]: assumes "in_codom R\\<^bsub>P\<^esub> y" obtains x where "P x" "R x y" using assms by blast subsubsection \Mappers\ 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_Agree.thy b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Agree.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Agree.thy @@ -0,0 +1,59 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Agreement\ +theory Binary_Relations_Agree + imports + Functions_Monotone +begin + +consts rel_agree_on :: "'a \ 'b \ bool" + +overloading + rel_agree_on_pred \ "rel_agree_on :: ('a \ bool) \ (('a \ 'b \ bool) \ bool) \ bool" +begin + definition "rel_agree_on_pred (P :: 'a \ bool) (\ :: ('a \ 'b \ bool) \ bool) \ + \R R' : \. R\\<^bsub>P\<^esub> = R'\\<^bsub>P\<^esub>" +end + +lemma rel_agree_onI [intro]: + assumes "\R R' x y. \ R \ \ R' \ P x \ R x y \ R' x y" + shows "rel_agree_on P \" + using assms unfolding rel_agree_on_pred_def by blast + +lemma rel_agree_onE: + assumes "rel_agree_on P \" + obtains "\R R'. \ R \ \ R' \ R\\<^bsub>P\<^esub> = R'\\<^bsub>P\<^esub>" + using assms unfolding rel_agree_on_pred_def by blast + +lemma rel_restrict_left_eq_if_rel_agree_onI: + assumes "rel_agree_on P \" + and "\ R" "\ R'" + shows "R\\<^bsub>P\<^esub> = R'\\<^bsub>P\<^esub>" + using assms by (blast elim: rel_agree_onE) + +lemma rel_agree_onD: + assumes "rel_agree_on P \" + and "\ R" "\ R'" + and "P x" + and "R x y" + shows "R' x y" + using assms by (blast elim: rel_agree_onE dest: fun_cong) + +lemma antimono_rel_agree_on: + "((\) \\<^sub>m (\) \ (\)) (rel_agree_on :: ('a \ bool) \ (('a \ 'b \ bool) \ bool) \ bool)" + by (intro mono_wrt_relI Fun_Rel_relI) (fastforce dest: rel_agree_onD) + +lemma le_if_in_dom_le_if_rel_agree_onI: + assumes "rel_agree_on A \" + and "\ R" "\ R'" + and "in_dom R \ A" + shows "R \ R'" + using assms by (auto dest: rel_agree_onD[where ?R="R"]) + +lemma eq_if_in_dom_le_if_rel_agree_onI: + assumes "rel_agree_on A \" + and "\ R" "\ R'" + and "in_dom R \ A" "in_dom R' \ A" + shows "R = R'" + using assms le_if_in_dom_le_if_rel_agree_onI by blast + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Extend.thy b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Extend.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Extend.thy @@ -0,0 +1,99 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Extensions\ +theory Binary_Relations_Extend + imports + Dependent_Binary_Relations +begin + +consts extend :: "'a \ 'b \ 'c \ 'c" + +definition "extend_rel x y R x' y' \ (x = x' \ y = y') \ R x' y'" +adhoc_overloading extend extend_rel + +lemma extend_leftI [iff]: "(extend x y R) x y" + unfolding extend_rel_def by blast + +lemma extend_rightI [intro]: + assumes "R x' y'" + shows "(extend x y R) x' y'" + unfolding extend_rel_def using assms by blast + +lemma extendE [elim]: + assumes "(extend x y R) x' y'" + obtains "x = x'" "y = y'" | "x \ x' \ y \ y'" "R x' y'" + using assms unfolding extend_rel_def by blast + +lemma extend_eq_self_if_rel [simp]: "R x y \ extend x y R = R" + by auto + +context + fixes x :: 'a and y :: 'b and R :: "'a \ 'b \ bool" +begin + +lemma in_dom_extend_eq: "in_dom (extend x y R) = in_dom R \ (=) x" + by (intro ext) auto + +lemma in_dom_extend_iff: "in_dom (extend x y R) x' \ in_dom R x' \ x = x'" + by (auto simp: in_dom_extend_eq) + +lemma codom_extend_eq: "in_codom (extend x y R) = in_codom R \ (=) y" + by (intro ext) auto + +lemma in_codom_extend_iff: "in_codom (extend x y R) y' \ in_codom R y' \ y = y'" + by (auto simp: codom_extend_eq) + +end + +lemma in_field_extend_eq: "in_field (extend x y R) = in_field R \ (=) x \ (=) y" + by (intro ext) auto + +lemma in_field_extend_iff: "in_field (extend x y R) z \ in_field R z \ z = x \ z = y" + by (auto simp: in_field_extend_eq) + +lemma mono_extend: "mono (extend x y :: ('a \ 'b \ bool) \ 'a \ 'b \ bool)" + by (intro monoI) force + +lemma dep_mono_dep_bin_rel_extend: + "((x : A) \\<^sub>m B x \\<^sub>m ({\}x : A'. B' x) \\<^sub>m ({\}x : A \ A'. (B \ B') x)) extend" + by fastforce + +consts glue :: "'a \ 'b" + +definition "glue_rel (\ :: ('a \ 'b \ bool) \ bool) x \ in_codom_on \ (\R. R x)" +adhoc_overloading glue glue_rel + +lemma glue_rel_eq_in_codom_on: "glue \ x = in_codom_on \ (\R. R x)" + unfolding glue_rel_def by simp + +lemma glueI [intro]: + assumes "\ R" + and "R x y" + shows "glue \ x y" + using assms unfolding glue_rel_def by blast + +lemma glueE [elim!]: + assumes "glue \ x y" + obtains R where "\ R" "R x y" + using assms unfolding glue_rel_def by blast + +lemma glue_bottom_eq [simp]: "glue (\ :: ('a \ 'b \ bool) \ bool) = \" + by (intro ext) auto + +lemma glue_eq_rel_eq_self [simp]: "glue ((=) R) = (R :: 'a \ 'b \ bool)" + by (intro ext) auto + +lemma glue_sup_eq_glue_sup_glue [simp]: "glue (A \ B) = glue A \ glue B" + supply glue_rel_eq_in_codom_on[simp] + by (rule ext) (urule in_codom_on_sup_pred_eq_in_codom_on_sup_in_codom_on) + +lemma mono_glue: "mono (glue :: (('a \ 'b \ bool) \ bool) \ ('a \ 'b \ bool))" + by auto + +lemma dep_bin_rel_glueI: + fixes \ defines [simp]: "D \ in_codom_on \ in_dom" + assumes "\R. \ R \ \A. ({\}x : A. B x) R" + shows "({\}x : D. B x) (glue \)" + using assms by (intro dep_bin_relI) auto + + +end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Dependent_Binary_Relations.thy b/thys/Transport/HOL_Basics/Binary_Relations/Dependent_Binary_Relations.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Dependent_Binary_Relations.thy @@ -0,0 +1,164 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Dependent Binary Relations\ +theory Dependent_Binary_Relations + imports + Binary_Relations_Agree +begin + +consts dep_bin_rel :: "'a \ ('b \ 'c) \ 'd \ bool" +consts bin_rel :: "'a \ 'b \ 'c \ bool" + +bundle bin_rel_syntax +begin +syntax "_dep_bin_rel" :: \idt \ 'a \ 'b \ 'c \ 'd\ ("{\}_ : _./ _" [41, 41, 40] 51) +notation bin_rel (infixl "{\}" 51) +end +bundle no_bin_rel_syntax +begin +no_syntax "_dep_bin_rel" :: \idt \ 'a \ 'b \ 'c \ 'd\ ("{\}_ : _./ _" [41, 41, 40] 51) +no_notation bin_rel (infixl "{\}" 51) +end +unbundle bin_rel_syntax +translations + "{\}x : A. B" \ "CONST dep_bin_rel A (\x. B)" + +definition "dep_bin_rel_pred (A :: 'a \ bool) (B :: 'a \ 'b \ bool) (R :: 'a \ 'b \ bool) \ + \x y. R x y \ A x \ B x y" +adhoc_overloading dep_bin_rel dep_bin_rel_pred + +definition "bin_rel_pred (A :: 'a \ bool) (B :: 'b \ bool) :: ('a \ 'b \ bool) \ bool \ + {\}(_ :: 'a) : A. B" +adhoc_overloading bin_rel bin_rel_pred + +lemma bin_rel_pred_eq_dep_bin_rel_pred: "A {\} B = {\}_ : A. B" + unfolding bin_rel_pred_def by auto + +lemma bin_rel_pred_eq_dep_bin_rel_pred_uhint [uhint]: + assumes "A \ A'" + and "B' \ (\_. B)" + shows "A {\} B \ {\}x : A'. B' x" + using assms by (simp add: bin_rel_pred_eq_dep_bin_rel_pred) + +lemma bin_rel_pred_iff_dep_bin_rel_pred: "(A {\} B) R \ ({\}_ : A. B) R" + unfolding bin_rel_pred_eq_dep_bin_rel_pred by auto + +lemma dep_bin_relI [intro]: + assumes "\x y. R x y \ A x" + and "\x y. R x y \ A x \ B x y" + shows "({\}x : A. B x) R" + using assms unfolding dep_bin_rel_pred_def by auto + +lemma dep_bin_rel_if_bin_rel_and: + assumes "\x y. R x y \ A x \ B x y" + shows "({\}x : A. B x) R" + using assms by auto + +lemma dep_bin_relE [elim]: + assumes "({\}x : A. B x) R" + and "R x y" + obtains "A x" "B x y" + using assms unfolding dep_bin_rel_pred_def by auto + +lemma dep_bin_relE': + assumes "({\}x : A. B x) R" + obtains "\x y. R x y \ A x \ B x y" + using assms by auto + +lemma bin_relI [intro]: + assumes "\x y. R x y \ A x" + and "\x y. R x y \ A x \ B y" + shows "(A {\} B) R" + using assms by (urule dep_bin_relI where chained = fact) + +lemma bin_rel_if_bin_rel_and: + assumes "\x y. R x y \ A x \ B y" + shows "(A {\} B) R" + using assms by (urule dep_bin_rel_if_bin_rel_and) + +lemma bin_relE [elim]: + assumes "(A {\} B) R" + and "R x y" + obtains "A x" "B y" + using assms by (urule (e) dep_bin_relE) + +lemma bin_relE': + assumes "(A {\} B) R" + obtains "\x y. R x y \ A x \ B y" + using assms by (urule (e) dep_bin_relE') + +lemma dep_bin_rel_cong [cong]: + "\A = A'; \x. A' x \ B x = B' x\ \ ({\}x : A. B x) = {\}x : A'. B' x" + by (intro ext iffI dep_bin_relI) fastforce+ + +lemma le_dep_bin_rel_if_le_dom: + assumes "A \ A'" + shows "({\}x : A. B x) \ ({\}x : A'. B x)" + using assms by (intro le_predI dep_bin_relI) auto + +lemma dep_bin_rel_covariant_codom: + assumes "({\}x : A. B x) R" + and "\x y. R x y \ A x \ B x y \ B' x y" + shows "({\}x : A. B' x) R" + using assms by (intro dep_bin_relI) auto + +lemma mono_dep_bin_rel: "((\) \\<^sub>m (\) \ (\) \ (\)) dep_bin_rel" + by (intro mono_wrt_relI Fun_Rel_relI dep_bin_relI) force + +lemma mono_bin_rel: "((\) \\<^sub>m (\) \ (\) \ (\)) ({\})" + by (intro mono_wrt_relI Fun_Rel_relI) auto + +lemma in_dom_le_if_dep_bin_rel: + assumes "({\}x : A. B x) R" + shows "in_dom R \ A" + using assms by auto + +lemma in_codom_le_in_codom_on_if_dep_bin_rel: + assumes "({\}x : A. B x) R" + shows "in_codom R \ in_codom_on A B" + using assms by fast + +lemma rel_restrict_left_eq_self_if_dep_bin_rel [simp]: + assumes "({\}x : A. B x) R" + shows "R\\<^bsub>A\<^esub> = R" + using assms rel_restrict_left_eq_self_if_in_dom_le by auto + +lemma dep_bin_rel_bottom_dom_iff_eq_bottom [iff]: "({\}x : \. B x) R \ R = \" + by fastforce + +lemma dep_bin_rel_bottom_codom_iff_eq_bottom [iff]: "({\}x : A. \) R \ R = \" + by fastforce + +lemma mono_bin_rel_dep_bin_rel_bin_rel_rel_comp: + "(A {\} B \\<^sub>m ({\}x : B. C x) \\<^sub>m A {\} in_codom_on B C) (\\)" + by fastforce + +lemma mono_dep_bin_rel_bin_rel_rel_inv: "(({\}x : A. B x) \\<^sub>m in_codom_on A B {\} A) rel_inv" + by force + +lemma mono_bin_rel_rel_inv: "(A {\} B \\<^sub>m B {\} A) rel_inv" + by auto + +lemma mono_dep_bin_rel_top_dep_bin_rel_inf_rel_restrict_left: + "(({\}x : A. B x) \\<^sub>m (P : \) \\<^sub>m ({\}x : A \ P. B x)) rel_restrict_left" + by fast + +lemma mono_dep_bin_rel_top_dep_bin_rel_inf_rel_restrict_right: + "(({\}x : A. B x) \\<^sub>m (P : \) \\<^sub>m ({\}x : A. (B x) \ P)) rel_restrict_right" + by fast + +lemma le_if_rel_agree_on_if_dep_bin_relI: + assumes "({\}x : A. B x) R" + and "rel_agree_on A \" + and "\ R" "\ R'" + shows "R \ R'" + using assms by (intro le_if_in_dom_le_if_rel_agree_onI in_dom_le_if_dep_bin_rel) + +lemma eq_if_rel_agree_on_if_dep_bin_relI: + assumes "({\}x : A. B x) R" "({\}x : A. B' x) R'" + and "rel_agree_on A \" + and "\ R" "\ R'" + shows "R = R'" + using assms by (intro eq_if_in_dom_le_if_rel_agree_onI in_dom_le_if_dep_bin_rel) + + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Clean_Function.thy b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Clean_Function.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Clean_Function.thy @@ -0,0 +1,252 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Clean Functions\ +theory Binary_Relations_Clean_Function + imports + Binary_Relations_Function_Base +begin + +text \Clean function relations may not contain further elements outside their specification.\ + +consts crel_dep_fun :: "'a \ ('b \ 'c) \ 'd \ bool" +consts crel_fun :: "'a \ 'b \ 'c \ bool" + +bundle crel_fun_syntax +begin +syntax + "_crel_fun" :: "'a \ 'b \ 'c \ bool" ("(_) \\<^sub>c (_)" [41, 40] 40) + "_crel_dep_fun" :: "idt \ 'a \ 'b \ 'c \ bool" ("'(_/ :/ _') \\<^sub>c (_)" [41, 41, 40] 40) +end +bundle no_crel_fun_syntax +begin +no_syntax + "_crel_fun" :: "'a \ 'b \ 'c \ bool" ("(_) \\<^sub>c (_)" [41, 40] 40) + "_crel_dep_fun" :: "idt \ 'a \ 'b \ 'c \ bool" ("'(_/ :/ _') \\<^sub>c (_)" [41, 41, 40] 40) +end +unbundle crel_fun_syntax +translations + "A \\<^sub>c B" \ "CONST crel_fun A B" + "(x : A) \\<^sub>c B" \ "CONST crel_dep_fun A (\x. B)" + +definition "crel_dep_fun_pred (A :: 'a \ bool) B R \ ((x : A) \ B x) R \ in_dom R = A" +adhoc_overloading crel_dep_fun crel_dep_fun_pred + +definition "crel_fun_pred (A :: 'a \ bool) B \ (((_ :: 'a) : A) \\<^sub>c B)" +adhoc_overloading crel_fun crel_fun_pred + +lemma crel_fun_eq_crel_dep_fun: + "(((A :: 'a \ bool) \\<^sub>c (B :: 'b \ bool)) :: ('a \ 'b \ bool) \ bool) = (((_ :: 'a) : A) \\<^sub>c B)" + by (simp add: crel_fun_pred_def) + +lemma crel_fun_eq_crel_dep_fun_uhint [uhint]: + assumes "(A :: 'a \ bool) \ A'" + and "B' \ (\(_ :: 'a). (B :: 'b \ bool))" + shows "(A \\<^sub>c B) \ ((x : A') \\<^sub>c B' x)" + using assms by (simp add: crel_fun_eq_crel_dep_fun) + +lemma crel_fun_iff_crel_dep_fun: + "((A :: 'a \ bool) \\<^sub>c (B :: 'b \ bool)) (R :: 'a \ 'b \ bool) \ (((_ :: 'a) : A) \\<^sub>c B) R" + by (simp add: crel_fun_pred_def) + +lemma crel_dep_funI [intro]: + assumes "((x : A) \ B x) R" + and "in_dom R \ A" + shows "((x : A) \\<^sub>c B x) R" + unfolding crel_dep_fun_pred_def using assms + by (intro conjI antisym le_in_dom_if_left_total_on) auto + +lemma crel_dep_funI': + assumes "left_total_on A R" + and "right_unique_on A R" + and "({\}x : A. B x) R" + shows "((x : A) \\<^sub>c B x) R" +proof (intro crel_dep_funI rel_dep_funI dep_mono_wrt_predI) + fix x assume "A x" + with assms obtain y where "B x y" "R x y" by auto + moreover with assms have "R`x = y" by (auto intro: eval_eq_if_right_unique_onI) + ultimately show "B x (R`x)" by simp +qed (use assms in auto) + +lemma crel_dep_funE: + assumes "((x : A) \\<^sub>c B x) R" + obtains "((x : A) \ B x) R" "in_dom R = A" + using assms unfolding crel_dep_fun_pred_def by auto + +lemma crel_dep_funE' [elim]: + notes crel_dep_funE[elim] + assumes "((x : A) \\<^sub>c B x) R" + obtains "((x : A) \ B x) R" "({\}x : A. B x) R" +proof + show "({\}x : A. B x) R" + proof (rule dep_bin_relI) + fix x y assume "R x y" "A x" + with assms have "R`x = y" "B x (R`x)" by auto + then show "B x y" by simp + qed (use assms in auto) +qed (use assms in auto) + +lemma crel_dep_fun_cong [cong]: + assumes "\x. A x \ A' x" + and "\x y. A' x \ B x y \ B' x y" + shows "((x : A) \\<^sub>c B x) = ((x : A') \\<^sub>c B' x)" + using assms by (intro ext) (auto intro!: crel_dep_funI elim!: crel_dep_funE) + +lemma in_dom_eq_if_crel_dep_fun [simp]: + assumes "((x : A) \\<^sub>c B x) R" + shows "in_dom R = A" + using assms by (auto elim: crel_dep_funE) + +lemma in_codom_le_in_codom_on_if_crel_dep_fun: + assumes "((x : A) \\<^sub>c B x) R" + shows "in_codom R \ in_codom_on A B" + using assms by fast + +lemma crel_funI [intro]: + assumes "(A \ B) R" + and "in_dom R \ A" + shows "(A \\<^sub>c B) R" + using assms by (urule crel_dep_funI) + +lemma crel_funI': + assumes "left_total_on A R" + and "right_unique_on A R" + and "(A {\} B) R" + shows "(A \\<^sub>c B) R" + using assms by (urule crel_dep_funI') + +lemma crel_funE: + assumes "(A \\<^sub>c B) R" + obtains "(A \ B) R" "in_dom R = A" + using assms by (urule (e) crel_dep_funE) + +lemma crel_funE' [elim]: + assumes "(A \\<^sub>c B) R" + obtains "(A \ B) R" "(A {\} B) R" + using assms by (urule (e) crel_dep_funE') + +lemma in_dom_eq_if_crel_fun [simp]: + assumes "(A \\<^sub>c B) R" + shows "in_dom R = A" + using assms by (urule in_dom_eq_if_crel_dep_fun) + +lemma eq_if_rel_if_rel_if_crel_dep_funI: + assumes "((x : A) \\<^sub>c B x) R" + and "R x y" "R x y'" + shows "y = y'" + using assms by (auto intro: eq_if_rel_if_rel_if_rel_dep_funI) + +lemma eval_eq_if_rel_if_crel_dep_funI [simp]: + assumes "((x : A) \\<^sub>c B x) R" + and "R x y" + shows "R`x = y" + using assms by (auto intro: eval_eq_if_rel_if_rel_dep_funI) + +lemma crel_dep_fun_relE: + assumes "((x : A) \\<^sub>c B x) R" + and "R x y" + obtains "A x" "B x y" "R`x = y" + using assms by (auto elim: rel_dep_fun_relE) + +lemma crel_dep_fun_relE': + assumes "((x : A) \\<^sub>c B x) R" + obtains "\x y. R x y \ A x \ B x y \ R`x = y" + using assms by (auto elim: crel_dep_fun_relE) + +lemma rel_restrict_left_eq_self_if_crel_dep_fun [simp]: + assumes "((x : A) \\<^sub>c B x) R" + shows "R\\<^bsub>A\<^esub> = R" + using assms by auto + +text \Note: clean function relations are not contravariant on their domain.\ + +lemma crel_dep_fun_covariant_codom: + assumes "((x : A) \\<^sub>c B x) R" + and "\x. A x \ B x (R`x) \ B' x (R`x)" + shows "((x : A) \\<^sub>c B' x) R" + using assms by (force intro: rel_dep_fun_covariant_codom) + +lemma comp_eq_eval_restrict_left_le_if_crel_dep_fun: + assumes [uhint]: "((x : A) \\<^sub>c B x) R" + shows "((=) \ eval R)\\<^bsub>A\<^esub> \ R" + supply rel_restrict_left_eq_self_if_crel_dep_fun[uhint] + by (urule comp_eq_eval_restrict_left_le_if_rel_dep_fun) (use assms in auto) + +lemma le_comp_eq_eval_restrict_left_if_rel_dep_fun: + assumes [uhint]: "((x : A) \\<^sub>c B x) R" + shows "R \ ((=) \ eval R)\\<^bsub>A\<^esub>" + supply rel_restrict_left_eq_self_if_crel_dep_fun[uhint] + by (urule restrict_left_le_comp_eq_eval_restrict_left_if_rel_dep_fun) (use assms in auto) + +corollary restrict_left_eq_comp_eq_eval_if_crel_dep_fun: + assumes "((x : A) \\<^sub>c B x) R" + shows "R = ((=) \ eval R)\\<^bsub>A\<^esub>" + using assms comp_eq_eval_restrict_left_le_if_crel_dep_fun + le_comp_eq_eval_restrict_left_if_rel_dep_fun + by (intro antisym) auto + +lemma eval_eq_if_crel_dep_fun_if_rel_dep_funI: + fixes R :: "'a \ 'b \ bool" + assumes "((x : A) \ B x) R" "((x : A') \\<^sub>c B' x) R'" + and "R \ R'" + and "A x" + shows "R`x = R'`x" +proof - + from assms have "A' x" by (blast elim: crel_dep_fun_relE) + with assms show ?thesis by (blast intro: eval_eq_if_rel_dep_funI) +qed + +lemma crel_dep_fun_ext: + assumes "((x : A) \\<^sub>c B x) R" "((x : A) \\<^sub>c B' x) R'" + and "\x. A x \ R`x = R'`x" + shows "R = R'" + using assms + by (intro eq_if_rel_agree_on_if_dep_bin_relI[where ?A=A and ?B=B and ?\="(=) R \ (=) R'"] + rel_agree_on_if_eval_eq_if_rel_dep_fun) + auto + +lemma eq_if_le_if_crel_dep_fun_if_rel_dep_fun: + assumes "((x : A) \ B x) R" "((x : A) \\<^sub>c B' x) R'" + and "R \ R'" + shows "R = R'" +proof (intro ext iffI) + fix x y assume "R' x y" + with assms have "R'`x = y" "A x" by auto + moreover with assms have "R`x = R'`x" by (blast intro: eval_eq_if_crel_dep_fun_if_rel_dep_funI) + ultimately show "R x y" using assms by (auto intro: rel_if_eval_eq_if_rel_dep_funI) +qed (use assms in auto) + +lemma ex_dom_crel_dep_fun_iff_crel_dep_fun_in_dom: + "(\(A :: 'a \ bool). ((x : A) \\<^sub>c B x) R) \ (((x : in_dom R) \\<^sub>c B x) R)" + by auto + +lemma crel_fun_bottom_bottom: "((\ :: 'a \ bool) \\<^sub>c A) (\ :: 'a \ 'b \ bool)" + by fastforce + +lemma crel_dep_fun_bottom_iff_eq_bottom [iff]: "((x : (\ :: 'a \ bool)) \\<^sub>c B x) R \ R = \" + by fastforce + +lemma mono_crel_dep_fun_top_crel_dep_fun_inf_rel_restrict_left: + "(((x : A) \\<^sub>c B x) \\<^sub>m (A' : \) \\<^sub>m (x : A \ A') \\<^sub>c B x) rel_restrict_left" + by (intro mono_wrt_predI dep_mono_wrt_predI crel_dep_funI' + (*TODO: should be solved by some type-checking automation*) + mono_right_unique_on_top_right_unique_on_inf_rel_restrict_left + [THEN dep_mono_wrt_predD, THEN dep_mono_wrt_predD] + mono_left_total_on_top_left_total_on_inf_rel_restrict_left + [THEN dep_mono_wrt_predD, THEN dep_mono_wrt_predD] + mono_dep_bin_rel_top_dep_bin_rel_inf_rel_restrict_left + [THEN mono_wrt_predD, THEN dep_mono_wrt_predD]) + auto + +lemma mono_rel_dep_fun_ge_crel_dep_fun_rel_restrict_left: + "(((x : A) \ B x) \\<^sub>m (A' : (\) A) \\<^sub>m (x : A') \\<^sub>c B x) rel_restrict_left" +proof (intro mono_wrt_predI dep_mono_wrt_predI crel_dep_funI) + fix A A' :: "'a \ bool" and B and R :: "'a \ 'b \ bool" assume "((x : A) \ B x) R" + with mono_rel_dep_fun_top_rel_dep_fun_inf_rel_restrict_left have "((x : A \ A') \ B x) R\\<^bsub>A'\<^esub>" + by force + moreover assume "A' \ A" + ultimately show "((x : A') \ B x) R\\<^bsub>A'\<^esub>" by (simp only: inf_absorb2) +qed auto + +lemma crel_dep_fun_eq_restrict: "((x : (A :: 'a \ bool)) \\<^sub>c (=) x) (=)\\<^bsub>A\<^esub>" + by fastforce + +end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function.thy b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function.thy @@ -0,0 +1,13 @@ +\<^marker>\creator "Kevin Kappelmann"\ +theory Binary_Relations_Function + imports + Binary_Relations_Clean_Function + Binary_Relations_Function_Base + Binary_Relations_Function_Composition + Binary_Relations_Function_Evaluation + Binary_Relations_Function_Extend + Binary_Relations_Function_Lambda +begin + + +end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Base.thy b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Base.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Base.thy @@ -0,0 +1,218 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Functions as Binary Relations\ +theory Binary_Relations_Function_Base + imports + Binary_Relations_Function_Evaluation + Binary_Relations_Left_Total +begin + +text \Function relations may contain further elements outside their specification.\ + +consts rel_dep_fun :: "'a \ ('b \ 'c) \ 'd \ bool" +consts rel_fun :: "'a \ 'b \ 'c \ bool" + +bundle rel_fun_syntax +begin +syntax + "_rel_fun" :: "'a \ 'b \ 'c \ bool" ("(_) \ (_)" [41, 40] 40) + "_rel_dep_fun" :: "idt \ 'a \ 'b \ 'c \ bool" ("'(_/ :/ _') \ (_)" [41, 41, 40] 40) +end +bundle no_rel_fun_syntax +begin +no_syntax + "_rel_fun" :: "'a \ 'b \ 'c \ bool" ("(_) \ (_)" [41, 40] 40) + "_rel_dep_fun" :: "idt \ 'a \ 'b \ 'c \ bool" ("'(_/ :/ _') \ (_)" [41, 41, 40] 40) +end +unbundle rel_fun_syntax +translations + "A \ B" \ "CONST rel_fun A B" + "(x : A) \ B" \ "CONST rel_dep_fun A (\x. B)" + +definition "rel_dep_fun_pred (A :: 'a \ bool) (B :: 'a \ 'b \ bool) (R :: 'a \ 'b \ bool) \ + left_total_on A R \ right_unique_on A R \ ((x : A) \\<^sub>m B x) (eval R)" +adhoc_overloading rel_dep_fun rel_dep_fun_pred + +definition "rel_fun_pred (A :: 'a \ bool) (B :: 'b \ bool) :: ('a \ 'b \ bool) \ bool \ + rel_dep_fun_pred A (\(_ :: 'a). B)" +adhoc_overloading rel_fun rel_fun_pred + +lemma rel_fun_eq_rel_dep_fun: + "(((A :: 'a \ bool) \ (B :: 'b \ bool)) :: ('a \ 'b \ bool) \ bool) = (((_ :: 'a) : A) \ B)" + by (simp add: rel_fun_pred_def) + +lemma rel_fun_eq_rel_dep_fun_uhint [uhint]: + assumes "(A :: 'a \ bool) \ A'" + and "B' \ (\(_ :: 'a). (B :: 'b \ bool))" + shows "(A \ B) \ ((x : A') \ B' x)" + using assms by (simp add: rel_fun_eq_rel_dep_fun) + +lemma rel_fun_iff_rel_dep_fun: + "((A :: 'a \ bool) \ (B :: 'b \ bool)) (R :: 'a \ 'b \ bool) \ (((_ :: 'a) : A) \ B) R" + by (simp add: rel_fun_pred_def) + +lemma rel_dep_funI [intro]: + assumes "left_total_on A R" + and "right_unique_on A R" + and "((x : A) \\<^sub>m B x) (eval R)" + shows "((x : A) \ B x) R" + using assms unfolding rel_dep_fun_pred_def by auto + +lemma rel_dep_funE [elim]: + assumes "((x : A) \ B x) R" + obtains "left_total_on A R" "right_unique_on A R" "((x : A) \\<^sub>m B x) (eval R)" + using assms unfolding rel_dep_fun_pred_def by auto + +lemma rel_dep_fun_cong [cong]: + assumes "\x. A x \ A' x" + and "\x y. A' x \ B x y \ B' x y" + shows "((x : A) \ B x) = ((x : A') \ B' x)" + using assms by (intro ext) (auto intro!: rel_dep_funI left_total_onI + dep_mono_wrt_predI intro: right_unique_onD elim!: rel_dep_funE) + +lemma rel_funI [intro]: + assumes "left_total_on A R" + and "right_unique_on A R" + and "(A \\<^sub>m B) (eval R)" + shows "(A \ B) R" + using assms by (urule rel_dep_funI) + +lemma rel_funE [elim]: + assumes "(A \ B) R" + obtains "left_total_on A R" "right_unique_on A R" "(A \\<^sub>m B) (eval R)" + using assms by (urule (e) rel_dep_funE) + +lemma mono_rel_dep_fun_mono_wrt_pred_eval: "(((x : A) \ B x) \\<^sub>m (x : A) \\<^sub>m B x) eval" + by auto + +lemma ex1_rel_right_if_rel_dep_funI: + assumes "((x : A) \ B x) R" + and "A x" + shows "\!y. R x y" + using assms by (blast dest: right_unique_onD) + +lemma eq_if_rel_if_rel_if_rel_dep_funI: + assumes "((x : A) \ B x) R" + and "A x" + and "R x y" "R x y'" + shows "y = y'" + using assms by (blast dest: right_unique_onD) + +lemma eval_eq_if_rel_if_rel_dep_funI [simp]: + assumes "((x : A) \ B x) R" + and "A x" + and "R x y" + shows "R`x = y" + using assms by (intro eq_if_rel_if_rel_if_rel_dep_funI[OF assms, symmetric]) + (blast intro!: rel_eval_if_ex1[where ?R=R] ex1_rel_right_if_rel_dep_funI) + +lemma rel_if_eval_eq_if_rel_dep_funI: + assumes "((x : A) \ B x) R" + and "A x" + and "R`x = y" + shows "R x y" + by (rule rel_if_eval_eq_if_in_dom_if_right_unique_on_eq[where ?R=R]) + (use assms in \blast dest: right_unique_onD\)+ + +corollary rel_eval_if_rel_dep_funI: + assumes "((x : A) \ B x) R" + and "A x" + shows "R x (R`x)" + using assms by (rule rel_if_eval_eq_if_rel_dep_funI) simp + +corollary rel_iff_eval_eq_if_rel_dep_funI: + assumes "((x : A) \ B x) R" + and "A x" + shows "R x y \ R`x = y" + using assms by (intro iffI; rule eval_eq_if_rel_if_rel_dep_funI rel_if_eval_eq_if_rel_dep_funI) + +lemma rel_dep_fun_relE: + assumes "((x : A) \ B x) R" + and "A x" + and "R x y" + obtains "B x y" "R`x = y" +proof + from assms show "R`x = y" by simp + with assms show "B x y" by blast +qed + +lemma rel_dep_fun_relE': + assumes "((x : A) \ B x) R" + obtains "\x y. A x \ R x y \ B x y \ R`x = y" + using assms by (auto elim: rel_dep_fun_relE) + +lemma rel_codom_if_rel_if_pred_if_rel_dep_fun: + assumes "((x : A) \ B x) R" + and "A x" + and "R x y" + shows "B x y" + using assms by (auto elim: rel_dep_fun_relE) + +lemma rel_dep_fun_contravariant_dom: + assumes "((x : A) \ B x) R" + and [dest]: "\x. A' x \ A x" + shows "((x : A') \ B x) R" + using assms by (fast intro!: rel_dep_funI dest: right_unique_onD) + +lemma rel_dep_fun_covariant_codom: + assumes "((x : A) \ B x) R" + and "\x. A x \ B x (R`x) \ B' x (R`x)" + shows "((x : A) \ B' x) R" + using assms by (fast dest: right_unique_onD) + +lemma rel_fun_in_codom_on_if_rel_dep_fun: + assumes "((x : A) \ B x) R" + shows "(A \ in_codom_on A B) R" + using assms by fastforce + +lemma comp_eq_eval_restrict_left_le_if_rel_dep_fun: + assumes "((x : A) \ B x) R" + shows "((=) \ eval R)\\<^bsub>A\<^esub> \ R\\<^bsub>A\<^esub>" + using assms by (intro le_relI) (force intro: rel_eval_if_rel_dep_funI) + +lemma restrict_left_le_comp_eq_eval_restrict_left_if_rel_dep_fun: + assumes "((x : A) \ B x) R" + shows "R\\<^bsub>A\<^esub> \ ((=) \ eval R)\\<^bsub>A\<^esub>" + using assms by (intro le_relI) force + +corollary restrict_left_eq_comp_eq_eval_if_rel_dep_fun: + assumes "((x : A) \ B x) R" + shows "R\\<^bsub>A\<^esub> = ((=) \ eval R)\\<^bsub>A\<^esub>" + using assms comp_eq_eval_restrict_left_le_if_rel_dep_fun + restrict_left_le_comp_eq_eval_restrict_left_if_rel_dep_fun + by (intro antisym) auto + +lemma eval_eq_if_rel_dep_funI: + fixes R :: "'a \ 'b \ bool" + assumes "((x : A) \ B x) R" "((x : A') \ B' x) R'" + and "R \ R'" + and "(A \ A') x" + shows "R`x = R'`x" +proof - + from assms have "R' x (R`x)" "R' x (R'`x)" by (auto intro: rel_eval_if_rel_dep_funI) + with assms show ?thesis by (intro eval_eq_if_rel_if_rel_dep_funI[symmetric]) force+ +qed + +lemma rel_agree_on_if_eval_eq_if_rel_dep_fun: + assumes crel_dep_fun: "\R. \ R \ \B. ((x : A) \ B x) R" + and "\R R' x. \ R \ \ R' \ A x \ R`x = R'`x" + shows "rel_agree_on A \" +proof (rule rel_agree_onI) + fix x y R R' assume hyps: "\ R" "\ R'" "A x" "R x y" + with crel_dep_fun have "y = R`x" by fastforce + also from assms hyps have "... = R'`x" by blast + finally have "y = R'`x" . + moreover from crel_dep_fun[OF \\ R'\] obtain B where "((x : A) \ B x) R'" by blast + ultimately show "R' x y" using \A x\ by (auto intro: rel_eval_if_rel_dep_funI) +qed + +lemma mono_rel_dep_fun_top_rel_dep_fun_inf_rel_restrict_left: + "(((x : A) \ B x) \\<^sub>m (A' : \) \\<^sub>m (x : A \ A') \ B x) rel_restrict_left" + by (intro mono_wrt_predI dep_mono_wrt_predI rel_dep_funI + (*TODO: should be solved by some type-checking automation*) + mono_right_unique_on_top_right_unique_on_inf_rel_restrict_left + [THEN dep_mono_wrt_predD, THEN dep_mono_wrt_predD] + mono_left_total_on_top_left_total_on_inf_rel_restrict_left + [THEN dep_mono_wrt_predD, THEN dep_mono_wrt_predD]) + auto + +end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Composition.thy b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Composition.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Composition.thy @@ -0,0 +1,49 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Composing Functions\ +theory Binary_Relations_Function_Composition + imports + Binary_Relations_Clean_Function + Restricted_Equality +begin + +lemma dep_bin_rel_eval_rel_comp_if_dep_bin_rel_if_crel_dep_fun: + assumes "((x : A) \\<^sub>c B x) R" + and "\x. A x \ ({\}y : B x. C x y) R'" + shows "({\}x : A. C x (R`x)) (R \\ R')" + using assms by force + +lemma crel_dep_fun_eval_rel_comp_if_rel_dep_fun_if_crel_dep_fun: + assumes "((x : A) \\<^sub>c B x) R" + and "\x. A x \ ((y : B x) \ C x y) R'" + shows "((x : A) \\<^sub>c C x (R`x)) (R \\ R')" +proof (intro crel_dep_funI' dep_bin_relI + (*TODO: should be solved by some type-checking automation*) + mono_left_total_on_comp[THEN dep_mono_wrt_predD, THEN mono_wrt_predD] + mono_right_unique_on_comp[THEN dep_mono_wrt_predD, THEN mono_wrt_predD]) + from assms show "left_total_on (in_codom R\\<^bsub>A\<^esub>) R'" by force + from assms show "right_unique_on (in_codom R\\<^bsub>A\<^esub>) R'" by (fast dest: right_unique_onD) +qed (use assms in \auto elim: rel_dep_fun_relE crel_dep_fun_relE\) + +lemma rel_comp_eval_eq_if_rel_dep_fun_if_crel_dep_funI [simp]: + assumes "((x : A) \\<^sub>c B x) R" + and "\x. A x \ ((y : B x) \ C x y) R'" + and "A x" + shows "(R \\ R')`x = R'`(R`x)" +proof (rule eval_eq_if_right_unique_onI) + from assms have "((x : A) \\<^sub>c C x (R`x)) (R \\ R')" + by (intro crel_dep_fun_eval_rel_comp_if_rel_dep_fun_if_crel_dep_fun) auto + then show "right_unique_on A (R \\ R')" by blast + from assms show "(R \\ R') x (R'`(R`x))" by (blast intro: rel_eval_if_rel_dep_funI) +qed (fact assms) + +lemma eq_restrict_comp_eq_if_crel_dep_fun [simp]: + assumes "((x : A) \\<^sub>c B x) R" + shows "(=\<^bsub>A\<^esub>) \\ R = R" + using assms by force + +lemma comp_eq_restrict_if_crel_dep_fun [simp]: + assumes "(A \\<^sub>c B) R" + shows "R \\ (=\<^bsub>B\<^esub>) = R" + using assms by force + +end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Evaluation.thy b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Evaluation.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Evaluation.thy @@ -0,0 +1,103 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Evaluation of Functions as Binary Relations\ +theory Binary_Relations_Function_Evaluation + imports + Binary_Relations_Right_Unique + Binary_Relations_Extend +begin + +consts eval :: "'a \ 'b \ 'c" + +definition "eval_rel R x \ THE y. R x y" +adhoc_overloading eval eval_rel + +bundle eval_syntax begin notation eval ("(_`_)" [999, 1000] 999) end +bundle no_eval_syntax begin no_notation eval ("(_`_)" [999, 1000] 999) end +unbundle eval_syntax + +lemma eval_eq_if_right_unique_onI: + assumes "right_unique_on P R" + and "P x" + and "R x y" + shows "R`x = y" + using assms unfolding eval_rel_def by (auto dest: right_unique_onD) + +lemma eval_eq_if_right_unique_on_eqI: + assumes "right_unique_on ((=) x) R" + and "R x y" + shows "R`x = y" + using assms by (auto intro: eval_eq_if_right_unique_onI) + +lemma rel_eval_if_ex1: + assumes "\!y. R x y" + shows "R x (R`x)" + using assms unfolding eval_rel_def by (rule theI') + +lemma rel_if_eval_eq_if_in_dom_if_right_unique_on_eq: + assumes "right_unique_on ((=) x) R" + and "in_dom R x" + and "R`x = y" + shows "R x y" + using assms by (blast intro: rel_eval_if_ex1[of R x] dest: right_unique_onD) + +corollary rel_eval_if_in_dom_if_right_unique_on_eq: + assumes "right_unique_on ((=) x) R" + and "in_dom R x" + shows "R x (R`x)" + using assms by (rule rel_if_eval_eq_if_in_dom_if_right_unique_on_eq) simp + +lemma eval_app_eq_eq [simp]: "(\x. (=) (f x))`x = f x" + by (auto intro: eval_eq_if_right_unique_onI) + +lemma extend_eval_eq_if_not_in_dom [simp]: + assumes "\(in_dom R x)" + shows "(extend x y R)`x = y" + using assms by (force intro: eval_eq_if_right_unique_on_eqI) + +corollary extend_bottom_eval_eq [simp]: + fixes x :: 'a and y :: 'b + shows "(extend x y (\ :: 'a \ 'b \ bool))`x = y" + by (intro extend_eval_eq_if_not_in_dom) auto + +lemma glue_eval_eqI: + assumes "right_unique_on P (glue \)" + and "\ R" + and "P x" + and "R x y" + shows "(glue \)`x = y" + using assms by (auto intro: eval_eq_if_right_unique_onI) + +lemma glue_eval_eq_evalI: + assumes "right_unique_on P (glue \)" + and "\ R" + and "P x" + and "in_dom R x" + shows "(glue \)`x = R`x" + using assms by (intro glue_eval_eqI[of P \ R]) + (auto intro: rel_if_eval_eq_if_in_dom_if_right_unique_on_eq[of x] dest: right_unique_onD) + +text \Note: the following rests on the definition of extend and eval:\ + +lemma extend_eval_eq_if_neq [simp]: + fixes R :: "'a \ 'b \ bool" + shows "x \ y \ (extend y z R)`x = R`x" + unfolding extend_rel_def eval_rel_def by auto + +lemma sup_eval_eq_if_not_in_dom_left [simp]: + fixes R S :: "'a \ 'b \ bool" + shows "\(in_dom S x) \ (R \ S)`x = R`x" + unfolding eval_rel_def by (cases "\y. S x y") auto + +lemma sup_eval_eq_if_not_in_dom_right [simp]: + fixes R S :: "'a \ 'b \ bool" + shows "\(in_dom R x) \ (R \ S)`x = S`x" + unfolding eval_rel_def by (cases "\y. R x y") auto + +lemma rel_restrict_left_eval_eq_if_pred [simp]: + fixes R :: "'a \ 'b \ bool" + assumes "P x" + shows "(R\\<^bsub>P\<^esub>)`x = R`x" + (*FIXME: proof relying on specific definitions; can we do better?*) + using assms unfolding eval_rel_def rel_restrict_left_pred_def by auto + +end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Extend.thy b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Extend.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Extend.thy @@ -0,0 +1,100 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Extending Functions\ +theory Binary_Relations_Function_Extend + imports + Binary_Relations_Clean_Function +begin + +lemma left_total_on_sup_eq_extend_if_left_total_on: + assumes "left_total_on A R" + shows "left_total_on (A \ (=) x) (extend x y R)" + using assms by fastforce + +lemma right_unique_on_sup_eq_extend_if_not_in_dom_if_right_unique_on: + assumes "right_unique_on A R" + and "\(in_dom R x)" + shows "right_unique_on (A \ (=) x) (extend x y R)" + using assms by (intro right_unique_onI) (auto dest: right_unique_onD elim!: extendE) + +lemma rel_dep_fun_extend_if_rel_dep_funI: + assumes "((x : A) \ B x) R" + and "\(in_dom R x)" + shows "((x' : A \ ((=) x)) \ (if x' = x then (=) y else B x')) (extend x y R)" + using assms by (intro rel_dep_funI left_total_on_sup_eq_extend_if_left_total_on + right_unique_on_sup_eq_extend_if_not_in_dom_if_right_unique_on dep_mono_wrt_predI) + auto + +lemma crel_dep_fun_extend_if_crel_dep_funI: + assumes "((x : A) \\<^sub>c B x) R" + and "\(in_dom R x)" + shows "((x' : A \ ((=) x)) \\<^sub>c (if x' = x then (=) y else B x')) (extend x y R)" + using assms by (intro crel_dep_funI rel_dep_fun_extend_if_rel_dep_funI) force+ + +context + fixes \ :: "('a \ 'b \ bool) \ bool" and A :: "('a \ 'b \ bool) \ 'a \ bool" and D + defines [simp]: "D \ in_codom_on \ A" +begin + +lemma rel_dep_fun_glue_if_right_unique_if_rel_dep_fun: + assumes funs: "\R. \ R \ ((x : A R) \ B x) R" + and runique: "right_unique_on D (glue \)" + shows "((x : D) \ B x) (glue \)" +proof (intro rel_dep_funI dep_mono_wrt_predI left_total_onI) + fix x assume "D x" + with funs obtain R where hyps: "\ R" "A R x" "((x : A R) \ B x) R" by auto + then have "B x (R`x)" by (auto elim: rel_dep_fun_relE) + moreover have "(glue \)`x = R`x" + proof (intro glue_eval_eqI) + show "\ R" by (fact hyps) + then have "A R \ D" by fastforce + with runique show "right_unique_on (A R) (glue \)" using antimono_right_unique_on by blast + qed (use hyps in \auto elim: rel_dep_fun_relE\) + ultimately show "B x (glue \`x)" by simp +qed (use assms in \fastforce+\) + +lemma crel_dep_fun_glue_if_right_unique_if_crel_dep_fun: + assumes "\R. \ R \ ((x : A R) \\<^sub>c B x) R" + and "right_unique_on D (glue \)" + shows "((x : D) \\<^sub>c B x) (glue \)" + using assms by (intro crel_dep_funI rel_dep_fun_glue_if_right_unique_if_rel_dep_fun) fastforce+ + +end + +lemma right_unique_on_sup_if_rel_agree_on_sup_if_right_unique_on: + assumes "right_unique_on P R" "right_unique_on P R'" + and "rel_agree_on (P \ in_dom R \ in_dom R') ((=) R \ (=) R')" + shows "right_unique_on P (R \ R')" +proof (intro right_unique_onI) + fix x y y' assume "P x" "(R \ R') x y" "(R \ R') x y'" + then obtain R1 R2 where rels: "R1 x y" "R2 x y'" and ors: "R1 = R \ R1 = R'" "R2 = R \ R2 = R'" by auto + then consider (neq) "R1 \ R2" | "R1 = R2" by auto + then show "y = y'" + proof cases + case neq + with rels ors obtain z z' where "R x z" "R' x z'" and yors: "y = z \ y = z'" "y' = z \ y' = z'" + by auto + with \P x\ have "R' x z" by (intro rel_agree_onD[where ?R=R and ?R'=R' and + ?P="P \ in_dom R \ in_dom R'" and ?\="(=) R \ (=) R'"] assms) + auto + with \P x\ \R' x z'\ assms have "z = z'" by (blast dest: right_unique_onD) + with yors show "y = y'" by auto + qed (use assms \P x\ rels ors in \auto dest: right_unique_onD\) +qed + +lemma crel_dep_fun_sup_if_eval_eq_if_crel_dep_fun: + assumes dep_funs: "((x : A) \\<^sub>c B x) R" "((x : A') \\<^sub>c B x) R'" + and "\x. A x \ A' x \ R`x = R'`x" + shows "((x : A \ A') \\<^sub>c B x) (R \ R')" +proof - + let ?A = "\S. if S = R then A else A'" + from dep_funs have "A = A'" if "R = R'" using that by (simp only: flip: + in_dom_eq_if_crel_dep_fun[OF dep_funs(1)] in_dom_eq_if_crel_dep_fun[OF dep_funs(2)]) + then have [uhint]: "in_codom_on ((=) R \ (=) R') ?A = A \ A'" + by (intro ext iffI) (fastforce split: if_splits)+ + show ?thesis by (urule (rr) crel_dep_fun_glue_if_right_unique_if_crel_dep_fun + right_unique_on_sup_if_rel_agree_on_sup_if_right_unique_on + rel_agree_on_if_eval_eq_if_rel_dep_fun) + (use assms in \auto 7 0 intro: rel_dep_fun_contravariant_dom dest: right_unique_onD\) +qed + +end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Lambda.thy b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Lambda.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Functions/Binary_Relations_Function_Lambda.thy @@ -0,0 +1,122 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Lambda Abstractions\ +theory Binary_Relations_Function_Lambda + imports Binary_Relations_Clean_Function +begin + +consts rel_lambda :: "'a \ ('b \ 'c) \ 'd" + +definition "rel_lambda_pred A f x y \ A x \ f x = y" +adhoc_overloading rel_lambda rel_lambda_pred + +bundle rel_lambda_syntax +begin +syntax + "_rel_lambda" :: "idt \ 'a \ 'b \ 'c" ("(2\_ : _./ _)" 60) +end +bundle no_rel_lambda_syntax +begin +no_syntax + "_rel_lambda" :: "idt \ 'a \ 'b \ 'c" ("(2\_ : _./ _)" 60) +end +unbundle rel_lambda_syntax +translations + "\x : A. f" \ "CONST rel_lambda A (\x. f)" + +lemma rel_lambdaI [intro]: + assumes "A x" + and "f x = y" + shows "(\x : A. f x) x y" + using assms unfolding rel_lambda_pred_def by auto + +lemma rel_lambda_appI: + assumes "A x" + shows "(\x : A. f x) x (f x)" + using assms by auto + +lemma rel_lambdaE [elim!]: + assumes "(\x : A. f x) x y" + obtains "A x" "y = f x" + using assms unfolding rel_lambda_pred_def by auto + +lemma rel_lambda_cong [cong]: + "\\x. A x \ A' x; \x. A' x \ f x = f' x\ \ (\x : A. f x) = \x : A'. f' x" + by (intro ext) auto + +lemma in_dom_rel_lambda_eq [simp]: "in_dom (\x : A. f x) = A" + by auto + +lemma in_codom_rel_lambda_eq_has_inverse_on [simp]: "in_codom (\x : A. f x) = has_inverse_on A f" + by fastforce + +lemma left_total_on_rel_lambda: "left_total_on A (\x : A. f x)" + by auto + +lemma right_unique_on_rel_lambda: "right_unique_on A (\x : A. f x)" + by auto + +lemma cdep_fun_rel_lambda: "((x : A) \\<^sub>c ((=) (f x))) (\x : A. f x)" + by (intro crel_dep_funI') auto + +text \Compare the following with @{thm mono_rel_dep_fun_mono_wrt_pred_eval}.\ + +lemma mono_wrt_pred_mono_crel_dep_fun_rel_lambda: + "(((x : A) \\<^sub>m B x) \\<^sub>m (x : A) \\<^sub>c B x) (rel_lambda A)" + by (intro mono_wrt_predI crel_dep_funI') auto + +lemma rel_lambda_eval_eq [simp]: "A x \ (\x : A. f x)`x = f x" + by (rule eval_eq_if_right_unique_onI) auto + +lemma app_eq_if_rel_lambda_eqI: + assumes "(\x : A. f x) = (\x : A. g x)" + and "A x" + shows "f x = g x" + using assms by (auto dest: fun_cong) + +lemma crel_dep_fun_inf_rel_lambda_inf_if_rel_dep_fun: + assumes "((x : A) \ B x) R" + shows "((x : A \ A') \\<^sub>c B x) (\x : A \ A'. R`x)" + using assms by force + +corollary crel_dep_fun_rel_lambda_if_le_if_rel_dep_fun: + assumes "((x : A) \ B x) R" + and [uhint]: "A' \ A" + shows "((x : A') \\<^sub>c B x) (\x : A'. R`x)" + supply inf_absorb2[uhint] + by (urule crel_dep_fun_inf_rel_lambda_inf_if_rel_dep_fun) (fact assms) + +lemma rel_lambda_ext: + assumes "((x : A) \\<^sub>c B x) R" + and "\x. A x \ f x = R`x" + shows "(\x : A. f x) = R" + using assms by (intro ext iffI) (auto intro!: rel_lambdaI intro: rel_eval_if_rel_dep_funI) + +lemma rel_lambda_eval_eq_if_crel_dep_fun [simp]: "((x : A) \\<^sub>c B x) R \ (\x : A. R`x) = R" + by (rule rel_lambda_ext) auto + +text \Every element of @{term "(x : A) \\<^sub>c (B x)"} may be expressed as a lambda abstraction.\ + +lemma eq_rel_lambda_if_crel_dep_funE: + assumes "((x : A) \\<^sub>c B x) R" + obtains f where "R = (\x : A. f x)" +proof + let ?f="(\x. R`x)" + from assms show "R = (\x : A. (\x. R`x) x)" by simp +qed + +lemma rel_restrict_left_eq_rel_lambda_if_le_if_rel_dep_fun: + assumes "((x : A) \ B x) R" + and "A' \ A" + shows "R\\<^bsub>A'\<^esub> = (\x : A'. R`x)" +proof - + from assms mono_rel_dep_fun_ge_crel_dep_fun_rel_restrict_left have "((x : A') \\<^sub>c B x) R\\<^bsub>A'\<^esub>" + by force + then show ?thesis + supply \A' \ A\[uhint] inf_absorb2[uhint] by (urule rel_lambda_ext[symmetric]) auto +qed + +lemma mono_rel_lambda: "mono (\A. \x : A. f x)" + by auto + + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/LBinary_Relations.thy b/thys/Transport/HOL_Basics/Binary_Relations/LBinary_Relations.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/LBinary_Relations.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/LBinary_Relations.thy @@ -1,15 +1,20 @@ \<^marker>\creator "Kevin Kappelmann"\ theory LBinary_Relations imports - Binary_Relation_Functions - Binary_Relations_Lattice + Binary_Relations_Function Binary_Relations_Order Binary_Relation_Properties + Binary_Relation_Functions + Binary_Relations_Agree + Binary_Relations_Extend + Binary_Relations_Lattice + Dependent_Binary_Relations Restricted_Equality + Reverse_Implies begin paragraph \Summary\ text \Basic concepts on binary relations.\ end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relation_Properties.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relation_Properties.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relation_Properties.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relation_Properties.thy @@ -1,19 +1,20 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Basic Properties\ theory Binary_Relation_Properties imports Binary_Relations_Antisymmetric Binary_Relations_Bi_Total Binary_Relations_Bi_Unique + Binary_Relations_Connected Binary_Relations_Injective Binary_Relations_Irreflexive Binary_Relations_Left_Total Binary_Relations_Reflexive Binary_Relations_Right_Unique Binary_Relations_Surjective Binary_Relations_Symmetric Binary_Relations_Transitive begin 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,75 +1,72 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Antisymmetric\ theory Binary_Relations_Antisymmetric imports Binary_Relation_Functions - HOL_Syntax_Bundles_Lattices - ML_Unification.ML_Unification_HOL_Setup - ML_Unification.Unify_Resolve_Tactics begin consts antisymmetric_on :: "'a \ 'b \ 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" + definition "antisymmetric_on_pred P R \ \x y : P. 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 consts antisymmetric :: "'a \ bool" overloading antisymmetric \ "antisymmetric :: ('a \ 'a \ bool) \ bool" begin definition "antisymmetric :: ('a \ 'a \ bool) \ _ \ antisymmetric_on (\ :: 'a \ bool)" end lemma antisymmetric_eq_antisymmetric_on: "(antisymmetric :: ('a \ 'a \ bool) \ _) = antisymmetric_on (\ :: 'a \ bool)" unfolding antisymmetric_def .. lemma antisymmetric_eq_antisymmetric_on_uhint [uhint]: "P \ (\ :: 'a \ bool) \ (antisymmetric :: ('a \ 'a \ bool) \ _) \ antisymmetric_on P" by (simp add: antisymmetric_eq_antisymmetric_on) lemma antisymmetricI [intro]: assumes "\x y. R x y \ R y x \ x = y" shows "antisymmetric R" using assms by (urule antisymmetric_onI) lemma antisymmetricD: assumes "antisymmetric R" and "R x y" "R y x" shows "x = y" using assms by (urule (d) antisymmetric_onD where chained = insert) simp_all lemma antisymmetric_on_if_antisymmetric: fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" 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 [iff]: "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_Connected.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Connected.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Connected.thy @@ -0,0 +1,65 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsubsection \Connected\ +theory Binary_Relations_Connected + imports + Binary_Relation_Functions +begin + +consts connected_on :: "'a \ 'b \ bool" + +overloading + connected_on_pred \ "connected_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool" +begin + definition "connected_on_pred P R \ \x y : P. x \ y \ R x y \ R y x" +end + +lemma connected_onI [intro]: + assumes "\x y. P x \ P y \ x \ y \ R x y \ R y x" + shows "connected_on P R" + using assms unfolding connected_on_pred_def by blast + +lemma connected_onE [elim]: + assumes "connected_on P R" + and "P x" "P y" + obtains "x = y" | "R x y" | "R y x" + using assms unfolding connected_on_pred_def by auto + +lemma connected_on_rel_inv_iff_connected_on [iff]: + "connected_on (P :: 'a \ bool) (R :: 'a \ 'a \ bool)\ \ connected_on P R" + by blast + +consts connected :: "'a \ bool" + +overloading + connected \ "connected :: ('a \ 'a \ bool) \ bool" +begin + definition "connected :: ('a \ 'a \ bool) \ bool \ connected_on (\ :: 'a \ bool)" +end + +lemma connected_eq_connected_on: + "(connected :: ('a \ 'a \ bool) \ _) = connected_on (\ :: 'a \ bool)" + unfolding connected_def .. + +lemma connected_eq_connected_on_uhint [uhint]: + "P \ (\ :: 'a \ bool) \ (connected :: ('a \ 'a \ bool) \ _) \ connected_on P" + by (simp add: connected_eq_connected_on) + +lemma connectedI [intro]: + assumes "\x y. x \ y \ R x y \ R y x" + shows "connected R" + using assms by (urule connected_onI) + +lemma connectedE [elim]: + assumes "connected R" + and "x \ y" + obtains "R x y" | "R y x" + using assms by (urule (e) connected_onE where chained = insert) auto + +lemma connected_on_if_connected: + fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" + assumes "connected R" + shows "connected_on P R" + using assms by (intro connected_onI) 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,152 +1,150 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Injective\ theory Binary_Relations_Injective imports - Reverse_Implies Functions_Monotone - ML_Unification.ML_Unification_HOL_Setup - ML_Unification.Unify_Resolve_Tactics + Reverse_Implies begin consts rel_injective_on :: "'a \ 'b \ 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'" + definition "rel_injective_on_pred P R \ \x x' : P. \ y. 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 lemma antimono_rel_injective_on: "((\) \\<^sub>m (\) \ (\)) (rel_injective_on :: ('a \ bool) \ ('a \ 'b \ bool) \ bool)" - by (intro dep_mono_wrt_relI) (auto dest: rel_injective_onD intro!: rel_injective_onI) + by (intro mono_wrt_relI) (auto dest: rel_injective_onD intro!: rel_injective_onI) consts rel_injective_at :: "'a \ 'b \ 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 lemma rel_injective_on_if_Fun_Rel_imp_if_rel_injective_at: - assumes "rel_injective_at Q R" + assumes "rel_injective_at (Q :: 'b \ bool) (R :: 'a \ 'b \ bool)" and "(R \ (\)) P Q" shows "rel_injective_on P R" using assms by (intro rel_injective_onI) (auto dest: rel_injective_atD) lemma rel_injective_at_if_Fun_Rel_rev_imp_if_rel_injective_on: - assumes "rel_injective_on P R" + assumes "rel_injective_on (P :: 'a \ bool) (R :: 'a \ 'b \ bool)" and "(R \ (\)) P Q" shows "rel_injective_at Q R" using assms by (intro rel_injective_atI) (auto dest: rel_injective_onD) consts rel_injective :: "'a \ bool" overloading rel_injective \ "rel_injective :: ('a \ 'b \ bool) \ bool" begin definition "(rel_injective :: ('a \ 'b \ bool) \ bool) \ rel_injective_on (\ :: 'a \ bool)" end lemma rel_injective_eq_rel_injective_on: "(rel_injective :: ('a \ 'b \ bool) \ _) = rel_injective_on (\ :: 'a \ bool)" unfolding rel_injective_def .. lemma rel_injective_eq_rel_injective_on_uhint [uhint]: assumes "P \ (\ :: 'a \ bool)" shows "rel_injective :: ('a \ 'b \ bool) \ _ \ rel_injective_on P" using assms by (simp add: rel_injective_eq_rel_injective_on) lemma rel_injectiveI [intro]: assumes "\x x' y. R x y \ R x' y \ x = x'" shows "rel_injective R" using assms by (urule rel_injective_onI) lemma rel_injectiveD: assumes "rel_injective R" and "R x y" "R x' y" shows "x = x'" using assms by (urule (d) rel_injective_onD where chained = insert) simp_all lemma rel_injective_eq_rel_injective_at: "(rel_injective :: ('a \ 'b \ bool) \ bool) = rel_injective_at (\ :: 'b \ bool)" by (intro ext iffI rel_injectiveI) (auto dest: rel_injective_atD rel_injectiveD) lemma rel_injective_eq_rel_injective_at_uhint [uhint]: assumes "P \ (\ :: 'b \ bool)" shows "rel_injective :: ('a \ 'b \ bool) \ bool \ rel_injective_at P" using assms by (simp add: rel_injective_eq_rel_injective_at) lemma rel_injective_on_if_rel_injective: fixes P :: "'a \ bool" and R :: "'a \ 'b \ bool" 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 lemma rel_injective_on_compI: fixes P :: "'a \ bool" assumes "rel_injective (R :: 'a \ 'b \ bool)" and "rel_injective_on (in_codom R \ in_dom S) (S :: 'b \ 'c \ bool)" shows "rel_injective_on P (R \\ S)" proof (rule rel_injective_onI) fix x x' y assume "P x" "P x'" "(R \\ S) x y" "(R \\ S) x' y" then obtain z z' where "R x z" "S z y" "R x' z'" "S z' y" by blast with assms have "z = z'" by (auto dest: rel_injective_onD) with \R x z\ \R x' z'\ assms show "x = x'" by (auto dest: rel_injectiveD) qed 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,98 +1,91 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Left Total\ theory Binary_Relations_Left_Total imports Functions_Monotone - ML_Unification.ML_Unification_HOL_Setup - ML_Unification.Unify_Resolve_Tactics begin consts left_total_on :: "'a \ 'b \ 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" + definition "left_total_on_pred P R \ \x : P. 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: +lemma le_in_dom_if_left_total_on: assumes "left_total_on P R" shows "P \ in_dom R" using assms by force lemma left_total_on_if_le_in_dom: assumes "P \ in_dom R" shows "left_total_on P R" using assms by fastforce lemma mono_left_total_on: "((\) \\<^sub>m (\) \ (\)) (left_total_on :: ('a \ bool) \ ('a \ 'b \ bool) \ bool)" - by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI) fastforce + by fastforce lemma le_in_dom_iff_left_total_on: "P \ in_dom R \ left_total_on P R" - using in_dom_if_left_total_on left_total_on_if_le_in_dom by auto + using le_in_dom_if_left_total_on left_total_on_if_le_in_dom by auto -lemma left_total_on_inf_restrict_leftI: - fixes P P' :: "'a \ bool" and R :: "'a \ 'b \ bool" - assumes "left_total_on P R" - shows "left_total_on (P \ P') R\\<^bsub>P'\<^esub>" - using assms by (intro left_total_onI) auto +lemma mono_left_total_on_top_left_total_on_inf_rel_restrict_left: + "((R : left_total_on P) \\<^sub>m (P' : \) \\<^sub>m left_total_on (P \ P')) rel_restrict_left" + by fast -lemma left_total_on_compI: - fixes P :: "'a \ bool" and R :: "'a \ 'b \ bool" - assumes "left_total_on P R" - and "left_total_on (in_codom (R\\<^bsub>P\<^esub>)) S" - shows "left_total_on P (R \\ S)" - using assms by (intro left_total_onI) force +lemma mono_left_total_on_comp: + "((R : left_total_on P) \\<^sub>m left_total_on (in_codom (R\\<^bsub>P\<^esub>)) \\<^sub>m left_total_on P) (\\)" + by fast consts left_total :: "'a \ bool" overloading left_total \ "left_total :: ('a \ 'b \ bool) \ bool" begin definition "(left_total :: ('a \ 'b \ bool) \ bool) \ left_total_on (\ :: 'a \ bool)" end lemma left_total_eq_left_total_on: "(left_total :: ('a \ 'b \ bool) \ _) = left_total_on (\ :: 'a \ bool)" unfolding left_total_def .. lemma left_total_eq_left_total_on_uhint [uhint]: assumes "P \ \ :: 'a \ bool" shows "left_total :: ('a \ 'b \ bool) \ _ \ left_total_on P" using assms by (simp add: left_total_eq_left_total_on) lemma left_totalI [intro]: assumes "\x. in_dom R x" shows "left_total R" using assms by (urule left_total_onI) lemma left_totalE: assumes "left_total R" obtains y where "R x y" using assms by (urule (e) left_total_onE where chained = insert) simp 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 \ 'b \ bool" 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,124 +1,125 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Reflexive\ theory Binary_Relations_Reflexive imports Functions_Monotone ML_Unification.ML_Unification_HOL_Setup ML_Unification.Unify_Resolve_Tactics begin consts reflexive_on :: "'a \ 'b \ 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 context fixes R :: "'a \ 'a \ bool" and P :: "'a \ bool" begin 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 R" by blast lemma mono_reflexive_on: "((\) \\<^sub>m (\) \ (\)) (reflexive_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool)" - by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI) fastforce + by fastforce lemma reflexive_on_if_le_pred_if_reflexive_on: fixes P' :: "'a \ bool" assumes "reflexive_on P R" and "P' \ P" shows "reflexive_on P' R" using assms by blast end lemma reflexive_on_sup_eq [simp]: "(reflexive_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool) (P \ 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_le: "reflexive_on (P :: 'a \ bool) (R :: 'a \ _) \ ((=)\\<^bsub>P\<^esub> \ R)" by blast consts reflexive :: "'a \ bool" overloading reflexive \ "reflexive :: ('a \ 'a \ bool) \ bool" begin definition "reflexive :: ('a \ 'a \ bool) \ bool \ reflexive_on (\ :: 'a \ bool)" end lemma reflexive_eq_reflexive_on: "(reflexive :: ('a \ 'a \ bool) \ _) = reflexive_on (\ :: 'a \ bool)" unfolding reflexive_def .. lemma reflexive_eq_reflexive_on_uhint [uhint]: "P \ (\ :: 'a \ bool) \ (reflexive :: ('a \ 'a \ bool) \ _) \ reflexive_on P" by (simp add: reflexive_eq_reflexive_on) lemma reflexiveI [intro]: assumes "\x. R x x" shows "reflexive R" using assms by (urule reflexive_onI) lemma reflexiveD: assumes "reflexive R" shows "R x x" using assms by (urule (d) reflexive_onD where chained = insert) simp lemma reflexive_on_if_reflexive: fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" 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 :: 'a \ 'a \ bool)\ \ reflexive R" by (blast dest: reflexiveD) lemma reflexive_iff_eq_le: "reflexive R \ ((=) \ R)" unfolding reflexive_eq_reflexive_on reflexive_on_iff_eq_restrict_le by auto paragraph \Instantiations\ lemma reflexive_eq: "reflexive (=)" by (rule reflexiveI) (rule refl) lemma reflexive_top: "reflexive (\ :: 'a \ 'a \ bool)" 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,172 +1,192 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Right Unique\ theory Binary_Relations_Right_Unique - imports Binary_Relations_Injective + imports + Binary_Relations_Injective + Binary_Relations_Extend begin consts right_unique_on :: "'a \ 'b \ 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'" + definition "right_unique_on_pred P R \ \x : P. \ y y'. 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 lemma antimono_right_unique_on: "((\) \\<^sub>m (\) \ (\)) (right_unique_on :: ('a \ bool) \ ('a \ 'b \ bool) \ bool)" - by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI) (fastforce dest: right_unique_onD) + by (fastforce dest: right_unique_onD) -lemma right_unique_on_compI: - fixes P :: "'a \ bool" and R :: "'a \ 'b \ bool" and S :: "'b \ 'c \ bool" - assumes "right_unique_on P R" - and "right_unique_on (in_codom (R\\<^bsub>P\<^esub>) \ in_dom S) S" - shows "right_unique_on P (R \\ S)" - using assms by (blast dest: right_unique_onD) +lemma mono_right_unique_on_top_right_unique_on_inf_rel_restrict_left: + "((R : right_unique_on P) \\<^sub>m (P' : \) \\<^sub>m right_unique_on (P \ P')) rel_restrict_left" + by (fast dest: right_unique_onD) + +lemma mono_right_unique_on_comp: + "((R : right_unique_on P) \\<^sub>m right_unique_on (in_codom (R\\<^bsub>P\<^esub>)) \\<^sub>m right_unique_on P) (\\)" + by (fast dest: right_unique_onD) + +context + fixes P :: "'a \ bool" and \ :: "('a \ 'b \ bool) \ bool" +begin + +lemma right_unique_on_glue_if_right_unique_on_sup: + assumes "\R R'. \ R \ \ R' \ right_unique_on P (R \ R')" + shows "right_unique_on P (glue \)" + using assms by (fastforce dest: right_unique_onD) + +lemma right_unique_on_if_right_unique_on_glue: + assumes "right_unique_on P (glue \)" + and "\ R" + shows "right_unique_on P R" + using assms by (intro right_unique_onI) (auto dest: right_unique_onD) + +end consts right_unique_at :: "'a \ 'b \ 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) lemma right_unique_on_if_Fun_Rel_imp_if_right_unique_at: - assumes "right_unique_at Q R" + assumes "right_unique_at (Q :: 'b \ bool) (R :: 'a \ 'b \ bool)" and "(R \ (\)) P Q" shows "right_unique_on P R" using assms by (intro right_unique_onI) (auto dest: right_unique_atD) lemma right_unique_at_if_Fun_Rel_rev_imp_if_right_unique_on: - assumes "right_unique_on P R" + assumes "right_unique_on (P :: 'a \ bool) (R :: 'a \ 'b \ bool)" and "(R \ (\)) P Q" shows "right_unique_at Q R" using assms by (intro right_unique_atI) (auto dest: right_unique_onD) consts right_unique :: "'a \ bool" overloading right_unique \ "right_unique :: ('a \ 'b \ bool) \ bool" begin definition "(right_unique :: ('a \ 'b \ bool) \ bool) \ right_unique_on (\ :: 'a \ bool)" end lemma right_unique_eq_right_unique_on: "(right_unique :: ('a \ 'b \ bool) \ _) = right_unique_on (\ :: 'a \ bool)" unfolding right_unique_def .. lemma right_unique_eq_right_unique_on_uhint [uhint]: assumes "P \ (\ :: 'a \ bool)" shows "right_unique :: ('a \ 'b \ bool) \ _ \ right_unique_on P" using assms by (simp only: right_unique_eq_right_unique_on) lemma right_uniqueI [intro]: assumes "\x y y'. R x y \ R x y' \ y = y'" shows "right_unique R" using assms by (urule right_unique_onI) lemma right_uniqueD: assumes "right_unique R" and "R x y" "R x y'" shows "y = y'" using assms by (urule (d) right_unique_onD where chained = insert) simp_all lemma right_unique_eq_right_unique_at: "(right_unique :: ('a \ 'b \ bool) \ bool) = right_unique_at (\ :: 'b \ bool)" by (intro ext iffI right_uniqueI) (auto dest: right_unique_atD right_uniqueD) lemma right_unique_eq_right_unique_at_uhint [uhint]: assumes "P \ (\ :: 'b \ bool)" shows "right_unique :: ('a \ 'b \ bool) \ _ \ right_unique_at P" using assms by (simp only: right_unique_eq_right_unique_at) lemma right_unique_on_if_right_unique: fixes P :: "'a \ bool" and R :: "'a \ 'b \ bool" 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 \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,108 +1,108 @@ \<^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 \ 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" + definition "rel_surjective_at_pred P R \ \y : P. 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: 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 lemma mono_rel_surjective_at: "((\) \\<^sub>m (\) \ (\)) (rel_surjective_at :: ('b \ bool) \ ('a \ 'b \ bool) \ bool)" - by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI) fastforce + by fastforce lemma rel_surjective_at_iff_le_codom: "rel_surjective_at (P :: 'b \ bool) (R :: 'a \ 'b \ bool) \ P \ in_codom R" by force lemma rel_surjective_at_compI: fixes P :: "'c \ bool" and R :: "'a \ 'b \ bool" and S :: "'b \ 'c \ bool" assumes surj_R: "rel_surjective_at (in_dom S) R" and surj_S: "rel_surjective_at P S" shows "rel_surjective_at P (R \\ S)" proof (rule rel_surjective_atI) fix y assume "P y" then obtain x where "S x y" using surj_S by auto moreover then have "in_dom S x" by auto moreover then obtain z where "R z x" using surj_R by auto ultimately show "in_codom (R \\ S) y" by blast qed consts rel_surjective :: "'a \ bool" overloading rel_surjective \ "rel_surjective :: ('b \ 'a \ bool) \ bool" begin definition "(rel_surjective :: ('b \ 'a \ bool) \ _) \ rel_surjective_at (\ :: 'a \ bool)" end lemma rel_surjective_eq_rel_surjective_at: "(rel_surjective :: ('b \ 'a \ bool) \ _) = rel_surjective_at (\ :: 'a \ bool)" unfolding rel_surjective_def .. lemma rel_surjective_eq_rel_surjective_at_uhint [uhint]: assumes "P \ (\ :: 'a \ bool)" shows "(rel_surjective :: ('b \ 'a \ bool) \ _) \ rel_surjective_at P" using assms by (simp add: rel_surjective_eq_rel_surjective_at) lemma rel_surjectiveI: assumes "\y. in_codom R y" shows "rel_surjective R" using assms by (urule rel_surjective_atI) lemma rel_surjectiveE: assumes "rel_surjective R" obtains x where "R x y" using assms by (urule (e) rel_surjective_atE where chained = insert) simp 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" by (urule rel_surjective_at_rel_inv_iff_left_total_on) lemma left_total_rel_inv_iff_rel_surjective [iff]: "left_total R\ \ rel_surjective R" by (urule left_total_on_rel_inv_iff_rel_surjective_at) lemma rel_surjective_at_if_surjective: fixes P :: "'a \ bool" and R :: "'b \ 'a \ bool" assumes "rel_surjective R" shows "rel_surjective_at P R" 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,117 +1,115 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Symmetric\ theory Binary_Relations_Symmetric imports Functions_Monotone - ML_Unification.ML_Unification_HOL_Setup - ML_Unification.Unify_Resolve_Tactics begin consts symmetric_on :: "'a \ 'b \ 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" + definition "symmetric_on_pred P R \ \x y : P. 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 \ 'a \ bool)" by (blast dest: symmetric_onD) lemma antimono_symmetric_on: "antimono (symmetric_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool)" by (intro antimonoI) (auto dest: symmetric_onD) lemma symmetric_on_if_le_pred_if_symmetric_on: fixes P' :: "'a \ bool" and R :: "'a \ 'a \ bool" assumes "symmetric_on P R" and "P' \ P" shows "symmetric_on P' R" using assms antimono_symmetric_on by blast consts symmetric :: "'a \ bool" overloading symmetric \ "symmetric :: ('a \ 'a \ bool) \ bool" begin definition "(symmetric :: ('a \ 'a \ bool) \ _) \ symmetric_on (\ :: 'a \ bool)" end lemma symmetric_eq_symmetric_on: "(symmetric :: ('a \ 'a \ bool) \ _) = symmetric_on (\ :: 'a \ bool)" unfolding symmetric_def .. lemma symmetric_eq_symmetric_on_uhint [uhint]: "P \ (\ :: 'a \ bool) \ (symmetric :: ('a \ 'a \ bool) \ _) \ symmetric_on P" by (simp add: symmetric_eq_symmetric_on) lemma symmetricI [intro]: assumes "\x y. R x y \ R y x" shows "symmetric R" using assms by (urule symmetric_onI) lemma symmetricD: assumes "symmetric R" and "R x y" shows "R y x" using assms by (urule (d) symmetric_onD where chained = insert) simp_all lemma symmetric_on_if_symmetric: fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" 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 :: 'a \ 'a \ bool)" 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 [iff]: "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 (\ :: 'a \ 'a \ bool)" 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,113 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Transitive\ theory Binary_Relations_Transitive imports Functions_Monotone - ML_Unification.ML_Unification_HOL_Setup - ML_Unification.Unify_Resolve_Tactics begin consts transitive_on :: "'a \ 'b \ 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" + definition "transitive_on_pred P R \ \x y z : P. 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 \ 'a \ bool)" by (auto intro!: transitive_onI dest: transitive_onD) lemma antimono_transitive_on: "antimono (transitive_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool)" by (intro antimonoI) (auto dest: transitive_onD) consts transitive :: "'a \ bool" overloading transitive \ "transitive :: ('a \ 'a \ bool) \ bool" begin definition "(transitive :: ('a \ 'a \ bool) \ bool) \ transitive_on (\ :: 'a \ bool)" end lemma transitive_eq_transitive_on: "(transitive :: ('a \ 'a \ bool) \ _) = transitive_on (\ :: 'a \ bool)" unfolding transitive_def .. lemma transitive_eq_transitive_on_uhint [uhint]: "P \ (\ :: 'a \ bool) \ (transitive :: ('a \ 'a \ bool) \ _) \ transitive_on P" by (simp add: transitive_eq_transitive_on) lemma transitiveI [intro]: assumes "\x y z. R x y \ R y z \ R x z" shows "transitive R" using assms by (urule transitive_onI) lemma transitiveD [dest]: assumes "transitive R" and "R x y" "R y z" shows "R x z" using assms by (urule (d) transitive_onD where chained = insert) simp_all lemma transitive_on_if_transitive: fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" 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" by (urule transitive_on_if_rel_comp_self_imp) (use assms in auto) 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 [iff]: "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 :: 'a \ 'a \ bool)" by fast paragraph \Instantiations\ lemma transitive_eq: "transitive (=)" by (rule transitiveI) (rule trans) lemma transitive_top: "transitive (\ :: 'a \ 'a \ bool)" by (rule transitiveI) auto end \ No newline at end of file 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,221 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Relators\ theory Function_Relators imports Binary_Relation_Functions Functions_Base Predicates_Lattice + ML_Unification.ML_Unification_HOL_Setup + ML_Unification.Unify_Resolve_Tactics 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)" +consts Dep_Fun_Rel :: "'a \ 'b \ 'c \ 'd \ bool" +consts Fun_Rel :: "'a \ 'b \ 'c \ 'd \ bool" 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) + "_Fun_Rel" :: "'a \ 'b \ 'c \ 'd \ bool" ("(_) \ (_)" [41, 40] 40) + "_Dep_Fun_Rel_rel" :: "idt \ idt \ 'a \ 'b \ 'c \ 'd \ bool" + ("'(_/ _/ \/ _') \ (_)" [41, 41, 41, 40] 40) + "_Dep_Fun_Rel_rel_if" :: "idt \ idt \ 'a \ bool \ 'b \ 'c \ 'd \ bool" + ("'(_/ _/ \/ _/ |/ _') \ (_)" [41, 41, 41, 40, 40] 40) + "_Dep_Fun_Rel_pred" :: "idt \ 'a \ 'b \ 'c \ 'd \ bool" + ("'(_/ :/ _') \ (_)" [41, 41, 40] 40) + "_Dep_Fun_Rel_pred_if" :: "idt \ 'a \ bool \ 'b \ 'c \ 'd \ bool" + ("'(_/ :/ _/ |/ _') \ (_)" [41, 41, 40, 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) + "_Fun_Rel" :: "'a \ 'b \ 'c \ 'd \ bool" ("(_) \ (_)" [41, 40] 40) + "_Dep_Fun_Rel_rel" :: "idt \ idt \ 'a \ 'b \ 'c \ 'd \ bool" + ("'(_/ _/ \/ _') \ (_)" [41, 41, 41, 40] 40) + "_Dep_Fun_Rel_rel_if" :: "idt \ idt \ 'a \ bool \ 'b \ 'c \ 'd \ bool" + ("'(_/ _/ \/ _/ |/ _') \ (_)" [41, 41, 41, 40, 40] 40) + "_Dep_Fun_Rel_pred" :: "idt \ 'a \ 'b \ 'c \ 'd \ bool" + ("'(_/ :/ _') \ (_)" [41, 41, 40] 40) + "_Dep_Fun_Rel_pred_if" :: "idt \ 'a \ bool \ 'b \ 'c \ 'd \ bool" + ("'(_/ :/ _/ |/ _') \ (_)" [41, 41, 40, 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)" + "R \ S" \ "CONST Fun_Rel R S" + "(x y \ R) \ S" \ "CONST Dep_Fun_Rel R (\x y. S)" + "(x y \ R | B) \ S" \ "CONST Dep_Fun_Rel R (\x y. CONST rel_if B S)" + "(x : P) \ R" \ "CONST Dep_Fun_Rel P (\x. R)" + "(x : P | B) \ R" \ "CONST Dep_Fun_Rel P (\x. CONST rel_if B R)" + +definition "Dep_Fun_Rel_rel R S f g \ \x y. R x y \ S x y (f x) (g y)" +adhoc_overloading Dep_Fun_Rel Dep_Fun_Rel_rel + +definition "Fun_Rel_rel (R :: 'a \ 'b \ bool) (S :: 'c \ 'd \ bool) \ + ((_ _ \ R) \ S) :: ('a \ 'c) \ ('b \ 'd) \ bool" +adhoc_overloading Fun_Rel Fun_Rel_rel + +definition "Dep_Fun_Rel_pred P R f g \ \x. P x \ R x (f x) (g x)" +adhoc_overloading Dep_Fun_Rel Dep_Fun_Rel_pred + +definition "Fun_Rel_pred (P :: 'a \ bool) (R :: 'b \ 'c \ bool) \ + ((_ : P) \ R) :: ('a \ 'b) \ ('a \ 'c) \ bool" +adhoc_overloading Fun_Rel Fun_Rel_pred + +lemma Fun_Rel_rel_eq_Dep_Fun_Rel_rel: + "((R :: 'a \ 'b \ bool) \ (S :: 'c \ 'd \ bool)) = ((_ _ \ R) \ S)" + unfolding Fun_Rel_rel_def by simp + +lemma Fun_Rel_rel_eq_Dep_Fun_Rel_rel_uhint [uhint]: + assumes "R \ R'" + and "S' \ (\(_ :: 'a) (_ :: 'b). S)" + shows "((R :: 'a \ 'b \ bool) \ (S :: 'c \ 'd \ bool)) = ((x y \ R') \ S' x y)" + using assms by (simp add: Fun_Rel_rel_eq_Dep_Fun_Rel_rel) + +lemma Fun_Rel_rel_iff_Dep_Fun_Rel_rel: + "((R :: 'a \ 'b \ bool) \ (S :: 'c \ 'd \ bool)) f g \ ((_ _ \ R) \ S) f g" + by (simp add: Fun_Rel_rel_eq_Dep_Fun_Rel_rel) + +lemma Fun_Rel_pred_eq_Dep_Fun_Rel_pred: + "((P :: 'a \ bool) \ (R :: 'b \ 'c \ bool)) = ((_ : P) \ R)" + unfolding Fun_Rel_pred_def by simp + +lemma Fun_Rel_pred_eq_Dep_Fun_Rel_pred_uhint [uhint]: + assumes "P \ P'" + and "R' \ (\(_ :: 'a). R)" + shows "((P :: 'a \ bool) \ (R :: 'b \ 'c \ bool)) = ((x : P') \ R' x)" + using assms by (simp add: Fun_Rel_pred_eq_Dep_Fun_Rel_pred) + +lemma Fun_Rel_pred_iff_Dep_Fun_Rel_pred: + "((P :: 'a \ bool) \ (R :: 'b \ 'c \ bool)) f g \ ((_ : P) \ R) (f :: 'a \ 'b) (g :: 'a \ 'c)" + by (simp add: Fun_Rel_pred_eq_Dep_Fun_Rel_pred) 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" + 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" +lemma Dep_Fun_Rel_relD [dest]: + 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" +lemma Dep_Fun_Rel_relE: + assumes "((x y \ R) \ S x y) f g" + obtains "\x y. R x y \ S x y (f x) (g y)" + using assms unfolding Dep_Fun_Rel_rel_def by blast + +lemma Fun_Rel_relI [intro]: + assumes "\x y. R x y \ S (f x) (g y)" + shows "(R \ S) f g" + using assms by (urule Dep_Fun_Rel_relI) + +lemma Fun_Rel_relD [dest]: + assumes "(R \ S) f g" and "R x y" - obtains "S x y (f x) (g y)" - using assms unfolding Dep_Fun_Rel_rel_def by blast + shows "S (f x) (g y)" + using assms by (urule Dep_Fun_Rel_relD) + +lemma Fun_Rel_relE: + assumes "((x y \ R) \ S x y) f g" + obtains "\x y. R x y \ S x y (f x) (g y)" + using assms by (urule (e) Dep_Fun_Rel_relE) lemma Dep_Fun_Rel_predI [intro]: assumes "\x. P x \ R x (f x) (g x)" - shows "([x \ P] \ R x) f g" + 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" +lemma Dep_Fun_Rel_predD [dest]: + 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)" +lemma Dep_Fun_Rel_predE: + assumes "((x : P) \ R x) f g" + obtains "\x. P x \ R x (f x) (g x)" using assms unfolding Dep_Fun_Rel_pred_def by blast +lemma Fun_Rel_predI [intro]: + assumes "\x. P x \ R (f x) (g x)" + shows "(P \ R) f g" + using assms by (urule Dep_Fun_Rel_predI) + +lemma Fun_Rel_predD [dest]: + assumes "(P \ R) f g" + and "P x" + shows "R (f x) (g x)" + using assms by (urule Dep_Fun_Rel_predD) + +lemma Fun_Rel_predE: + assumes "(P \ R) f g" + obtains "\x. P x \ R (f x) (g x)" + using assms by (urule (e) Dep_Fun_Rel_predE) + lemma rel_inv_Dep_Fun_Rel_rel_eq [simp]: - "([x y \ R] \ S x y)\ = ([y x \ R\] \ (S x y)\)" + fixes R :: "'a \ 'b \ bool" and S :: "'a \ 'b \ 'c \ 'd \ bool" + shows "((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)\)" + fixes P :: "'a \ bool" and R :: "'a \ 'b \ 'c \ bool" + shows "((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)" + fixes P :: "'a \ bool" and R :: "'a \ 'b \ 'c \ bool" + shows "((x : P) \ R x) = ((x (_ :: 'a) \ (((\) 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)" + assumes "((x y \ R) \ S x y) f g" + and "\x y. R x y \ ((x' y' \ T x y) \ U x y 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) \ 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)" + 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)" + assumes "((x : P) \ R x) f g" + and "\x. P x \ ((x' y' \ S x) \ T x x' y') f' g'" + and "\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)" + 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>" + 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 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>)" + 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 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)\"] + 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/Functions_Base.thy b/thys/Transport/HOL_Basics/Functions/Functions_Base.thy --- a/thys/Transport/HOL_Basics/Functions/Functions_Base.thy +++ b/thys/Transport/HOL_Basics/Functions/Functions_Base.thy @@ -1,68 +1,147 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Functions\ subsection \Basic Functions\ theory Functions_Base - imports HOL_Basics_Base + imports + Binary_Relation_Functions begin definition "id x \ x" lemma id_eq_self [simp]: "id x = x" unfolding id_def .. -definition "comp f g x \ f (g x)" +consts comp :: "'a \ 'b \ 'c" bundle comp_syntax begin notation comp (infixl "\" 55) end bundle no_comp_syntax begin no_notation comp (infixl "\" 55) end unbundle comp_syntax +definition "comp_fun f g x \ f (g x)" +adhoc_overloading comp comp_fun + lemma comp_eq [simp]: "(f \ g) x = f (g x)" - unfolding comp_def .. + unfolding comp_fun_def .. lemma id_comp_eq [simp]: "id \ f = f" by (rule ext) simp lemma comp_id_eq [simp]: "f \ id = f" by (rule ext) simp +lemma bottom_comp_eq [simp]: "\ \ f = \" by auto +lemma top_comp_eq [simp]: "\ \ f = \" by auto + + definition "dep_fun_map f g h x \ g x (f x) (h (f x))" -abbreviation "fun_map f g h \ dep_fun_map f (\_ _. g) h" +definition "fun_map f g h \ dep_fun_map f (\_ _. g) h" bundle dep_fun_map_syntax begin syntax "_fun_map" :: "('a \ 'b) \ ('c \ 'd) \ ('b \ 'c) \ - ('a \ 'd)" ("(_) \ (_)" [41, 40] 40) + ('a \ 'd)" ("(_) \ (_)" [41, 40] 40) "_dep_fun_map" :: "idt \ ('a \ 'b) \ ('c \ 'd) \ ('b \ 'c) \ - ('a \ 'd)" ("[_/ : / _] \ (_)" [41, 41, 40] 40) + ('a \ 'd)" ("'(_/ : / _') \ (_)" [41, 41, 40] 40) end bundle no_dep_fun_map_syntax begin no_syntax "_fun_map" :: "('a \ 'b) \ ('c \ 'd) \ ('b \ 'c) \ - ('a \ 'd)" ("(_) \ (_)" [41, 40] 40) + ('a \ 'd)" ("(_) \ (_)" [41, 40] 40) "_dep_fun_map" :: "idt \ ('a \ 'b) \ ('c \ 'd) \ ('b \ 'c) \ - ('a \ 'd)" ("[_/ : / _] \ (_)" [41, 41, 40] 40) + ('a \ 'd)" ("'(_/ : / _') \ (_)" [41, 41, 40] 40) end unbundle dep_fun_map_syntax translations - "f \ g" \ "CONST fun_map f g" - "[x : f] \ g" \ "CONST dep_fun_map f (\x. g)" + "f \ g" \ "CONST fun_map f g" + "(x : f) \ g" \ "CONST dep_fun_map f (\x. g)" -lemma dep_fun_map_eq [simp]: "([x : f] \ g x) h x = g x (f x) (h (f x))" +lemma fun_map_eq_dep_fun_map: "(f \ g) = ((_ : f) \ (\_. g))" + unfolding fun_map_def by simp + +lemma fun_map_eq_dep_fun_map_uhint [uhint]: + assumes "f \ f'" + and "g' \ (\_ _. g)" + shows "(f \ g) = ((x : f') \ g' x)" + using assms by (simp add: fun_map_eq_dep_fun_map) + +lemma dep_fun_map_eq [simp]: "((x : f) \ g x) h x = g x (f x) (h (f x))" unfolding dep_fun_map_def .. -lemma fun_map_eq_comp [simp]: "(f \ g) h = g \ h \ f" - by fastforce +lemma fun_map_eq_comp [simp]: "(f \ g) h = g \ h \ f" + by (intro ext) (urule dep_fun_map_eq) -lemma fun_map_eq [simp]: "(f \ g) h x = g (h (f x))" +lemma fun_map_eq [simp]: "(f \ g) h x = g (h (f x))" unfolding fun_map_eq_comp by simp lemma fun_map_id_eq_comp [simp]: "fun_map id = (\)" by (intro ext) simp -lemma fun_map_id_eq_comp' [simp]: "(f \ id) h = h \ f" +lemma fun_map_id_eq_comp' [simp]: "(f \ id) h = h \ f" by (intro ext) simp +consts has_inverse_on :: "'a \ 'b \ 'c \ bool" + +definition "has_inverse_on_pred (P :: 'a \ bool) f \ in_codom_on P ((=) \ f)" +adhoc_overloading has_inverse_on has_inverse_on_pred + +lemma has_inverse_on_pred_eq_in_codom_on: "has_inverse_on P = in_codom_on P \ ((\) (=))" + unfolding has_inverse_on_pred_def by auto + +lemma has_inverse_onI [intro]: + assumes "P x" + shows "has_inverse_on P f (f x)" + using assms unfolding has_inverse_on_pred_def by auto + +lemma has_inverse_on_if_pred_if_eq: + assumes "y = f x" + and "P x" + shows "has_inverse_on P f y" + using assms by auto + +lemma has_inverse_onE [elim]: + assumes "has_inverse_on P f y" + obtains x where "P x" "y = f x" + using assms unfolding has_inverse_on_pred_def by auto + +context + notes has_inverse_on_pred_eq_in_codom_on[simp] +begin + +lemma has_inverse_on_bottom_eq [simp]: "has_inverse_on \ = \" + by (urule in_codom_bottom_pred_eq_bottom) + +lemma has_inverse_on_eq_pred_eq_eq_app [simp]: "has_inverse_on ((=) P) f = (=) (f P)" + by (urule in_codom_on_eq_pred_eq) + +lemma has_inverse_on_sup_eq_has_inverse_on_sup_has_inverse_on [simp]: + "has_inverse_on (P \ Q) = has_inverse_on P \ has_inverse_on Q" + supply ext[simp] by (urule in_codom_on_sup_pred_eq_in_codom_on_sup_in_codom_on) + +end + +definition "has_inverse \ has_inverse_on \" + +lemma has_inverse_eq_has_inverse_on_top: "has_inverse = has_inverse_on \" + unfolding has_inverse_def by simp + +lemma has_inverse_eq_has_inverse_on_top_uhint [uhint]: + assumes "P \ \" + shows "has_inverse \ has_inverse_on P" + using assms by (simp add: has_inverse_eq_has_inverse_on_top) + +lemma has_inverse_iff_has_inverse_on_top: "has_inverse f y \ has_inverse_on \ f y" + by (simp add: has_inverse_eq_has_inverse_on_top) + +lemma has_inverseI [intro]: + assumes "y = f x" + shows "has_inverse f y" + using assms by (urule has_inverse_on_if_pred_if_eq) simp + +lemma has_inverseE [elim]: + assumes "has_inverse f y" + obtains x where "y = f x" + using assms by (urule (e) has_inverse_onE) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Functions/Functions_Restrict.thy b/thys/Transport/HOL_Basics/Functions/Functions_Restrict.thy --- a/thys/Transport/HOL_Basics/Functions/Functions_Restrict.thy +++ b/thys/Transport/HOL_Basics/Functions/Functions_Restrict.thy @@ -1,40 +1,46 @@ \<^marker>\creator "Kevin Kappelmann"\ theory Functions_Restrict - imports HOL.HOL + imports HOL_Basics_Base begin consts fun_restrict :: "'a \ 'b \ 'a" -overloading - fun_restrict_pred \ "fun_restrict :: ('a \ 'b) \ ('a \ bool) \ 'a \ 'b" -begin - definition "fun_restrict_pred f P x \ if P x then f x else undefined" -end - bundle fun_restrict_syntax begin notation fun_restrict ("(_)\(\<^bsub>_\<^esub>)" [1000]) end bundle no_fun_restrict_syntax begin no_notation fun_restrict ("(_)\(\<^bsub>_\<^esub>)" [1000]) end +definition "fun_restrict_pred f P x \ if P x then f x else undefined" +adhoc_overloading fun_restrict fun_restrict_pred + context includes fun_restrict_syntax begin lemma fun_restrict_eq [simp]: assumes "P x" shows "f\\<^bsub>P\<^esub> x = f x" using assms unfolding fun_restrict_pred_def by auto lemma fun_restrict_eq_if_not [simp]: assumes "\(P x)" shows "f\\<^bsub>P\<^esub> x = undefined" using assms unfolding fun_restrict_pred_def by auto +lemma fun_restrict_eq_if: "f\\<^bsub>P\<^esub> x = (if P x then f x else undefined)" + by auto + +lemma fun_restrict_cong [cong]: + assumes "\x. P x \ P' x" + and "\x. P' x \ f x = g x" + shows "f\\<^bsub>P\<^esub> = g\\<^bsub>P'\<^esub>" + using assms by (intro ext) (auto simp: fun_restrict_eq_if) + end end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Functions/LFunctions.thy b/thys/Transport/HOL_Basics/Functions/LFunctions.thy --- a/thys/Transport/HOL_Basics/Functions/LFunctions.thy +++ b/thys/Transport/HOL_Basics/Functions/LFunctions.thy @@ -1,13 +1,13 @@ \<^marker>\creator "Kevin Kappelmann"\ theory LFunctions imports - Functions_Base Function_Properties Function_Relators + Functions_Restrict begin paragraph \Summary\ text \Basic concepts on functions.\ end 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,182 +1,171 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Bijections\ theory Functions_Bijection imports Functions_Inverse Functions_Monotone begin consts bijection_on :: "'a \ 'b \ 'c \ 'd \ bool" -overloading - bijection_on_pred \ "bijection_on :: ('a \ bool) \ ('b \ bool) \ ('a \ 'b) \ ('b \ 'a) \ bool" +definition "bijection_on_pred (P :: 'a \ bool) (Q :: 'b \ bool) f g \ + (P \\<^sub>m Q) f \ + (Q \\<^sub>m P) g \ + inverse_on P f g \ + inverse_on Q g f" +adhoc_overloading bijection_on bijection_on_pred + +context + fixes P :: "'a \ bool" and Q :: "'b \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" 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" + assumes "(P \\<^sub>m Q) f" + and "(Q \\<^sub>m P) g" and "inverse_on P f g" - and "inverse_on P' g f" - shows "bijection_on P P' f g" + and "inverse_on Q g f" + shows "bijection_on P Q f g" using assms unfolding bijection_on_pred_def by blast lemma bijection_onE [elim]: - 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" + assumes "bijection_on P Q f g" + obtains "(P \\<^sub>m Q) f" "(Q \\<^sub>m P) g" + "inverse_on P f g" "inverse_on Q g f" using assms unfolding bijection_on_pred_def by blast -context - fixes P :: "'a \ bool" and P' :: "'b \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" -begin - lemma mono_wrt_pred_if_bijection_on_left: - assumes "bijection_on P P' f g" - shows "([P] \\<^sub>m P') f" + assumes "bijection_on P Q f g" + shows "(P \\<^sub>m Q) 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" + assumes "bijection_on P Q f g" + shows "(Q \\<^sub>m P) g" using assms by (elim bijection_onE) lemma bijection_on_pred_right: - assumes "bijection_on P P' f g" + assumes "bijection_on P Q f g" and "P x" - shows "P' (f x)" + shows "Q (f x)" using assms by blast lemma bijection_on_pred_left: - assumes "bijection_on P P' f g" - and "P' y" + assumes "bijection_on P Q f g" + and "Q y" shows "P (g y)" using assms by blast lemma inverse_on_if_bijection_on_left_right: - assumes "bijection_on P P' f g" + assumes "bijection_on P Q 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" + assumes "bijection_on P Q f g" + shows "inverse_on Q g f" using assms by (elim bijection_onE) lemma bijection_on_left_right_eq_self: - assumes "bijection_on P P' f g" + assumes "bijection_on P Q 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" + assumes "bijection_on P Q f g" + and "Q 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 - end context - fixes P :: "'a \ bool" and P' :: "'b \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" + fixes P :: "'a \ bool" and Q :: "'b \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" begin +lemma bijection_on_right_left_if_bijection_on_left_right: + assumes "bijection_on P Q f g" + shows "bijection_on Q P g f" + using assms by auto + lemma injective_on_if_bijection_on_left: - assumes "bijection_on P P' f g" + assumes "bijection_on P Q 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" + assumes "bijection_on P Q f g" + shows "injective_on Q g" by (intro injective_on_if_inverse_on) (fact inverse_on_if_bijection_on_right_left[OF assms]) end lemma bijection_on_compI: - fixes P :: "'a \ bool" and P' :: "'b \ bool" and P'' :: "'c \ bool" - and f :: "'a \ 'b" and g :: "'b \ 'a" and f' :: "'b \ 'c" and g' :: "'c \ 'b" + fixes P :: "'a \ bool" and P' :: "'b \ bool" and Q :: "'c \ bool" assumes "bijection_on P P' f g" - and "bijection_on P' P'' f' g'" - shows "bijection_on P P'' (f' \ f) (g \ g')" + and "bijection_on P' Q f' g'" + shows "bijection_on P Q (f' \ f) (g \ g')" using assms by (intro bijection_onI) (auto intro: dep_mono_wrt_pred_comp_dep_mono_wrt_pred_compI' inverse_on_compI - elim!: bijection_onE) + elim!: bijection_onE simp: mono_wrt_pred_eq_dep_mono_wrt_pred) consts bijection :: "'a \ 'b \ bool" -overloading - bijection \ "bijection :: ('a \ 'b) \ ('b \ 'a) \ bool" -begin - definition "(bijection :: ('a \ 'b) \ ('b \ 'a) \ bool) \ - bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool)" -end +definition "(bijection_rel :: ('a \ 'b) \ ('b \ 'a) \ bool) \ + bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool)" +adhoc_overloading bijection bijection_rel lemma bijection_eq_bijection_on: "(bijection :: ('a \ 'b) \ ('b \ 'a) \ bool) = bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool)" - unfolding bijection_def .. + unfolding bijection_rel_def .. lemma bijection_eq_bijection_on_uhint [uhint]: assumes "P \ (\ :: 'a \ bool)" and "Q \ (\ :: 'b \ bool)" shows "(bijection :: ('a \ 'b) \ ('b \ 'a) \ bool) = bijection_on P Q" using assms by (simp add: bijection_eq_bijection_on) context - fixes P :: "'a \ bool" and P' :: "'b \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" + fixes P :: "'a \ bool" and Q :: "'b \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" begin lemma bijectionI [intro]: assumes "inverse f g" and "inverse g f" shows "bijection f g" by (urule bijection_onI) (simp | urule assms)+ lemma bijectionE [elim]: assumes "bijection f g" obtains "inverse f g" "inverse g f" using assms by (urule (e) 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) end -context - fixes P :: "'a \ bool" and P' :: "'b \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" -begin - lemma bijection_right_left_if_bijection_left_right: + fixes f :: "'a \ 'b" and g :: "'b \ 'a" assumes "bijection f g" shows "bijection g f" using assms by auto paragraph \Instantiations\ -lemma bijection_on_self_id: "bijection_on P P (id :: 'a \ 'a) (id :: 'a \ 'a)" - by (intro bijection_onI inverse_onI dep_mono_wrt_predI) simp_all +lemma bijection_on_self_id: "bijection_on (P :: 'a \ bool) P id id" + by (intro bijection_onI inverse_onI mono_wrt_predI) simp_all -end 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,71 +1,70 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Injective\ theory Functions_Injective imports + Bounded_Quantifiers Functions_Base HOL_Syntax_Bundles_Lattices - ML_Unification.ML_Unification_HOL_Setup - ML_Unification.Unify_Resolve_Tactics begin consts injective_on :: "'a \ 'b \ 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'" + definition "injective_on_pred P f \ \x x' : P. 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 consts injective :: "'a \ bool" overloading injective \ "injective :: ('a \ 'b) \ bool" begin definition "(injective :: ('a \ 'b) \ bool) \ injective_on (\ :: 'a \ bool)" end lemma injective_eq_injective_on: "(injective :: ('a \ 'b) \ bool) = injective_on (\ :: 'a \ bool)" unfolding injective_def .. lemma injective_eq_injective_on_uhint [uhint]: assumes "P \ (\ :: 'a \ bool)" shows "injective :: ('a \ 'b) \ bool \ injective_on P" using assms by (simp add: injective_eq_injective_on) lemma injectiveI [intro]: assumes "\x x'. f x = f x' \ x = x'" shows "injective f" using assms by (urule injective_onI) lemma injectiveD: assumes "injective f" and "f x = f x'" shows "x = x'" using assms by (urule (d) injective_onD where chained = insert) simp_all 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,93 +1,99 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Inverse\ theory Functions_Inverse imports Functions_Injective - Functions_Monotone + Binary_Relations_Function_Evaluation begin consts inverse_on :: "'a \ 'b \ 'c \ 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" + definition "inverse_on_pred P f g \ \x : P. 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 \ 'b) (g :: 'b \ 'a)" 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 lemma inverse_on_compI: fixes P :: "'a \ bool" and P' :: "'b \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" and f' :: "'b \ 'c" and g' :: "'c \ 'b" assumes "inverse_on P f g" and "inverse_on P' f' g'" - and "([P] \\<^sub>m P') f" + and "(P \\<^sub>m P') f" shows "inverse_on P (f' \ f) (g \ g')" using assms by (intro inverse_onI) (auto dest!: inverse_onD) consts inverse :: "'a \ 'b \ bool" overloading inverse \ "inverse :: ('a \ 'b) \ ('b \ 'a) \ bool" begin definition "(inverse :: ('a \ 'b) \ ('b \ 'a) \ bool) \ inverse_on (\ :: 'a \ bool)" end lemma inverse_eq_inverse_on: "(inverse :: ('a \ 'b) \ ('b \ 'a) \ bool) = inverse_on (\ :: 'a \ bool)" unfolding inverse_def .. lemma inverse_eq_inverse_on_uhint [uhint]: assumes "P \ \ :: 'a \ bool" shows "inverse :: ('a \ 'b) \ ('b \ 'a) \ bool \ inverse_on P" using assms by (simp add: inverse_eq_inverse_on) lemma inverseI [intro]: assumes "\x. g (f x) = x" shows "inverse f g" using assms by (urule inverse_onI) lemma inverseD: assumes "inverse f g" shows "g (f x) = x" using assms by (urule (d) inverse_onD where chained = insert) simp_all lemma inverse_on_if_inverse: fixes P :: "'a \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" assumes "inverse f g" shows "inverse_on P f g" using assms by (intro inverse_onI) (blast dest: inverseD) -lemma inverse_THE_eq_if_injective: +lemma right_unique_eq_app_if_injective: assumes "injective f" - shows "inverse f (\z. THE y. z = f y)" - using assms injectiveD by fastforce + shows "right_unique (\y x. y = f x)" + using assms by (auto dest: injectiveD) -lemma injective_inverseE: +lemma inverse_eval_eq_app_if_injective: + assumes "injective f" + shows "inverse f (eval (\y x. y = f x))" + by (urule inverseI eval_eq_if_right_unique_onI right_unique_eq_app_if_injective)+ + (use assms in simp_all) + +lemma inverse_if_injectiveE: assumes "injective (f :: 'a \ 'b)" obtains g :: "'b \ 'a" where "inverse f g" - using assms inverse_THE_eq_if_injective by blast + using assms inverse_eval_eq_app_if_injective by blast 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,350 +1,475 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Monotonicity\ theory Functions_Monotone imports Binary_Relations_Order_Base Function_Relators - Predicates + Predicate_Functions + Predicates_Order 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)" +consts dep_mono_wrt :: "'a \ 'b \ 'c \ bool" +consts mono_wrt :: "'a \ 'b \ 'c \ bool" 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) - "_dep_mono_wrt_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'b \ bool) \ - ('a \ 'b) \ bool" ("[_/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) + "_mono_wrt" :: "'a \ 'b \ 'c \ bool" ("(_) \\<^sub>m (_)" [41, 40] 40) + "_dep_mono_wrt_rel" :: "idt \ idt \ 'a \ 'b \ 'c \ bool" + ("'(_/ _/ \/ _') \\<^sub>m (_)" [41, 41, 41, 40] 40) + "_dep_mono_wrt_rel_if" :: "idt \ idt \ 'a \ bool \ 'b \ 'c \ bool" + ("'(_/ _/ \/ _/ |/ _') \\<^sub>m (_)" [41, 41, 41, 40, 40] 40) + "_dep_mono_wrt_pred" :: "idt \ 'a \ 'b \ 'c \ bool" ("'(_/ :/ _') \\<^sub>m (_)" [41, 41, 40] 40) + "_dep_mono_wrt_pred_if" :: "idt \ 'a \ bool \ 'b \ 'c \ bool" + ("'(_/ :/ _/ |/ _') \\<^sub>m (_)" [41, 41, 40, 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) + "_mono_wrt" :: "'a \ 'b \ 'c \ bool" ("(_) \\<^sub>m (_)" [41, 40] 40) + "_dep_mono_wrt_rel" :: "idt \ idt \ 'a \ 'b \ 'c \ bool" + ("'(_/ _/ \/ _') \\<^sub>m (_)" [41, 41, 41, 40] 40) + "_dep_mono_wrt_rel_if" :: "idt \ idt \ 'a \ bool \ 'b \ 'c \ bool" + ("'(_/ _/ \/ _/ |/ _') \\<^sub>m (_)" [41, 41, 41, 40, 40] 40) + "_dep_mono_wrt_pred" :: "idt \ 'a \ 'b \ 'c \ bool" ("'(_/ :/ _') \\<^sub>m (_)" [41, 41, 40] 40) + "_dep_mono_wrt_pred_if" :: "idt \ 'a \ bool \ 'b \ 'c \ bool" + ("'(_/ :/ _/ |/ _') \\<^sub>m (_)" [41, 41, 40, 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 pred_if B Q)" + "R \\<^sub>m S" \ "CONST mono_wrt R S" + "(x y \ R) \\<^sub>m S" \ "CONST dep_mono_wrt R (\x y. S)" + "(x y \ R | B) \\<^sub>m S" \ "CONST dep_mono_wrt R (\x y. CONST rel_if B S)" + "(x : P) \\<^sub>m Q" \ "CONST dep_mono_wrt P (\x. Q)" + "(x : P | B) \\<^sub>m Q" \ "CONST dep_mono_wrt P (\x. CONST pred_if B Q)" + + +definition "dep_mono_wrt_rel (R :: 'a \ 'a \ bool) + (S :: 'a \ 'a \ 'b \ 'b \ bool) (f :: 'a \ 'b) \ ((x y \ R) \ S x y) f f" +adhoc_overloading dep_mono_wrt dep_mono_wrt_rel + +definition "mono_wrt_rel (R :: 'a \ 'a \ bool) (S :: 'b \ 'b \ bool) \ + ((_ _ \ R) \\<^sub>m S) :: ('a \ 'b) \ bool" +adhoc_overloading mono_wrt mono_wrt_rel + +definition "dep_mono_wrt_pred (P :: 'a \ bool) (Q :: 'a \ 'b \ bool) (f :: 'a \ 'b) \ + ((x : P) \ (\(_ :: 'b). Q x)) f f" +adhoc_overloading dep_mono_wrt dep_mono_wrt_pred + +definition "mono_wrt_pred (P :: 'a \ bool) (Q :: 'b \ bool) \ + (((_ :: 'a) : P) \\<^sub>m Q) :: ('a \ 'b) \ bool" +adhoc_overloading mono_wrt mono_wrt_pred + +lemma mono_wrt_rel_eq_dep_mono_wrt_rel: + "((R :: 'a \ 'a \ bool) \\<^sub>m (S :: 'b \ 'b \ bool)) = ((_ _ \ R) \\<^sub>m S)" + unfolding mono_wrt_rel_def by simp + +lemma mono_wrt_rel_eq_dep_mono_wrt_rel_uhint [uhint]: + assumes "R \ R'" + and "S' \ (\(_ :: 'a) (_ :: 'a). S)" + shows "((R :: 'a \ 'a \ bool) \\<^sub>m (S :: 'b \ 'b \ bool)) = ((x y \ R') \\<^sub>m S' x y)" + using assms by (simp add: mono_wrt_rel_eq_dep_mono_wrt_rel) + +lemma mono_wrt_rel_iff_dep_mono_wrt_rel: + "((R :: 'a \ 'a \ bool) \\<^sub>m (S :: 'b \ 'b \ bool)) f \ + dep_mono_wrt R (\(_ :: 'a) (_ :: 'a). S) (f :: 'a \ 'b)" + by (simp add: mono_wrt_rel_eq_dep_mono_wrt_rel) + +lemma mono_wrt_pred_eq_dep_mono_wrt_pred: + "((P :: 'a \ bool) \\<^sub>m (Q :: 'b \ bool)) = (((_ :: 'a) : P) \\<^sub>m Q)" + unfolding mono_wrt_pred_def by simp + +lemma mono_wrt_pred_eq_dep_mono_wrt_pred_uhint [uhint]: + assumes "P \ P'" + and "Q' \ (\(_ :: 'a). Q)" + shows "((P :: 'a \ bool) \\<^sub>m (Q :: 'b \ bool)) = (((x : P') \\<^sub>m Q' x) :: ('a \ 'b) \ bool)" + using assms by (simp add: mono_wrt_pred_eq_dep_mono_wrt_pred) + +lemma mono_wrt_pred_iff_dep_mono_wrt_pred: + "((P :: 'a \ bool) \\<^sub>m (Q :: 'b \ bool)) f \ (((_ :: 'a) : P) \\<^sub>m Q) (f :: 'a \ 'b)" + by (simp add: mono_wrt_pred_eq_dep_mono_wrt_pred) 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" + 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" +lemma dep_mono_wrt_relE: + assumes "((x y \ R) \\<^sub>m S x y) f" 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" +lemma dep_mono_wrt_relD [dest]: + 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 + using assms unfolding dep_mono_wrt_rel_def by blast + +lemma mono_wrt_relI [intro]: + assumes "\x y. R x y \ S (f x) (f y)" + shows "(R \\<^sub>m S) f" + using assms by (urule dep_mono_wrt_relI) + +lemma mono_wrt_relE: + assumes "(R \\<^sub>m S) f" + obtains "\x y. R x y \ S (f x) (f y)" + using assms by (urule (e) dep_mono_wrt_relE) + +lemma mono_wrt_relD [dest]: + assumes "(R \\<^sub>m S) f" + and "R x y" + shows "S (f x) (f y)" + using assms by (urule dep_mono_wrt_relD) lemma dep_mono_wrt_predI [intro]: assumes "\x. P x \ Q x (f x)" - shows "([x \ P] \\<^sub>m Q x) f" + 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" +lemma dep_mono_wrt_predE: + assumes "((x : P) \\<^sub>m Q x) f" 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" +lemma dep_mono_wrt_predD [dest]: + assumes "((x : P) \\<^sub>m Q x) f" and "P x" shows "Q x (f x)" - using assms by blast + using assms unfolding dep_mono_wrt_pred_def by blast + +lemma mono_wrt_predI [intro]: + assumes "\x. P x \ Q (f x)" + shows "(P \\<^sub>m Q) f" + using assms by (urule dep_mono_wrt_predI) + +lemma mono_wrt_predE: + assumes "(P \\<^sub>m Q) f" + obtains "\x. P x \ Q (f x)" + using assms by (urule (e) dep_mono_wrt_predE) + +lemma mono_wrt_predD [dest]: + assumes "(P \\<^sub>m Q) f" + and "P x" + shows "Q (f x)" + using assms by (urule dep_mono_wrt_predD) + +context + fixes R :: "'a \ 'a \ bool" and S :: "'a \ 'a \ 'b \ 'b \ bool" and f :: "'a \ 'b" + and P :: "'a \ bool" and Q :: "'a \ 'b \ bool" +begin 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" + 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" + 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" + 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" + 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" + "((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" + "((x : P) \ (\(_ :: 'b). 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)" + "((y x \ R\) \\<^sub>m (S x y)\) = ((x y \ R) \\<^sub>m S x y)" by (intro ext) force lemma in_dom_if_rel_if_dep_mono_wrt_rel: - assumes "([x y \ R] \\<^sub>m S x y) f" + 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 +end + +context + fixes R :: "'a \ 'a \ bool" and f :: "'a \ 'b" +begin + 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" + 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" + 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 + 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 + shows "(in_field R \\<^sub>m in_field S) f" + using assms by fast 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" + 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 + using assms by (intro 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 + shows "(P \\<^sub>m Q) f" + using assms by (intro 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" +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 (\))" +end -definition "antimono \ ((\) \\<^sub>m (\))" +definition "mono :: (('a :: ord) \ ('b :: ord)) \ bool + \ (((\) :: 'a \ 'a \ bool) \\<^sub>m ((\) :: 'b \ 'b \ bool))" + +lemma mono_eq_mono_wrt_le [simp]: "(mono :: (('a :: ord) \ ('b :: ord)) \ bool) = + (((\) :: 'a \ 'a \ bool) \\<^sub>m ((\) :: 'b \ 'b \ bool))" + unfolding mono_def by simp + +lemma mono_eq_mono_wrt_le_uhint [uhint]: + assumes "R \ (\) :: 'a \ 'a \ bool" + and "S \ (\) :: 'b \ 'b \ bool" + shows "mono :: (('a :: ord) \ ('b :: ord)) \ bool \ (R \\<^sub>m S)" + using assms by simp + +lemma mono_iff_mono_wrt_le [iff]: "mono f \ ((\) \\<^sub>m (\)) f" by simp lemma monoI [intro]: assumes "\x y. x \ y \ f x \ f y" shows "mono f" - unfolding mono_def using assms by blast + using assms by (urule mono_wrt_relI) lemma monoE [elim]: assumes "mono f" obtains "\x y. x \ y \ f x \ f y" - using assms unfolding mono_def by blast + using assms by (urule (e) mono_wrt_relE) lemma monoD: assumes "mono f" and "x \ y" shows "f x \ f y" - using assms by blast + using assms by (urule mono_wrt_relD) + +definition "antimono :: (('a :: ord) \ ('b :: ord)) \ bool + \ (((\) :: 'a \ 'a \ bool) \\<^sub>m ((\) :: 'b \ 'b \ bool))" + +lemma antimono_eq_mono_wrt_le_ge [simp]: "(antimono :: (('a :: ord) \ ('b :: ord)) \ bool) = + (((\) :: 'a \ 'a \ bool) \\<^sub>m ((\) :: 'b \ 'b \ bool))" + unfolding antimono_def by simp + +lemma antimono_eq_mono_wrt_le_ge_uhint [uhint]: + assumes "R \ (\) :: 'a \ 'a \ bool" + and "S \ (\) :: 'b \ 'b \ bool" + shows "antimono :: (('a :: ord) \ ('b :: ord)) \ bool \ (R \\<^sub>m S)" + using assms by simp + +lemma antimono_iff_mono_wrt_le_ge [iff]: "antimono f \ ((\) \\<^sub>m (\)) f" by simp lemma antimonoI [intro]: - assumes "\x y. x \ y \ f y \ f x" + assumes "\x y. x \ y \ f x \ f y" shows "antimono f" - unfolding antimono_def using assms by blast + by (urule mono_wrt_relI) (urule assms) lemma antimonoE [elim]: assumes "antimono f" - obtains "\x y. x \ y \ f y \ f x" - using assms unfolding antimono_def by blast + obtains "\x y. x \ y \ f x \ f y" + using assms by (urule (e) mono_wrt_relE) lemma antimonoD: assumes "antimono f" and "x \ y" - shows "f y \ f x" - using assms by blast + shows "f x \ f y" + using assms by (urule mono_wrt_relD) -lemma antimono_Dep_Fun_Rel_rel_left: "antimono (\R. [x y \ R] \ S x y)" +lemma antimono_Dep_Fun_Rel_rel_left: "antimono (\(R :: 'a \ 'b \ bool). ((x y \ R) \ S x y))" by (intro antimonoI) auto -lemma antimono_Dep_Fun_Rel_pred_left: "antimono (\P. [x \ P] \ Q x)" +lemma antimono_Dep_Fun_Rel_pred_left: "antimono (\(P :: 'a \ bool). ((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)" +lemma antimono_dep_mono_wrt_rel_left: "antimono (\(R :: 'a \ 'a \ bool). ((x y \ R) \\<^sub>m S x y))" by (intro antimonoI) blast -lemma antimono_dep_mono_wrt_pred_left: "antimono (\P. [x \ P] \\<^sub>m Q x)" +lemma antimono_dep_mono_wrt_pred_left: "antimono (\(P :: 'a \ bool). ((x : P) \\<^sub>m Q x))" 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" + fixes R R' :: "'a \ 'b \ bool" + assumes "((x y \ R) \ S x y) f g" + and "R' \ R" + shows "((x y \ R) \ 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" + fixes P P' :: "'a \ bool" + assumes "((x : P) \ Q x) f g" + and "P' \ P" + shows "((x : P') \ 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" + fixes R R' :: "'a \ 'a \ bool" + assumes "((x y \ R) \\<^sub>m S x y) f" + and "R' \ R" + shows "((x y \ R') \\<^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" + fixes P P' :: "'a \ bool" + assumes "((x : P) \\<^sub>m Q x) f" + and "P' \ P" + shows "((x : P') \\<^sub>m Q x) f" using assms by blast -lemma mono_Dep_Fun_Rel_rel_right: "mono (\S. [x y \ R] \ S x y)" +lemma mono_Dep_Fun_Rel_rel_right: "mono (\(S :: 'a \ 'b \ 'c \ 'd \ bool). ((x y \ R) \ S x y))" by (intro monoI) blast -lemma mono_Dep_Fun_Rel_pred_right: "mono (\Q. [x \ P] \ Q x)" +lemma mono_Dep_Fun_Rel_pred_right: "mono (\(Q :: 'a \ 'b \ 'c \ bool). ((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)" +lemma mono_dep_mono_wrt_rel_right: "mono (\(S :: 'a \ 'a \ 'b \ 'b \ bool). ((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)" +lemma mono_dep_mono_wrt_pred_right: "mono (\(Q :: 'a \ 'b \ bool). ((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" + 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" + 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" + 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" + 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" + 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" + 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" + 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" + 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'" + 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)" + shows "((x y \ R) \\<^sub>m U x y (f x) (f y)) (f' \ f)" 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)" + 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'" + 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)" + shows "((x : P) \\<^sub>m (\y. S x (f x) (f x) y y)) (f' \ f)" 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'" + 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)" + shows "((x : P) \\<^sub>m S x (f x)) (f' \ f)" 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)" + 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) -lemma mono_wrt_rel_top [iff]: "(R \\<^sub>m \) f" by auto -lemma mono_wrt_pred_top [iff]: "([P] \\<^sub>m \) f" by auto +lemma mono_wrt_rel_top [iff]: + fixes R :: "'a \ 'a \ bool" and f :: "'a \ 'b" + shows "(R \\<^sub>m (\ :: 'b \ 'b \ bool)) f" + by auto + +lemma mono_wrt_pred_top [iff]: + fixes P :: "'a \ bool" and f :: "'a \ 'b" + shows "(P \\<^sub>m (\ :: 'b \ bool)) f" + by auto 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_wrt_rel_self_id: + fixes R :: "'a \ 'a \ bool" + shows "(R \\<^sub>m R) (id :: 'a \ 'a)" + by auto + +lemma mono_wrt_pred_self_id: + fixes P :: "'a \ bool" + shows "(P \\<^sub>m P) (id :: 'a \ 'a)" + 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 +lemma mono_rel_comp: "((\) \\<^sub>m (\) \ (\)) (\\)" by (intro mono_wrt_relI Fun_Rel_relI le_relI) fast +lemma mono_rel_inv: "mono rel_inv" by (intro monoI) fast lemma mono_rel_restrict_left: "((\) \\<^sub>m (\) \ (\)) (rel_restrict_left :: ('a \ 'b \ bool) \ ('a \ bool) \ 'a \ 'b \ bool)" - by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI le_relI) fastforce + by (intro mono_wrt_relI Fun_Rel_relI le_relI) fastforce lemma mono_rel_restrict_right: "((\) \\<^sub>m (\) \ (\)) (rel_restrict_right :: ('a \ 'b \ bool) \ ('b \ bool) \ 'a \ 'b \ bool)" - by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI le_relI) fastforce + by (intro mono_wrt_relI Fun_Rel_relI le_relI) fastforce +lemma mono_ball: "((\) \\<^sub>m (\) \ (\)) ball" by fast +lemma mono_bex: "((\) \\<^sub>m (\) \ (\)) bex" by 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,63 +1,66 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Surjective\ theory Functions_Surjective imports - HOL_Syntax_Bundles_Lattices - ML_Unification.ML_Unification_HOL_Setup - ML_Unification.Unify_Resolve_Tactics + Functions_Base begin consts surjective_at :: "'a \ 'b \ 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)" + definition "surjective_at_pred (P :: 'a \ bool) (f :: 'b \ 'a) \ \y : P. has_inverse f y" end lemma surjective_atI [intro]: - assumes "\y. P y \ \x. y = f x" + assumes "\y. P y \ has_inverse f y" 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 +lemma has_inverse_if_pred_if_surjective_at: + assumes "surjective_at P f" + and "P y" + shows "has_inverse f y" + using assms by blast + consts surjective :: "'a \ bool" overloading surjective \ "surjective :: ('b \ 'a) \ bool" begin definition "(surjective :: ('b \ 'a) \ bool) \ surjective_at (\ :: 'a \ bool)" end lemma surjective_eq_surjective_at: "(surjective :: ('b \ 'a) \ bool) = surjective_at (\ :: 'a \ bool)" unfolding surjective_def .. lemma surjective_eq_surjective_at_uhint [uhint]: assumes "P \ \ :: 'a \ bool" shows "surjective :: ('b \ 'a) \ bool \ surjective_at P" using assms by (simp add: surjective_eq_surjective_at) lemma surjectiveI [intro]: - assumes "\y. \x. y = f x" + assumes "\y. has_inverse f y" shows "surjective f" using assms by (urule surjective_atI) lemma surjectiveE: assumes "surjective f" - obtains x where "y = f x" - using assms by (urule (e) surjective_atE where chained = insert) simp + obtains "\y. has_inverse f y" + using assms unfolding surjective_eq_surjective_at by (force dest: has_inverseI) lemma surjective_at_if_surjective: fixes P :: "'a \ bool" and f :: "'b \ '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.thy b/thys/Transport/HOL_Basics/Galois/Galois.thy --- a/thys/Transport/HOL_Basics/Galois/Galois.thy +++ b/thys/Transport/HOL_Basics/Galois/Galois.thy @@ -1,14 +1,13 @@ \<^marker>\creator "Kevin Kappelmann"\ theory Galois imports Galois_Equivalences Galois_Relator begin paragraph \Summary\ text \We define the concept of (partial) Galois connections, Galois equivalences, and the Galois relator. For details refer to \<^cite>\"transport"\.\ - end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Galois/Galois_Property.thy b/thys/Transport/HOL_Basics/Galois/Galois_Property.thy --- a/thys/Transport/HOL_Basics/Galois/Galois_Property.thy +++ b/thys/Transport/HOL_Basics/Galois/Galois_Property.thy @@ -1,70 +1,70 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Property\ theory Galois_Property imports Half_Galois_Property begin context galois_prop begin definition "galois_prop \ ((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) \ ((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>))" notation galois_prop.galois_prop (infix "\" 50) lemma galois_propI [intro]: assumes "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" and "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" unfolding galois_prop_def using assms by auto lemma galois_propI': assumes "\x y. in_dom (\\<^bsub>L\<^esub>) x \ in_codom (\\<^bsub>R\<^esub>) y \ x \\<^bsub>L\<^esub> r y \ l x \\<^bsub>R\<^esub> y" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" - using assms by blast + using assms by (blast elim: galois_rel.left_GaloisE) lemma galois_propE [elim]: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" obtains "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" using assms unfolding galois_prop_def by auto interpretation g : galois_prop S T f g for S T f g. lemma galois_prop_eq_half_galois_prop_left_rel_inf_half_galois_prop_right: "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) = ((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) \ ((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>))" by (intro ext) auto lemma galois_prop_left_rel_right_iff_left_right_rel: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" and "in_dom (\\<^bsub>L\<^esub>) x" "in_codom (\\<^bsub>R\<^esub>) y" shows "x \\<^bsub>L\<^esub> r y \ l x \\<^bsub>R\<^esub> y" using assms by blast lemma rel_inv_galois_prop_eq_galois_prop_rel_inv [simp]: "((\\<^bsub>R\<^esub>) \ (\\<^bsub>L\<^esub>))\ = ((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>))" by (intro ext) blast corollary galois_prop_rel_inv_iff_galois_prop [iff]: "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) f g \ ((\\<^bsub>R\<^esub>) \ (\\<^bsub>L\<^esub>)) g f" by auto end context galois begin lemma galois_prop_left_right_if_transitive_if_deflationary_on_if_inflationary_on_if_mono_wrt_rel: 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>) \" and "deflationary_on (in_codom (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>) \" and "transitive (\\<^bsub>L\<^esub>)" "transitive (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_propI half_galois_prop_left_left_right_if_transitive_if_deflationary_on_if_mono_wrt_rel half_galois_prop_right_left_right_if_transitive_if_inflationary_on_if_mono_wrt_rel) end 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 + using assms by fast 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 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 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 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/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,285 +1,286 @@ \<^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 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 (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 :: 'c \ bool) (R :: 'c \ 'c \ bool). (inflationary_on P R\ :: ('c \ 'c) \ bool) \ deflationary_on P R" and "\(P :: 'c \ bool) (R :: 'c \ 'c \ bool). (deflationary_on P R\ :: ('c \ 'c) \ bool) \ inflationary_on P R" and "\(P :: 'b \ bool). reflexive_on P (\\<^bsub>R\<^esub>) \ reflexive_on P (\\<^bsub>R\<^esub>)" and "\(R :: 'a \ 'a \ bool). 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) + galois_prop.half_galois_prop_right_rel_inv_iff_half_galois_prop_left + mono_wrt_rel_eq_dep_mono_wrt_rel) 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,79 +1,79 @@ \<^marker>\creator "Kevin Kappelmann"\ -subsection \Alignment With Definitions from HOL-Algebra\ +subsection \Alignment With Galois 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 rel_restrict_right_eq[simp] and rel_restrict_leftI[intro!] 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 (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 (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 rel_restrict_right_eq[simp] and rel_restrict_leftI[intro!] rel_restrict_leftE[elim!] in_codom_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_Algebra_Alignment_Orders.thy b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Algebra_Alignment_Orders.thy --- a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Algebra_Alignment_Orders.thy +++ b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Algebra_Alignment_Orders.thy @@ -1,71 +1,71 @@ \<^marker>\creator "Kevin Kappelmann"\ -subsection \Alignment With Definitions from HOL-Algebra\ +subsection \Alignment With Order Definitions from HOL-Algebra\ theory HOL_Algebra_Alignment_Orders imports "HOL-Algebra.Order" HOL_Alignment_Orders begin named_theorems HOL_Algebra_order_alignment context equivalence begin lemma reflexive_on_carrier [HOL_Algebra_order_alignment]: "reflexive_on (carrier S) (.=)" by blast lemma transitive_on_carrier [HOL_Algebra_order_alignment]: "transitive_on (carrier S) (.=)" using trans by blast lemma preorder_on_carrier [HOL_Algebra_order_alignment]: "preorder_on (carrier S) (.=)" using reflexive_on_carrier transitive_on_carrier by blast lemma symmetric_on_carrier [HOL_Algebra_order_alignment]: "symmetric_on (carrier S) (.=)" using sym by blast lemma partial_equivalence_rel_on_carrier [HOL_Algebra_order_alignment]: "partial_equivalence_rel_on (carrier S) (.=)" using transitive_on_carrier symmetric_on_carrier by blast lemma equivalence_rel_on_carrier [HOL_Algebra_order_alignment]: "equivalence_rel_on (carrier S) (.=)" using reflexive_on_carrier partial_equivalence_rel_on_carrier by blast end lemma equivalence_iff_equivalence_rel_on_carrier [HOL_Algebra_order_alignment]: "equivalence S \ equivalence_rel_on (carrier S) (.=\<^bsub>S\<^esub>)" using equivalence.equivalence_rel_on_carrier by (blast dest: intro!: equivalence.intro dest: symmetric_onD transitive_onD) context partial_order begin lemma reflexive_on_carrier [HOL_Algebra_order_alignment]: "reflexive_on (carrier L) (\)" by blast lemma transitive_on_carrier [HOL_Algebra_order_alignment]: "transitive_on (carrier L) (\)" using le_trans by blast lemma preorder_on_carrier [HOL_Algebra_order_alignment]: "preorder_on (carrier L) (\)" using reflexive_on_carrier transitive_on_carrier by blast lemma antisymmetric_on_carrier [HOL_Algebra_order_alignment]: "antisymmetric_on (carrier L) (\)" by blast lemma partial_order_on_carrier [HOL_Algebra_order_alignment]: "partial_order_on (carrier L) (\)" using preorder_on_carrier antisymmetric_on_carrier by blast 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,386 +1,385 @@ \<^marker>\creator "Kevin Kappelmann"\ -subsection \Alignment With Definitions from HOL.Main\ +subsection \Alignment With Binary Relation 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 \ 'a \ bool) \ _ \ antisymmetric_on (mem_of S)" end lemma antisymmetric_on_set_eq_antisymmetric_on_pred [simp]: "(antisymmetric_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool) = antisymmetric_on (mem_of S)" unfolding antisymmetric_on_set_def by simp lemma antisymmetric_on_set_eq_antisymmetric_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "antisymmetric_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool \ antisymmetric_on P" using assms by simp lemma antisymmetric_on_set_iff_antisymmetric_on_pred [iff]: "antisymmetric_on (S :: 'a set) (R :: 'a \ 'a \ bool) \ 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 \ 'b \ bool) \ _ \ rel_injective_on (mem_of S)" definition "rel_injective_at_set (S :: 'a set) :: ('b \ 'a \ bool) \ _ \ 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 \ 'b \ bool) \ bool) = rel_injective_on (mem_of S)" unfolding rel_injective_on_set_def by simp lemma rel_injective_on_set_eq_rel_injective_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "rel_injective_on (S :: 'a set) :: ('a \ 'b \ bool) \ bool \ rel_injective_on P" using assms by simp lemma rel_injective_on_set_iff_rel_injective_on_pred [iff]: "rel_injective_on (S :: 'a set) (R :: 'a \ 'b \ bool) \ 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_eq_rel_injective_at_pred_uhint [uhint]: assumes "P \ mem_of S" shows "rel_injective_at (S :: 'a set) :: ('b \ 'a \ bool) \ bool \ rel_injective_at P" using assms 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 \ 'a \ bool) \ bool \ irreflexive_on (mem_of S)" end lemma irreflexive_on_set_eq_irreflexive_on_pred [simp]: "(irreflexive_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool) = irreflexive_on (mem_of S)" unfolding irreflexive_on_set_def by simp lemma irreflexive_on_set_eq_irreflexive_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "irreflexive_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool \ irreflexive_on P" using assms by simp lemma irreflexive_on_set_iff_irreflexive_on_pred [iff]: "irreflexive_on (S :: 'a set) (R :: 'a \ 'a \ bool) \ 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 \ 'b \ bool) \ _ \ 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 \ 'b \ bool) \ bool) = left_total_on (mem_of S)" unfolding left_total_on_set_def by simp lemma left_total_on_set_eq_left_total_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "left_total_on (S :: 'a set) :: ('a \ 'b \ bool) \ bool \ left_total_on P" using assms by simp lemma left_total_on_set_iff_left_total_on_pred [iff]: "left_total_on (S :: 'a set) (R :: 'a \ 'b \ bool) \ 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 \ 'a \ bool) \ _ \ 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_eq_reflexive_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "reflexive_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool \ reflexive_on P" using assms 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 \ 'b \ bool) \ _ \ right_unique_on (mem_of S)" definition "right_unique_at_set (S :: 'a set) :: ('b \ 'a \ bool) \ _ \ 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 \ 'b \ bool) \ bool) = right_unique_on (mem_of S)" unfolding right_unique_on_set_def by simp lemma right_unique_on_set_eq_right_unique_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "right_unique_on (S :: 'a set) :: ('a \ 'b \ bool) \ bool \ right_unique_on P" using assms by simp lemma right_unique_on_set_iff_right_unique_on_pred [iff]: "right_unique_on (S :: 'a set) (R :: 'a \ 'b \ bool) \ 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 \ bool) \ _ \ 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) \ bool) = rel_surjective_at (mem_of S)" unfolding rel_surjective_at_set_def by simp lemma rel_surjective_at_set_eq_rel_surjective_at_pred_uhint [uhint]: assumes "P \ mem_of S" shows "rel_surjective_at (S :: 'a set) :: ('b \ 'a \ bool) \ bool \ rel_surjective_at P" using assms by simp lemma rel_surjective_at_set_iff_rel_surjective_at_pred [iff]: "rel_surjective_at (S :: 'a set) (R :: 'b \ 'a \ bool) \ 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 \ 'a \ bool) \ _ \ 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_eq_symmetric_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "symmetric_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool \ symmetric_on P" using assms 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 \ 'a \ bool) \ _ \ 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_eq_transitive_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "transitive_on (S :: 'a set) :: ('a \ 'a \ bool) \ bool \ transitive_on P" using assms 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) subparagraph \Bi-Total\ lemma bi_total_on_set_eq_bi_total_on_pred [simp]: "(bi_total_on (S :: 'a set) (T :: 'b set) :: ('a \ 'b \ bool) \ bool) = bi_total_on (mem_of S) (mem_of T)" by auto lemma bi_total_on_set_eq_bi_total_on_pred_uhint [uhint]: assumes "P \ mem_of S" and "Q \ mem_of T" shows "bi_total_on (S :: 'a set) (T :: 'b set) :: ('a \ 'b \ bool) \ bool \ bi_total_on P Q" using assms by simp lemma bi_total_on_set_iff_bi_total_on_pred [iff]: "bi_total_on (S :: 'a set) (T :: 'b set) (R :: 'a \ 'b \ bool) \ bi_total_on (mem_of S) (mem_of T) R" by simp lemma Transfer_bi_total_eq_bi_total [HOL_bin_rel_alignment]: "Transfer.bi_total = Binary_Relations_Bi_Total.bi_total" unfolding bi_total_alt_def by (auto simp add: HOL_bin_rel_alignment) subparagraph \Bi-Unique\ lemma bi_unique_on_set_eq_bi_unique_on_pred [simp]: "(bi_unique_on (S :: 'a set) :: ('a \ 'b \ bool) \ bool) = bi_unique_on (mem_of S)" by auto lemma bi_unique_on_set_eq_bi_unique_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "bi_unique_on (S :: 'a set) :: ('a \ 'b \ bool) \ bool \ bi_unique_on P" using assms by simp lemma bi_unique_on_set_iff_bi_unique_on_pred [iff]: "bi_unique_on (S :: 'a set) (R :: 'a \ 'b \ bool) \ bi_unique_on (mem_of S) R" by simp lemma Transfer_bi_unique_eq_bi_unique [HOL_bin_rel_alignment]: "Transfer.bi_unique = Binary_Relations_Bi_Unique.bi_unique" unfolding bi_unique_alt_def by (auto simp add: HOL_bin_rel_alignment) 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 +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 eq_onp_eq_eq_restrict [HOL_bin_rel_alignment]: "eq_onp = rel_restrict_left (=)" unfolding eq_onp_def by (intro ext) auto -overloading - rel_restrict_left_set \ "rel_restrict_left :: ('a \ 'b \ bool) \ 'a set \ 'a \ 'b \ bool" - rel_restrict_right_set \ "rel_restrict_right :: ('a \ 'b \ bool) \ 'b set \ 'a \ 'b \ bool" -begin - definition "rel_restrict_left_set (R :: 'a \ 'b \ bool) (S :: 'a set) \ R\\<^bsub>mem_of S\<^esub>" - definition "rel_restrict_right_set (R :: 'a \ 'b \ bool) (S :: 'b set) \ R\\<^bsub>mem_of S\<^esub>" -end +definition "rel_restrict_left_set (R :: 'a \ 'b \ bool) (S :: 'a set) \ R\\<^bsub>mem_of S\<^esub>" +adhoc_overloading rel_restrict_left rel_restrict_left_set + +definition "rel_restrict_right_set (R :: 'a \ 'b \ bool) (S :: 'b set) \ R\\<^bsub>mem_of S\<^esub>" +adhoc_overloading rel_restrict_right rel_restrict_right_set lemma rel_restrict_left_set_eq_restrict_left_pred [simp]: - "(R\\<^bsub>S :: 'a set\<^esub> :: 'a \ 'b \ bool) = R\\<^bsub>mem_of S\<^esub>" + "R\\<^bsub>S\<^esub> = R\\<^bsub>mem_of S\<^esub>" unfolding rel_restrict_left_set_def by simp lemma rel_restrict_left_set_eq_restrict_left_pred_uhint [uhint]: - assumes "P \ mem_of S" - shows "(R\\<^bsub>S :: 'a set\<^esub> :: 'a \ 'b \ bool) = R\\<^bsub>P\<^esub>" + assumes "R \ R'" + and "P \ mem_of S" + shows "R\\<^bsub>S\<^esub> \ R'\\<^bsub>P\<^esub>" using assms 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" +lemma restrict_left_set_iff_restrict_left_pred [iff]: "R\\<^bsub>S\<^esub> x y \ R\\<^bsub>mem_of S\<^esub> x y" by simp lemma rel_restrict_right_set_eq_restrict_right_pred [simp]: - "(R\\<^bsub>S :: 'b set\<^esub> :: 'a \ 'b \ bool) = R\\<^bsub>mem_of S\<^esub>" + "R\\<^bsub>S\<^esub> = R\\<^bsub>mem_of S\<^esub>" unfolding rel_restrict_right_set_def by simp lemma rel_restrict_right_set_eq_restrict_right_pred_uhint [uhint]: - assumes "P \ mem_of S" - shows "(R\\<^bsub>S :: 'b set\<^esub> :: 'a \ 'b \ bool) = R\\<^bsub>P\<^esub>" + assumes "R \ R'" + and "P \ mem_of S" + shows "R\\<^bsub>S\<^esub> \ R'\\<^bsub>P\<^esub>" using assms 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" +lemma restrict_right_set_iff_restrict_right_pred [iff]: "R\\<^bsub>S\<^esub> 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,180 +1,175 @@ \<^marker>\creator "Kevin Kappelmann"\ -subsection \Alignment With Definitions from HOL.Main\ +subsection \Alignment With Function 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 +definition "bijection_on_set (S :: 'a set) (S' :: 'b set) :: ('a \ 'b) \ ('b \ 'a) \ bool \ + bijection_on (mem_of S) (mem_of S')" +adhoc_overloading bijection_on bijection_on_set lemma bijection_on_set_eq_bijection_on_pred [simp]: - "(bijection_on (S :: 'a set) (S' :: 'b set) :: ('a \ 'b) \ ('b \ 'a) \ _) = - bijection_on (mem_of S) (mem_of S')" + "bijection_on (S :: 'a set) (S' :: 'b set) = bijection_on (mem_of S) (mem_of S')" unfolding bijection_on_set_def by simp lemma bijection_on_set_eq_bijection_on_pred_uhint [uhint]: assumes "P \ mem_of S" and "Q \ mem_of S'" - shows "bijection_on (S :: 'a set) (S' :: 'b set) :: ('a \ 'b) \ ('b \ 'a) \ _ \ bijection_on P Q" + shows "bijection_on S S' \ bijection_on P Q" using assms by simp lemma bijection_on_set_iff_bijection_on_pred [iff]: "bijection_on (S :: 'a set) (S' :: 'b set) (f :: 'a \ 'b) (g :: 'b \ 'a) \ bijection_on (mem_of S) (mem_of S') f g" by simp lemma bij_betw_bijection_onE: assumes "bij_betw (f :: 'a \ 'b) S S'" obtains g :: "'b \ 'a" 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 :: 'a \ 'b) (g :: 'b \ 'a)" 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 :: 'a \ 'b) S S' \ (\(g :: 'b \ 'a). 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_eq_injective_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "injective_on (S :: 'a set) :: ('a \ 'b) \ _ \ injective_on P" using assms 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) \ ('b \ 'a) \ bool \ inverse_on (mem_of S)" end lemma inverse_on_set_eq_inverse_on_pred [simp]: "(inverse_on (S :: 'a set) :: ('a \ 'b) \ ('b \ 'a) \ _) = inverse_on (mem_of S)" unfolding inverse_on_set_def by simp lemma inverse_on_set_eq_inverse_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "inverse_on (S :: 'a set) :: ('a \ 'b) \ ('b \ 'a) \ _ \ inverse_on P" using assms by simp lemma inverse_on_set_iff_inverse_on_pred [iff]: "inverse_on (S :: 'a set) (f :: 'a \ 'b) (g :: 'b \ 'a) \ 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>" 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_eq_surjective_at_pred_uhint [uhint]: assumes "P \ mem_of S" shows "surjective_at (S :: 'a set) :: ('b \ 'a) \ _ \ surjective_at P" using assms 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" +lemma rel_fun_eq_Fun_Rel_rel [HOL_fun_alignment]: "BNF_Def.rel_fun = Fun_Rel" by (intro ext) (auto dest: rel_funD) - end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Orders.thy b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Orders.thy --- a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Orders.thy +++ b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Orders.thy @@ -1,143 +1,143 @@ \<^marker>\creator "Kevin Kappelmann"\ -subsection \Alignment With Definitions from HOL\ +subsection \Alignment With Order Definitions from HOL\ theory HOL_Alignment_Orders imports "HOL-Library.Preorder" HOL_Alignment_Binary_Relations HOL_Syntax_Bundles_Orders Orders begin named_theorems HOL_order_alignment paragraph \Functions\ definition "rel R x y \ (x, y) \ R" lemma rel_of_eq [simp]: "rel = (\R x y. (x, y) \ R)" unfolding rel_def by simp lemma rel_of_iff [iff]: "rel R x y \ (x, y) \ R" by simp subparagraph \Bi-Related\ overloading bi_related_set \ "bi_related :: 'a rel \ 'a \ 'a \ bool" begin definition "bi_related_set (S :: 'a rel) \ bi_related (rel S) :: 'a \ 'a \ bool" end lemma bi_related_set_eq_bi_related_pred [simp]: - "(bi_related (S :: 'a rel) :: 'a \ 'a \ bool) = bi_related (rel S)" + "((\\<^bsub>S :: 'a rel\<^esub>) :: 'a \ 'a \ bool) = (\\<^bsub>rel S\<^esub>)" unfolding bi_related_set_def by simp lemma bi_related_set_eq_bi_related_pred_uhint [uhint]: assumes "R \ rel S" - shows "bi_related (S :: 'a rel) :: 'a \ 'a \ bool \ bi_related R" + shows "(\\<^bsub>S :: 'a rel\<^esub>) :: 'a \ 'a \ bool \ (\\<^bsub>R\<^esub>)" using assms by simp lemma bi_related_set_iff_bi_related_pred [iff]: "(x :: 'a) \\<^bsub>(S :: 'a rel)\<^esub> (y :: 'a) \ x \\<^bsub>rel S\<^esub> y" by simp lemma (in preorder_equiv) equiv_eq_bi_related [HOL_order_alignment]: "equiv = bi_related (\)" by (intro ext) (auto intro: equiv_antisym dest: equivD1 equivD2) subparagraph \Inflationary\ overloading inflationary_on_set \ "inflationary_on :: 'a set \ ('a \ 'b \ bool) \ ('a \ 'b) \ bool" begin definition "inflationary_on_set (S :: 'a set) :: ('a \ 'b \ bool) \ ('a \ 'b) \ bool \ inflationary_on (mem_of S)" end lemma inflationary_on_set_eq_inflationary_on_pred [simp]: "(inflationary_on (S :: 'a set) :: ('a \ 'b \ bool) \ ('a \ 'b) \ bool) = inflationary_on (mem_of S)" unfolding inflationary_on_set_def by simp lemma inflationary_on_set_eq_inflationary_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "inflationary_on (S :: 'a set) :: ('a \ 'b \ bool) \ ('a \ 'b) \ bool \ inflationary_on P" using assms by simp lemma inflationary_on_set_iff_inflationary_on_pred [iff]: "inflationary_on (S :: 'a set) (R :: 'a \ 'b \ bool) (f :: 'a \ 'b) \ inflationary_on (mem_of S) R f" by simp subparagraph \Deflationary\ overloading deflationary_on_set \ "deflationary_on :: 'a set \ ('b \ 'a \ bool) \ ('a \ 'b) \ bool" begin definition "deflationary_on_set (S :: 'a set) :: ('b \ 'a \ bool) \ ('a \ 'b) \ bool \ deflationary_on (mem_of S)" end lemma deflationary_on_set_eq_deflationary_on_pred [simp]: "(deflationary_on (S :: 'a set) :: ('b \ 'a \ bool) \ ('a \ 'b) \ bool) = deflationary_on (mem_of S)" unfolding deflationary_on_set_def by simp lemma deflationary_on_set_eq_deflationary_on_pred_uhint [uhint]: assumes "P \ mem_of S" shows "deflationary_on (S :: 'a set) :: ('b \ 'a \ bool) \ ('a \ 'b) \ bool \ deflationary_on P" using assms by simp lemma deflationary_on_set_iff_deflationary_on_pred [iff]: "deflationary_on (S :: 'a set) (R :: 'b \ 'a \ bool) (f :: 'a \ 'b) \ deflationary_on (mem_of S) R f" by simp subparagraph \Idempotent\ overloading idempotent_on_set \ "idempotent_on :: 'a set \ ('a \ 'a \ bool) \ ('a \ 'a) \ bool" begin definition "idempotent_on_set (S :: 'a set) :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool \ idempotent_on (mem_of S)" end lemma idempotent_on_set_eq_idempotent_on_pred [simp]: "(idempotent_on (S :: 'a set) :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool) = idempotent_on (mem_of S)" unfolding idempotent_on_set_def by simp lemma idempotent_on_set_iff_idempotent_on_pred [iff]: "idempotent_on (S :: 'a set) (R :: 'a \ 'a \ bool) (f :: 'a \ 'a) \ idempotent_on (mem_of S) R f" by simp paragraph \Properties\ subparagraph \Equivalence Relations\ lemma equiv_eq_equivalence_rel [HOL_order_alignment]: "equivp = equivalence_rel" by (intro ext) (fastforce intro!: equivpI simp: HOL_bin_rel_alignment reflexive_eq_reflexive_on elim!: equivpE) subparagraph \Partial Equivalence Relations\ lemma part_equiv_eq_partial_equivalence_rel_if_rel [HOL_order_alignment]: assumes "R x y" shows "part_equivp R = partial_equivalence_rel R" using assms by (fastforce intro!: part_equivpI simp: HOL_bin_rel_alignment elim!: part_equivpE) subparagraph \Partial Orders\ lemma (in order) partial_order [HOL_order_alignment]: "partial_order (\)" using order_refl order_trans order_antisym by blast subparagraph \Preorders\ lemma (in partial_preordering) preorder [HOL_order_alignment]: "preorder (\<^bold>\)" using refl trans by blast lemma partial_preordering_eq [HOL_order_alignment]: "partial_preordering = Preorders.preorder" by (intro ext) (auto intro: partial_preordering.intro dest: partial_preordering.trans partial_preordering.refl reflexiveD) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Predicates.thy b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Predicates.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Predicates.thy @@ -0,0 +1,39 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Alignment With Predicate Definitions from HOL\ +theory HOL_Alignment_Predicates + imports + Main + HOL_Mem_Of + Predicates +begin + +named_theorems HOL_predicate_alignment + +subparagraph \Quantifiers\ + +adhoc_overloading ball Ball + +lemma Ball_eq_ball_pred [HOL_predicate_alignment]: "\\<^bsub>A\<^esub> = \\<^bsub>mem_of A\<^esub>" by auto + +lemma Ball_eq_ball_pred_uhint [uhint]: + assumes "P \ mem_of A" + shows "\\<^bsub>A\<^esub> = \\<^bsub>P\<^esub>" + using assms by (simp add: Ball_eq_ball_pred) + +lemma Ball_iff_ball_pred [HOL_predicate_alignment]: "(\x : A. Q x) \ (\x : mem_of A. Q x)" + by (simp add: Ball_eq_ball_pred) + +adhoc_overloading bex Bex + +lemma Bex_eq_bex_pred [HOL_predicate_alignment]: "\\<^bsub>A\<^esub> = \\<^bsub>mem_of A\<^esub>" by fast + +lemma Bex_eq_bex_pred_uhint [uhint]: + assumes "P \ mem_of A" + shows "\\<^bsub>A\<^esub> = \\<^bsub>P\<^esub>" + using assms by (simp add: Bex_eq_bex_pred) + +lemma Bex_iff_bex_pred [HOL_predicate_alignment]: "(\x : A. Q x) \ (\x : mem_of A. Q x)" + by (simp add: Bex_eq_bex_pred) + + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignments.thy b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignments.thy --- a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignments.thy +++ b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignments.thy @@ -1,13 +1,14 @@ \<^marker>\creator "Kevin Kappelmann"\ section \HOL Alignments\ theory HOL_Alignments imports HOL_Alignment_Binary_Relations HOL_Alignment_Functions HOL_Alignment_Orders + HOL_Alignment_Predicates begin paragraph \Summary\ text \Alignment of concepts with HOL counterparts\ end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/HOL_Basics.thy b/thys/Transport/HOL_Basics/HOL_Basics.thy --- a/thys/Transport/HOL_Basics/HOL_Basics.thy +++ b/thys/Transport/HOL_Basics/HOL_Basics.thy @@ -1,30 +1,30 @@ \<^marker>\creator "Kevin Kappelmann"\ section \HOL-Basics\ theory HOL_Basics imports LBinary_Relations LFunctions Galois Orders Predicates begin paragraph \Summary\ text \Library on top of HOL axioms, as required for Transport \<^cite>\"transport"\. Requires \<^emph>\only\ the HOL axioms, nothing else. Includes: \<^enum> Basic concepts on binary relations, relativised properties, and restricted equalities e.g. @{term "left_total_on"} and @{term "eq_restrict"}. -\<^enum> Basic concepts on functions, relativised properties, and generalised relators, +\<^enum> Basic concepts on functions, relativised properties, and relators, e.g. @{term "injective_on"} and @{term "dep_mono_wrt_pred"}. \<^enum> Basic concepts on orders and relativised order-theoretic properties, e.g. @{term "partial_equivalence_rel_on"}. \<^enum> Galois connections, Galois equivalences, order equivalences, and other related concepts on order functors, e.g. @{term "galois.galois_equivalence"}. \<^enum> Basic concepts on predicates. \<^enum> Syntax bundles for HOL @{dir "HOL_Syntax_Bundles"}. \<^enum> Alignments for concepts that have counterparts in the HOL library - see @{dir "HOL_Alignments"}.\ end diff --git a/thys/Transport/HOL_Basics/HOL_Basics_Base.thy b/thys/Transport/HOL_Basics/HOL_Basics_Base.thy --- a/thys/Transport/HOL_Basics/HOL_Basics_Base.thy +++ b/thys/Transport/HOL_Basics/HOL_Basics_Base.thy @@ -1,8 +1,9 @@ \<^marker>\creator "Kevin Kappelmann"\ chapter \HOL-Basics\ theory HOL_Basics_Base imports HOL.HOL + Adhoc_Overloading begin end diff --git a/thys/Transport/HOL_Basics/HOL_Mem_Of.thy b/thys/Transport/HOL_Basics/HOL_Mem_Of.thy --- a/thys/Transport/HOL_Basics/HOL_Mem_Of.thy +++ b/thys/Transport/HOL_Basics/HOL_Mem_Of.thy @@ -1,18 +1,25 @@ \<^marker>\creator "Kevin Kappelmann"\ theory HOL_Mem_Of imports HOL.Set + ML_Unification.ML_Unification_HOL_Setup begin definition "mem_of A x \ x \ A" lemma mem_of_eq: "mem_of = (\A x. x \ A)" unfolding mem_of_def by simp lemma mem_of_iff [iff]: "mem_of A x \ x \ A" unfolding mem_of_def by simp +lemma mem_of_eq_mem_uhint [uhint]: + assumes "x \ x'" + and "A \ A'" + shows "mem_of A x \ x' \ A'" + using assms by simp + lemma mem_of_UNIV_eq_top [simp]: "mem_of UNIV = \" by auto lemma mem_of_empty_eq_bot [simp]: "mem_of {} = \" by auto end \ No newline at end of file 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,125 +1,125 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Closure Operators\ theory Closure_Operators imports Order_Functions_Base begin consts idempotent_on :: "'a \ 'b \ 'c \ bool" overloading idempotent_on_pred \ "idempotent_on :: ('a \ bool) \ ('a \ 'a \ bool) \ ('a \ 'a) \ bool" begin definition "idempotent_on_pred (P :: 'a \ bool) (R :: 'a \ 'a \ bool) (f :: 'a \ 'a) \ rel_equivalence_on P (rel_map f R) f" end context fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" and f :: "'a \ 'a" begin 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_pred_def using assms by fastforce lemma idempotent_onE [elim]: assumes "idempotent_on P R f" and "P x" obtains "f x \\<^bsub>R\<^esub> f (f x)" using assms unfolding idempotent_on_pred_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_pred_def by simp lemma idempotent_onD: assumes "idempotent_on P R f" and "P x" shows "f x \\<^bsub>R\<^esub> f (f x)" using assms by blast end consts idempotent :: "'a \ 'b \ bool" overloading idempotent \ "idempotent :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool" begin definition "(idempotent :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool) \ idempotent_on (\ :: 'a \ bool)" end lemma idempotent_eq_idempotent_on: "(idempotent :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool) = idempotent_on (\ :: 'a \ bool)" unfolding idempotent_def .. lemma idempotent_eq_idempotent_on_uhint [uhint]: assumes "P \ (\ :: 'a \ bool)" shows "idempotent :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool \ idempotent_on P" using assms by (simp add: idempotent_eq_idempotent_on) context fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" and f :: "'a \ 'a" begin lemma idempotentI [intro]: assumes "\x. f x \\<^bsub>R\<^esub> f (f x)" shows "idempotent R f" using assms by (urule idempotent_onI) lemma idempotentD [dest]: assumes "idempotent R f" shows "f x \\<^bsub>R\<^esub> f (f x)" using assms by (urule (e) idempotent_onE where chained = insert) simp lemma idempotent_on_if_idempotent: assumes "idempotent R f" shows "idempotent_on P R f" using assms by (intro idempotent_onI) auto end consts closure_operator :: "'a \ 'b \ bool" overloading closure_operator \ "closure_operator :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool" begin definition "closure_operator (R :: 'a \ 'a \ bool) :: ('a \ 'a) \ bool \ (R \\<^sub>m R) \ inflationary_on (in_field R) R \ idempotent_on (in_field R) R" end 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" + assumes "closure_operator (R :: 'a \ 'a \ bool) f" shows "(R \\<^sub>m R) f" using assms by (elim closure_operatorE) context fixes R :: "'a \ 'a \ bool" and f :: "'a \ 'a" begin 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 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,513 +1,513 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Functions On Orders\ subsubsection \Basics\ theory Order_Functions_Base imports Functions_Monotone Binary_Relations_Antisymmetric Binary_Relations_Symmetric Preorders begin subparagraph \Bi-Relation\ consts bi_related :: "'a \ 'b \ 'b \ bool" overloading bi_related \ "bi_related :: ('a \ 'a \ bool) \ 'a \ 'a \ bool" begin definition "bi_related (R :: 'a \ 'a \ bool) x y \ R x y \ R y x" end (*Note: we are not using (\\) as infix here because it would produce an ambiguous -grammar whenever using a of the form "definition c \ t"*) +grammar whenever using an expression of the form "definition c \ t"*) bundle bi_related_syntax begin syntax "_bi_related" :: "'a \ 'b \ '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 \ 'b \ '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 context fixes P :: "'a \ bool" and R S :: "'a \ 'a \ bool" and x y :: 'a begin lemma symmetric_bi_related [iff]: "symmetric ((\\<^bsub>R\<^esub>) :: 'a \ 'a \ bool)" by (intro symmetricI) blast lemma reflexive_bi_related_if_reflexive [intro]: assumes "reflexive R" shows "reflexive ((\\<^bsub>R\<^esub>) :: 'a \ 'a \ bool)" using assms by (intro reflexiveI) (blast dest: reflexiveD) lemma transitive_bi_related_if_transitive [intro]: assumes "transitive R" shows "transitive ((\\<^bsub>R\<^esub>) :: 'a \ 'a \ bool)" using assms by (intro transitiveI bi_relatedI) auto lemma mono_bi_related: "mono (bi_related :: ('a \ 'a \ bool) \ 'a \ 'a \ bool)" 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_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') end lemma bi_related_le_eq_if_antisymmetric_on_in_field: assumes "antisymmetric_on (in_field R) (R :: 'a \ 'a \ bool)" shows "((\\<^bsub>R\<^esub>) :: 'a \ 'a \ bool) \ (=)" using assms by (intro le_relI eq_if_bi_related_if_in_field_le_if_antisymmetric_on) blast+ subparagraph \Inflationary\ consts inflationary_on :: "'a \ 'b \ 'c \ bool" overloading inflationary_on_pred \ "inflationary_on :: ('a \ bool) \ ('a \ 'b \ bool) \ ('a \ 'b) \ bool" begin text \Often also called "extensive".\ - definition "inflationary_on_pred P (R :: 'a \ 'b \ bool) (f :: 'a \ 'b) \ \x. P x \ R x (f x)" + definition "inflationary_on_pred P (R :: 'a \ 'b \ bool) (f :: 'a \ 'b) \ \x : P. 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 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 mono_inflationary_on_rel: "((\) \\<^sub>m (\) \ (\)) (inflationary_on :: ('a \ bool) \ ('a \ 'b \ bool) \ ('a \ 'b) \ bool)" - by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI) auto + by (intro mono_wrt_relI Fun_Rel_relI) auto context fixes P P' :: "'a \ bool" and R :: "'a \ 'b \ bool" and f :: "'a \ 'b" begin (*FIXME: should be automatically derivable from above monotonicity lemma*) lemma inflationary_on_if_le_pred_if_inflationary_on: assumes "inflationary_on P R f" and "P' \ P" 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 end lemma inflationary_on_sup_eq [simp]: "(inflationary_on :: ('a \ bool) \ ('a \ 'b \ bool) \ ('a \ 'b) \ bool) (P \ Q) = inflationary_on P \ inflationary_on Q" by (intro ext iffI inflationary_onI) (auto intro: inflationary_on_if_le_pred_if_inflationary_on) consts inflationary :: "'a \ 'b \ bool" overloading inflationary \ "inflationary :: ('a \ 'b \ bool) \ ('a \ 'b) \ bool" begin definition "(inflationary :: ('a \ 'b \ bool) \ ('a \ 'b) \ bool) \ inflationary_on (\ :: 'a \ bool)" end lemma inflationary_eq_inflationary_on: "(inflationary :: ('a \ 'b \ bool) \ ('a \ 'b) \ bool) = inflationary_on (\ :: 'a \ bool)" unfolding inflationary_def .. lemma inflationary_eq_inflationary_on_uhint [uhint]: assumes "P \ \ :: 'a \ bool" shows "inflationary :: ('a \ 'b \ bool) \ ('a \ 'b) \ bool \ inflationary_on P" using assms by (simp add: inflationary_eq_inflationary_on) lemma inflationaryI [intro]: assumes "\x. R x (f x)" shows "inflationary R f" using assms by (urule inflationary_onI) lemma inflationaryD: assumes "inflationary R f" shows "R x (f x)" using assms by (urule (d) inflationary_onD where chained = insert) simp lemma inflationary_on_if_inflationary: fixes P :: "'a \ bool" and R :: "'a \ 'b \ bool" and f :: "'a \ 'b" 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\ consts deflationary_on :: "'a \ 'b \ 'c \ bool" overloading deflationary_on_pred \ "deflationary_on :: ('a \ bool) \ ('b \ 'a \ bool) \ ('a \ 'b) \ bool" begin definition "deflationary_on_pred (P :: 'a \ bool) (R :: 'b \ 'a \ bool) :: ('a \ 'b) \ bool \ inflationary_on P R\" end context fixes P :: "'a \ bool" and R :: "'b \ 'a \ bool" and f :: "'a \ 'b" begin lemma deflationary_on_eq_inflationary_on_rel_inv: "(deflationary_on P R :: ('a \ 'b) \ bool) = inflationary_on P R\" unfolding deflationary_on_pred_def .. declare deflationary_on_eq_inflationary_on_rel_inv[symmetric, 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 end context fixes P P' :: "'a \ bool" and R :: "'b \ 'a \ bool" and f :: "'a \ 'b" begin corollary deflationary_on_rel_inv_eq_inflationary_on [simp]: "(deflationary_on P (S :: 'a \ 'b \ bool)\ :: ('a \ 'b) \ bool) = inflationary_on P S" unfolding deflationary_on_eq_inflationary_on_rel_inv by simp lemma deflationary_on_eq_dep_mono_wrt_pred_rel_inv: - "(deflationary_on P R :: ('a \ 'b) \ bool) = ([x \ P] \\<^sub>m R\ x)" + "(deflationary_on P R :: ('a \ 'b) \ bool) = ((x : P) \\<^sub>m R\ x)" by blast 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 mono_deflationary_on: "((\) \\<^sub>m (\) \ (\)) (deflationary_on :: ('a \ bool) \ ('b \ 'a \ bool) \ ('a \ 'b) \ bool)" by blast (*FIXME: should be automatically derivable from above monotonicity lemma*) lemma deflationary_on_if_le_pred_if_deflationary_on: assumes "deflationary_on P R f" and "P' \ P" shows "deflationary_on P' R f" using assms by blast 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) \ ('b \ 'a \ bool) \ ('a \ 'b) \ bool) (P \ Q) = deflationary_on P \ deflationary_on Q" unfolding deflationary_on_eq_inflationary_on_rel_inv by auto end consts deflationary :: "'a \ 'b \ bool" overloading deflationary \ "deflationary :: ('b \ 'a \ bool) \ ('a \ 'b) \ bool" begin definition "(deflationary :: ('b \ 'a \ bool) \ ('a \ 'b) \ bool) \ deflationary_on (\ :: 'a \ bool)" end lemma deflationary_eq_deflationary_on: "(deflationary :: ('b \ 'a \ bool) \ ('a \ 'b) \ bool) = deflationary_on (\ :: 'a \ bool)" unfolding deflationary_def .. lemma deflationary_eq_deflationary_on_uhint [uhint]: assumes "P \ \ :: 'a \ bool" shows "deflationary :: ('b \ 'a \ bool) \ ('a \ 'b) \ bool \ deflationary_on P" using assms by (simp add: deflationary_eq_deflationary_on) lemma deflationaryI [intro]: assumes "\x. R (f x) x" shows "deflationary R f" using assms by (urule deflationary_onI) lemma deflationaryD: assumes "deflationary R f" shows "R (f x) x" using assms by (urule (d) deflationary_onD where chained = insert) simp lemma deflationary_on_if_deflationary: fixes P :: "'a \ bool" and R :: "'b \ 'a \ bool" and f :: "'a \ 'b" 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 context fixes P P' :: "'a \ bool" and R :: "'a \ 'a \ bool" and f :: "'a \ 'a" begin 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 "x \\<^bsub>R\<^esub> f 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\ :: ('a \ 'a) \ bool) = rel_equivalence_on P R" by (intro ext) fastforce lemma mono_rel_equivalence_on: "((\) \\<^sub>m (\) \ (\)) (rel_equivalence_on :: ('a \ bool) \ ('a \ 'a \ bool) \ ('a \ 'a) \ bool)" by blast (*FIXME: should be automatically derivable from above monotonicity lemma*) lemma rel_equivalence_on_if_le_pred_if_rel_equivalence_on: 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 \ 'a \ bool) \ ('a \ 'a) \ bool) (P \ 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 "(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 :: ('a \ 'a) \ bool) = 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 :: ('a \ 'a) \ bool) = rel_equivalence_on P R" using assms by (simp add: deflationary_on_eq_inflationary_on_rel_inv rel_equivalence_on_eq) end consts rel_equivalence :: "'a \ 'b \ bool" overloading rel_equivalence \ "rel_equivalence :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool" begin definition "(rel_equivalence :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool) \ rel_equivalence_on (\ :: 'a \ bool)" end lemma rel_equivalence_eq_rel_equivalence_on: "(rel_equivalence :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool) = rel_equivalence_on (\ :: 'a \ bool)" unfolding rel_equivalence_def .. lemma rel_equivalence_eq_rel_equivalence_on_uhint [uhint]: assumes "P \ \ :: 'a \ bool" shows "rel_equivalence :: ('a \ 'a \ bool) \ ('a \ 'a) \ bool \ rel_equivalence_on P" using assms by (simp add: rel_equivalence_eq_rel_equivalence_on) context fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" and f :: "'a \ 'a" begin lemma rel_equivalenceI [intro]: assumes "inflationary R f" and "deflationary R f" shows "rel_equivalence R f" using assms by (urule rel_equivalence_onI) lemma rel_equivalenceE [elim]: assumes "rel_equivalence R f" obtains "inflationary R f" "deflationary R f" using assms by (urule (e) rel_equivalence_onE) lemma inflationary_if_rel_equivalence: assumes "rel_equivalence R f" shows "inflationary R f" using assms by (rule rel_equivalenceE) lemma deflationary_if_rel_equivalence: assumes "rel_equivalence R f" shows "deflationary R f" using assms by (rule rel_equivalenceE) lemma rel_equivalence_on_if_rel_equivalence: 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 by blast end end diff --git a/thys/Transport/HOL_Basics/Orders/Functors/Order_Equivalences.thy b/thys/Transport/HOL_Basics/Orders/Functors/Order_Equivalences.thy --- a/thys/Transport/HOL_Basics/Orders/Functors/Order_Equivalences.thy +++ b/thys/Transport/HOL_Basics/Orders/Functors/Order_Equivalences.thy @@ -1,96 +1,95 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Equivalences\ theory Order_Equivalences imports Order_Functors_Base Partial_Equivalence_Relations Preorders begin context order_functors begin definition "order_equivalence \ ((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l \ ((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r \ rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \ \ rel_equivalence_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>) \" notation order_functors.order_equivalence (infix "\\<^sub>o" 50) lemma order_equivalenceI [intro]: 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_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" and "rel_equivalence_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>) \" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" unfolding order_equivalence_def using assms by blast lemma order_equivalenceE [elim]: assumes "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" obtains "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" "rel_equivalence_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>) \" using assms unfolding order_equivalence_def by blast interpretation of : order_functors S T f g for S T f g . lemma rel_inv_order_equivalence_eq_order_equivalence [simp]: "((\\<^bsub>R\<^esub>) \\<^sub>o (\\<^bsub>L\<^esub>))\ = ((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>))" - by (intro ext) - (auto intro!: of.order_equivalenceI simp: of.flip_unit_eq_counit) + by (intro ext) (auto intro!: of.order_equivalenceI simp: of.flip_unit_eq_counit) corollary order_equivalence_right_left_iff_order_equivalence_left_right: "((\\<^bsub>R\<^esub>) \\<^sub>o (\\<^bsub>L\<^esub>)) r l \ ((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" by (simp flip: rel_inv_order_equivalence_eq_order_equivalence) text \Due to the symmetry given by @{thm "order_equivalence_right_left_iff_order_equivalence_left_right"}, for any theorem on @{term "(\\<^bsub>L\<^esub>)"}, we obtain a corresponding theorem on @{term "(\\<^bsub>R\<^esub>)"} by flipping the roles of the two functors. As such, in what follows, we do not explicitly state these free theorems but users can obtain them as needed by creating a flipped interpretation of @{locale order_functors}.\ lemma order_equivalence_rel_inv_eq_order_equivalence [simp]: "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) = ((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>))" - by (intro ext) (auto intro!: of.order_equivalenceI) + by (intro ext) (auto 0 4 intro!: of.order_equivalenceI) lemma in_codom_left_eq_in_dom_left_if_order_equivalence: assumes "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" shows "in_codom (\\<^bsub>L\<^esub>) = in_dom (\\<^bsub>L\<^esub>)" using assms by (elim order_equivalenceE) (rule in_codom_eq_in_dom_if_rel_equivalence_on_in_field) corollary preorder_on_in_field_left_if_transitive_if_order_equivalence: assumes "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" and "transitive (\\<^bsub>L\<^esub>)" shows "preorder_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" using assms by (elim order_equivalenceE) (rule preorder_on_in_field_if_transitive_if_rel_equivalence_on) lemma order_equivalence_partial_equivalence_rel_not_reflexive_not_transitive: assumes "\(y :: 'b) y'. y \ y'" shows "\(L :: 'a \ 'a \ bool) (R :: 'b \ 'b \ bool) l r. (L \\<^sub>o R) l r \ partial_equivalence_rel L \ \(reflexive_on (in_field R) R) \ \(transitive_on (in_field R) R)" proof - from assms obtain cy cy' where "(cy :: 'b) \ cy'" by blast let ?cx = "undefined :: 'a" let ?L = "\x x'. ?cx = x \ x = x'" and ?R = "\y y'. (y = cy \ y = cy') \ (y' = cy \ y' = cy') \ (y \ cy' \ y' \ cy')" and ?l = "\(a :: 'a). cy" and ?r = "\(b :: 'b). ?cx" have "(?L \\<^sub>o ?R)?l ?r" using \cy \ cy'\ by (intro of.order_equivalenceI) (auto 0 4) moreover have "partial_equivalence_rel ?L" by blast moreover have "\(transitive_on (in_field ?R) ?R)" and "\(reflexive_on (in_field ?R) ?R)" using \cy \ cy'\ by auto ultimately show "?thesis" by blast qed end end \ No newline at end of file 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" + 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 + using assms by (intro 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) 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) 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) (fastforce 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/Linear_Orders.thy b/thys/Transport/HOL_Basics/Orders/Linear_Orders.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Orders/Linear_Orders.thy @@ -0,0 +1,57 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Linear Orders\ +theory Linear_Orders + imports + Binary_Relations_Connected + Partial_Orders +begin + +definition "linear_order_on \ partial_order_on \ connected_on" + +lemma linear_order_onI [intro]: + assumes "partial_order_on P R" + and "connected_on P R" + shows "linear_order_on P R" + using assms unfolding linear_order_on_def by auto + +lemma linear_order_onE [elim]: + assumes "linear_order_on P R" + obtains "partial_order_on P R" "connected_on P R" + using assms unfolding linear_order_on_def by auto + +definition "(linear_order :: ('a \ 'a \ bool) \ bool) \ linear_order_on (\ :: 'a \ bool)" + +lemma linear_order_eq_linear_order_on: + "(linear_order :: ('a \ 'a \ bool) \ bool) = linear_order_on (\ :: 'a \ bool)" + unfolding linear_order_def .. + +lemma linear_order_eq_linear_order_on_uhint [uhint]: + assumes "P \ \ :: 'a \ bool" + shows "(linear_order :: ('a \ 'a \ bool) \ bool) \ linear_order_on P" + using assms by (simp add: linear_order_eq_linear_order_on) + +context + fixes R :: "'a \ 'a \ bool" +begin + +lemma linear_orderI [intro]: + assumes "partial_order R" + and "connected R" + shows "linear_order R" + using assms by (urule linear_order_onI) + +lemma linear_orderE [elim]: + assumes "linear_order R" + obtains "partial_order R" "connected R" + using assms by (urule (e) linear_order_onE) + +lemma linear_order_on_if_linear_order: + fixes P :: "'a \ bool" + assumes "linear_order R" + shows "linear_order_on P R" + using assms by (elim linear_orderE) + (intro linear_order_onI partial_order_on_if_partial_order connected_on_if_connected) + +end + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Orders/Orders.thy b/thys/Transport/HOL_Basics/Orders/Orders.thy --- a/thys/Transport/HOL_Basics/Orders/Orders.thy +++ b/thys/Transport/HOL_Basics/Orders/Orders.thy @@ -1,17 +1,18 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Orders\ theory Orders imports Equivalence_Relations + Linear_Orders Order_Functions Order_Functors Partial_Equivalence_Relations Partial_Orders Preorders begin paragraph \Summary\ text \Basic order-theoretic concepts.\ end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Predicates/Bounded_Definite_Description.thy b/thys/Transport/HOL_Basics/Predicates/Bounded_Definite_Description.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Predicates/Bounded_Definite_Description.thy @@ -0,0 +1,48 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Bounded Definite Description\ +theory Bounded_Definite_Description + imports + Bounded_Quantifiers +begin + +consts bthe :: "'a \ ('b \ bool) \ 'b" + +bundle bounded_the_syntax +begin +syntax "_bthe" :: "[idt, 'a, bool] \ 'b" ("(3THE _ : _./ _)" [0, 0, 10] 10) +end +bundle no_bounded_the_syntax +begin +no_syntax "_bthe" :: "[idt, 'a, bool] \ 'b" ("(3THE _ : _./ _)" [0, 0, 10] 10) +end +unbundle bounded_the_syntax +translations "THE x : P. Q" \ "CONST bthe P (\x. Q)" + +definition "bthe_pred P Q \ The (P \ Q)" +adhoc_overloading bthe bthe_pred + +lemma bthe_eqI [intro]: + assumes "Q a" + and "P a" + and "\x. \P x; Q x\ \ x = a" + shows "(THE x : P. Q x) = a" + unfolding bthe_pred_def by (auto intro: assms) + +lemma pred_bthe_if_ex1E: + assumes "\!x : P. Q x" + obtains "P (THE x : P. Q x)" "Q (THE x : P. Q x)" + unfolding bthe_pred_def inf_fun_def using theI'[OF assms[unfolded bex1_pred_def]] + by auto + +lemma pred_btheI: + assumes "\!x : P. Q x" + shows "P (THE x : P. Q x)" + using assms by (elim pred_bthe_if_ex1E) + +lemma pred_btheI': + assumes "\!x : P. Q x" + shows "Q (THE x : P. Q x)" + using assms by (elim pred_bthe_if_ex1E) + + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Predicates/Bounded_Quantifiers.thy b/thys/Transport/HOL_Basics/Predicates/Bounded_Quantifiers.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Predicates/Bounded_Quantifiers.thy @@ -0,0 +1,262 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Bounded Quantifiers\ +theory Bounded_Quantifiers + imports + HOL_Basics_Base + Predicates_Lattice + ML_Unification.ML_Unification_HOL_Setup +begin + +consts ball :: "'a \ ('b \ bool) \ bool" + +bundle ball_syntax +begin +syntax + "_ball" :: \[idts, 'a, bool] \ bool\ ("(2\_ : _./ _)" 10) + "_ball2" :: \[idts, 'a, bool] \ bool\ +notation ball ("\(\<^bsub>_\<^esub>)") +end +bundle no_ball_syntax +begin +no_syntax + "_ball" :: \[idts, 'a, bool] \ bool\ ("(2\_ : _./ _)" 10) + "_ball2" :: \[idts, 'a, bool] \ bool\ +no_notation ball ("\(\<^bsub>_\<^esub>)") +end +unbundle ball_syntax +translations + "\x xs : P. Q" \ "CONST ball P (\x. _ball2 xs P Q)" + "_ball2 x P Q" \ "\x : P. Q" + "\x : P. Q" \ "CONST ball P (\x. Q)" + +consts bex :: "'a \ ('b \ bool) \ bool" + +bundle bex_syntax +begin +syntax + "_bex" :: \[idts, 'a, bool] \ bool\ ("(2\_ : _./ _)" 10) + "_bex2" :: \[idts, 'a, bool] \ bool\ +notation bex ("\(\<^bsub>_\<^esub>)") +end +bundle no_bex_syntax +begin +no_syntax + "_bex" :: \[idts, 'a, bool] \ bool\ ("(2\_ : _./ _)" 10) + "_bex2" :: \[idts, 'a, bool] \ bool\ +no_notation bex ("\(\<^bsub>_\<^esub>)") +end +unbundle bex_syntax +translations + "\x xs : P. Q" \ "CONST bex P (\x. _bex2 xs P Q)" + "_bex2 x P Q" \ "\x : P. Q" + "\x : P. Q" \ "CONST bex P (\x. Q)" + +consts bex1 :: "'a \ ('b \ bool) \ bool" + +bundle bex1_syntax +begin +syntax + "_bex1" :: \[idts, 'a, bool] \ bool\ ("(2\!_ : _./ _)" 10) + "_bex12" :: \[idts, 'a, bool] \ bool\ +notation bex1 ("\!(\<^bsub>_\<^esub>)") +end +bundle no_bex1_syntax +begin +no_syntax + "_bex1" :: \[idts, 'a, bool] \ bool\ ("(2\!_ : _./ _)" 10) + "_bex12" :: \[idts, 'a, bool] \ bool\ +no_notation bex1 ("\!(\<^bsub>_\<^esub>)") +end +unbundle bex1_syntax +translations + "\!x xs : P. Q" \ "CONST bex1 P (\x. _bex12 xs P Q)" + "_bex12 x P Q" \ "\!x : P. Q" + "\!x : P. Q" \ "CONST bex1 P (\x. Q)" + +bundle bounded_quantifier_syntax begin unbundle ball_syntax bex_syntax bex1_syntax end +bundle no_bounded_quantifier_syntax begin unbundle no_ball_syntax no_bex_syntax no_bex1_syntax end + +definition "ball_pred P Q \ \x. P x \ Q x" +adhoc_overloading ball ball_pred + +definition "bex_pred P Q \ \x. P x \ Q x" +adhoc_overloading bex bex_pred + +definition "bex1_pred P Q \ \!x. P x \ Q x" +adhoc_overloading bex1 bex1_pred + +(*copied from HOL.Set.thy*) +simproc_setup defined_ball ("\x : P. Q x \ U x") = \ + K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms ball_pred_def})) +\ +simproc_setup defined_bex ("\x : P. Q x \ U x") = \ + K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms bex_pred_def})) +\ + +lemma ballI [intro!]: + assumes "\x. P x \ Q x" + shows "\x : P. Q x" + using assms unfolding ball_pred_def by blast + +lemma ballE [elim]: + assumes "\x : P. Q x" + obtains "\x. P x \ Q x" + using assms unfolding ball_pred_def by blast + +lemma ballE': + assumes "\x : P. Q x" + obtains "\(P x)" | "P x" "Q x" + using assms by blast + +lemma ballD: "\x : P. Q x \ P x \ Q x" + by blast + +lemma ball_cong: "\P = P'; \x. P' x \ Q x \ Q' x\ \ (\x : P. Q x) \ (\x : P'. Q' x)" + by auto + +lemma ball_cong_simp [cong]: + "\P = P'; \x. P' x =simp=> Q x \ Q' x\ \ (\x : P. Q x) \ (\x : P'. Q' x)" + unfolding simp_implies_def by (rule ball_cong) + +(*copied from HOL.Set.thy*) +ML \ +structure Simpdata = +struct + open Simpdata + val mksimps_pairs = [(\<^const_name>\ball_pred\, @{thms ballD})] @ mksimps_pairs +end +open Simpdata +\ +declaration \fn _ => Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))\ + +lemma atomize_ball: "(\x. P x \ Q x) \ Trueprop (\x : P. Q x)" + by (simp only: ball_pred_def atomize_all atomize_imp) + +declare atomize_ball[symmetric, rulify] +declare atomize_ball[symmetric, defn] + +lemma bexI [intro!]: + (*better argument order: Q often determines the choice for x*) + assumes "\x. Q x \ P x" + shows "\x : P. Q x" + using assms unfolding bex_pred_def by blast + +lemma bexE [elim!]: + assumes "\x : P. Q x" + obtains x where "P x" "Q x" + using assms unfolding bex_pred_def by blast + +lemma bexD: + assumes "\x : P. Q x" + shows "\x. P x \ Q x" + using assms by blast + +lemma bex_cong: + "\P = P'; \x. P' x \ Q x \ Q' x\ \ (\x : P. Q x) \ (\x : P'. Q' x)" + by blast + +lemma bex_cong_simp [cong]: + "\P = P'; \x. P' x =simp=> Q x \ Q' x\ \ (\x : P. Q x) \ (\x : P'. Q' x)" + unfolding simp_implies_def by (rule bex_cong) + +lemma bex1I [intro]: + (*better argument order: Q often determines the choice for x*) + assumes "\!x. Q x \ P x" + shows "\!x : P. Q x" + using assms unfolding bex1_pred_def by blast + +lemma bex1D [dest!]: + assumes "\!x : P. Q x" + shows "\!x. P x \ Q x" + using assms unfolding bex1_pred_def by blast + +lemma bex1_cong: "\P = P'; \x. P x \ Q x \ Q' x\ \ (\!x : P. Q x) \ (\!x : P'. Q' x)" + by blast + +lemma bex1_cong_simp [cong]: + "\P = P'; \x. P x =simp=> Q x \ Q' x\ \ (\!x : P. Q x) \ (\!x : P'. Q' x)" + unfolding simp_implies_def by (rule bex1_cong) + +lemma ball_iff_ex_pred [iff]: "(\x : P. Q) \ ((\x. P x) \ Q)" + by auto + +lemma bex_iff_ex_and [iff]: "(\x : P. Q) \ ((\x. P x) \ Q)" + by blast + +lemma ball_eq_imp_iff_imp [iff]: "(\x : P. x = y \ Q x) \ (P y \ Q y)" + by blast + +lemma ball_eq_imp_iff_imp' [iff]: "(\x : P. y = x \ Q x) \ (P y \ Q y)" + by blast + +lemma bex_eq_iff_pred [iff]: "(\x : P. x = y) \ P y" + by blast + +lemma bex_eq_iff_pred' [iff]: "(\x : P. y = x) \ P y" + by blast + +lemma bex_eq_and_iff_pred [iff]: "(\x : P. x = y \ Q x) \ P y \ Q y" + by blast + +lemma bex_eq_and_iff_pred' [iff]: "(\x : P. y = x \ Q x) \ P y \ Q y" + by blast + +lemma ball_and_iff_ball_and_ball: "(\x : P. Q x \ U x) \ (\x : P. Q x) \ (\x : P. U x)" + by auto + +lemma bex_or_iff_bex_or_bex: "(\x : P. Q x \ U x) \ (\x : P. Q x) \ (\x : P. U x)" + by auto + +lemma ball_or_iff_ball_or [iff]: "(\x : P. Q x \ U) \ ((\x : P. Q x) \ U)" + by auto + +lemma ball_or_iff_or_ball [iff]: "(\x : P. Q \ U x) \ (Q \ (\x : P. U x))" + by auto + +lemma ball_imp_iff_imp_ball [iff]: "(\x : P. Q \ U x) \ (Q \ (\x : P. U x))" + by auto + +lemma bex_and_iff_bex_and [iff]: "(\x : P. Q x \ U) \ ((\x : P. Q x) \ U)" + by auto + +lemma bex_and_iff_and_bex [iff]: "(\x : P. Q \ U x) \ (Q \ (\x : P. U x))" + by auto + +lemma ball_imp_iff_bex_imp [iff]: "(\x : P. Q x \ U) \ ((\x : P. Q x) \ U)" + by auto + +lemma not_ball_iff_bex_not [iff]: "(\(\x : P. Q x)) \ (\x : P. \(Q x))" + by auto + +lemma not_bex_iff_ball_not [iff]: "(\(\x : P. Q x)) \ (\x : P. \(Q x))" + by auto + +lemma bex1_iff_bex_and [iff]: "(\!x : P. Q) \ ((\!x. P x) \ Q)" + by auto + +lemma ball_bottom_eq_top [simp]: "\\<^bsub>\\<^esub> = \" by auto +lemma bex_bottom_eq_bottom [simp]: "\\<^bsub>\\<^esub> = \" by fastforce +lemma bex1_bottom_eq_bottom [simp]: "\!\<^bsub>\\<^esub> = \" by fastforce + +lemma ball_top_eq_all [simp]: "\\<^bsub>\\<^esub> = All" by fastforce + +lemma ball_top_eq_all_uhint [uhint]: + assumes "P \ (\ :: 'a \ bool)" + shows "\\<^bsub>P\<^esub> \ All" + using assms by simp + +lemma bex_top_eq_ex [simp]: "\\<^bsub>\\<^esub> = Ex" by fastforce + +lemma bex_top_eq_ex_uhint [uhint]: + assumes "P \ (\ :: 'a \ bool)" + shows "\\<^bsub>P\<^esub> \ Ex" + using assms by simp + +lemma bex1_top_eq_ex1 [simp]: "\!\<^bsub>\\<^esub> = Ex1" by fastforce + +lemma bex1_top_eq_ex1_uhint [uhint]: + assumes "P \ (\ :: 'a \ bool)" + shows "\!\<^bsub>P\<^esub> \ Ex1" + using assms by simp + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Predicates/Predicate_Functions.thy b/thys/Transport/HOL_Basics/Predicates/Predicate_Functions.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Predicates/Predicate_Functions.thy @@ -0,0 +1,52 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Functions on Predicates\ +theory Predicate_Functions + imports + Functions_Base +begin + +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 + +consts pred_if :: "bool \ 'a \ 'a" + +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 + +definition "pred_if_pred B P x \ B \ P x" +adhoc_overloading pred_if pred_if_pred + +lemma pred_if_eq_pred_if_pred [simp]: + assumes "B" + shows "(pred_if B P) = P" + unfolding pred_if_pred_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_pred_def using assms by blast + +lemma pred_if_if_impI [intro]: + assumes "B \ P x" + shows "(pred_if B P) x" + unfolding pred_if_pred_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_pred_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/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,55 +1,16 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Predicates\ theory Predicates imports - Functions_Base + Bounded_Definite_Description + Bounded_Quantifiers + Predicate_Functions + Predicates_Lattice 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/HOL_Basics/Predicates/Predicates_Order.thy b/thys/Transport/HOL_Basics/Predicates/Predicates_Order.thy --- a/thys/Transport/HOL_Basics/Predicates/Predicates_Order.thy +++ b/thys/Transport/HOL_Basics/Predicates/Predicates_Order.thy @@ -1,26 +1,27 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Orders\ theory Predicates_Order imports HOL.Orderings begin lemma le_predI [intro]: assumes "\x. P x \ Q x" shows "P \ Q" using assms by (rule predicate1I) lemma le_predD [dest]: assumes "P \ Q" and "P x" shows "Q x" using assms by (rule predicate1D) lemma le_predE: assumes "P \ Q" and "P x" obtains "Q x" using assms by blast +declare le_boolD[dest] and le_boolI[intro] 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,51 +1,53 @@ chapter AFP 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/Adhoc_Overloading" "HOL_Basics/Binary_Relations" + "HOL_Basics/Binary_Relations/Functions" "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_Connection.thy b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Connection.thy --- a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Connection.thy +++ b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Connection.thy @@ -1,54 +1,53 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Connection\ theory Transport_Compositions_Agree_Galois_Connection imports Transport_Compositions_Agree_Monotone Transport_Compositions_Agree_Galois_Property begin context transport_comp_agree begin interpretation flip : transport_comp_agree R2 L2 r2 l2 R1 L1 r1 l1 . lemma galois_connectionI: assumes galois: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" "((\\<^bsub>L2\<^esub>) \ (\\<^bsub>R2\<^esub>)) l2 r2" and mono_L1_L2_l1: "\x y. x \\<^bsub>L1\<^esub> y \ l1 x \\<^bsub>R1\<^esub> l1 y \ l1 x \\<^bsub>L2\<^esub> l1 y" and mono_R2_R1_r2: "\x y. x \\<^bsub>R2\<^esub> y \ r2 x \\<^bsub>L2\<^esub> r2 y \ r2 x \\<^bsub>R1\<^esub> r2 y" - and "([in_dom (\\<^bsub>L1\<^esub>)] \ [in_codom (\\<^bsub>R2\<^esub>)] \ (\)) - (rel_bimap l1 r2 (\\<^bsub>R1\<^esub>)) (rel_bimap l1 r2 (\\<^bsub>L2\<^esub>))" + and "(in_dom (\\<^bsub>L1\<^esub>) \ in_codom (\\<^bsub>R2\<^esub>) \ (\)) (rel_bimap l1 r2 (\\<^bsub>R1\<^esub>)) (rel_bimap l1 r2 (\\<^bsub>L2\<^esub>))" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" proof - - from galois mono_L1_L2_l1 have "([in_dom (\\<^bsub>L1\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" - by (intro dep_mono_wrt_predI) (blast elim!: in_domE g1.galois_connectionE) + from galois mono_L1_L2_l1 have "(in_dom (\\<^bsub>L1\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" + by (intro mono_wrt_predI) (blast elim!: in_domE g1.galois_connectionE) moreover from galois mono_R2_R1_r2 - have "([in_codom (\\<^bsub>R2\<^esub>)] \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" - by (intro dep_mono_wrt_predI) (blast elim!: in_codomE g2.galois_connectionE) + have "(in_codom (\\<^bsub>R2\<^esub>) \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" + by (intro mono_wrt_predI) (blast elim!: in_codomE g2.galois_connectionE) ultimately show ?thesis using assms by (intro galois_connectionI galois_propI mono_wrt_rel_leftI flip.mono_wrt_rel_leftI) auto qed lemma galois_connectionI': assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" "((\\<^bsub>L2\<^esub>) \ (\\<^bsub>R2\<^esub>)) l2 r2" and "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) l1" "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) r2" - and "([in_dom (\\<^bsub>L1\<^esub>)] \ [in_codom (\\<^bsub>R2\<^esub>)] \ (\)) + and "(in_dom (\\<^bsub>L1\<^esub>) \ in_codom (\\<^bsub>R2\<^esub>) \ (\)) (rel_bimap l1 r2 (\\<^bsub>R1\<^esub>)) (rel_bimap l1 r2 (\\<^bsub>L2\<^esub>))" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_connectionI) auto end context transport_comp_same begin corollary galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" "((\\<^bsub>R1\<^esub>) \ (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (rule galois_connectionI) auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Equivalence.thy b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Equivalence.thy --- a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Equivalence.thy +++ b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Equivalence.thy @@ -1,56 +1,56 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Equivalence\ theory Transport_Compositions_Agree_Galois_Equivalence imports Transport_Compositions_Agree_Galois_Connection begin context transport_comp_agree begin interpretation flip : transport_comp_agree R2 L2 r2 l2 R1 L1 r1 l1 . lemma galois_equivalenceI: assumes galois: "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" "((\\<^bsub>L2\<^esub>) \\<^sub>G (\\<^bsub>R2\<^esub>)) l2 r2" and mono_L1_L2_l1: "\x y. x \\<^bsub>L1\<^esub> y \ l1 x \\<^bsub>R1\<^esub> l1 y \ l1 x \\<^bsub>L2\<^esub> l1 y" and mono_R2_R1_r2: "\x y. x \\<^bsub>R2\<^esub> y \ r2 x \\<^bsub>L2\<^esub> r2 y \ r2 x \\<^bsub>R1\<^esub> r2 y" - and "([in_dom (\\<^bsub>L1\<^esub>)] \ [in_codom (\\<^bsub>R2\<^esub>)] \ (\)) + and "(in_dom (\\<^bsub>L1\<^esub>) \ in_codom (\\<^bsub>R2\<^esub>) \ (\)) (rel_bimap l1 r2 (\\<^bsub>R1\<^esub>)) (rel_bimap l1 r2 (\\<^bsub>L2\<^esub>))" - and mono_iff2: "([in_dom (\\<^bsub>R2\<^esub>)] \ [in_codom (\\<^bsub>L1\<^esub>)] \ (\)) + and mono_iff2: "(in_dom (\\<^bsub>R2\<^esub>) \ in_codom (\\<^bsub>L1\<^esub>) \ (\)) (rel_bimap r2 l1 (\\<^bsub>R1\<^esub>)) (rel_bimap r2 l1 (\\<^bsub>L2\<^esub>))" shows "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" proof - - from galois mono_L1_L2_l1 have "([in_codom (\\<^bsub>L1\<^esub>)] \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" - by (intro dep_mono_wrt_predI) blast - moreover from galois mono_R2_R1_r2 have "([in_dom (\\<^bsub>R2\<^esub>)] \\<^sub>m in_dom (\\<^bsub>R1\<^esub>)) r2" - by (intro dep_mono_wrt_predI) blast - moreover from mono_iff2 have "([in_dom (\\<^bsub>R2\<^esub>)] \ [in_codom (\\<^bsub>L1\<^esub>)] \ (\)) + from galois mono_L1_L2_l1 have "(in_codom (\\<^bsub>L1\<^esub>) \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" + by (intro mono_wrt_predI) blast + moreover from galois mono_R2_R1_r2 have "(in_dom (\\<^bsub>R2\<^esub>) \\<^sub>m in_dom (\\<^bsub>R1\<^esub>)) r2" + by (intro mono_wrt_predI) blast + moreover from mono_iff2 have "(in_dom (\\<^bsub>R2\<^esub>) \ in_codom (\\<^bsub>L1\<^esub>) \ (\)) (rel_bimap r2 l1 (\\<^bsub>L2\<^esub>)) (rel_bimap r2 l1 (\\<^bsub>R1\<^esub>))" by blast ultimately show ?thesis using assms by (intro galois_equivalenceI galois_connectionI flip.galois_propI) auto qed lemma galois_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" "((\\<^bsub>L2\<^esub>) \\<^sub>G (\\<^bsub>R2\<^esub>)) l2 r2" and "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) l1" "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) r2" - and "([in_dom (\\<^bsub>L1\<^esub>)] \ [in_codom (\\<^bsub>R2\<^esub>)] \ (\)) + and "(in_dom (\\<^bsub>L1\<^esub>) \ in_codom (\\<^bsub>R2\<^esub>) \ (\)) (rel_bimap l1 r2 (\\<^bsub>R1\<^esub>)) (rel_bimap l1 r2 (\\<^bsub>L2\<^esub>))" - and "([in_dom (\\<^bsub>R2\<^esub>)] \ [in_codom (\\<^bsub>L1\<^esub>)] \ (\)) + and "(in_dom (\\<^bsub>R2\<^esub>) \ in_codom (\\<^bsub>L1\<^esub>) \ (\)) (rel_bimap r2 l1 (\\<^bsub>R1\<^esub>)) (rel_bimap r2 l1 (\\<^bsub>L2\<^esub>))" shows "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_equivalenceI) auto end context transport_comp_same begin lemma galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" "((\\<^bsub>R1\<^esub>) \\<^sub>G (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" using assms by (rule galois_equivalenceI) auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Property.thy b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Property.thy --- a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Property.thy +++ b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Galois_Property.thy @@ -1,51 +1,51 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Property\ theory Transport_Compositions_Agree_Galois_Property imports Transport_Compositions_Agree_Base begin context transport_comp_agree begin lemma galois_propI: assumes galois1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and galois2: "((\\<^bsub>L2\<^esub>) \ (\\<^bsub>R2\<^esub>)) l2 r2" - and mono_l1: "([in_dom (\\<^bsub>L1\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" - and mono_r2: "([in_codom (\\<^bsub>R2\<^esub>)] \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" - and agree: "([in_dom (\\<^bsub>L1\<^esub>)] \ [in_codom (\\<^bsub>R2\<^esub>)] \ (\)) + and mono_l1: "(in_dom (\\<^bsub>L1\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" + and mono_r2: "(in_codom (\\<^bsub>R2\<^esub>) \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" + and agree: "(in_dom (\\<^bsub>L1\<^esub>) \ in_codom (\\<^bsub>R2\<^esub>) \ (\)) (rel_bimap l1 r2 (\\<^bsub>R1\<^esub>)) (rel_bimap l1 r2 (\\<^bsub>L2\<^esub>))" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" proof (rule galois_prop.galois_propI') fix x y assume "in_dom (\\<^bsub>L\<^esub>) x" "in_codom (\\<^bsub>R\<^esub>) y" with mono_r2 mono_l1 have "in_dom (\\<^bsub>L2\<^esub>) (l1 x)" "in_codom (\\<^bsub>R1\<^esub>) (r2 y)" by auto have "x \\<^bsub>L\<^esub> r y \ x \\<^bsub>L1\<^esub> r1 (r2 y)" by simp also from galois1 \in_dom (\\<^bsub>L1\<^esub>) x\ \in_codom (\\<^bsub>R1\<^esub>) (r2 y)\ have "... \ l1 x \\<^bsub>R1\<^esub> r2 y" by (rule g1.galois_prop_left_rel_right_iff_left_right_rel) also from agree \in_dom (\\<^bsub>L1\<^esub>) x\ \in_codom (\\<^bsub>R2\<^esub>) y\ have "... \ l1 x \\<^bsub>L2\<^esub> r2 y" by fastforce also from galois2 \in_dom (\\<^bsub>L2\<^esub>) (l1 x)\ \in_codom (\\<^bsub>R2\<^esub>) y\ have "... \ l x \\<^bsub>R2\<^esub> y" unfolding l_def by (simp add: g2.galois_prop_left_rel_right_iff_left_right_rel) finally show "x \\<^bsub>L\<^esub> r y \ l x \\<^bsub>R\<^esub> y" . qed end context transport_comp_same begin corollary galois_propI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>R1\<^esub>) \ (\\<^bsub>R2\<^esub>)) l2 r2" - and "([in_dom (\\<^bsub>L1\<^esub>)] \\<^sub>m in_dom (\\<^bsub>R1\<^esub>)) l1" - and "([in_codom (\\<^bsub>R2\<^esub>)] \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" + and "(in_dom (\\<^bsub>L1\<^esub>) \\<^sub>m in_dom (\\<^bsub>R1\<^esub>)) l1" + and "(in_codom (\\<^bsub>R2\<^esub>) \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (rule galois_propI) auto end end \ No newline at end of file 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" + 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 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 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 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_Monotone.thy b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Monotone.thy --- a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Monotone.thy +++ b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Monotone.thy @@ -1,30 +1,30 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Monotonicity\ theory Transport_Compositions_Agree_Monotone imports Transport_Compositions_Agree_Base begin context transport_comp_agree begin lemma mono_wrt_rel_leftI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "\x y. x \\<^bsub>L1\<^esub> y \ l1 x \\<^bsub>R1\<^esub> l1 y \ l1 x \\<^bsub>L2\<^esub> l1 y" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" - unfolding left_eq_comp using assms by (rule dep_mono_wrt_rel_compI) + unfolding left_eq_comp using assms by (urule dep_mono_wrt_rel_compI) end context transport_comp_same begin lemma mono_wrt_rel_leftI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" using assms by (rule mono_wrt_rel_leftI) 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,132 +1,132 @@ \<^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" + 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 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 "\(P :: 'i \ bool) (R :: 'j \ 'i \ bool). (inflationary_on P R\ :: ('i \ 'j) \ bool) \ deflationary_on P R" and "\(R :: 'i \ 'i \ bool). transitive R\ \ transitive R" and "\R. in_field R\ \ in_field R" - by simp_all + by (simp_all add: mono_wrt_rel_eq_dep_mono_wrt_rel) 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) + mono_wrt_rel_leftI flip.mono_wrt_rel_leftI 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,589 +1,590 @@ \<^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 + shows "(in_dom (\\<^bsub>L\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" + using assms by (intro 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 + shows "(in_codom (\\<^bsub>L\<^esub>) \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" + using assms by (intro 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 - +thm mono_rel_comp 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) + using mono_rel_comp R_S_le by blast + also have "... = S \\ (R \\ R)" by (simp flip: rel_comp_assoc) + also have "... \ (S \\ R)" using mono_rel_comp R_R_le by blast 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 + using mono_rel_comp S_R_le by (auto simp flip: rel_comp_assoc) + also have "... \ (R \\ S)" using mono_rel_comp 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_Property.thy b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Property.thy --- a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Property.thy +++ b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Property.thy @@ -1,195 +1,196 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Property\ theory Transport_Compositions_Generic_Galois_Property 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" and "flip.t1.counit \ \\<^sub>2" by (simp_all only: t1.flip_unit_eq_counit t2.flip_counit_eq_unit) lemma half_galois_prop_left_left_rightI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and deflationary_counit1: "deflationary_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and trans_R1: "transitive (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "reflexive_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" 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>)" - and mono_in_codom_r2: "([in_codom (\\<^bsub>R\<^esub>)] \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" + and mono_in_codom_r2: "(in_codom (\\<^bsub>R\<^esub>) \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" shows "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" proof (rule half_galois_prop_leftI) fix x z assume "x \<^bsub>L\<^esub>\ z" then show "l x \\<^bsub>R\<^esub> z" proof (intro right_rel_if_left_relI) from \x \<^bsub>L\<^esub>\ z\ show "in_codom (\\<^bsub>R2\<^esub>) z" by blast fix y assume "y \\<^bsub>R1\<^esub> l1 (r z)" moreover have "l1 (r z) \\<^bsub>R1\<^esub> r2 z" proof - from mono_in_codom_r2 \x \<^bsub>L\<^esub>\ z\ have "in_codom (\\<^bsub>R1\<^esub>) (r2 z)" by blast with deflationary_counit1 show "l1 (r z) \\<^bsub>R1\<^esub> r2 z" by auto qed ultimately show "y \\<^bsub>R1\<^esub> r2 z" using trans_R1 by blast next 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 qed (insert assms, auto) qed lemma half_galois_prop_left_left_rightI': assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and deflationary_counit1: "deflationary_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and trans_R1: "transitive (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" 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>)" - and mono_in_codom_r2: "([in_codom (\\<^bsub>R\<^esub>)] \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" + and mono_in_codom_r2: "(in_codom (\\<^bsub>R\<^esub>) \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" shows "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" proof (rule half_galois_prop_leftI) fix x z assume "x \<^bsub>L\<^esub>\ z" then show "l x \\<^bsub>R\<^esub> z" proof (intro right_rel_if_left_relI') from \x \<^bsub>L\<^esub>\ z\ show "in_codom (\\<^bsub>R2\<^esub>) z" by blast fix y assume "y \\<^bsub>R1\<^esub> l1 (r z)" moreover have "l1 (r z) \\<^bsub>R1\<^esub> r2 z" proof - from mono_in_codom_r2 \x \<^bsub>L\<^esub>\ z\ have "in_codom (\\<^bsub>R1\<^esub>) (r2 z)" by blast with deflationary_counit1 show "l1 (r z) \\<^bsub>R1\<^esub> r2 z" by auto qed ultimately show "y \\<^bsub>R1\<^esub> r2 z" using trans_R1 by blast next assume "in_dom (\\<^bsub>L2\<^esub>) (l1 x)" with refl_L2 have "l1 x \\<^bsub>L2\<^esub> l1 x" by blast with \((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2\ show "in_codom (\\<^bsub>L2\<^esub>) (l1 x)" "l x \\<^bsub>R2\<^esub> l2 (l1 x)" by auto qed (insert assms, auto) qed lemma half_galois_prop_right_left_rightI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and inflationary_counit1: "inflationary_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "((\\<^bsub>R2\<^esub>) \<^sub>h\ (\\<^bsub>L2\<^esub>)) r2 l2" and inflationary_unit2: "inflationary_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and trans_L2: "transitive (\\<^bsub>L2\<^esub>)" - and mono_in_dom_l1: "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" + and mono_in_dom_l1: "(in_dom (\\<^bsub>L\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" 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>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" proof (rule half_galois_prop_rightI) fix x z assume "x \\<^bsub>R\<^esub> z" then show "x \\<^bsub>L\<^esub> r z" proof (intro flip.right_rel_if_left_relI) fix y assume "r2 (l x) \\<^bsub>L2\<^esub> y" moreover have "l1 x \\<^bsub>L2\<^esub> r2 (l x)" proof - from mono_in_dom_l1 \x \\<^bsub>R\<^esub> z\ have "in_dom (\\<^bsub>L2\<^esub>) (l1 x)" by blast with inflationary_unit2 show "l1 x \\<^bsub>L2\<^esub> r2 (l x)" by auto qed ultimately show "l1 x \\<^bsub>L2\<^esub> y" using trans_L2 by blast fix y assume "l1 x \\<^bsub>R1\<^esub> y" with \((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1\ \x \\<^bsub>R\<^esub> z\ show "x \\<^bsub>L1\<^esub> r1 y" by blast next assume "in_codom (\\<^bsub>R1\<^esub>) (r2 z)" with inflationary_counit1 show "r2 z \\<^bsub>R1\<^esub> l1 (r z)" by auto from \((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1\ \in_codom (\\<^bsub>R1\<^esub>) (r2 z)\ show "in_codom (\\<^bsub>L1\<^esub>) (r z)" - by (auto intro: in_codom_if_rel_if_dep_mono_wrt_rel) + by (auto intro: in_codom_if_rel_if_dep_mono_wrt_rel + simp: mono_wrt_rel_eq_dep_mono_wrt_rel) qed (insert assms, auto elim: galois_rel.left_GaloisE) qed lemma half_galois_prop_right_left_rightI': assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and inflationary_unit1: "inflationary_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and inflationary_counit1: "\y z. y \\<^bsub>R1\<^esub> r2 z \ y \\<^bsub>R1\<^esub> l1 (r z)" and "in_dom (\\<^bsub>R1\<^esub>) \ in_codom (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \<^sub>h\ (\\<^bsub>L2\<^esub>)) r2 l2" and inflationary_unit2: "inflationary_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and trans_L2: "transitive (\\<^bsub>L2\<^esub>)" - and mono_in_dom_l1: "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" + and mono_in_dom_l1: "(in_dom (\\<^bsub>L\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" 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 "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" proof (rule half_galois_prop_rightI) fix x z assume "x \\<^bsub>R\<^esub> z" then show "x \\<^bsub>L\<^esub> r z" proof (intro flip.right_rel_if_left_relI') from \x \\<^bsub>R\<^esub> z\ inflationary_unit1 show "x \\<^bsub>L1\<^esub> r1 (l1 x)" by (fastforce elim: galois_rel.left_GaloisE) fix y assume "y \\<^bsub>R1\<^esub> r2 z" with inflationary_counit1 show "y \\<^bsub>R1\<^esub> l1 (r z)" by auto next fix y from mono_in_dom_l1 \x \\<^bsub>R\<^esub> z\ have "in_dom (\\<^bsub>L2\<^esub>) (l1 x)" by blast with inflationary_unit2 have "l1 x \\<^bsub>L2\<^esub> r2 (l x)" by auto moreover assume "r2 (l x) \\<^bsub>L2\<^esub> y" ultimately show "l1 x \\<^bsub>L2\<^esub> y" using trans_L2 by blast qed (insert assms, auto elim: galois_rel.left_GaloisE) qed lemma galois_prop_left_rightI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "rel_equivalence_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "transitive (\\<^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 "inflationary_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "preorder_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "middle_compatible_codom" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_propI half_galois_prop_left_left_rightI half_galois_prop_right_left_rightI flip.mono_in_codom_left_rel_left1_if_in_codom_rel_comp_le mono_in_dom_left_rel_left1_if_in_dom_rel_comp_le in_dom_right1_left2_right1_le_if_right1_left2_right1_le) (auto elim!: preorder_on_in_fieldE intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom) lemma galois_prop_left_rightI': assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "inflationary_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and rel_equiv_counit1: "rel_equivalence_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and trans_R1: "transitive (\\<^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 "inflationary_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "preorder_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "middle_compatible_dom" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" proof (rule galois_propI) show "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" using assms by (intro half_galois_prop_left_left_rightI' flip.mono_in_codom_left_rel_left1_if_in_codom_rel_comp_le flip.in_codom_right1_left2_right1_le_if_right1_left2_right1_le) (auto elim!: rel_equivalence_onE preorder_on_in_fieldE intro: 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) have "y \\<^bsub>R1\<^esub> l1 (r1 (r2 z))" if "y \\<^bsub>R1\<^esub> r2 z" for y z proof - note \y \\<^bsub>R1\<^esub> r2 z\ moreover with rel_equiv_counit1 have "r2 z \\<^bsub>R1\<^esub> \\<^sub>1 (r2 z)" by blast ultimately show ?thesis using trans_R1 by auto qed moreover have "in_dom (\\<^bsub>R1\<^esub>) \ in_codom (\\<^bsub>R1\<^esub>)" proof - from rel_equiv_counit1 trans_R1 have "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" by (intro reflexive_on_in_field_if_transitive_if_rel_equivalence_on) auto then show ?thesis by (simp only: in_codom_eq_in_dom_if_reflexive_on_in_field) qed ultimately show "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" using assms by (intro half_galois_prop_right_left_rightI' mono_in_dom_left_rel_left1_if_in_dom_rel_comp_le) auto qed end end \ No newline at end of file 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" + 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 (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>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) +proof (rule 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 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) +proof (rule 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" + 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" + 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" + 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 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 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 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" + 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" + 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 + with mono_in_codom_l1 show ?thesis by (intro 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 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" + 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" + 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 + with mono_in_dom_l1 show ?thesis by (intro 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/Compositions/Generic/Transport_Compositions_Generic_Order_Equivalence.thy b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Order_Equivalence.thy --- a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Order_Equivalence.thy +++ b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Order_Equivalence.thy @@ -1,331 +1,331 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Order Equivalence\ theory Transport_Compositions_Generic_Order_Equivalence imports Transport_Compositions_Generic_Monotone begin context transport_comp begin context begin interpretation flip : transport_comp R2 L2 r2 l2 R1 L1 r1 l1 . subsubsection \Unit\ paragraph \Inflationary\ lemma inflationary_on_in_dom_unitI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and inflationary_unit1: "inflationary_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and inflationary_counit1: "inflationary_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and refl_R1: "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and inflationary_unit2: "inflationary_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" 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" + and mono_in_dom_l1: "(in_dom (\\<^bsub>L\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" and in_codom_rel_comp_le: "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom ((\\<^bsub>R1\<^esub>))" shows "inflationary_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" proof (rule inflationary_onI) fix x assume "in_dom (\\<^bsub>L\<^esub>) x" show "x \\<^bsub>L\<^esub> \ x" proof (rule left_relI) from \in_dom (\\<^bsub>L\<^esub>) x\ \((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1\ have "in_dom (\\<^bsub>R1\<^esub>) (l1 x)" by blast with refl_R1 have "l1 x \\<^bsub>R1\<^esub> l1 x" by blast moreover from \in_dom (\\<^bsub>L\<^esub>) x\ have "in_dom (\\<^bsub>L1\<^esub>) x" by blast moreover note inflationary_unit1 ultimately show "x \<^bsub>L1\<^esub>\ l1 x" by (intro t1.left_GaloisI) auto from \in_dom (\\<^bsub>L\<^esub>) x\ mono_in_dom_l1 have "in_dom (\\<^bsub>L2\<^esub>) (l1 x)" by blast with inflationary_unit2 show "l1 x \\<^bsub>L2\<^esub> r2 (l x)" by auto show "r2 (l x) \<^bsub>R1\<^esub>\ \ x" proof (rule flip.t2.left_GaloisI) from refl_L2 \in_dom (\\<^bsub>L2\<^esub>) (l1 x)\ have "l1 x \\<^bsub>L2\<^esub> l1 x" by blast with in_codom_rel_comp_le \l1 x \\<^bsub>R1\<^esub> l1 x\ \l1 x \\<^bsub>L2\<^esub> r2 (l x)\ have "in_codom (\\<^bsub>R1\<^esub>) (r2 (l x))" by blast with \((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1\ show "in_codom (\\<^bsub>L1\<^esub>) (\ x)" - by (auto intro: in_codom_if_rel_if_dep_mono_wrt_rel) + by (auto intro: in_codom_if_rel_if_dep_mono_wrt_rel simp: mono_wrt_rel_eq_dep_mono_wrt_rel) from \in_codom (\\<^bsub>R1\<^esub>) (r2 (l x))\ inflationary_counit1 show "r2 (l x) \\<^bsub>R1\<^esub> l1 (\ x)" by auto qed qed qed lemma inflationary_on_in_codom_unitI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and inflationary_unit1: "inflationary_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and inflationary_counit1: "inflationary_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and refl_R1: "reflexive_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and inflationary_unit2: "inflationary_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" 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" + and mono_in_codom_l1: "(in_codom (\\<^bsub>L\<^esub>) \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" and in_codom_rel_comp_le: "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom ((\\<^bsub>R1\<^esub>))" shows "inflationary_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" proof (rule inflationary_onI) fix x assume "in_codom (\\<^bsub>L\<^esub>) x" show "x \\<^bsub>L\<^esub> \ x" proof (rule left_relI) from \in_codom (\\<^bsub>L\<^esub>) x\ have "in_codom (\\<^bsub>L1\<^esub>) x" "in_codom (\\<^bsub>R1\<^esub>) (l1 x)" by blast+ with inflationary_unit1 show "x \<^bsub>L1\<^esub>\ l1 x" by (intro t1.left_GaloisI) auto from mono_in_codom_l1 \in_codom (\\<^bsub>L\<^esub>) x\ have "in_codom (\\<^bsub>L2\<^esub>) (l1 x)" by blast with inflationary_unit2 show "l1 x \\<^bsub>L2\<^esub> r2 (l x)" by auto show "r2 (l x) \<^bsub>R1\<^esub>\ \ x" proof (rule flip.t2.left_GaloisI) from refl_L2 \in_codom (\\<^bsub>L2\<^esub>) (l1 x)\ have "l1 x \\<^bsub>L2\<^esub> l1 x" by blast moreover from refl_R1 \in_codom (\\<^bsub>R1\<^esub>) (l1 x)\ have "l1 x \\<^bsub>R1\<^esub> l1 x" by blast moreover note in_codom_rel_comp_le \l1 x \\<^bsub>L2\<^esub> r2 (l x)\ ultimately have "in_codom (\\<^bsub>R1\<^esub>) (r2 (l x))" by blast with \((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1\ show "in_codom (\\<^bsub>L1\<^esub>) (\ x)" - by (auto intro: in_codom_if_rel_if_dep_mono_wrt_rel) + by (auto intro: in_codom_if_rel_if_dep_mono_wrt_rel simp: mono_wrt_rel_eq_dep_mono_wrt_rel) from \in_codom (\\<^bsub>R1\<^esub>) (r2 (l x))\ inflationary_counit1 show "r2 (l x) \\<^bsub>R1\<^esub> l1 (\ x)" by auto qed qed qed corollary inflationary_on_in_field_unitI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "inflationary_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "inflationary_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "inflationary_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "reflexive_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" - and "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" - and "([in_codom (\\<^bsub>L\<^esub>)] \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" + and "(in_dom (\\<^bsub>L\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" + and "(in_codom (\\<^bsub>L\<^esub>) \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom ((\\<^bsub>R1\<^esub>))" shows "inflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" proof - from assms have "inflationary_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" by (intro inflationary_on_in_dom_unitI) (auto intro: inflationary_on_if_le_pred_if_inflationary_on reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) moreover from assms have "inflationary_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" by (intro inflationary_on_in_codom_unitI) (auto intro: inflationary_on_if_le_pred_if_inflationary_on 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 text \Deflationary\ lemma deflationary_on_in_dom_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and in_dom_R1_le_in_codom_R1: "in_dom (\\<^bsub>R1\<^esub>) \ in_codom (\\<^bsub>R1\<^esub>)" and deflationary_L2: "deflationary_on (in_dom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" 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" + and mono_in_dom_l1: "(in_dom (\\<^bsub>L\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" and in_dom_rel_comp_le: "in_dom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_dom ((\\<^bsub>R1\<^esub>))" shows "deflationary_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" proof (rule deflationary_onI) fix x assume "in_dom (\\<^bsub>L\<^esub>) x" show "\ x \\<^bsub>L\<^esub> x" proof (rule left_relI) from refl_L1 \in_dom (\\<^bsub>L\<^esub>) x\ have "x \\<^bsub>L1\<^esub> x" by blast moreover with \((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1\ have "l1 x \\<^bsub>R1\<^esub> l1 x" by blast ultimately show "l1 x \<^bsub>R1\<^esub>\ x" by auto from mono_in_dom_l1 \in_dom (\\<^bsub>L\<^esub>) x\ have "in_dom (\\<^bsub>L2\<^esub>) (l1 x)" by blast with deflationary_L2 show "r2 (l x) \\<^bsub>L2\<^esub> l1 x" by auto show "\ x \<^bsub>L1\<^esub>\ r2 (l x)" proof (rule t1.left_GaloisI) from refl_L2 \in_dom (\\<^bsub>L2\<^esub>) (l1 x)\ have "l1 x \\<^bsub>L2\<^esub> l1 x" by blast with in_dom_rel_comp_le \r2 (l x) \\<^bsub>L2\<^esub> l1 x\ \l1 x \\<^bsub>R1\<^esub> l1 x\ have "in_dom (\\<^bsub>R1\<^esub>) (r2 (l x))" by blast with \((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1\ have "in_dom (\\<^bsub>L1\<^esub>) (\ x)" - by (auto intro: in_dom_if_rel_if_dep_mono_wrt_rel) + by (auto intro: in_dom_if_rel_if_dep_mono_wrt_rel simp: mono_wrt_rel_eq_dep_mono_wrt_rel) with refl_L1 show "\ x \\<^bsub>L1\<^esub> r1 (r2 (l x))" by (auto intro: in_field_if_in_codom) from \in_dom (\\<^bsub>R1\<^esub>) (r2 (l x))\ in_dom_R1_le_in_codom_R1 show "in_codom (\\<^bsub>R1\<^esub>) (r2 (l x))" by blast qed qed qed lemma deflationary_on_in_codom_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and refl_L1: "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and in_dom_R1_le_in_codom_R1: "in_dom (\\<^bsub>R1\<^esub>) \ in_codom (\\<^bsub>R1\<^esub>)" and deflationary_L2: "deflationary_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" 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" + and mono_in_codom_l1: "(in_codom (\\<^bsub>L\<^esub>) \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" and in_dom_rel_comp_le: "in_dom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_dom ((\\<^bsub>R1\<^esub>))" shows "deflationary_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" proof (rule deflationary_onI) fix x assume "in_codom (\\<^bsub>L\<^esub>) x" show "\ x \\<^bsub>L\<^esub> x" proof (rule left_relI) from refl_L1 \in_codom (\\<^bsub>L\<^esub>) x\ have "x \\<^bsub>L1\<^esub> x" by blast moreover with \((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1\ have "l1 x \\<^bsub>R1\<^esub> l1 x" by blast ultimately show "l1 x \<^bsub>R1\<^esub>\ x" by auto from mono_in_codom_l1 \in_codom (\\<^bsub>L\<^esub>) x\ have "in_codom (\\<^bsub>L2\<^esub>) (l1 x)" by blast with deflationary_L2 show "r2 (l x) \\<^bsub>L2\<^esub> l1 x" by auto show "\ x \<^bsub>L1\<^esub>\ r2 (l x)" proof (rule t1.left_GaloisI) from refl_L2 \in_codom (\\<^bsub>L2\<^esub>) (l1 x)\ have "l1 x \\<^bsub>L2\<^esub> l1 x" by blast with in_dom_rel_comp_le \r2 (l x) \\<^bsub>L2\<^esub> l1 x\ \l1 x \\<^bsub>R1\<^esub> l1 x\ have "in_dom (\\<^bsub>R1\<^esub>) (r2 (l x))" by blast with in_dom_R1_le_in_codom_R1 show "in_codom (\\<^bsub>R1\<^esub>) (r2 (l x))" by blast with \((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1\ have "in_codom (\\<^bsub>L1\<^esub>) (\ x)" - by (auto intro: in_codom_if_rel_if_dep_mono_wrt_rel) + by (auto intro: in_codom_if_rel_if_dep_mono_wrt_rel simp: mono_wrt_rel_eq_dep_mono_wrt_rel) with refl_L1 show "\ x \\<^bsub>L1\<^esub> r1 (r2 (l x))" by auto qed qed qed corollary deflationary_on_in_field_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "in_dom (\\<^bsub>R1\<^esub>) \ in_codom (\\<^bsub>R1\<^esub>)" and "deflationary_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "reflexive_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" - and "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" - and "([in_codom (\\<^bsub>L\<^esub>)] \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" + and "(in_dom (\\<^bsub>L\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" + and "(in_codom (\\<^bsub>L\<^esub>) \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" and "in_dom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_dom ((\\<^bsub>R1\<^esub>))" shows "deflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" proof - from assms have "deflationary_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" by (intro deflationary_on_in_dom_unitI) (auto intro: deflationary_on_if_le_pred_if_deflationary_on reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) moreover from assms have "deflationary_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" by (intro deflationary_on_in_codom_unitI) (auto intro: deflationary_on_if_le_pred_if_deflationary_on 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 text \Relational Equivalence\ corollary rel_equivalence_on_in_field_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "inflationary_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "inflationary_on (in_codom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "rel_equivalence_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "reflexive_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" - and "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" - and "([in_codom (\\<^bsub>L\<^esub>)] \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" + and "(in_dom (\\<^bsub>L\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" + and "(in_codom (\\<^bsub>L\<^esub>) \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" and "in_dom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_dom ((\\<^bsub>R1\<^esub>))" and "in_codom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_codom ((\\<^bsub>R1\<^esub>))" 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 simp only: in_codom_eq_in_dom_if_reflexive_on_in_field) subsubsection \Counit\ text \Corresponding lemmas for the counit can be obtained by flipping the interpretation of the locale, i.e. \ interpretation flip : transport_comp R2 L2 r2 l2 R1 L1 r1 l1 rewrites "flip.t2.unit \ \\<^sub>1" and "flip.t2.counit \ \\<^sub>1" and "flip.t1.unit \ \\<^sub>2" and "flip.t1.counit \ \\<^sub>2" and "flip.unit \ \" and "flip.counit \ \" unfolding transport_comp.transport_defs by (auto simp: order_functors.flip_counit_eq_unit) \ \ end subsubsection \Order Equivalence\ interpretation flip : transport_comp R2 L2 r2 l2 R1 L1 r1 l1 rewrites "flip.t2.unit \ \\<^sub>1" and "flip.t2.counit \ \\<^sub>1" and "flip.t1.unit \ \\<^sub>2" and "flip.t1.counit \ \\<^sub>2" and "flip.counit \ \" and "flip.unit \ \" by (simp_all only: order_functors.flip_counit_eq_unit) lemma order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "inflationary_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and rel_equiv_counit1: "rel_equivalence_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \<^sub>h\ (\\<^bsub>L2\<^esub>)) r2 l2" and rel_equiv_unit2: "rel_equivalence_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "inflationary_on (in_field (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>) \\<^sub>2" and "reflexive_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>)" and middle_compatible: "middle_compatible_codom" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" proof (rule order_equivalenceI) show "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" using rel_equiv_unit2 \((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1\ \((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2\ middle_compatible by (intro mono_wrt_rel_leftI) auto show "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" using rel_equiv_counit1 \((\\<^bsub>R2\<^esub>) \<^sub>h\ (\\<^bsub>L2\<^esub>)) r2 l2\ \((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1\ middle_compatible by (intro flip.mono_wrt_rel_leftI) (auto intro: inflationary_on_if_le_pred_if_inflationary_on in_field_if_in_codom) from middle_compatible have in_dom_rel_comp_les: "in_dom ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>)) \ in_dom (\\<^bsub>L2\<^esub>)" "in_dom ((\\<^bsub>L2\<^esub>) \\ (\\<^bsub>R1\<^esub>) \\ (\\<^bsub>L2\<^esub>)) \ in_dom ((\\<^bsub>R1\<^esub>))" by (auto intro: in_dom_right1_left2_right1_le_if_right1_left2_right1_le flip.in_dom_right1_left2_right1_le_if_right1_left2_right1_le) - moreover then have "([in_dom (\\<^bsub>L\<^esub>)] \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" - and "([in_codom (\\<^bsub>L\<^esub>)] \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" + moreover then have "(in_dom (\\<^bsub>L\<^esub>) \\<^sub>m in_dom (\\<^bsub>L2\<^esub>)) l1" + and "(in_codom (\\<^bsub>L\<^esub>) \\<^sub>m in_codom (\\<^bsub>L2\<^esub>)) l1" using \((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1\ middle_compatible by (auto intro: mono_in_dom_left_rel_left1_if_in_dom_rel_comp_le mono_in_codom_left_rel_left1_if_in_codom_rel_comp_le) ultimately show "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms by (intro rel_equivalence_on_in_field_unitI) (auto intro: inflationary_on_if_le_pred_if_inflationary_on intro!: in_field_if_in_codom) note in_dom_rel_comp_les - moreover then have "([in_dom (\\<^bsub>R\<^esub>)] \\<^sub>m in_dom (\\<^bsub>R1\<^esub>)) r2" - and "([in_codom (\\<^bsub>R\<^esub>)] \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" + moreover then have "(in_dom (\\<^bsub>R\<^esub>) \\<^sub>m in_dom (\\<^bsub>R1\<^esub>)) r2" + and "(in_codom (\\<^bsub>R\<^esub>) \\<^sub>m in_codom (\\<^bsub>R1\<^esub>)) r2" using \((\\<^bsub>R2\<^esub>) \<^sub>h\ (\\<^bsub>L2\<^esub>)) r2 l2\ middle_compatible by (auto intro!: flip.mono_in_dom_left_rel_left1_if_in_dom_rel_comp_le flip.mono_in_codom_left_rel_left1_if_in_codom_rel_comp_le) ultimately show "rel_equivalence_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>) \" using assms by (intro flip.rel_equivalence_on_in_field_unitI) (auto intro: inflationary_on_if_le_pred_if_inflationary_on intro!: in_field_if_in_codom) qed corollary order_equivalence_if_order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>o (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>L2\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>)" and "middle_compatible_codom" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalenceI) (auto elim!: t1.order_equivalenceE t2.order_equivalenceE rel_equivalence_onE intro!: reflexive_on_in_field_if_transitive_if_rel_equivalence_on t1.half_galois_prop_left_left_right_if_transitive_if_deflationary_on_if_mono_wrt_rel flip.t1.half_galois_prop_left_left_right_if_transitive_if_deflationary_on_if_mono_wrt_rel intro: deflationary_on_if_le_pred_if_deflationary_on in_field_if_in_codom) corollary order_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 "((\\<^bsub>L2\<^esub>) \\<^sub>G (\\<^bsub>R2\<^esub>)) l2 r2" and "reflexive_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>)" and "middle_compatible_codom" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalenceI) (auto elim!: t1.galois_equivalenceE t2.galois_equivalenceE intro!: t1.inflationary_on_unit_if_reflexive_on_if_galois_equivalence flip.t1.inflationary_on_unit_if_reflexive_on_if_galois_equivalence t2.rel_equivalence_on_unit_if_reflexive_on_if_galois_equivalence flip.t2.rel_equivalence_on_unit_if_reflexive_on_if_galois_equivalence) 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,171 +1,171 @@ \<^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] +lemmas related_Fun_Rel_combI = Fun_Rel_relD[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\ structure A = Standard_Mixed_Unification \ 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 |> Type_Unification.e_unify Unification_Util.unify_types), prems_unifier = SOME (Transport_Mixed_Unification.first_higherp_decomp_comb_higher_unify |> Unification_Combinator.norm_unifier Envir_Normalisation.beta_norm_term_unif), normalisers = SOME Transport_Mixed_Unification.norms_first_higherp_decomp_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_decomp_comb_higher_unify) + (Unification_Util.inst_norm_term' + Transport_Mixed_Unification.norms_first_higherp_decomp_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] *) + [ML_Krattr \Drule.rearrange_prems [1] #> Drule.rearrange_prems [2,3]\, 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,225 +1,225 @@ \<^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" + 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" + 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 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 (intro transitive_onI) (force dest: transitive_onD) end context fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" begin 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 end 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" + 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)" + 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 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)" + "((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 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,411 +1,411 @@ (* Title: Transport/transport.ML Author: Kevin Kappelmann Prototype for Transport. See README.md for future work. *) (*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 +val simp_rhs = Simplifier.rewrite #> Conversion_Util.rhs_conv #> Conversion_Util.apply_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_decomp_comb_higher_unify Transport_Mixed_Unification.first_higherp_decomp_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) + |> (Simplifier.rewrite ctxt |> Conversion_Util.apply_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_Order_Pattern_Unification.e_match Unification_Util.match_types (Higher_Order_Pattern_Decomp_Unification.e_match unif Unification_Combinator.fail_match) Unification_Combinator.fail_match |> Type_Unification.e_match Unification_Util.match_types) binders in Unify_Resolve_Base.unify_resolve_any_tac Higher_Order_Pattern_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_Dep_Fun_Rel_Examples.thy b/thys/Transport/Transport/Examples/Transport_Dep_Fun_Rel_Examples.thy --- a/thys/Transport/Transport/Examples/Transport_Dep_Fun_Rel_Examples.thy +++ b/thys/Transport/Transport/Examples/Transport_Dep_Fun_Rel_Examples.thy @@ -1,85 +1,79 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Example Transports for Dependent Function Relator\ theory Transport_Dep_Fun_Rel_Examples imports Transport_Prototype Transport_Syntax + HOL_Alignment_Binary_Relations "HOL-Library.IArray" begin paragraph \Summary\ text \Dependent function relator examples from \<^cite>\"transport"\. Refer to the paper for more details.\ context includes galois_rel_syntax transport_syntax notes 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] + [ML_Krattr \Drule.rearrange_prems [1] #> Drule.rearrange_prems [2,3]\, per_intro] begin interpretation transport L R l r for L R l r . abbreviation "Zpos \ ((=\<^bsub>(\)(0 :: int)\<^esub>) :: int \ _)" lemma Zpos_per [per_intro]: "(Zpos \\<^bsub>PER\<^esub> (=)) nat int" by fastforce lemma sub_parametric [trp_in_dom]: - "([i _ \ Zpos] \ [j _ \ Zpos | j \ i] \ Zpos) (-) (-)" + "((i _ \ Zpos) \ (j _ \ Zpos | j \ i) \ Zpos) (-) (-)" by fastforce trp_term nat_sub :: "nat \ nat \ nat" where x = "(-) :: int \ _" - and L = "[i _ \ Zpos] \ [j _ \ Zpos | j \ i] \ Zpos" - and R = "[n _ \ (=)] \ [m _ \ (=)| m \ n] \ (=)" + and L = "(i _ \ Zpos) \ (j _ \ Zpos | j \ i) \ Zpos" + and R = "(n _ \ (=)) \ (m _ \ (=)| m \ n) \ (=)" (*fastforce discharges the remaining side-conditions*) - by (trp_prover) fastforce+ + by trp_prover fastforce+ thm nat_sub_app_eq text \Note: as of now, @{command trp_term} does not rewrite the Galois relator of dependent function relators.\ thm nat_sub_related' abbreviation "LRel \ list_all2" abbreviation "IARel \ rel_iarray" -lemma transp_eq_transitive: "transp = transitive" - by (auto intro: transpI dest: transpD) -lemma symp_eq_symmetric: "symp = symmetric" - by (auto intro: sympI dest: sympD symmetricD) - lemma [per_intro]: assumes "partial_equivalence_rel R" shows "(LRel R \\<^bsub>PER\<^esub> IARel R) IArray.IArray IArray.list_of" using assms by (fastforce simp flip: transp_eq_transitive symp_eq_symmetric intro: list.rel_transp list.rel_symp iarray.rel_transp iarray.rel_symp elim: iarray.rel_cases)+ lemma [trp_in_dom]: - "([xs _ \ LRel R] \ [i _ \ (=) | i < length xs] \ R) (!) (!)" + "((xs _ \ LRel R) \ (i _ \ (=) | i < length xs) \ R) (!) (!)" by (fastforce simp: list_all2_lengthD list_all2_nthD2) context fixes R :: "'a \ 'a \ bool" assumes [per_intro]: "partial_equivalence_rel R" begin interpretation Rper : transport_partial_equivalence_rel_id R by unfold_locales per_prover declare Rper.partial_equivalence_rel_equivalence [per_intro] trp_term iarray_index where x = "(!) :: 'a list \ _" - and L = "([xs _ \ LRel R] \ [i _ \ (=) | i < length xs] \ R)" - and R = "([xs _ \ IARel R] \ [i _ \ (=) | i < IArray.length xs] \ R)" - by (trp_prover) + and L = "((xs _ \ LRel R) \ (i _ \ (=) | i < length xs) \ R)" + and R = "((xs _ \ IARel R) \ (i _ \ (=) | i < IArray.length xs) \ R)" + by trp_prover (*fastforce discharges the remaining side-conditions*) (fastforce simp: list_all2_lengthD elim: iarray.rel_cases)+ end end end diff --git a/thys/Transport/Transport/Examples/Transport_Lists_Sets_Examples.thy b/thys/Transport/Transport/Examples/Transport_Lists_Sets_Examples.thy --- a/thys/Transport/Transport/Examples/Transport_Lists_Sets_Examples.thy +++ b/thys/Transport/Transport/Examples/Transport_Lists_Sets_Examples.thy @@ -1,167 +1,167 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Example Transports Between Lists and Sets\ theory Transport_Lists_Sets_Examples imports Transport_Prototype Transport_Syntax "HOL-Library.FSet" begin paragraph \Summary\ text \Introductory examples from \<^cite>\"transport"\. Transports between lists and (finite) sets. Refer to the paper for more details.\ context includes galois_rel_syntax transport_syntax begin paragraph \Introductory examples from paper\ text \Left and right relations.\ definition "LFSL xs xs' \ fset_of_list xs = fset_of_list xs'" abbreviation (input) "(LFSR :: 'a fset \ _) \ (=)" definition "LSL xs xs' \ set xs = set xs'" abbreviation (input) "(LSR :: 'a set \ _) \ (=\<^bsub>finite :: 'a set \ bool\<^esub>)" interpretation t : transport LSL R l r for LSL R l r . text \Proofs of equivalences.\ lemma list_fset_PER [per_intro]: "(LFSL \\<^bsub>PER\<^esub> LFSR) fset_of_list sorted_list_of_fset" unfolding LFSL_def by fastforce lemma list_set_PER [per_intro]: "(LSL \\<^bsub>PER\<^esub> LSR) set sorted_list_of_set" unfolding LSL_def by fastforce text \We can rewrite the Galois relators in the following theorems to the relator of the paper.\ definition "LFS xs s \ fset_of_list xs = s" definition "LS xs s \ set xs = s" lemma LFSL_Galois_eq_LFS: "(\<^bsub>LFSL\<^esub>\\<^bsub>LFSR sorted_list_of_fset\<^esub>) \ LFS" unfolding LFS_def LFSL_def by (intro eq_reflection ext) (auto) lemma LFSR_Galois_eq_inv_LFS: "(\<^bsub>LFSR\<^esub>\\<^bsub>LFSL fset_of_list\<^esub>) \ LFS\" unfolding LFS_def LFSL_def by (intro eq_reflection ext) (auto) lemma LSL_Galois_eq_LS: "(\<^bsub>LSL\<^esub>\\<^bsub>LSR sorted_list_of_set\<^esub>) \ LS" unfolding LS_def LSL_def by (intro eq_reflection ext) (auto) declare LFSL_Galois_eq_LFS[trp_relator_rewrite, trp_uhint] LFSR_Galois_eq_inv_LFS[trp_relator_rewrite, trp_uhint] LSL_Galois_eq_LS[trp_relator_rewrite, trp_uhint] definition "max_list xs \ foldr max xs (0 :: nat)" text \Proof of parametricity for @{term max_list}.\ lemma max_max_list_removeAll_eq_maxlist: assumes "x \ set xs" shows "max x (max_list (removeAll x xs)) = max_list xs" unfolding max_list_def using assms by (induction xs) (simp_all, (metis max.left_idem removeAll_id max.left_commute)+) lemma max_list_parametric [trp_in_dom]: "(LSL \ (=)) max_list max_list" -proof (intro Dep_Fun_Rel_relI) +proof (intro Fun_Rel_relI) fix xs xs' :: "nat list" assume "LSL xs xs'" then have "finite (set xs)" "set xs = set xs'" unfolding LSL_def by auto then show "max_list xs = max_list xs'" proof (induction "set xs" arbitrary: xs xs' rule: finite_induct) case (insert x F) then have "F = set (removeAll x xs)" by auto moreover from insert have "... = set (removeAll x xs')" by auto ultimately have "max_list (removeAll x xs) = max_list (removeAll x xs')" (is "?lhs = ?rhs") using insert by blast then have "max x ?lhs = max x ?rhs" by simp then show ?case using insert max_max_list_removeAll_eq_maxlist insertI1 by metis qed auto qed lemma LFSL_eq_LSL: "LFSL \ LSL" unfolding LFSL_def LSL_def by (intro eq_reflection ext) (auto simp: fset_of_list_elem) lemma max_list_parametricfin [trp_in_dom]: "(LFSL \ (=)) max_list max_list" using max_list_parametric by (simp only: LFSL_eq_LSL) text \Transport from lists to finite sets.\ trp_term max_fset :: "nat fset \ nat" where x = max_list and L = "(LFSL \ (=))" by trp_prover text \Use @{command print_theorems} to show all theorems. Here's the correctness theorem:\ lemma "(LFS \ (=)) max_list max_fset" by (trp_hints_resolve max_fset_related') lemma [trp_in_dom]: "(LFSR \ (=)) max_fset max_fset" by simp text \Transport from lists to sets.\ trp_term max_set :: "nat set \ nat" where x = max_list by trp_prover lemma "(LS \ (=)) max_list max_set" by (trp_hints_resolve max_set_related') text \The registration of symmetric equivalence rules is not done by default as of now, but that would not be a problem in principle.\ lemma list_fset_PER_sym [per_intro]: "(LFSR \\<^bsub>PER\<^esub> LFSL) sorted_list_of_fset fset_of_list" by (subst transport.partial_equivalence_rel_equivalence_right_left_iff_partial_equivalence_rel_equivalence_left_right) (fact list_fset_PER) text \Transport from finite sets to lists.\ trp_term max_list' :: "nat list \ nat" where x = max_fset by trp_prover lemma "(LFS\ \ (=)) max_fset max_list'" by (trp_hints_resolve max_list'_related') text \Transporting higher-order functions.\ lemma map_parametric [trp_in_dom]: "(((=) \ (=)) \ LSL \ LSL) map map" - unfolding LSL_def by (intro Dep_Fun_Rel_relI) simp + unfolding LSL_def by (intro Fun_Rel_relI) simp lemma [trp_uhint]: "P \ (=) \ P \ (=) \ (=)" by simp lemma [trp_uhint]: "P \ \ \ (=\<^bsub>P :: 'a \ bool\<^esub>) \ ((=) :: 'a \ _)" by simp (*sorted_list_of_fset requires a linorder, but in theory, we could use a different transport function to avoid that constraint*) trp_term map_set :: "('a :: linorder \ 'b) \ 'a set \ ('b :: linorder) set" where x = "map :: ('a :: linorder \ 'b) \ 'a list \ ('b :: linorder) list" by trp_prover lemma "(((=) \ (=)) \ LS \ LS) map map_set" by (trp_hints_resolve map_set_related') lemma filter_parametric [trp_in_dom]: "(((=) \ (\)) \ LSL \ LSL) filter filter" - unfolding LSL_def by (intro Dep_Fun_Rel_relI) simp + unfolding LSL_def by (intro Fun_Rel_relI) simp trp_term filter_set :: "('a :: linorder \ bool) \ 'a set \ 'a set" where x = "filter :: ('a :: linorder \ bool) \ 'a list \ 'a list" by trp_prover lemma "(((=) \ (=)) \ LS \ LS) filter filter_set" by (trp_hints_resolve filter_set_related') lemma append_parametric [trp_in_dom]: "(LSL \ LSL \ LSL) append append" - unfolding LSL_def by (intro Dep_Fun_Rel_relI) simp + unfolding LSL_def by (intro Fun_Rel_relI) simp trp_term append_set :: "('a :: linorder) set \ 'a set \ 'a set" where x = "append :: ('a :: linorder) list \ 'a list \ 'a list" by trp_prover lemma "(LS \ LS \ LS) append append_set" by (trp_hints_resolve append_set_related') text \The prototype also provides a simplified definition.\ lemma "append_set s s' \ set (sorted_list_of_set s) \ set (sorted_list_of_set s')" by (fact append_set_app_eq) lemma "finite s \ finite s' \ append_set s s' = s \ s'" by (auto simp: append_set_app_eq) end end 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 force + with assms have "\ x \\<^bsub>L\<^esub> \ x'" by fastforce 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 + by (intro t.order_equivalenceI 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,208 @@ \<^marker>\creator "Kevin Kappelmann"\ theory Transport_Typedef imports "HOL-Library.FSet" Transport_Typedef_Base Transport_Prototype Transport_Syntax + HOL_Alignment_Functions 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]: "(\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 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) fastforce + by (intro 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) + by (intro 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 + by (intro 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] + using transfer_raw[simplified HOL_fun_alignment 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/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,106 @@ \<^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_Dep_Fun_Rel (R :: 'a \ 'a \ bool) (S :: 'a \ 'a \ 'b \ 'b \ bool) \ + ((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_Fun_Rel_rel" :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ + ('a \ 'b) \ 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) + ('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) + ('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_Fun_Rel_rel" :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ + ('a \ 'b) \ 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) + ('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) + ('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)" + "(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)" +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>))" + 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" + 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>))" + 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 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,276 +1,277 @@ \<^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>\" by (urule symmetric_on_if_le_pred_if_symmetric_on) (use assms in \urule Refl_Rel_symmetric_onI\, simp) 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 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) blast + using assms by (intro 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) + by (simp_all add: galois_prop.half_galois_prop_left_rel_inv_iff_half_galois_prop_right + mono_wrt_rel_eq_dep_mono_wrt_rel) 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 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 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 :: 'c \ bool) (R :: 'd \ 'c \ bool). (inflationary_on P R\ :: ('c \ 'd) \ bool) \ deflationary_on P R" - by simp_all + by (simp_all add: mono_wrt_rel_eq_dep_mono_wrt_rel) 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 "((_ 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" + 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" + 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 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" + 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_Base.thy b/thys/Transport/Transport/Functions/Transport_Functions_Base.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Base.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Base.thy @@ -1,470 +1,470 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Transport For Functions\ subsection \Basic Setup\ theory Transport_Functions_Base imports Monotone_Function_Relator Transport_Base begin paragraph \Summary\ text \Basic setup for closure proofs. We introduce locales for the syntax, the dependent relator, the non-dependent relator, the monotone dependent relator, and the monotone non-dependent relator.\ definition "flip2 f x1 x2 x3 x4 \ f x2 x1 x4 x3" lemma flip2_eq: "flip2 f x1 x2 x3 x4 = f x2 x1 x4 x3" unfolding flip2_def by simp lemma flip2_eq_rel_inv [simp]: "flip2 R x y = (R y x)\" by (intro ext) (simp only: flip2_eq rel_inv_iff_rel) lemma flip2_flip2_eq_self [simp]: "flip2 (flip2 f) = f" by (intro ext) (simp add: flip2_eq) lemma flip2_eq_flip2_iff_eq [iff]: "flip2 f = flip2 g \ f = g" unfolding flip2_def by (intro iffI ext) (auto dest: fun_cong) paragraph \Dependent Function Relator\ locale transport_Dep_Fun_Rel_syntax = t1 : transport L1 R1 l1 r1 + dfro1 : hom_Dep_Fun_Rel_orders L1 L2 + dfro2 : hom_Dep_Fun_Rel_orders R1 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 :: "'a2 \ 'a1 \ 'b1 \ 'b2" and r2 :: "'a1 \ 'a2 \ 'b2 \ 'b1" begin notation L1 (infix "\\<^bsub>L1\<^esub>" 50) notation R1 (infix "\\<^bsub>R1\<^esub>" 50) notation t1.ge_left (infix "\\<^bsub>L1\<^esub>" 50) notation t1.ge_right (infix "\\<^bsub>R1\<^esub>" 50) notation t1.left_Galois (infix "\<^bsub>L1\<^esub>\" 50) notation t1.ge_Galois_left (infix "\\<^bsub>L1\<^esub>" 50) notation t1.right_Galois (infix "\<^bsub>R1\<^esub>\" 50) notation t1.ge_Galois_right (infix "\\<^bsub>R1\<^esub>" 50) notation t1.right_ge_Galois (infix "\<^bsub>R1\<^esub>\" 50) notation t1.Galois_right (infix "\\<^bsub>R1\<^esub>" 50) notation t1.left_ge_Galois (infix "\<^bsub>L1\<^esub>\" 50) notation t1.Galois_left (infix "\\<^bsub>L1\<^esub>" 50) notation t1.unit ("\\<^sub>1") notation t1.counit ("\\<^sub>1") notation L2 ("(\\<^bsub>L2 (_) (_)\<^esub>)" 50) notation R2 ("(\\<^bsub>R2 (_) (_)\<^esub>)" 50) notation dfro1.right_infix ("(_) \\<^bsub>L2 (_) (_)\<^esub> (_)" [51,51,51,51] 50) notation dfro2.right_infix ("(_) \\<^bsub>R2 (_) (_)\<^esub> (_)" [51,51,51,51] 50) notation dfro1.o.ge_right ("(\\<^bsub>L2 (_) (_)\<^esub>)" 50) notation dfro2.o.ge_right ("(\\<^bsub>R2 (_) (_)\<^esub>)" 50) notation dfro1.ge_right_infix ("(_) \\<^bsub>L2 (_) (_)\<^esub> (_)" [51,51,51,51] 50) notation dfro2.ge_right_infix ("(_) \\<^bsub>R2 (_) (_)\<^esub> (_)" [51,51,51,51] 50) notation l2 ("l2\<^bsub>(_) (_)\<^esub>") notation r2 ("r2\<^bsub>(_) (_)\<^esub>") sublocale t2 : transport "(\\<^bsub>L2 x (r1 x')\<^esub>)" "(\\<^bsub>R2 (l1 x) x'\<^esub>)" "l2\<^bsub>x' x\<^esub>" "r2\<^bsub>x x'\<^esub>" for x x' . notation t2.left_Galois ("(\<^bsub>L2 (_) (_)\<^esub>\)" 50) notation t2.right_Galois ("(\<^bsub>R2 (_) (_)\<^esub>\)" 50) abbreviation "left2_Galois_infix y x x' y' \ (\<^bsub>L2 x x'\<^esub>\) y y'" notation left2_Galois_infix ("(_) \<^bsub>L2 (_) (_)\<^esub>\ (_)" [51,51,51,51] 50) abbreviation "right2_Galois_infix y' x x' y \ (\<^bsub>R2 x x'\<^esub>\) y' y" notation right2_Galois_infix ("(_) \<^bsub>R2 (_) (_)\<^esub>\ (_)" [51,51,51,51] 50) notation t2.ge_Galois_left ("(\\<^bsub>L2 (_) (_)\<^esub>)" 50) notation t2.ge_Galois_right ("(\\<^bsub>R2 (_) (_)\<^esub>)" 50) abbreviation (input) "ge_Galois_left_left2_infix y' x x' y \ (\\<^bsub>L2 x x'\<^esub>) y' y" notation ge_Galois_left_left2_infix ("(_) \\<^bsub>L2 (_) (_)\<^esub> (_)" [51,51,51,51] 50) abbreviation (input) "ge_Galois_left_right2_infix y x x' y' \ (\\<^bsub>R2 x x'\<^esub>) y y'" notation ge_Galois_left_right2_infix ("(_) \\<^bsub>R2 (_) (_)\<^esub> (_)" [51,51,51,51] 50) notation t2.right_ge_Galois ("(\<^bsub>R2 (_) (_)\<^esub>\)" 50) notation t2.left_ge_Galois ("(\<^bsub>L2 (_) (_)\<^esub>\)" 50) abbreviation "left2_ge_Galois_left_infix y x x' y' \ (\<^bsub>L2 x x'\<^esub>\) y y'" notation left2_ge_Galois_left_infix ("(_) \<^bsub>L2 (_) (_)\<^esub>\ (_)" [51,51,51,51] 50) abbreviation "right2_ge_Galois_left_infix y' x x' y \ (\<^bsub>R2 x x'\<^esub>\) y' y" notation right2_ge_Galois_left_infix ("(_) \<^bsub>R2 (_) (_)\<^esub>\ (_)" [51,51,51,51] 50) notation t2.Galois_right ("(\\<^bsub>R2 (_) (_)\<^esub>)" 50) notation t2.Galois_left ("(\\<^bsub>L2 (_) (_)\<^esub>)" 50) abbreviation (input) "Galois_left2_infix y' x x' y \ (\\<^bsub>L2 x x'\<^esub>) y' y" notation Galois_left2_infix ("(_) \\<^bsub>L2 (_) (_)\<^esub> (_)" [51,51,51,51] 50) abbreviation (input) "Galois_right2_infix y x x' y' \ (\\<^bsub>R2 x x'\<^esub>) y y'" notation Galois_right2_infix ("(_) \\<^bsub>R2 (_) (_)\<^esub> (_)" [51,51,51,51] 50) abbreviation "t2_unit x x' \ t2.unit x' x" notation t2_unit ("\\<^bsub>2 (_) (_)\<^esub>") abbreviation "t2_counit x x' \ t2.counit x' x" notation t2_counit ("\\<^bsub>2 (_) (_)\<^esub>") end locale transport_Dep_Fun_Rel = transport_Dep_Fun_Rel_syntax 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 :: "'a2 \ 'a1 \ 'b1 \ 'b2" and r2 :: "'a1 \ 'a2 \ 'b2 \ 'b1" begin -definition "L \ [x1 x2 \ (\\<^bsub>L1\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)" +definition "L \ (x1 x2 \ (\\<^bsub>L1\<^esub>)) \ (\\<^bsub>L2 x1 x2\<^esub>)" -lemma left_rel_eq_Dep_Fun_Rel: "L = ([x1 x2 \ (\\<^bsub>L1\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>))" +lemma left_rel_eq_Dep_Fun_Rel: "L = ((x1 x2 \ (\\<^bsub>L1\<^esub>)) \ (\\<^bsub>L2 x1 x2\<^esub>))" unfolding L_def .. -definition "l \ ([x' : r1] \ l2 x')" +definition "l \ ((x' : r1) \ l2 x')" -lemma left_eq_dep_fun_map: "l = ([x' : r1] \ l2 x')" +lemma left_eq_dep_fun_map: "l = ((x' : r1) \ l2 x')" unfolding l_def .. lemma left_eq [simp]: "l f x' = l2\<^bsub>x' (r1 x')\<^esub> (f (r1 x'))" unfolding left_eq_dep_fun_map by simp context begin interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . abbreviation "R \ flip.L" abbreviation "r \ flip.l" -lemma right_rel_eq_Dep_Fun_Rel: "R = ([x1' x2' \ (\\<^bsub>R1\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>))" +lemma right_rel_eq_Dep_Fun_Rel: "R = ((x1' x2' \ (\\<^bsub>R1\<^esub>)) \ (\\<^bsub>R2 x1' x2'\<^esub>))" unfolding flip.L_def .. -lemma right_eq_dep_fun_map: "r = ([x : l1] \ r2 x)" +lemma right_eq_dep_fun_map: "r = ((x : l1) \ r2 x)" unfolding flip.l_def .. end lemma right_eq [simp]: "r g x = r2\<^bsub>x (l1 x)\<^esub> (g (l1 x))" unfolding right_eq_dep_fun_map by simp lemmas transport_defs = left_rel_eq_Dep_Fun_Rel left_eq_dep_fun_map right_rel_eq_Dep_Fun_Rel right_eq_dep_fun_map 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 "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ f x1 \\<^bsub>L2 x1 x2\<^esub> f' x2" shows "f \\<^bsub>L\<^esub> f'" unfolding left_rel_eq_Dep_Fun_Rel using assms by blast lemma left_relE [elim]: assumes "f \\<^bsub>L\<^esub> f'" and "x1 \\<^bsub>L1\<^esub> x2" obtains "f x1 \\<^bsub>L2 x1 x2\<^esub> f' x2" using assms unfolding left_rel_eq_Dep_Fun_Rel by blast interpretation flip_inv : transport_Dep_Fun_Rel "(\\<^bsub>R1\<^esub>)" "(\\<^bsub>L1\<^esub>)" r1 l1 "flip2 R2" "flip2 L2" r2 l2 . lemma flip_inv_right_eq_ge_left: "flip_inv.R = (\\<^bsub>L\<^esub>)" unfolding left_rel_eq_Dep_Fun_Rel flip_inv.right_rel_eq_Dep_Fun_Rel by (simp only: rel_inv_Dep_Fun_Rel_rel_eq flip2_eq_rel_inv[symmetric, of "L2"]) interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma flip_inv_left_eq_ge_right: "flip_inv.L \ (\\<^bsub>R\<^esub>)" unfolding flip.flip_inv_right_eq_ge_left . subparagraph \Useful Rewritings for Dependent Relation\ lemma left_rel2_unit_eqs_left_rel2I: assumes "\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 "\x. x \\<^bsub>L1\<^esub> x \ (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) \ (\\<^bsub>L2 x x\<^esub>)" and "x \\<^bsub>L1\<^esub> x" and "x \\<^bsub>L1\<^esub> \\<^sub>1 x" shows "(\\<^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>)" using assms by (auto intro!: antisym) lemma left2_eq_if_bi_related_if_monoI: - assumes mono_L2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" + assumes mono_L2: "((x1 x2 \ (\\<^bsub>L1\<^esub>)) \\<^sub>m (x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3) \ (\)) L2" and "x1 \\<^bsub>L1\<^esub> x2" and "x1 \\<^bsub>L1\<^esub> x3" and "x2 \\<^bsub>L1\<^esub> x4" and trans_L1: "transitive (\\<^bsub>L1\<^esub>)" shows "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x3 x4\<^esub>)" proof (intro antisym) from \x1 \\<^bsub>L1\<^esub> x3\ \x2 \\<^bsub>L1\<^esub> x4\ have "x3 \\<^bsub>L1\<^esub> x1" "x2 \\<^bsub>L1\<^esub> x4" by auto with \x1 \\<^bsub>L1\<^esub> x2\ mono_L2 show "(\\<^bsub>L2 x1 x2\<^esub>) \ (\\<^bsub>L2 x3 x4\<^esub>)" by blast from \x1 \\<^bsub>L1\<^esub> x3\ \x2 \\<^bsub>L1\<^esub> x4\ have "x1 \\<^bsub>L1\<^esub> x3" "x4 \\<^bsub>L1\<^esub> x2" by auto moreover from \x3 \\<^bsub>L1\<^esub> x1\ \x1 \\<^bsub>L1\<^esub> x2\ \x2 \\<^bsub>L1\<^esub> x4\ have "x3 \\<^bsub>L1\<^esub> x4" using trans_L1 by blast ultimately show "(\\<^bsub>L2 x3 x4\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" using mono_L2 by blast qed end paragraph \Function Relator\ locale transport_Fun_Rel_syntax = tdfrs : transport_Dep_Fun_Rel_syntax 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 :: "'b1 \ 'b1 \ bool" and R2 :: "'b2 \ 'b2 \ bool" and l2 :: "'b1 \ 'b2" and r2 :: "'b2 \ 'b1" begin notation L1 (infix "\\<^bsub>L1\<^esub>" 50) notation R1 (infix "\\<^bsub>R1\<^esub>" 50) notation tdfrs.t1.ge_left (infix "\\<^bsub>L1\<^esub>" 50) notation tdfrs.t1.ge_right (infix "\\<^bsub>R1\<^esub>" 50) notation tdfrs.t1.left_Galois (infix "\<^bsub>L1\<^esub>\" 50) notation tdfrs.t1.ge_Galois_left (infix "\\<^bsub>L1\<^esub>" 50) notation tdfrs.t1.right_Galois (infix "\<^bsub>R1\<^esub>\" 50) notation tdfrs.t1.ge_Galois_right (infix "\\<^bsub>R1\<^esub>" 50) notation tdfrs.t1.right_ge_Galois (infix "\<^bsub>R1\<^esub>\" 50) notation tdfrs.t1.Galois_right (infix "\\<^bsub>R1\<^esub>" 50) notation tdfrs.t1.left_ge_Galois (infix "\<^bsub>L1\<^esub>\" 50) notation tdfrs.t1.Galois_left (infix "\\<^bsub>L1\<^esub>" 50) notation tdfrs.t1.unit ("\\<^sub>1") notation tdfrs.t1.counit ("\\<^sub>1") notation L2 (infix "\\<^bsub>L2\<^esub>" 50) notation R2 (infix "\\<^bsub>R2\<^esub>" 50) notation tdfrs.t2.ge_left (infix "\\<^bsub>L2\<^esub>" 50) notation tdfrs.t2.ge_right (infix "\\<^bsub>R2\<^esub>" 50) notation tdfrs.t2.left_Galois (infix "\<^bsub>L2\<^esub>\" 50) notation tdfrs.t2.ge_Galois_left (infix "\\<^bsub>L2\<^esub>" 50) notation tdfrs.t2.right_Galois (infix "\<^bsub>R2\<^esub>\" 50) notation tdfrs.t2.ge_Galois_right (infix "\\<^bsub>R2\<^esub>" 50) notation tdfrs.t2.right_ge_Galois (infix "\<^bsub>R2\<^esub>\" 50) notation tdfrs.t2.Galois_right (infix "\\<^bsub>R2\<^esub>" 50) notation tdfrs.t2.left_ge_Galois (infix "\<^bsub>L2\<^esub>\" 50) notation tdfrs.t2.Galois_left (infix "\\<^bsub>L2\<^esub>" 50) notation tdfrs.t2.unit ("\\<^sub>2") notation tdfrs.t2.counit ("\\<^sub>2") end locale transport_Fun_Rel = transport_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 :: "'b1 \ 'b1 \ bool" and R2 :: "'b2 \ 'b2 \ bool" and l2 :: "'b1 \ 'b2" and r2 :: "'b2 \ 'b1" begin (*FIXME: we have to repeat the Galois syntax here since tdfr already contains a Galois instance, blocking a galois sublocale interpretation here*) notation tdfr.L ("L") notation tdfr.R ("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.left_Galois (infix "\<^bsub>L\<^esub>\" 50) notation tdfr.ge_Galois_left (infix "\\<^bsub>L\<^esub>" 50) notation tdfr.right_Galois (infix "\<^bsub>R\<^esub>\" 50) notation tdfr.ge_Galois_right (infix "\\<^bsub>R\<^esub>" 50) notation tdfr.right_ge_Galois (infix "\<^bsub>R\<^esub>\" 50) notation tdfr.Galois_right (infix "\\<^bsub>R\<^esub>" 50) notation tdfr.left_ge_Galois (infix "\<^bsub>L\<^esub>\" 50) notation tdfr.Galois_left (infix "\\<^bsub>L\<^esub>" 50) notation tdfr.unit ("\") notation tdfr.counit ("\") lemma left_rel_eq_Fun_Rel: "(\\<^bsub>L\<^esub>) = ((\\<^bsub>L1\<^esub>) \ (\\<^bsub>L2\<^esub>))" - unfolding tdfr.left_rel_eq_Dep_Fun_Rel by simp + by (urule tdfr.left_rel_eq_Dep_Fun_Rel) -lemma left_eq_fun_map: "l = (r1 \ l2)" +lemma left_eq_fun_map: "l = (r1 \ l2)" by (intro ext) simp interpretation flip : transport_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma right_rel_eq_Fun_Rel: "(\\<^bsub>R\<^esub>) = ((\\<^bsub>R1\<^esub>) \ (\\<^bsub>R2\<^esub>))" unfolding flip.left_rel_eq_Fun_Rel .. -lemma right_eq_fun_map: "r = (l1 \ r2)" +lemma right_eq_fun_map: "r = (l1 \ r2)" unfolding flip.left_eq_fun_map .. lemmas transport_defs = left_rel_eq_Fun_Rel right_rel_eq_Fun_Rel left_eq_fun_map right_eq_fun_map end paragraph \Monotone Dependent Function Relator\ locale transport_Mono_Dep_Fun_Rel = 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 :: "'a2 \ 'a1 \ 'b1 \ 'b2" and r2 :: "'a1 \ 'a2 \ 'b2 \ 'b1" begin definition "L \ tdfr.L\<^sup>\" lemma left_rel_eq_tdfr_left_Refl_Rel: "L = tdfr.L\<^sup>\" unfolding L_def .. -lemma left_rel_eq_Mono_Dep_Fun_Rel: "L = ([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\ (\\<^bsub>L2 x1 x2\<^esub>))" +lemma left_rel_eq_Mono_Dep_Fun_Rel: "L = ((x1 x2 \ (\\<^bsub>L1\<^esub>)) \\ (\\<^bsub>L2 x1 x2\<^esub>))" unfolding left_rel_eq_tdfr_left_Refl_Rel tdfr.left_rel_eq_Dep_Fun_Rel by simp lemma left_rel_eq_tdfr_left_rel_if_reflexive_on: assumes "reflexive_on (in_field tdfr.L) tdfr.L" shows "L = tdfr.L" unfolding left_rel_eq_tdfr_left_Refl_Rel using assms by (rule Refl_Rel_eq_self_if_reflexive_on) abbreviation "l \ tdfr.l" lemma left_eq_tdfr_left: "l = tdfr.l" .. interpretation flip : transport_Mono_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . abbreviation "R \ flip.L" lemma right_rel_eq_tdfr_right_Refl_Rel: "R = tdfr.R\<^sup>\" unfolding flip.left_rel_eq_tdfr_left_Refl_Rel .. -lemma right_rel_eq_Mono_Dep_Fun_Rel: "R = ([y1 y2 \ (\\<^bsub>R1\<^esub>)] \\ (\\<^bsub>R2 y1 y2\<^esub>))" +lemma right_rel_eq_Mono_Dep_Fun_Rel: "R = ((y1 y2 \ (\\<^bsub>R1\<^esub>)) \\ (\\<^bsub>R2 y1 y2\<^esub>))" unfolding flip.left_rel_eq_Mono_Dep_Fun_Rel .. lemma right_rel_eq_tdfr_right_rel_if_reflexive_on: assumes "reflexive_on (in_field tdfr.R) tdfr.R" shows "R = tdfr.R" using assms by (rule flip.left_rel_eq_tdfr_left_rel_if_reflexive_on) abbreviation "r \ tdfr.r" lemma right_eq_tdfr_right: "r = tdfr.r" .. lemmas transport_defs = left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel 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) end paragraph \Monotone Function Relator\ locale transport_Mono_Fun_Rel = transport_Fun_Rel_syntax L1 R1 l1 r1 L2 R2 l2 r2 + tfr : transport_Fun_Rel L1 R1 l1 r1 L2 R2 l2 r2 + tpdfr : transport_Mono_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 :: "'b1 \ 'b1 \ bool" and R2 :: "'b2 \ 'b2 \ bool" and l2 :: "'b1 \ 'b2" and r2 :: "'b2 \ 'b1" begin (*FIXME: we have to repeat the Galois syntax here since tdfr already contains a Galois instance, blocking a galois sublocale interpretation here*) notation tpdfr.L ("L") notation tpdfr.R ("R") abbreviation "l \ tpdfr.l" abbreviation "r \ tpdfr.r" notation tpdfr.L (infix "\\<^bsub>L\<^esub>" 50) notation tpdfr.R (infix "\\<^bsub>R\<^esub>" 50) notation tpdfr.ge_left (infix "\\<^bsub>L\<^esub>" 50) notation tpdfr.ge_right (infix "\\<^bsub>R\<^esub>" 50) notation tpdfr.left_Galois (infix "\<^bsub>L\<^esub>\" 50) notation tpdfr.ge_Galois_left (infix "\\<^bsub>L\<^esub>" 50) notation tpdfr.right_Galois (infix "\<^bsub>R\<^esub>\" 50) notation tpdfr.ge_Galois_right (infix "\\<^bsub>R\<^esub>" 50) notation tpdfr.right_ge_Galois (infix "\<^bsub>R\<^esub>\" 50) notation tpdfr.Galois_right (infix "\\<^bsub>R\<^esub>" 50) notation tpdfr.left_ge_Galois (infix "\<^bsub>L\<^esub>\" 50) notation tpdfr.Galois_left (infix "\\<^bsub>L\<^esub>" 50) notation tpdfr.unit ("\") notation tpdfr.counit ("\") lemma left_rel_eq_Mono_Fun_Rel: "(\\<^bsub>L\<^esub>) = ((\\<^bsub>L1\<^esub>) \\ (\\<^bsub>L2\<^esub>))" unfolding tpdfr.left_rel_eq_Mono_Dep_Fun_Rel by simp -lemma left_eq_fun_map: "l = (r1 \ l2)" +lemma left_eq_fun_map: "l = (r1 \ l2)" unfolding tfr.left_eq_fun_map .. interpretation flip : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma right_rel_eq_Mono_Fun_Rel: "(\\<^bsub>R\<^esub>) = ((\\<^bsub>R1\<^esub>) \\ (\\<^bsub>R2\<^esub>))" unfolding flip.left_rel_eq_Mono_Fun_Rel .. -lemma right_eq_fun_map: "r = (l1 \ r2)" +lemma right_eq_fun_map: "r = (l1 \ r2)" unfolding flip.left_eq_fun_map .. lemmas transport_defs = tpdfr.transport_defs 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" + 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>)" + ((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>)" + ((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>)" + 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>)" + 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" + 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" + ((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>)" + 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 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 + 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" + 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>)" + ((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>)" + ((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>)" + 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 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>)" + 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" + 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" + ((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>)" + 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 + 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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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>)" + ((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" + ((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" + ((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' \ (\\<^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 "((_ 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,352 +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) + 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 (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 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) + 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 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 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) + 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 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>)" + (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>)" + (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_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>)" + (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_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>)" + (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>)" + (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>)" + (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 "((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" + (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>)" + (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 "((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" + (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>)" + (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 \ (\\<^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" + 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_Property.thy b/thys/Transport/Transport/Functions/Transport_Functions_Galois_Property.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Galois_Property.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Galois_Property.thy @@ -1,474 +1,472 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Property\ theory Transport_Functions_Galois_Property imports Transport_Functions_Monotone begin paragraph \Dependent Function Relator\ context transport_Dep_Fun_Rel begin context begin interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma left_right_rel_if_left_rel_rightI: 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 refl_R1: "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and half_galois_prop_left2: "\x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)) (l2\<^bsub> x' (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and R2_le1: "\x'. x' \\<^bsub>R1\<^esub> x' \ (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) \ (\\<^bsub>R2 x' x'\<^esub>)" and R2_le2: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and ge_L2_r2_le2: "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (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_R2: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" and "g \\<^bsub>R\<^esub> g" and "f \\<^bsub>L\<^esub> r g" shows "l f \\<^bsub>R\<^esub> g" proof (rule flip.left_relI) fix x1' x2' assume [iff]: "x1' \\<^bsub>R1\<^esub> x2'" with refl_R1 have [iff]: "x1' \\<^bsub>R1\<^esub> x1'" by auto with mono_r1 half_galois_prop_left1 have [iff]: "\\<^sub>1 x1' \\<^bsub>R1\<^esub> x1'" by (intro t1.counit_rel_if_right_rel_if_half_galois_prop_left_if_mono_wrt_rel) with refl_R1 have "\\<^sub>1 x1' \\<^bsub>R1\<^esub> \\<^sub>1 x1'" by blast with \g \\<^bsub>R\<^esub> g\ have "g (\\<^sub>1 x1') \\<^bsub>R2 (\\<^sub>1 x1') (\\<^sub>1 x1')\<^esub> g (\\<^sub>1 x1')" by blast with R2_le2 have "g (\\<^sub>1 x1') \\<^bsub>R2 (\\<^sub>1 x1') x1'\<^esub> g (\\<^sub>1 x1')" by blast let ?x1 = "r1 x1'" from \f \\<^bsub>L\<^esub> r g\ \x1' \\<^bsub>R1\<^esub> x1'\ have "f ?x1 \\<^bsub>L2 ?x1 ?x1\<^esub> r g ?x1" using mono_r1 by blast then have "f ?x1 \\<^bsub>L2 ?x1 ?x1\<^esub> r2\<^bsub>?x1 (\\<^sub>1 x1')\<^esub> (g (\\<^sub>1 x1'))" by simp with ge_L2_r2_le2 have "f ?x1 \\<^bsub>L2 ?x1 ?x1\<^esub> r2\<^bsub>?x1 x1'\<^esub> (g (\\<^sub>1 x1'))" using \_ \\<^bsub>R2 (\\<^sub>1 x1') x1'\<^esub> g (\\<^sub>1 x1')\ by blast with half_galois_prop_left2 have "l2\<^bsub> x1' ?x1\<^esub> (f ?x1) \\<^bsub>R2 (\\<^sub>1 x1') x1'\<^esub> g (\\<^sub>1 x1')" using \_ \\<^bsub>R2 (\\<^sub>1 x1') x1'\<^esub> g (\\<^sub>1 x1')\ by auto moreover from \g \\<^bsub>R\<^esub> g\ \\\<^sub>1 x1' \\<^bsub>R1\<^esub> x1'\ have "... \\<^bsub>R2 (\\<^sub>1 x1') x1'\<^esub> g x1'" by blast ultimately have "l2\<^bsub> x1' ?x1\<^esub> (f ?x1) \\<^bsub>R2 (\\<^sub>1 x1') x1'\<^esub> g x1'" using trans_R2 by blast with R2_le1 R2_le2 have "l2\<^bsub> x1' ?x1\<^esub> (f ?x1) \\<^bsub>R2 x1' x2'\<^esub> g x1'" by blast moreover from \g \\<^bsub>R\<^esub> g\ \x1' \\<^bsub>R1\<^esub> x2'\ have "... \\<^bsub>R2 x1' x2'\<^esub> g x2'" by blast ultimately have "l2\<^bsub> x1' ?x1\<^esub> (f ?x1) \\<^bsub>R2 x1' x2'\<^esub> g x2'" using trans_R2 by blast then show "l f x1' \\<^bsub>R2 x1' x2'\<^esub> g x2'" by simp qed lemma left_right_rel_if_left_rel_right_ge_left2_assmI: assumes mono_r1: "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" - and "([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) + and "((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 "x' \\<^bsub>R1\<^esub> x'" and "in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y'" shows "(\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y') \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub> y')" - using dep_mono_wrt_relD[OF mono_r1 \x' \\<^bsub>R1\<^esub> x'\] assms(2-4,6) - by (blast dest!: t1.half_galois_prop_leftD) + using mono_wrt_relD[OF mono_r1] assms(2-) by blast 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.t1.counit \ \\<^sub>1" and "\R x y. (flip2 R x y)\ \ R y x" and "\R. in_dom R\ \ in_codom R" and "\R x1 x2. in_codom (flip2 R x1 x2) \ in_dom (R x2 x1)" and "\R S. (R\ \\<^sub>m S\) \ (R \\<^sub>m S)" and "\R S x1 x2 x1' x2'. (flip2 R x1 x2 \<^sub>h\ flip2 S x1' x2') \ (S x2' x1' \\<^sub>h R x2 x1)\" and "\R S. (R\ \<^sub>h\ S\) \ (S \\<^sub>h R)\" and "\x1 x2 x3 x4. flip2 L2 x1 x2 \ flip2 L2 x3 x4 \ (\\<^bsub>L2 x2 x1\<^esub>) \ (\\<^bsub>L2 x4 x3\<^esub>)" and "\(R :: 'z \ 'z \ bool) (P :: 'z \ bool). reflexive_on P R\ \ reflexive_on P R" and "\R x1 x2. transitive (flip2 R x1 x2 :: 'z \ 'z \ bool) \ transitive (R x2 x1)" - and "\x x. ([in_dom (\\<^bsub>L2 x' (\\<^sub>1 x')\<^esub>)] \ flip2 R2 (l1 x') (l1 x')) - \ ([in_dom (\\<^bsub>L2 x' (\\<^sub>1 x')\<^esub>)] \ (\\<^bsub>R2 (l1 x') (l1 x')\<^esub>))\" + and "\x x. ((in_dom (\\<^bsub>L2 x' (\\<^sub>1 x')\<^esub>)) \ flip2 R2 (l1 x') (l1 x')) + \ ((in_dom (\\<^bsub>L2 x' (\\<^sub>1 x')\<^esub>)) \ (\\<^bsub>R2 (l1 x') (l1 x')\<^esub>))\" by (simp_all add: flip_inv_left_eq_ge_right flip_inv_right_eq_ge_left t1.flip_counit_eq_unit - galois_prop.rel_inv_half_galois_prop_right_eq_half_galois_prop_left_rel_inv) + galois_prop.rel_inv_half_galois_prop_right_eq_half_galois_prop_left_rel_inv + mono_wrt_rel_eq_dep_mono_wrt_rel Fun_Rel_pred_eq_Dep_Fun_Rel_pred) lemma left_rel_right_if_left_right_relI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) \\<^sub>h (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (r2\<^bsub>x (l1 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 y. x \\<^bsub>L1\<^esub> x \ in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (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 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "f \\<^bsub>L\<^esub> f" and "l f \\<^bsub>R\<^esub> g" shows "f \\<^bsub>L\<^esub> r g" using assms by (intro flip_inv.left_right_rel_if_left_rel_rightI[simplified rel_inv_iff_rel]) lemma left_rel_right_if_left_right_rel_le_right2_assmI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>))\ r1 l1" - and "([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 "((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>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" and "x \\<^bsub>L1\<^esub> x" and "in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y" shows "(\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y)" using assms by (intro flip_inv.left_right_rel_if_left_rel_right_ge_left2_assmI [simplified rel_inv_iff_rel]) auto end lemma left_rel_right_iff_left_right_relI: 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' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)) (l2\<^bsub> x' (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) \\<^sub>h (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (r2\<^bsub>x (l1 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 y. x \\<^bsub>L1\<^esub> x \ in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y)" and "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (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>)" and "f \\<^bsub>L\<^esub> f" and "g \\<^bsub>R\<^esub> g" shows "f \\<^bsub>L\<^esub> r g \ l f \\<^bsub>R\<^esub> g" using assms by (intro iffI left_right_rel_if_left_rel_rightI) (auto intro!: left_rel_right_if_left_right_relI) lemma half_galois_prop_left2_if_half_galois_prop_left2_if_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x' \\<^bsub>R1\<^esub> x'" shows "((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)) (l2\<^bsub> x' (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" using assms by (auto intro: t1.right_left_Galois_if_right_relI) lemma half_galois_prop_right2_if_half_galois_prop_right2_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 "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>h (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x \\<^bsub>L1\<^esub> x" shows "((\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) \\<^sub>h (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (r2\<^bsub>x (l1 x)\<^esub>)" by (auto intro!: assms t1.left_Galois_left_if_left_relI) lemma left_rel_right_iff_left_right_relI': 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 galois_prop2: "\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>)" + ((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>)" + ((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>)" and "f \\<^bsub>L\<^esub> f" and "g \\<^bsub>R\<^esub> g" shows "f \\<^bsub>L\<^esub> r g \ l f \\<^bsub>R\<^esub> g" proof - from galois_prop2 have "((\\<^bsub>L2 x (r1 x')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" "((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>h (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" if "x \<^bsub>L1\<^esub>\ x'" for x x' using \x \<^bsub>L1\<^esub>\ x'\ by blast+ with assms show ?thesis by (intro left_rel_right_iff_left_right_relI left_right_rel_if_left_rel_right_ge_left2_assmI left_rel_right_if_left_right_rel_le_right2_assmI half_galois_prop_left2_if_half_galois_prop_left2_if_left_GaloisI half_galois_prop_right2_if_half_galois_prop_right2_if_left_GaloisI) auto qed lemma left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI: 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 antimono_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" + "((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" shows "\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>)" proof - fix x1 x2 assume "x1 \\<^bsub>L1\<^esub> x2" with galois_conn1 refl_L1 have "x1 \\<^bsub>L1\<^esub> x1" "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast intro: t1.rel_unit_if_left_rel_if_half_galois_prop_right_if_mono_wrt_rel)+ moreover with refl_L1 have "x2 \\<^bsub>L1\<^esub> x2" "\\<^sub>1 x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by auto moreover note dep_mono_wrt_relD[OF antimono_L2 \x1 \\<^bsub>L1\<^esub> x2\] and dep_mono_wrt_relD[OF antimono_L2 \x1 \\<^bsub>L1\<^esub> x1\] ultimately show "(\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" "(\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" using \x1 \\<^bsub>L1\<^esub> x2\ by auto qed lemma left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI: 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 mono_R2: - "([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" + "((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" shows "\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>)" proof - fix x1' x2' assume "x1' \\<^bsub>R1\<^esub> x2'" with galois_conn1 refl_R1 have "x2' \\<^bsub>R1\<^esub> x2'" "\\<^sub>1 x1' \\<^bsub>R1\<^esub> x1'" by (blast intro: t1.counit_rel_if_right_rel_if_half_galois_prop_left_if_mono_wrt_rel)+ moreover with refl_R1 have "x1' \\<^bsub>R1\<^esub> x1'" "\\<^sub>1 x1' \\<^bsub>R1\<^esub> \\<^sub>1 x1'" by auto moreover note dep_mono_wrt_relD[OF mono_R2 \\\<^sub>1 x1' \\<^bsub>R1\<^esub> x1'\] and dep_mono_wrt_relD[OF mono_R2 \x1' \\<^bsub>R1\<^esub> x1'\] ultimately show "(\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" "(\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" using \x1' \\<^bsub>R1\<^esub> x2'\ by auto qed corollary left_rel_right_iff_left_right_rel_if_monoI: 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 \ (\\<^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>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 "\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>)" + ((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>)" + ((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>)" and "f \\<^bsub>L\<^esub> f" and "g \\<^bsub>R\<^esub> g" shows "f \\<^bsub>L\<^esub> r g \ l f \\<^bsub>R\<^esub> g" using assms by (intro left_rel_right_iff_left_right_relI' left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI 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) end paragraph \Function Relator\ context transport_Fun_Rel begin corollary left_right_rel_if_left_rel_rightI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \<^sub>h\ (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>R2\<^esub>)" and "g \\<^bsub>R\<^esub> g" and "f \\<^bsub>L\<^esub> r g" shows "l f \\<^bsub>R\<^esub> g" using assms by (intro tdfr.left_right_rel_if_left_rel_rightI) simp_all corollary left_rel_right_if_left_right_relI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>h (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>L2\<^esub>)" and "f \\<^bsub>L\<^esub> f" and "l f \\<^bsub>R\<^esub> g" shows "f \\<^bsub>L\<^esub> r g" using assms by (intro tdfr.left_rel_right_if_left_right_relI) simp_all corollary left_rel_right_iff_left_right_relI: 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>)" and "f \\<^bsub>L\<^esub> f" and "g \\<^bsub>R\<^esub> g" shows "f \\<^bsub>L\<^esub> r g \ l f \\<^bsub>R\<^esub> g" using assms by (intro tdfr.left_rel_right_iff_left_right_relI) auto end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin lemma half_galois_prop_left_left_rightI: assumes "(tdfr.L \\<^sub>m tdfr.R) l" 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_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)) (l2\<^bsub> x' (r1 x')\<^esub>) (r2\<^bsub>(r1 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' y'. x' \\<^bsub>R1\<^esub> x' \ in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (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>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel using assms by (intro half_galois_prop_leftI[unfolded left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel] Refl_Rel_app_leftI[where ?f=l] tdfr.left_right_rel_if_left_rel_rightI) (auto elim!: galois_rel.left_GaloisE) lemma half_galois_prop_right_left_rightI: 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 "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) \\<^sub>h (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (r2\<^bsub>x (l1 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 y. x \\<^bsub>L1\<^esub> x \ in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (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 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel using assms by (intro half_galois_prop_rightI[unfolded left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel] Refl_Rel_app_rightI[where ?f=r] tdfr.left_rel_right_if_left_right_relI) (auto elim!: galois_rel.left_GaloisE in_codomE Refl_RelE intro!: in_fieldI) corollary galois_prop_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' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)) (l2\<^bsub> x' (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) \\<^sub>h (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (r2\<^bsub>x (l1 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 y. x \\<^bsub>L1\<^esub> x \ in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y)" and "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (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>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_propI half_galois_prop_left_left_rightI half_galois_prop_right_left_rightI) auto corollary galois_prop_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 galois_prop2: "\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>)" + ((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>)" + ((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" proof - from galois_prop2 have "((\\<^bsub>L2 x (r1 x')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" "((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>h (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" if "x \<^bsub>L1\<^esub>\ x'" for x x' using \x \<^bsub>L1\<^esub>\ x'\ by blast+ with assms show ?thesis by (intro galois_prop_left_rightI tdfr.left_right_rel_if_left_rel_right_ge_left2_assmI tdfr.left_rel_right_if_left_right_rel_le_right2_assmI tdfr.half_galois_prop_left2_if_half_galois_prop_left2_if_left_GaloisI tdfr.half_galois_prop_right2_if_half_galois_prop_right2_if_left_GaloisI) auto qed corollary galois_prop_left_right_if_mono_if_galois_propI: 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_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 \ (\\<^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>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 "\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>)" + ((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>)" + ((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_prop_left_rightI' 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) text \Note that we could further rewrite @{thm "galois_prop_left_right_if_mono_if_galois_propI"}, as we will do later for Galois connections, by applying @{thm "tdfr.mono_wrt_rel_leftI"} and @{thm "tdfr.mono_wrt_rel_rightI"} to the first premises. However, this is not really helpful here. Moreover, the resulting theorem will not result in a useful lemma for the flipped instance of @{locale transport_Dep_Fun_Rel} since @{thm "tdfr.mono_wrt_rel_leftI"} and @{thm "tdfr.mono_wrt_rel_rightI"} are not flipped dual but only flipped-inversed dual.\ end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin lemma half_galois_prop_left_left_rightI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>L2\<^esub>) \<^sub>h\ (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>R2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" - using assms - by (intro tpdfr.half_galois_prop_left_left_rightI tfr.mono_wrt_rel_leftI) + using assms by (intro tpdfr.half_galois_prop_left_left_rightI tfr.mono_wrt_rel_leftI) simp_all interpretation flip : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma half_galois_prop_right_left_rightI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "((\\<^bsub>L2\<^esub>) \\<^sub>h (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>L2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" - using assms - by (intro tpdfr.half_galois_prop_right_left_rightI flip.tfr.mono_wrt_rel_leftI) + using assms by (intro tpdfr.half_galois_prop_right_left_rightI flip.tfr.mono_wrt_rel_leftI) simp_all corollary galois_prop_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_propI half_galois_prop_left_left_rightI half_galois_prop_right_left_rightI) auto 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,859 @@ \<^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" + 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" + 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" + 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" + 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" + 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" + 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>)" + ((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>)" + ((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 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) + 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 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 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>)" + 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 \ (\\<^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" + 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 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" + "((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 : \) \\<^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" + 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 : \) \\<^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" + 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 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 : \) \\<^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" + 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 \ (\\<^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" + 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 \ (\\<^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" + 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 "((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" + 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>\))" + 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 rel_restrict_leftI 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 \ (\\<^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>\))" + 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>) + 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") + ((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>) + ((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 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 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: 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 + by (urule tdfr.Dep_Fun_Rel_left_Galois_if_left_GaloisI assms | simp)+ 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 + by (urule tdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI assms | simp)+ 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 + by (urule tdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI' assms | simp)+ 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 + by (urule tdfr.left_Galois_iff_Dep_Fun_Rel_left_GaloisI assms | simp)+ 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 + by (urule tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI assms + | simp)+ 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 + by (urule tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq assms | simp)+ 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" + 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" + 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 \ (\\<^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" + 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 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") + 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) + 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 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 : \) \\<^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") + 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 : \) \\<^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>" + 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 rel_restrict_leftI 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 \ (\\<^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" + 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 \ (\\<^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>" + 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 rel_restrict_leftI 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 \ (\\<^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" + 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 \ (\\<^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>" + 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 rel_restrict_leftI 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 "((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" + 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>" + 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 rel_restrict_leftI 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 \ (\\<^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>\))" + 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 "((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>\))" + 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>) + 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>" + ((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 + by (urule tpdfr.Dep_Fun_Rel_left_Galois_if_left_GaloisI assms | simp)+ 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 + by (urule tpdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI flip.tfr.mono_wrt_rel_leftI)+ + (urule assms | simp)+ 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 rel_restrict_leftI 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 + by (urule tpdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq assms | simp)+ end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions_Monotone.thy b/thys/Transport/Transport/Functions/Transport_Functions_Monotone.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Monotone.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Monotone.thy @@ -1,174 +1,171 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Monotonicity\ theory Transport_Functions_Monotone imports Transport_Functions_Base 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 . lemma mono_wrt_rel_leftI: assumes mono_r1: "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and mono_l2: "\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 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 ge_R2_l2_le2: "\x1' x2' y. x1' \\<^bsub>R1\<^esub> x2' \ in_codom (\\<^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>x2' (r1 x2')\<^esub> y)" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" -proof (intro dep_mono_wrt_relI flip.left_relI) +proof (intro mono_wrt_relI flip.left_relI) fix f1 f2 x1' x2' assume [iff]: "x1' \\<^bsub>R1\<^esub> x2'" with mono_r1 have "r1 x1' \\<^bsub>L1\<^esub> r1 x2'" (is "?x1 \\<^bsub>L1\<^esub> ?x2") by blast moreover assume "f1 \\<^bsub>L\<^esub> f2" ultimately have "f1 ?x1 \\<^bsub>L2 ?x1 ?x2\<^esub> f2 ?x2" (is "?y1 \\<^bsub>L2 ?x1 ?x2\<^esub> ?y2") by blast with mono_l2 have "l2\<^bsub>x2' ?x1\<^esub> ?y1 \\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub> l2\<^bsub>x2' ?x1\<^esub> ?y2" by blast with R2_le1 have "l2\<^bsub>x2' ?x1\<^esub> ?y1 \\<^bsub>R2 x1' x2'\<^esub> l2\<^bsub>x2' ?x1\<^esub> ?y2" by blast with R2_l2_le1 have "l2\<^bsub>x1' ?x1\<^esub> ?y1 \\<^bsub>R2 x1' x2'\<^esub> l2\<^bsub>x2' ?x1\<^esub> ?y2" using \?y1 \\<^bsub>L2 ?x1 ?x2\<^esub> _\ by blast with ge_R2_l2_le2 have "l2\<^bsub>x1' ?x1\<^esub> ?y1 \\<^bsub>R2 x1' x2'\<^esub> l2\<^bsub>x2' ?x2\<^esub> ?y2" using \_ \\<^bsub>L2 ?x1 ?x2\<^esub> ?y2\ by blast then show "l f1 x1' \\<^bsub>R2 x1' x2'\<^esub> l f2 x2'" by simp qed lemma mono_wrt_rel_left_in_dom_mono_left_assm: - assumes "([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>)" + assumes "(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 "transitive (\\<^bsub>R2 x1' x2'\<^esub>)" and "x1' \\<^bsub>R1\<^esub> x2'" and "in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) y" shows "(\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub> y) \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x1' (r1 x1')\<^esub> y)" using assms by blast lemma mono_wrt_rel_left_in_codom_mono_left_assm: - assumes "([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>)" + assumes "(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 "transitive (\\<^bsub>R2 x1' x2'\<^esub>)" and "x1' \\<^bsub>R1\<^esub> x2'" and "in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) y" shows "(\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub> y) \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub> y)" using assms by blast lemma mono_wrt_rel_left_if_transitiveI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" 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>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') 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>)" + (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>)" + (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 "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" using assms by (intro mono_wrt_rel_leftI mono_wrt_rel_left_in_dom_mono_left_assm mono_wrt_rel_left_in_codom_mono_left_assm) auto lemma mono_wrt_rel_left2_if_mono_wrt_rel_left2_if_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "\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>)" shows "\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>)" - using assms by (intro dep_mono_wrt_relI) fastforce + using assms by (intro mono_wrt_relI) fastforce interpretation flip_inv : transport_Dep_Fun_Rel "(\\<^bsub>R1\<^esub>)" "(\\<^bsub>L1\<^esub>)" r1 l1 "flip2 R2" "flip2 L2" r2 l2 rewrites "flip_inv.R \ (\\<^bsub>L\<^esub>)" and "flip_inv.L \ (\\<^bsub>R\<^esub>)" and "flip_inv.t1.counit \ \\<^sub>1" and "\R x y. (flip2 R x y)\ \ R y x" and "\R x1 x2. in_dom (flip2 R x1 x2) \ in_codom (R x2 x1)" and "\R x1 x2. in_codom (flip2 R x1 x2) \ in_dom (R x2 x1)" and "\R S. (R\ \\<^sub>m S\) \ (R \\<^sub>m S)" 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 x3 x4. flip2 L2 x1 x2 \ flip2 L2 x3 x4 \ (\\<^bsub>L2 x2 x1\<^esub>) \ (\\<^bsub>L2 x4 x3\<^esub>)" and "\x1' x2' y1 y2. flip_inv.dfro2.right_infix y1 x1' x2' \ flip_inv.dfro2.right_infix y2 x1' x2' \ (\\<^bsub>L2 x2' x1'\<^esub>) y1 \ (\\<^bsub>L2 x2' x1'\<^esub>) y2" - and "\P x1 x2. ([P] \ flip2 L2 x1 x2) \ ([P] \ (\\<^bsub>L2 x2 x1\<^esub>))" - and "\P R. ([P] \ R\) \ ([P] \ R)\" + and "\(P :: 'f \ bool) x1 x2. (P \ flip2 L2 x1 x2) \ (P \ (\\<^bsub>L2 x2 x1\<^esub>))" + and "\(P :: 'f \ bool) (R :: 'g \ 'g \ bool). (P \ R\) \ (P \ R)\" and "\x1 x2. transitive (flip2 L2 x1 x2) \ transitive (\\<^bsub>L2 x2 x1\<^esub>)" by (simp_all add: flip_inv_left_eq_ge_right flip_inv_right_eq_ge_left - t1.flip_counit_eq_unit del: rel_inv_iff_rel) + t1.flip_counit_eq_unit mono_wrt_rel_eq_dep_mono_wrt_rel Fun_Rel_pred_eq_Dep_Fun_Rel_pred + del: rel_inv_iff_rel) lemma mono_wrt_rel_rightI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" 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 "\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 "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_dom (\\<^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>x1 (l1 x1)\<^esub> y')" shows "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" using assms by (intro flip_inv.mono_wrt_rel_leftI[simplified rel_inv_iff_rel]) lemma mono_wrt_rel_right_if_transitiveI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" 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 "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^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 "\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 "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" using assms by (intro flip_inv.mono_wrt_rel_left_if_transitiveI [simplified rel_inv_iff_rel]) lemma mono_wrt_rel_right2_if_mono_wrt_rel_right2_if_left_GaloisI: assumes assms1: "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" 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>)" shows "\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>)" proof - show "((\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) \\<^sub>m (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>)" if "x1 \\<^bsub>L1\<^esub> x2" for x1 x2 proof - from \x1 \\<^bsub>L1\<^esub> x2\ have "x1 \<^bsub>L1\<^esub>\ l1 x2" using assms1 by (intro t1.left_Galois_left_if_left_relI) blast with mono_r2 show ?thesis by auto qed qed end paragraph \Function Relator\ context transport_Fun_Rel begin lemma mono_wrt_rel_leftI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" using assms by (intro tdfr.mono_wrt_rel_leftI) simp_all end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin lemmas mono_wrt_rel_leftI = mono_wrt_rel_Refl_Rel_Refl_Rel_if_mono_wrt_rel [of tdfr.L tdfr.R l, folded transport_defs] end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin lemmas mono_wrt_rel_leftI = tpdfr.mono_wrt_rel_leftI[OF tfr.mono_wrt_rel_leftI] 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,429 @@ \<^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 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) (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 :: 'z \ 'z \ bool). transitive R\ \ transitive R" and "\(R :: 'z \ 'z \ bool). symmetric R\ \ symmetric R" and "\(R :: 'z \ 'z \ bool). 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 + by (auto intro!: Eq_TrueI simp: t1.flip_unit_eq_counit galois_prop.half_galois_prop_right_rel_inv_iff_half_galois_prop_left + mono_wrt_rel_eq_dep_mono_wrt_rel 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 "((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 "((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 "((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 "((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 "((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 "((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,727 +1,728 @@ \<^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 "\(R :: 'z \ 'z \ bool) (P :: 'z \ bool). (inflationary_on P R\ :: ('z \ 'z) \ bool) \ deflationary_on P R" and "\(P :: 'b2 \ bool) x. (inflationary_on P (flip2 R2 x x) :: ('b2 \ 'b2) \ bool) \ 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 \ 'z \ bool) (P :: 'z \ bool). reflexive_on P R\ \ reflexive_on P R" and "\(R :: 'z \ 'z \ bool). 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) + galois_prop.rel_inv_half_galois_prop_right_eq_half_galois_prop_left_rel_inv + mono_wrt_rel_eq_dep_mono_wrt_rel) 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 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>)" + 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) + 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 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) + 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 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 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>)" + 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) + 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 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) + 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 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" + "((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" + "((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" + "((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" + "((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 force+ + with pre_equiv1 have "r1 x' \<^bsub>L1\<^esub>\ x'" "r1 x' \<^bsub>L1\<^esub>\ \\<^sub>1 x'" by fastforce+ 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 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 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>)" + ((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>)" + ((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>)" + ((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>)" + ((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 (intro t1.left_GaloisI) (auto elim!: t1.preorder_equivalence_order_equivalenceE t1.order_equivalenceE bi_relatedE) 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 elim!: le_relE) 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 "((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 "((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 \ (\\<^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" + 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/Functions/Transport_Functions_Relation_Simplifications.thy b/thys/Transport/Transport/Functions/Transport_Functions_Relation_Simplifications.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Relation_Simplifications.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Relation_Simplifications.thy @@ -1,102 +1,102 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Simplification of Left and Right Relations\ theory Transport_Functions_Relation_Simplifications imports Transport_Functions_Order_Base Transport_Functions_Galois_Equivalence begin paragraph \Dependent Function Relator\ context transport_Dep_Fun_Rel begin text \Due to @{thm "transport_Mono_Dep_Fun_Rel.left_rel_eq_tdfr_left_rel_if_reflexive_on"}, we can apply all results from @{locale "transport_Mono_Dep_Fun_Rel"} to @{locale "transport_Dep_Fun_Rel"} whenever @{term "(\\<^bsub>L\<^esub>)"} and @{term "(\\<^bsub>R\<^esub>)"} are reflexive.\ lemma reflexive_on_in_field_left_rel2_le_assmI: assumes refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" - and mono_L2: "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" + and mono_L2: "((x1 : \) \\<^sub>m (x2 x3 \ (\\<^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 reflexive_on_in_field_mono_assm_left2I: - assumes mono_L2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" + assumes mono_L2: "((x1 x2 \ (\\<^bsub>L1\<^esub>)) \\<^sub>m (x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3) \ (\)) L2" and refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" - shows "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" + shows "((x1 : \) \\<^sub>m (x2 x3 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2) \\<^sub>m (\)) L2" proof (intro dep_mono_wrt_predI dep_mono_wrt_relI rel_if_if_impI) fix x1 x2 x3 assume "x1 \\<^bsub>L1\<^esub> x2" "x2 \\<^bsub>L1\<^esub> x3" with refl_L1 have "x1 \\<^bsub>L1\<^esub> x1" by blast from Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_L2 \x1 \\<^bsub>L1\<^esub> x1\] \x2 \\<^bsub>L1\<^esub> x3\] \x1 \\<^bsub>L1\<^esub> x2\ show "(\\<^bsub>L2 x1 x2\<^esub>) \ (\\<^bsub>L2 x1 x3\<^esub>)" by blast qed lemma reflexive_on_in_field_left_if_equivalencesI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "preorder_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^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 (x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3) \ (\)) L2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>L2 x1 x2\<^esub>)" shows "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" using assms by (intro reflexive_on_in_field_leftI left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI reflexive_on_in_field_left_rel2_le_assmI reflexive_on_in_field_mono_assm_left2I) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin lemma left_rel_eq_tdfr_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^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 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>L2 x1 x2\<^esub>)" shows "(\\<^bsub>L\<^esub>) = tdfr.L" using assms by (intro left_rel_eq_tdfr_left_rel_if_reflexive_on tdfr.reflexive_on_in_field_leftI) auto lemma left_rel_eq_tdfr_leftI_if_equivalencesI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "preorder_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^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 (x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3) \ (\)) L2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>L2 x1 x2\<^esub>)" shows "(\\<^bsub>L\<^esub>) = tdfr.L" using assms by (intro left_rel_eq_tdfr_left_rel_if_reflexive_on tdfr.reflexive_on_in_field_left_if_equivalencesI) auto end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin lemma left_rel_eq_tfr_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" shows "(\\<^bsub>L\<^esub>) = tfr.tdfr.L" using assms by (intro tpdfr.left_rel_eq_tdfr_leftI) auto 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 (fold ext[of 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 (fold ext[of 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 rel_restrict_leftE, rule rel_restrict_leftI, assumption, rule assms, assumption+) done lemma Fpred_top: "Fpred \ \ \ x" apply (subst Fpred_def) apply (rule Frel_refl_strong; subst 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_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) + by (intro ext iffI conjI Set.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 rel_restrict_leftI, rule refl, drule rev_subsetD, assumption, elim CollectE case_prodE 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 rel_restrict_leftE) apply (elim FpredE) apply (rule Frel_mono_strong, assumption; rule rel_restrict_leftI, assumption+) apply (rule rel_restrict_leftI) apply (rule Frel_mono_strong, assumption; erule 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_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 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 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 (rule 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_Order_Equivalence.thy b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Equivalence.thy --- a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Equivalence.thy +++ b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Equivalence.thy @@ -1,189 +1,189 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Order Equivalence\ theory Transport_Natural_Functors_Order_Equivalence imports Transport_Natural_Functors_Base begin context fixes R1 :: "'a \ 'a \ bool" and R2 :: "'b \ 'b \ bool" and R3 :: "'c \ 'c \ bool" and f1 :: "'a \ 'a" and f2 :: "'b \ 'b" and f3 :: "'c \ 'c" and R :: "('d, 'a, 'b, 'c) F \ ('d, 'a, 'b, 'c) F \ bool" and f :: "('d, 'a, 'b, 'c) F \ ('d, 'a, 'b, 'c) F" defines "R \ Frel R1 R2 R3" and "f \ Fmap f1 f2 f3" begin lemma inflationary_on_in_dom_FrelI: assumes "inflationary_on (in_dom R1) R1 f1" and "inflationary_on (in_dom R2) R2 f2" and "inflationary_on (in_dom R3) R3 f3" shows "inflationary_on (in_dom R) R f" apply (unfold R_def f_def) apply (rule inflationary_onI) apply (subst (asm) in_dom_Frel_eq_Fpred_in_dom) apply (erule FpredE) apply (subst Frel_Fmap_eq2) apply (rule Frel_refl_strong) apply (rule inflationary_onD[where ?R=R1] inflationary_onD[where ?R=R2] inflationary_onD[where ?R=R3], rule assms, assumption+)+ done lemma inflationary_on_in_codom_FrelI: assumes "inflationary_on (in_codom R1) R1 f1" and "inflationary_on (in_codom R2) R2 f2" and "inflationary_on (in_codom R3) R3 f3" shows "inflationary_on (in_codom R) R f" apply (unfold R_def f_def) apply (rule inflationary_onI) apply (subst (asm) in_codom_Frel_eq_Fpred_in_codom) apply (erule FpredE) apply (subst Frel_Fmap_eq2) apply (rule Frel_refl_strong) apply (rule inflationary_onD[where ?R=R1] inflationary_onD[where ?R=R2] inflationary_onD[where ?R=R3], rule assms, assumption+)+ done end context fixes R1 :: "'a \ 'a \ bool" and R2 :: "'b \ 'b \ bool" and R3 :: "'c \ 'c \ bool" and f1 :: "'a \ 'a" and f2 :: "'b \ 'b" and f3 :: "'c \ 'c" and R :: "('d, 'a, 'b, 'c) F \ ('d, 'a, 'b, 'c) F \ bool" and f :: "('d, 'a, 'b, 'c) F \ ('d, 'a, 'b, 'c) F" defines "R \ Frel R1 R2 R3" and "f \ Fmap f1 f2 f3" begin lemma inflationary_on_in_field_FrelI: assumes "inflationary_on (in_field R1) R1 f1" and "inflationary_on (in_field R2) R2 f2" and "inflationary_on (in_field R3) R3 f3" shows "inflationary_on (in_field R) R f" apply (unfold R_def f_def) apply (subst in_field_eq_in_dom_sup_in_codom) apply (subst inflationary_on_sup_eq) apply (unfold inf_apply) apply (subst inf_bool_def) apply (rule conjI; rule inflationary_on_in_dom_FrelI inflationary_on_in_codom_FrelI; rule inflationary_on_if_le_pred_if_inflationary_on, rule assms, rule le_predI, rule in_field_if_in_dom in_field_if_in_codom, assumption) done lemma deflationary_on_in_dom_FrelI: assumes "deflationary_on (in_dom R1) R1 f1" and "deflationary_on (in_dom R2) R2 f2" and "deflationary_on (in_dom R3) R3 f3" shows "deflationary_on (in_dom R) R f" apply (unfold R_def f_def) apply (subst deflationary_on_eq_inflationary_on_rel_inv) apply (subst in_codom_rel_inv_eq_in_dom[symmetric]) apply (unfold Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (rule inflationary_on_in_codom_FrelI; subst deflationary_on_eq_inflationary_on_rel_inv[symmetric], subst in_codom_rel_inv_eq_in_dom, rule assms) done lemma deflationary_on_in_codom_FrelI: assumes "deflationary_on (in_codom R1) R1 f1" and "deflationary_on (in_codom R2) R2 f2" and "deflationary_on (in_codom R3) R3 f3" shows "deflationary_on (in_codom R) R f" apply (unfold R_def f_def) apply (subst deflationary_on_eq_inflationary_on_rel_inv) apply (subst in_dom_rel_inv_eq_in_codom[symmetric]) apply (unfold Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (rule inflationary_on_in_dom_FrelI; subst deflationary_on_eq_inflationary_on_rel_inv[symmetric], subst in_dom_rel_inv_eq_in_codom, rule assms) done end context fixes R1 :: "'a \ 'a \ bool" and R2 :: "'b \ 'b \ bool" and R3 :: "'c \ 'c \ bool" and f1 :: "'a \ 'a" and f2 :: "'b \ 'b" and f3 :: "'c \ 'c" and R :: "('d, 'a, 'b, 'c) F \ ('d, 'a, 'b, 'c) F \ bool" and f :: "('d, 'a, 'b, 'c) F \ ('d, 'a, 'b, 'c) F" defines "R \ Frel R1 R2 R3" and "f \ Fmap f1 f2 f3" begin lemma deflationary_on_in_field_FrelI: assumes "deflationary_on (in_field R1) R1 f1" and "deflationary_on (in_field R2) R2 f2" and "deflationary_on (in_field R3) R3 f3" shows "deflationary_on (in_field R) R f" apply (unfold R_def f_def) apply (subst deflationary_on_eq_inflationary_on_rel_inv) - apply (subst in_field_rel_inv_eq[symmetric]) + apply (subst in_field_rel_inv_eq_in_field[symmetric]) apply (unfold Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (rule inflationary_on_in_field_FrelI; subst deflationary_on_eq_inflationary_on_rel_inv[symmetric], - subst in_field_rel_inv_eq, + subst in_field_rel_inv_eq_in_field, rule assms) done lemma rel_equivalence_on_in_field_FrelI: assumes "rel_equivalence_on (in_field R1) R1 f1" and "rel_equivalence_on (in_field R2) R2 f2" and "rel_equivalence_on (in_field R3) R3 f3" shows "rel_equivalence_on (in_field R) R f" apply (unfold R_def f_def) apply (subst rel_equivalence_on_eq) apply (unfold inf_apply) apply (subst inf_bool_def) apply (insert assms) apply (elim rel_equivalence_onE) apply (rule conjI) apply (rule inflationary_on_in_field_FrelI; assumption) apply (fold R_def f_def) apply (rule deflationary_on_in_field_FrelI; assumption) done end context transport_natural_functor begin lemmas inflationary_on_in_field_unitI = inflationary_on_in_field_FrelI [of L1 "\\<^sub>1" L2 "\\<^sub>2" L3 "\\<^sub>3", folded transport_defs unit_eq_Fmap] lemmas deflationary_on_in_field_unitI = deflationary_on_in_field_FrelI [of L1 "\\<^sub>1" L2 "\\<^sub>2" L3 "\\<^sub>3", folded transport_defs unit_eq_Fmap] lemmas rel_equivalence_on_in_field_unitI = rel_equivalence_on_in_field_FrelI [of L1 "\\<^sub>1" L2 "\\<^sub>2" L3 "\\<^sub>3", folded transport_defs unit_eq_Fmap] interpretation flip : transport_natural_functor R1 L1 r1 l1 R2 L2 r2 l2 R3 L3 r3 l3 rewrites "flip.unit \ \" and "flip.t1.unit \ \\<^sub>1" and "flip.t2.unit \ \\<^sub>2" and "flip.t3.unit \ \\<^sub>3" by (simp_all only: order_functors.flip_counit_eq_unit) lemma order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^sub>o (\\<^bsub>R2\<^esub>)) l2 r2" and "((\\<^bsub>L3\<^esub>) \\<^sub>o (\\<^bsub>R3\<^esub>)) l3 r3" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" apply (insert assms) apply (elim order_functors.order_equivalenceE) apply (rule order_equivalenceI; rule mono_wrt_rel_leftI flip.mono_wrt_rel_leftI rel_equivalence_on_in_field_unitI flip.rel_equivalence_on_in_field_unitI; assumption) done end end \ No newline at end of file