diff --git a/thys/Amortized_Complexity/ROOT b/thys/Amortized_Complexity/ROOT --- a/thys/Amortized_Complexity/ROOT +++ b/thys/Amortized_Complexity/ROOT @@ -1,23 +1,24 @@ chapter AFP session Amortized_Complexity = "HOL-Library" + options [timeout = 600] sessions + "HOL-Data_Structures" Pairing_Heap Skew_Heap Splay_Tree theories Amortized_Framework0 Amortized_Examples Skew_Heap_Analysis Splay_Tree_Analysis Splay_Tree_Analysis_Optimal Splay_Heap_Analysis Pairing_Heap_Tree_Analysis Pairing_Heap_Tree_Analysis2 Pairing_Heap_List1_Analysis Pairing_Heap_List1_Analysis2 Pairing_Heap_List2_Analysis document_files "root.tex" "root.bib" diff --git a/thys/Amortized_Complexity/Skew_Heap_Analysis.thy b/thys/Amortized_Complexity/Skew_Heap_Analysis.thy --- a/thys/Amortized_Complexity/Skew_Heap_Analysis.thy +++ b/thys/Amortized_Complexity/Skew_Heap_Analysis.thy @@ -1,201 +1,199 @@ (* Author: Tobias Nipkow *) section "Skew Heap Analysis" theory Skew_Heap_Analysis imports Complex_Main Skew_Heap.Skew_Heap Amortized_Framework + "HOL-Data_Structures.Define_Time_Function" Priority_Queue_ops_merge begin text\The following proof is a simplified version of the one by Kaldewaij and Schoenmakers~\<^cite>\"KaldewaijS-IPL91"\.\ text \right-heavy:\ definition rh :: "'a tree => 'a tree => nat" where "rh l r = (if size l < size r then 1 else 0)" text \Function \\\ in \<^cite>\"KaldewaijS-IPL91"\: number of right-heavy nodes on left spine.\ fun lrh :: "'a tree \ nat" where "lrh Leaf = 0" | "lrh (Node l _ r) = rh l r + lrh l" text \Function \\\ in \<^cite>\"KaldewaijS-IPL91"\: number of not-right-heavy nodes on right spine.\ fun rlh :: "'a tree \ nat" where "rlh Leaf = 0" | "rlh (Node l _ r) = (1 - rh l r) + rlh r" lemma Gexp: "2 ^ lrh t \ size t + 1" by (induction t) (auto simp: rh_def) corollary Glog: "lrh t \ log 2 (size1 t)" by (metis Gexp le_log2_of_power size1_size) lemma Dexp: "2 ^ rlh t \ size t + 1" by (induction t) (auto simp: rh_def) corollary Dlog: "rlh t \ log 2 (size1 t)" by (metis Dexp le_log2_of_power size1_size) -function T_merge :: "'a::linorder tree \ 'a tree \ nat" where -"T_merge Leaf t = 1" | -"T_merge t Leaf = 1" | -"T_merge (Node l1 a1 r1) (Node l2 a2 r2) = - (if a1 \ a2 then T_merge (Node l2 a2 r2) r1 else T_merge (Node l1 a1 r1) r2) + 1" -by pat_completeness auto -termination -by (relation "measure (\(x, y). size x + size y)") auto +time_fun merge fun \ :: "'a tree \ int" where "\ Leaf = 0" | "\ (Node l _ r) = \ l + \ r + rh l r" lemma \_nneg: "\ t \ 0" by (induction t) auto lemma plus_log_le_2log_plus: "\ x > 0; y > 0; b > 1 \ \ log b x + log b y \ 2 * log b (x + y)" by(subst mult_2; rule add_mono; auto) lemma rh1: "rh l r \ 1" by(simp add: rh_def) lemma amor_le_long: "T_merge t1 t2 + \ (merge t1 t2) - \ t1 - \ t2 \ lrh(merge t1 t2) + rlh t1 + rlh t2 + 1" proof (induction t1 t2 rule: merge.induct) case 1 thus ?case by simp next case 2 thus ?case by simp next case (3 l1 a1 r1 l2 a2 r2) show ?case proof (cases "a1 \ a2") case True let ?t1 = "Node l1 a1 r1" let ?t2 = "Node l2 a2 r2" let ?m = "merge ?t2 r1" have "T_merge ?t1 ?t2 + \ (merge ?t1 ?t2) - \ ?t1 - \ ?t2 = T_merge ?t2 r1 + 1 + \ ?m + \ l1 + rh ?m l1 - \ ?t1 - \ ?t2" using True by (simp) also have "\ = T_merge ?t2 r1 + 1 + \ ?m + rh ?m l1 - \ r1 - rh l1 r1 - \ ?t2" by simp also have "\ \ lrh ?m + rlh ?t2 + rlh r1 + rh ?m l1 + 2 - rh l1 r1" using "3.IH"(1)[OF True] by linarith also have "\ = lrh ?m + rlh ?t2 + rlh r1 + rh ?m l1 + 1 + (1 - rh l1 r1)" using rh1[of l1 r1] by (simp) also have "\ = lrh ?m + rlh ?t2 + rlh ?t1 + rh ?m l1 + 1" by (simp) also have "\ = lrh (merge ?t1 ?t2) + rlh ?t1 + rlh ?t2 + 1" using True by(simp) finally show ?thesis . next case False with 3 show ?thesis by auto qed qed lemma amor_le: "T_merge t1 t2 + \ (merge t1 t2) - \ t1 - \ t2 \ lrh(merge t1 t2) + rlh t1 + rlh t2 + 1" by(induction t1 t2 rule: merge.induct)(auto) lemma a_merge: "T_merge t1 t2 + \(merge t1 t2) - \ t1 - \ t2 \ 3 * log 2 (size1 t1 + size1 t2) + 1" (is "?l \ _") proof - have "?l \ lrh(merge t1 t2) + rlh t1 + rlh t2 + 1" using amor_le[of t1 t2] by arith also have "\ = real(lrh(merge t1 t2)) + rlh t1 + rlh t2 + 1" by simp also have "\ = real(lrh(merge t1 t2)) + (real(rlh t1) + rlh t2) + 1" by simp also have "rlh t1 \ log 2 (size1 t1)" by(rule Dlog) also have "rlh t2 \ log 2 (size1 t2)" by(rule Dlog) also have "lrh (merge t1 t2) \ log 2 (size1(merge t1 t2))" by(rule Glog) also have "size1(merge t1 t2) = size1 t1 + size1 t2 - 1" by(simp add: size1_size size_merge) also have "log 2 (size1 t1 + size1 t2 - 1) \ log 2 (size1 t1 + size1 t2)" by(simp add: size1_size) also have "log 2 (size1 t1) + log 2 (size1 t2) \ 2 * log 2 (real(size1 t1) + (size1 t2))" by(rule plus_log_le_2log_plus) (auto simp: size1_size) finally show ?thesis by(simp) qed +text \Command \time_fun\ does not work for @{const skew_heap.insert} and @{const skew_heap.del_min} +because they are the result of a locale and not what they seem. +However, their manual definition is trivial:\ + definition T_insert :: "'a::linorder \ 'a tree \ int" where -"T_insert a t = T_merge (Node Leaf a Leaf) t + 1" +"T_insert a t = T_merge (Node Leaf a Leaf) t" -lemma a_insert: "T_insert a t + \(skew_heap.insert a t) - \ t \ 3 * log 2 (size1 t + 2) + 2" +lemma a_insert: "T_insert a t + \(skew_heap.insert a t) - \ t \ 3 * log 2 (size1 t + 2) + 1" using a_merge[of "Node Leaf a Leaf" "t"] by (simp add: numeral_eq_Suc T_insert_def rh_def) definition T_del_min :: "('a::linorder) tree \ int" where -"T_del_min t = (case t of Leaf \ 1 | Node t1 a t2 \ T_merge t1 t2 + 1)" +"T_del_min t = (case t of Leaf \ 0 | Node t1 a t2 \ T_merge t1 t2)" -lemma a_del_min: "T_del_min t + \(skew_heap.del_min t) - \ t \ 3 * log 2 (size1 t + 2) + 2" +lemma a_del_min: "T_del_min t + \(skew_heap.del_min t) - \ t \ 3 * log 2 (size1 t + 2) + 1" proof (cases t) case Leaf thus ?thesis by (simp add: T_del_min_def) next case (Node t1 _ t2) have [arith]: "log 2 (2 + (real (size t1) + real (size t2))) \ log 2 (4 + (real (size t1) + real (size t2)))" by simp from Node show ?thesis using a_merge[of t1 t2] - by (simp add: size1_size T_del_min_def rh_def) + by (simp add: size1_size rh_def T_del_min_def) qed subsubsection "Instantiation of Amortized Framework" lemma T_merge_nneg: "T_merge t1 t2 \ 0" by(induction t1 t2 rule: T_merge.induct) auto fun exec :: "'a::linorder op \ 'a tree list \ 'a tree" where "exec Empty [] = Leaf" | "exec (Insert a) [t] = skew_heap.insert a t" | "exec Del_min [t] = skew_heap.del_min t" | "exec Merge [t1,t2] = merge t1 t2" fun cost :: "'a::linorder op \ 'a tree list \ nat" where "cost Empty [] = 1" | "cost (Insert a) [t] = T_merge (Node Leaf a Leaf) t + 1" | "cost Del_min [t] = (case t of Leaf \ 1 | Node t1 a t2 \ T_merge t1 t2 + 1)" | "cost Merge [t1,t2] = T_merge t1 t2" fun U where "U Empty [] = 1" | "U (Insert _) [t] = 3 * log 2 (size1 t + 2) + 2" | "U Del_min [t] = 3 * log 2 (size1 t + 2) + 2" | "U Merge [t1,t2] = 3 * log 2 (size1 t1 + size1 t2) + 1" interpretation Amortized where arity = arity and exec = exec and inv = "\_. True" and cost = cost and \ = \ and U = U proof (standard, goal_cases) case 1 show ?case by simp next case (2 t) show ?case using \_nneg[of t] by linarith next case (3 ss f) show ?case proof (cases f) case Empty thus ?thesis using 3(2) by (auto) next case [simp]: (Insert a) obtain t where [simp]: "ss = [t]" using 3(2) by (auto) thus ?thesis using a_merge[of "Node Leaf a Leaf" "t"] by (simp add: numeral_eq_Suc insert_def rh_def T_merge_nneg) next case [simp]: Del_min obtain t where [simp]: "ss = [t]" using 3(2) by (auto) thus ?thesis proof (cases t) case Leaf with Del_min show ?thesis by simp next case (Node t1 _ t2) have [arith]: "log 2 (2 + (real (size t1) + real (size t2))) \ log 2 (4 + (real (size t1) + real (size t2)))" by simp from Del_min Node show ?thesis using a_merge[of t1 t2] by (simp add: size1_size T_merge_nneg) qed next case [simp]: Merge obtain t1 t2 where "ss = [t1,t2]" using 3(2) by (auto simp: numeral_eq_Suc) thus ?thesis using a_merge[of t1 t2] by (simp add: T_merge_nneg) qed qed end diff --git a/thys/Amortized_Complexity/Splay_Tree_Analysis.thy b/thys/Amortized_Complexity/Splay_Tree_Analysis.thy --- a/thys/Amortized_Complexity/Splay_Tree_Analysis.thy +++ b/thys/Amortized_Complexity/Splay_Tree_Analysis.thy @@ -1,432 +1,432 @@ subsection "Splay Tree Analysis" theory Splay_Tree_Analysis imports Splay_Tree_Analysis_Base Amortized_Framework begin subsubsection "Analysis of splay" definition A_splay :: "'a::linorder \ 'a tree \ real" where "A_splay a t = T_splay a t + \(splay a t) - \ t" text\The following lemma is an attempt to prove a generic lemma that covers both zig-zig cases. However, the lemma is not as nice as one would like. Hence it is used only once, as a demo. Ideally the lemma would involve function @{const A_splay}, but that is impossible because this involves @{const splay} and thus depends on the ordering. We would need a truly symmetric version of @{const splay} that takes the ordering as an explicit argument. Then we could define all the symmetric cases by one final equation \<^prop>\splay2 less t = splay2 (\x y. Not (less x y)) (mirror t)\. This would simplify the code and the proofs.\ lemma zig_zig: fixes lx x rx lb b rb a ra u lb1 lb2 defines [simp]: "X == Node lx (x) rx" defines[simp]: "B == Node lb b rb" defines [simp]: "t == Node B a ra" defines [simp]: "A' == Node rb a ra" defines [simp]: "t' == Node lb1 u (Node lb2 b A')" assumes hyps: "lb \ \\" and IH: "T_splay x lb + \ lb1 + \ lb2 - \ lb \ 2 * \ lb - 3 * \ X + 1" and prems: "size lb = size lb1 + size lb2 + 1" "X \ subtrees lb" shows "T_splay x lb + \ t' - \ t \ 3 * (\ t - \ X)" proof - define B' where [simp]: "B' = Node lb2 b A'" have "T_splay x lb + \ t' - \ t = T_splay x lb + \ lb1 + \ lb2 - \ lb + \ B' + \ A' - \ B" using prems by(auto simp: A_splay_def size_if_splay algebra_simps in_set_tree_if split: tree.split) also have "\ \ 2 * \ lb + \ B' + \ A' - \ B - 3 * \ X + 1" using IH prems(2) by(auto simp: algebra_simps) also have "\ \ \ lb + \ B' + \ A' - 3 * \ X + 1" by(simp) also have "\ \ \ B' + 2 * \ t - 3 * \ X " using prems ld_ld_1_less[of "size1 lb" "size1 A'"] by(simp add: size_if_splay) also have "\ \ 3 * \ t - 3 * \ X" using prems by(simp add: size_if_splay) finally show ?thesis by simp qed lemma A_splay_ub: "\ bst t; Node l x r : subtrees t \ \ A_splay x t \ 3 * (\ t - \(Node l x r)) + 1" proof(induction x t rule: splay.induct) case 1 thus ?case by simp next case 2 thus ?case by (auto simp: A_splay_def) next case 4 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 5 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 7 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 10 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 12 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 13 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case (3 x b A B CD) (* A readable proof *) let ?t = "\\A, x, B\, b, CD\" let ?t' = "\A, x, \B, b, CD\\" have *: "l = A \ r = B" using "3.prems" by(fastforce dest: in_set_tree_if) have "A_splay x ?t = 1 + \ ?t' - \ ?t" using "3.hyps" by (simp add: A_splay_def) also have "\ = 1 + \ ?t' + \ \B, b, CD\ - \ ?t - \ \A, x, B\" by(simp) also have "\ = 1 + \ \B, b, CD\ - \ \A, x, B\" by(simp) also have "\ \ 1 + \ ?t - \(Node A x B)" using log_le_cancel_iff[of 2 "size1(Node B b CD)" "size1 ?t"] by (simp) also have "\ \ 1 + 3 * (\ ?t - \(Node A x B))" using log_le_cancel_iff[of 2 "size1(Node A x B)" "size1 ?t"] by (simp) finally show ?case using * by simp next case (9 b x AB C D) (* An automatic proof *) let ?A = "\AB, b, \C, x, D\\" have "x \ set_tree AB" using "9.prems"(1) by auto with 9 show ?case using log_le_cancel_iff[of 2 "size1(Node AB b C)" "size1 ?A"] log_le_cancel_iff[of 2 "size1(Node C x D)" "size1 ?A"] by (auto simp: A_splay_def algebra_simps simp del:log_le_cancel_iff) next case (6 x a b A B C) hence *: "\l, x, r\ \ subtrees A" by(fastforce dest: in_set_tree_if) obtain A1 a' A2 where sp: "splay x A = Node A1 a' A2" using splay_not_Leaf[OF \A \ Leaf\] by blast let ?X = "Node l x r" let ?AB = "Node A a B" let ?ABC = "Node ?AB b C" let ?A' = "Node A1 a' A2" let ?BC = "Node B b C" let ?A2BC = "Node A2 a ?BC" let ?A1A2BC = "Node A1 a' ?A2BC" have 0: "\ ?A1A2BC = \ ?ABC" using sp by(simp add: size_if_splay) have 1: "\ ?A1A2BC - \ ?ABC = \ A1 + \ A2 + \ ?A2BC + \ ?BC - \ A - \ ?AB" using 0 by (simp) have "A_splay x ?ABC = T_splay x A + 1 + \ ?A1A2BC - \ ?ABC" using "6.hyps" sp by(simp add: A_splay_def) also have "\ = T_splay x A + 1 + \ A1 + \ A2 + \ ?A2BC + \ ?BC - \ A - \ ?AB" using 1 by simp also have "\ = T_splay x A + \ ?A' - \ ?A' - \ A + \ ?A2BC + \ ?BC - \ ?AB + 1" by(simp) also have "\ = A_splay x A + \ ?A2BC + \ ?BC - \ ?AB - \ ?A' + 1" using sp by(simp add: A_splay_def) also have "\ \ 3 * \ A + \ ?A2BC + \ ?BC - \ ?AB - \ ?A' - 3 * \ ?X + 2" using "6.IH" "6.prems"(1) * by(simp) also have "\ = 2 * \ A + \ ?A2BC + \ ?BC - \ ?AB - 3 * \ ?X + 2" using sp by(simp add: size_if_splay) also have "\ < \ A + \ ?A2BC + \ ?BC - 3 * \ ?X + 2" by(simp) also have "\ < \ ?A2BC + 2 * \ ?ABC - 3 * \ ?X + 1" using sp ld_ld_1_less[of "size1 A" "size1 ?BC"] by(simp add: size_if_splay) also have "\ < 3 * \ ?ABC - 3 * \ ?X + 1" using sp by(simp add: size_if_splay) finally show ?case by simp next case (8 a x b B A C) hence *: "\l, x, r\ \ subtrees B" by(fastforce dest: in_set_tree_if) obtain B1 b' B2 where sp: "splay x B = Node B1 b' B2" using splay_not_Leaf[OF \B \ Leaf\] by blast let ?X = "Node l x r" let ?AB = "Node A a B" let ?ABC = "Node ?AB b C" let ?B' = "Node B1 b' B2" let ?AB1 = "Node A a B1" let ?B2C = "Node B2 b C" let ?AB1B2C = "Node ?AB1 b' ?B2C" have 0: "\ ?AB1B2C = \ ?ABC" using sp by(simp add: size_if_splay) have 1: "\ ?AB1B2C - \ ?ABC = \ B1 + \ B2 + \ ?AB1 + \ ?B2C - \ B - \ ?AB" using 0 by (simp) have "A_splay x ?ABC = T_splay x B + 1 + \ ?AB1B2C - \ ?ABC" using "8.hyps" sp by(simp add: A_splay_def) also have "\ = T_splay x B + 1 + \ B1 + \ B2 + \ ?AB1 + \ ?B2C - \ B - \ ?AB" using 1 by simp also have "\ = T_splay x B + \ ?B' - \ ?B' - \ B + \ ?AB1 + \ ?B2C - \ ?AB + 1" by simp also have "\ = A_splay x B + \ ?AB1 + \ ?B2C - \ ?AB - \ ?B' + 1" using sp by (simp add: A_splay_def) also have "\ \ 3 * \ B + \ ?AB1 + \ ?B2C - \ ?AB - \ ?B' - 3 * \ ?X + 2" using "8.IH" "8.prems"(1) * by(simp) also have "\ = 2 * \ B + \ ?AB1 + \ ?B2C - \ ?AB - 3 * \ ?X + 2" using sp by(simp add: size_if_splay) also have "\ < \ B + \ ?AB1 + \ ?B2C - 3 * \ ?X + 2" by(simp) also have "\ < \ B + 2 * \ ?ABC - 3 * \ ?X + 1" using sp ld_ld_1_less[of "size1 ?AB1" "size1 ?B2C"] by(simp add: size_if_splay) also have "\ < 3 * \ ?ABC - 3 * \ ?X + 1" by(simp) finally show ?case by simp next case (11 b x c C A D) hence *: "\l, x, r\ \ subtrees C" by(fastforce dest: in_set_tree_if) obtain C1 c' C2 where sp: "splay x C = Node C1 c' C2" using splay_not_Leaf[OF \C \ Leaf\] by blast let ?X = "Node l x r" let ?CD = "Node C c D" let ?ACD = "Node A b ?CD" let ?C' = "Node C1 c' C2" let ?C2D = "Node C2 c D" let ?AC1 = "Node A b C1" have "A_splay x ?ACD = A_splay x C + \ ?C2D + \ ?AC1 - \ ?CD - \ ?C' + 1" using "11.hyps" sp by(auto simp: A_splay_def size_if_splay algebra_simps split: tree.split) also have "\ \ 3 * \ C + \ ?C2D + \ ?AC1 - \ ?CD - \ ?C' - 3 * \ ?X + 2" using "11.IH" "11.prems"(1) * by(auto simp: algebra_simps) also have "\ = 2 * \ C + \ ?C2D + \ ?AC1 - \ ?CD - 3 * \ ?X + 2" using sp by(simp add: size_if_splay) also have "\ \ \ C + \ ?C2D + \ ?AC1 - 3 * \ ?X + 2" by(simp) also have "\ \ \ C + 2 * \ ?ACD - 3 * \ ?X + 1" using sp ld_ld_1_less[of "size1 ?C2D" "size1 ?AC1"] by(simp add: size_if_splay algebra_simps) also have "\ \ 3 * \ ?ACD - 3 * \ ?X + 1" by(simp) finally show ?case by simp next case (14 a x b CD A B) hence 0: "x \ set_tree B \ x \ set_tree A" using "14.prems"(1) \b by(auto) hence 1: "x \ set_tree CD" using "14.prems" \b \a by (auto) obtain C c D where sp: "splay x CD = Node C c D" using splay_not_Leaf[OF \CD \ Leaf\] by blast from zig_zig[of CD x D C l r _ b B a A] 14 sp 0 show ?case by(auto simp: A_splay_def size_if_splay algebra_simps) (* The explicit version: have "A_splay x ?A = A_splay x ?R + \ ?B' + \ ?A' - \ ?B - \ ?R' + 1" using "14.prems" 1 sp by(auto simp: A_splay_def size_if_splay algebra_simps split: tree.split) also have "\ \ 3 * \ ?R + \ ?B' + \ ?A' - \ ?B - \ ?R' - 3 * \ ?X + 2" using 14 0 by(auto simp: algebra_simps) also have "\ = 2 * \ rb + \ ?B' + \ ?A' - \ ?B - 3 * \ ?X + 2" using sp by(simp add: size_if_splay) also have "\ \ \ ?R + \ ?B' + \ ?A' - 3 * \ ?X + 2" by(simp) also have "\ \ \ ?B' + 2 * \ ?A - 3 * \ ?X + 1" using sp ld_ld_1_less[of "size1 ?R" "size1 ?A'"] by(simp add: size_if_splay algebra_simps) also have "\ \ 3 * \ ?A - 3 * \ ?X + 1" using sp by(simp add: size_if_splay) finally show ?case by simp *) qed lemma A_splay_ub2: assumes "bst t" "x : set_tree t" shows "A_splay x t \ 3 * (\ t - 1) + 1" proof - from assms(2) obtain l r where N: "Node l x r : subtrees t" by (metis set_treeE) have "A_splay x t \ 3 * (\ t - \(Node l x r)) + 1" by(rule A_splay_ub[OF assms(1) N]) also have "\ \ 3 * (\ t - 1) + 1" by(simp add: field_simps) finally show ?thesis . qed lemma A_splay_ub3: assumes "bst t" shows "A_splay x t \ 3 * \ t + 1" proof cases assume "t = Leaf" thus ?thesis by(simp add: A_splay_def) next assume "t \ Leaf" from ex_in_set_tree[OF this assms] obtain x' where a': "x' \ set_tree t" "splay x' t = splay x t" "T_splay x' t = T_splay x t" by blast show ?thesis using A_splay_ub2[OF assms a'(1)] by(simp add: A_splay_def a') qed subsubsection "Analysis of insert" lemma amor_insert: assumes "bst t" shows "T_insert x t + \(Splay_Tree.insert x t) - \ t \ 4 * log 2 (size1 t) + 2" (is "?l \ ?r") proof cases - assume "t = Leaf" thus ?thesis by(simp add: T_insert_def) + assume "t = Leaf" thus ?thesis by(simp) next assume "t \ Leaf" then obtain l e r where [simp]: "splay x t = Node l e r" by (metis tree.exhaust splay_Leaf_iff) let ?t = "real(T_splay x t)" let ?Plr = "\ l + \ r" let ?Ps = "\ t" let ?slr = "real(size1 l) + real(size1 r)" let ?LR = "log 2 (1 + ?slr)" have opt: "?t + \ (splay x t) - ?Ps \ 3 * log 2 (real (size1 t)) + 1" using A_splay_ub3[OF \bst t\, simplified A_splay_def, of x] by (simp) from less_linear[of e x] show ?thesis proof (elim disjE) assume "e=x" have nneg: "log 2 (1 + real (size t)) \ 0" by simp thus ?thesis using \t \ Leaf\ opt \e=x\ - apply(simp add: T_insert_def algebra_simps) using nneg by arith + apply(simp add: algebra_simps) using nneg by arith next let ?L = "log 2 (real(size1 l) + 1)" assume "e < x" hence "e \ x" by simp hence "?l = (?t + ?Plr - ?Ps) + ?L + ?LR" - using \t \ Leaf\ \e by(simp add: T_insert_def) + using \t \ Leaf\ \e by(simp) also have "?t + ?Plr - ?Ps \ 2 * log 2 ?slr + 1" using opt size_splay[of x t,symmetric] by(simp) also have "?L \ log 2 ?slr" by(simp) also have "?LR \ log 2 ?slr + 1" proof - have "?LR \ log 2 (2 * ?slr)" by (simp add:) also have "\ \ log 2 ?slr + 1" by (simp add: log_mult del:distrib_left_numeral) finally show ?thesis . qed finally show ?thesis using size_splay[of x t,symmetric] by (simp) next let ?R = "log 2 (2 + real(size r))" assume "x < e" hence "e \ x" by simp hence "?l = (?t + ?Plr - ?Ps) + ?R + ?LR" - using \t \ Leaf\ \x < e\ by(simp add: T_insert_def) + using \t \ Leaf\ \x < e\ by(simp) also have "?t + ?Plr - ?Ps \ 2 * log 2 ?slr + 1" using opt size_splay[of x t,symmetric] by(simp) also have "?R \ log 2 ?slr" by(simp) also have "?LR \ log 2 ?slr + 1" proof - have "?LR \ log 2 (2 * ?slr)" by (simp add:) also have "\ \ log 2 ?slr + 1" by (simp add: log_mult del:distrib_left_numeral) finally show ?thesis . qed finally show ?thesis using size_splay[of x t, symmetric] by simp qed qed subsubsection "Analysis of delete" definition A_splay_max :: "'a::linorder tree \ real" where "A_splay_max t = T_splay_max t + \(splay_max t) - \ t" lemma A_splay_max_ub: "t \ Leaf \ A_splay_max t \ 3 * (\ t - 1) + 1" proof(induction t rule: splay_max.induct) case 1 thus ?case by (simp) next case (2 A) thus ?case using one_le_log_cancel_iff[of 2 "size1 A + 1"] by (simp add: A_splay_max_def del: one_le_log_cancel_iff) next case (3 l b rl c rr) show ?case proof cases assume "rr = Leaf" thus ?thesis using one_le_log_cancel_iff[of 2 "1 + size1 rl"] one_le_log_cancel_iff[of 2 "1 + size1 l + size1 rl"] log_le_cancel_iff[of 2 "size1 l + size1 rl" "1 + size1 l + size1 rl"] by (auto simp: A_splay_max_def field_simps simp del: log_le_cancel_iff one_le_log_cancel_iff) next assume "rr \ Leaf" then obtain l' u r' where sp: "splay_max rr = Node l' u r'" using splay_max_Leaf_iff tree.exhaust by blast hence 1: "size rr = size l' + size r' + 1" using size_splay_max[of rr,symmetric] by(simp) let ?C = "Node rl c rr" let ?B = "Node l b ?C" let ?B' = "Node l b rl" let ?C' = "Node ?B' c l'" have "A_splay_max ?B = A_splay_max rr + \ ?B' + \ ?C' - \ rr - \ ?C + 1" using "3.prems" sp 1 by(auto simp add: A_splay_max_def) also have "\ \ 3 * (\ rr - 1) + \ ?B' + \ ?C' - \ rr - \ ?C + 2" using 3 \rr \ Leaf\ by auto also have "\ = 2 * \ rr + \ ?B' + \ ?C' - \ ?C - 1" by simp also have "\ \ \ rr + \ ?B' + \ ?C' - 1" by simp also have "\ \ 2 * \ ?B + \ ?C' - 2" using ld_ld_1_less[of "size1 ?B'" "size1 rr"] by(simp add:) also have "\ \ 3 * \ ?B - 2" using 1 by simp finally show ?case by simp qed qed lemma A_splay_max_ub3: "A_splay_max t \ 3 * \ t + 1" proof cases assume "t = Leaf" thus ?thesis by(simp add: A_splay_max_def) next assume "t \ Leaf" show ?thesis using A_splay_max_ub[OF \t \ Leaf\] by(simp) qed lemma amor_delete: assumes "bst t" shows "T_delete a t + \(Splay_Tree.delete a t) - \ t \ 6 * log 2 (size1 t) + 2" proof (cases t) - case Leaf thus ?thesis by(simp add: Splay_Tree.delete_def T_delete_def) + case Leaf thus ?thesis by(simp add: Splay_Tree.delete_def) next case [simp]: (Node ls x rs) then obtain l e r where sp[simp]: "splay a (Node ls x rs) = Node l e r" by (metis tree.exhaust splay_Leaf_iff) let ?t = "real(T_splay a t)" let ?Plr = "\ l + \ r" let ?Ps = "\ t" let ?slr = "real(size1 l) + real(size1 r)" let ?LR = "log 2 (1 + ?slr)" let ?lslr = "log 2 (real (size ls) + (real (size rs) + 2))" have "?lslr \ 0" by simp have opt: "?t + \ (splay a t) - ?Ps \ 3 * log 2 (real (size1 t)) + 1" using A_splay_ub3[OF \bst t\, simplified A_splay_def, of a] by (simp add: field_simps) show ?thesis proof (cases "e=a") case False thus ?thesis - using opt apply(simp add: Splay_Tree.delete_def T_delete_def field_simps) + using opt apply(simp add: Splay_Tree.delete_def field_simps) using \?lslr \ 0\ by arith next case [simp]: True show ?thesis proof (cases l) case Leaf have 1: "log 2 (real (size r) + 2) \ 0" by(simp) show ?thesis - using Leaf opt apply(simp add: Splay_Tree.delete_def T_delete_def field_simps) + using Leaf opt apply(simp add: Splay_Tree.delete_def field_simps) using 1 \?lslr \ 0\ by arith next case (Node ll y lr) then obtain l' y' r' where [simp]: "splay_max (Node ll y lr) = Node l' y' r'" using splay_max_Leaf_iff tree.exhaust by blast have "bst l" using bst_splay[OF \bst t\, of a] by simp have "\ r' \ 0" apply (induction r') by (auto) have optm: "real(T_splay_max l) + \ (splay_max l) - \ l \ 3 * \ l + 1" using A_splay_max_ub3[of l, simplified A_splay_max_def] by (simp add: field_simps Node) have 1: "log 2 (2+(real(size l')+real(size r))) \ log 2 (2+(real(size l)+real(size r)))" using size_splay_max[of l] Node by simp have 2: "log 2 (2 + (real (size l') + real (size r'))) \ 0" by simp have 3: "log 2 (size1 l' + size1 r) \ log 2 (size1 l' + size1 r') + log 2 ?slr" apply simp using 1 2 by arith have 4: "log 2 (real(size ll) + (real(size lr) + 2)) \ ?lslr" using size_if_splay[OF sp] Node by simp show ?thesis using add_mono[OF opt optm] Node 3 - apply(simp add: Splay_Tree.delete_def T_delete_def field_simps) + apply(simp add: Splay_Tree.delete_def field_simps) using 4 \\ r' \ 0\ by arith qed qed qed subsubsection "Overall analysis" fun U where "U Empty [] = 1" | "U (Splay _) [t] = 3 * log 2 (size1 t) + 1" | "U (Insert _) [t] = 4 * log 2 (size1 t) + 3" | "U (Delete _) [t] = 6 * log 2 (size1 t) + 3" interpretation Amortized where arity = arity and exec = exec and inv = bst and cost = cost and \ = \ and U = U proof (standard, goal_cases) case (1 ss f) show ?case proof (cases f) case Empty thus ?thesis using 1 by auto next case (Splay a) then obtain t where "ss = [t]" "bst t" using 1 by auto with Splay bst_splay[OF \bst t\, of a] show ?thesis by (auto split: tree.split) next case (Insert a) then obtain t where "ss = [t]" "bst t" using 1 by auto with bst_splay[OF \bst t\, of a] Insert show ?thesis by (auto simp: splay_bstL[OF \bst t\] splay_bstR[OF \bst t\] split: tree.split) next case (Delete a) then obtain t where "ss = [t]" "bst t" using 1 by auto with 1 Delete show ?thesis by(simp add: bst_delete) qed next case (2 t) thus ?case by (induction t) auto next case (3 ss f) show ?case (is "?l \ ?r") proof(cases f) case Empty thus ?thesis using 3(2) by(simp add: A_splay_def) next case (Splay a) then obtain t where "ss = [t]" "bst t" using 3 by auto thus ?thesis using Splay A_splay_ub3[OF \bst t\] by(simp add: A_splay_def) next case [simp]: (Insert a) then obtain t where [simp]: "ss = [t]" and "bst t" using 3 by auto thus ?thesis using amor_insert[of t a] by auto next case [simp]: (Delete a) then obtain t where [simp]: "ss = [t]" and "bst t" using 3 by auto thus ?thesis using amor_delete[of t a] by auto qed qed end diff --git a/thys/Amortized_Complexity/Splay_Tree_Analysis_Base.thy b/thys/Amortized_Complexity/Splay_Tree_Analysis_Base.thy --- a/thys/Amortized_Complexity/Splay_Tree_Analysis_Base.thy +++ b/thys/Amortized_Complexity/Splay_Tree_Analysis_Base.thy @@ -1,139 +1,119 @@ section "Splay Tree" subsection "Basics" theory Splay_Tree_Analysis_Base imports Lemmas_log Splay_Tree.Splay_Tree + "HOL-Data_Structures.Define_Time_Function" begin declare size1_size[simp] abbreviation "\ t == log 2 (size1 t)" fun \ :: "'a tree \ real" where "\ Leaf = 0" | "\ (Node l a r) = \ (Node l a r) + \ l + \ r" -fun T_splay :: "'a::linorder \ 'a tree \ nat" where -"T_splay x Leaf = 1" | -"T_splay x (Node AB b CD) = - (case cmp x b of - EQ \ 1 | - LT \ (case AB of - Leaf \ 1 | - Node A a B \ - (case cmp x a of EQ \ 1 | - LT \ if A = Leaf then 1 else T_splay x A + 1 | - GT \ if B = Leaf then 1 else T_splay x B + 1)) | - GT \ (case CD of - Leaf \ 1 | - Node C c D \ - (case cmp x c of EQ \ 1 | - LT \ if C = Leaf then 1 else T_splay x C + 1 | - GT \ if D = Leaf then 1 else T_splay x D + 1)))" +time_fun cmp +time_fun splay equations splay.simps(1) splay_code lemma T_splay_simps[simp]: "T_splay a (Node l a r) = 1" "x T_splay x (Node Leaf b CD) = 1" "a T_splay a (Node (Node A a B) b CD) = 1" "x x T_splay x (Node (Node A a B) b CD) = (if A = Leaf then 1 else T_splay x A + 1)" "x a T_splay x (Node (Node A a B) b CD) = (if B = Leaf then 1 else T_splay x B + 1)" "b T_splay x (Node AB b Leaf) = 1" "b T_splay a (Node AB b (Node C a D)) = 1" "b x T_splay x (Node AB b (Node C c D)) = (if C=Leaf then 1 else T_splay x C + 1)" "b c T_splay x (Node AB b (Node C c D)) = (if D=Leaf then 1 else T_splay x D + 1)" by auto declare T_splay.simps(2)[simp del] -definition T_insert :: "'a::linorder \ 'a tree \ nat" where -"T_insert x t = (if t = Leaf then 0 else T_splay x t)" +time_fun insert -fun T_splay_max :: "'a::linorder tree \ nat" where -"T_splay_max Leaf = 1" | -"T_splay_max (Node A a Leaf) = 1" | -"T_splay_max (Node A a (Node B b C)) = (if C=Leaf then 1 else T_splay_max C + 1)" +lemma T_insert_simp: "T_insert x t = (if t = Leaf then 0 else T_splay x t)" +by(auto split: tree.split) -definition T_delete :: "'a::linorder \ 'a tree \ nat" where -"T_delete x t = - (if t = Leaf then 0 - else T_splay x t + - (case splay x t of - Node l a r \ if x \ a then 0 else if l = Leaf then 0 else T_splay_max l))" +time_fun splay_max + +time_fun delete lemma ex_in_set_tree: "t \ Leaf \ bst t \ \x' \ set_tree t. splay x' t = splay x t \ T_splay x' t = T_splay x t" proof(induction x t rule: splay.induct) case (6 x b c A) hence "splay x A \ Leaf" by simp then obtain A1 u A2 where [simp]: "splay x A = Node A1 u A2" by (metis tree.exhaust) have "b < c" "bst A" using "6.prems" by auto from "6.IH"[OF \A \ Leaf\ \bst A\] obtain x' where "x' \ set_tree A" "splay x' A = splay x A" "T_splay x' A = T_splay x A" by blast moreover hence "x'x \x \b \bst A\ by force next case (8 a x c B) hence "splay x B \ Leaf" by simp then obtain B1 u B2 where [simp]: "splay x B = Node B1 u B2" by (metis tree.exhaust) have "a < c" "bst B" using "8.prems" by auto from "8.IH"[OF \B \ Leaf\ \bst B\] obtain x' where "x' \ set_tree B" "splay x' B = splay x B" "T_splay x' B = T_splay x B" by blast moreover hence "ax \a \a \bst B\ by force next case (11 b x c C) hence "splay x C \ Leaf" by simp then obtain C1 u C2 where [simp]: "splay x C = Node C1 u C2" by (metis tree.exhaust) have "b < c" "bst C" using "11.prems" by auto from "11.IH"[OF \C \ Leaf\ \bst C\] obtain x' where "x' \ set_tree C" "splay x' C = splay x C" "T_splay x' C = T_splay x C" by blast moreover hence "bb \x \b \bst C\ by force next case (14 a x c D) hence "splay x D \ Leaf" by simp then obtain D1 u D2 where [simp]: "splay x D = Node D1 u D2" by (metis tree.exhaust) have "a < c" "bst D" using "14.prems" by auto from "14.IH"[OF \D \ Leaf\ \bst D\] obtain x' where "x' \ set_tree D" "splay x' D = splay x D" "T_splay x' D = T_splay x D" by blast moreover hence "ca \c \a \bst D\ by force qed (auto simp: le_less) datatype 'a op = Empty | Splay 'a | Insert 'a | Delete 'a fun arity :: "'a::linorder op \ nat" where "arity Empty = 0" | "arity (Splay x) = 1" | "arity (Insert x) = 1" | "arity (Delete x) = 1" fun exec :: "'a::linorder op \ 'a tree list \ 'a tree" where "exec Empty [] = Leaf" | "exec (Splay x) [t] = splay x t" | "exec (Insert x) [t] = Splay_Tree.insert x t" | "exec (Delete x) [t] = Splay_Tree.delete x t" fun cost :: "'a::linorder op \ 'a tree list \ nat" where "cost Empty [] = 1" | "cost (Splay x) [t] = T_splay x t" | "cost (Insert x) [t] = T_insert x t" | "cost (Delete x) [t] = T_delete x t" end diff --git a/thys/Amortized_Complexity/Splay_Tree_Analysis_Optimal.thy b/thys/Amortized_Complexity/Splay_Tree_Analysis_Optimal.thy --- a/thys/Amortized_Complexity/Splay_Tree_Analysis_Optimal.thy +++ b/thys/Amortized_Complexity/Splay_Tree_Analysis_Optimal.thy @@ -1,495 +1,495 @@ subsection "Splay Tree Analysis (Optimal)" theory Splay_Tree_Analysis_Optimal imports Splay_Tree_Analysis_Base Amortized_Framework "HOL-Library.Sum_of_Squares" begin text\This analysis follows Schoenmakers~\<^cite>\"Schoenmakers-IPL93"\.\ subsubsection "Analysis of splay" locale Splay_Analysis = fixes \ :: real and \ :: real assumes a1[arith]: "\ > 1" assumes A1: "\1 \ x; 1 \ y; 1 \ z\ \ (x+y) * (y+z) powr \ \ (x+y) powr \ * (x+y+z)" assumes A2: "\1 \ l'; 1 \ r'; 1 \ lr; 1 \ r\ \ \ * (l'+r') * (lr+r) powr \ * (lr+r'+r) powr \ \ (l'+r') powr \ * (l'+lr+r') powr \ * (l'+lr+r'+r)" assumes A3: "\1 \ l'; 1 \ r'; 1 \ ll; 1 \ r\ \ \ * (l'+r') * (l'+ll) powr \ * (r'+r) powr \ \ (l'+r') powr \ * (l'+ll+r') powr \ * (l'+ll+r'+r)" begin lemma nl2: "\ ll \ 1; lr \ 1; r \ 1 \ \ log \ (ll + lr) + \ * log \ (lr + r) \ \ * log \ (ll + lr) + log \ (ll + lr + r)" apply(rule powr_le_cancel_iff[THEN iffD1, OF a1]) apply(simp add: powr_add mult.commute[of \] powr_powr[symmetric] A1) done definition \ :: "'a tree \ 'a tree \ real" where "\ t1 t2 = \ * log \ (size1 t1 + size1 t2)" fun \ :: "'a tree \ real" where "\ Leaf = 0" | "\ (Node l _ r) = \ l + \ r + \ l r" definition A :: "'a::linorder \ 'a tree \ real" where "A a t = T_splay a t + \(splay a t) - \ t" lemma A_simps[simp]: "A a (Node l a r) = 1" "a A a (Node (Node ll a lr) b r) = \ lr r - \ lr ll + 1" "b A a (Node l b (Node rl a rr)) = \ rl l - \ rr rl + 1" by(auto simp add: A_def \_def algebra_simps) lemma A_ub: "\ bst t; Node la a ra : subtrees t \ \ A a t \ log \ ((size1 t)/(size1 la + size1 ra)) + 1" proof(induction a t rule: splay.induct) case 1 thus ?case by simp next case 2 thus ?case by auto next case 4 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 5 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 7 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 10 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 12 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case 13 hence False by(fastforce dest: in_set_tree_if) thus ?case .. next case (3 b a lb rb ra) have "b \ set_tree ra" using "3.prems"(1) by auto thus ?case using "3.prems"(1,2) nl2[of "size1 lb" "size1 rb" "size1 ra"] by (auto simp: \_def log_divide algebra_simps) next case (9 a b la lb rb) have "b \ set_tree la" using "9.prems"(1) by auto thus ?case using "9.prems"(1,2) nl2[of "size1 rb" "size1 lb" "size1 la"] by (auto simp add: \_def log_divide algebra_simps) next case (6 x b a lb rb ra) hence 0: "x \ set_tree rb \ x \ set_tree ra" using "6.prems"(1) by auto hence 1: "x \ set_tree lb" using "6.prems" \x by (auto) then obtain lu u ru where sp: "splay x lb = Node lu u ru" using "6.prems"(1,2) by(cases "splay x lb") auto have "b < a" using "6.prems"(1,2) by (auto) let ?lu = "real (size1 lu)" let ?ru = "real (size1 ru)" let ?rb = "real(size1 rb)" let ?r = "real(size1 ra)" have "1 + log \ (?lu + ?ru) + \ * log \ (?rb + ?r) + \ * log \ (?rb + ?ru + ?r) \ \ * log \ (?lu + ?ru) + \ * log \ (?lu + ?rb + ?ru) + log \ (?lu + ?rb + ?ru + ?r)" (is "?L\?R") proof(rule powr_le_cancel_iff[THEN iffD1, OF a1]) show "\ powr ?L \ \ powr ?R" using A2[of ?lu ?ru ?rb ?r] by(simp add: powr_add add_ac mult.commute[of \] powr_powr[symmetric]) qed thus ?case using 6 0 1 sp by(auto simp: A_def \_def size_if_splay algebra_simps log_divide) next case (8 b x a rb lb ra) hence 0: "x \ set_tree lb \ x \ set_tree ra" using "8.prems"(1) by(auto) hence 1: "x \ set_tree rb" using "8.prems" \b \x by (auto) then obtain lu u ru where sp: "splay x rb = Node lu u ru" using "8.prems"(1,2) by(cases "splay x rb") auto let ?lu = "real (size1 lu)" let ?ru = "real (size1 ru)" let ?lb = "real(size1 lb)" let ?r = "real(size1 ra)" have "1 + log \ (?lu + ?ru) + \ * log \ (?lu + ?lb) + \ * log \ (?ru + ?r) \ \ * log \ (?lu + ?ru) + \ * log \ (?lu + ?lb + ?ru) + log \ (?lu + ?lb + ?ru + ?r)" (is "?L\?R") proof(rule powr_le_cancel_iff[THEN iffD1, OF a1]) show "\ powr ?L \ \ powr ?R" using A3[of ?lu ?ru ?lb ?r] by(simp add: powr_add mult.commute[of \] powr_powr[symmetric]) qed thus ?case using 8 0 1 sp by(auto simp add: A_def size_if_splay \_def log_divide algebra_simps) next case (11 a x b lb la rb) hence 0: "x \ set_tree rb \ x \ set_tree la" using "11.prems"(1) by (auto) hence 1: "x \ set_tree lb" using "11.prems" \a \x by (auto) then obtain lu u ru where sp: "splay x lb = Node lu u ru" using "11.prems"(1,2) by(cases "splay x lb") auto let ?lu = "real (size1 lu)" let ?ru = "real (size1 ru)" let ?l = "real(size1 la)" let ?rb = "real(size1 rb)" have "1 + log \ (?lu + ?ru) + \ * log \ (?lu + ?l) + \ * log \ (?ru + ?rb) \ \ * log \ (?lu + ?ru) + \ * log \ (?lu + ?ru + ?rb) + log \ (?lu + ?l + ?ru + ?rb)" (is "?L\?R") proof(rule powr_le_cancel_iff[THEN iffD1, OF a1]) show "\ powr ?L \ \ powr ?R" using A3[of ?ru ?lu ?rb ?l] by(simp add: powr_add mult.commute[of \] powr_powr[symmetric]) (simp add: algebra_simps) qed thus ?case using 11 0 1 sp by(auto simp add: A_def size_if_splay \_def log_divide algebra_simps) next case (14 a x b rb la lb) hence 0: "x \ set_tree lb \ x \ set_tree la" using "14.prems"(1) by(auto) hence 1: "x \ set_tree rb" using "14.prems" \a \b by (auto) then obtain lu u ru where sp: "splay x rb = Node lu u ru" using "14.prems"(1,2) by(cases "splay x rb") auto let ?la = "real(size1 la)" let ?lb = "real(size1 lb)" let ?lu = "real (size1 lu)" let ?ru = "real (size1 ru)" have "1 + log \ (?lu + ?ru) + \ * log \ (?la + ?lb) + \ * log \ (?lu + ?la + ?lb) \ \ * log \ (?lu + ?ru) + \ * log \ (?lu + ?lb + ?ru) + log \ (?lu + ?lb + ?ru + ?la)" (is "?L\?R") proof(rule powr_le_cancel_iff[THEN iffD1, OF a1]) show "\ powr ?L \ \ powr ?R" using A2[of ?ru ?lu ?lb ?la] by(simp add: powr_add add_ac mult.commute[of \] powr_powr[symmetric]) qed thus ?case using 14 0 1 sp by(auto simp add: A_def size_if_splay \_def log_divide algebra_simps) qed lemma A_ub2: assumes "bst t" "a : set_tree t" shows "A a t \ log \ ((size1 t)/2) + 1" proof - from assms(2) obtain la ra where N: "Node la a ra : subtrees t" by (metis set_treeE) have "A a t \ log \ ((size1 t)/(size1 la + size1 ra)) + 1" by(rule A_ub[OF assms(1) N]) also have "\ \ log \ ((size1 t)/2) + 1" by(simp add: field_simps) finally show ?thesis by simp qed lemma A_ub3: assumes "bst t" shows "A a t \ log \ (size1 t) + 1" proof cases assume "t = Leaf" thus ?thesis by(simp add: A_def) next assume "t \ Leaf" from ex_in_set_tree[OF this assms] obtain a' where a': "a' \ set_tree t" "splay a' t = splay a t" "T_splay a' t = T_splay a t" by blast have [arith]: "log \ 2 > 0" by simp show ?thesis using A_ub2[OF assms a'(1)] by(simp add: A_def a' log_divide) qed definition Am :: "'a::linorder tree \ real" where "Am t = T_splay_max t + \(splay_max t) - \ t" lemma Am_simp3': "\ c Leaf\ \ Am (Node l c (Node rl b rr)) = (case splay_max rr of Node rrl _ rrr \ Am rr + \ rrl (Node l c rl) + \ l rl - \ rl rr - \ rrl rrr + 1)" by(auto simp: Am_def \_def size_if_splay_max algebra_simps neq_Leaf_iff split: tree.split) lemma Am_ub: "\ bst t; t \ Leaf \ \ Am t \ log \ ((size1 t)/2) + 1" proof(induction t rule: splay_max.induct) case 1 thus ?case by (simp) next case 2 thus ?case by (simp add: Am_def) next case (3 l b rl c rr) show ?case proof cases assume "rr = Leaf" thus ?thesis using nl2[of 1 "size1 rl" "size1 l"] log_le_cancel_iff[of \ 2 "2 + real(size rl)"] by(auto simp: Am_def \_def log_divide field_simps simp del: log_le_cancel_iff) next assume "rr \ Leaf" then obtain l' u r' where sp: "splay_max rr = Node l' u r'" using splay_max_Leaf_iff tree.exhaust by blast hence 1: "size rr = size l' + size r' + 1" using size_splay_max[of rr] by(simp) let ?l = "real (size1 l)" let ?rl = "real (size1 rl)" let ?l' = "real (size1 l')" let ?r' = "real (size1 r')" have "1 + log \ (?l' + ?r') + \ * log \ (?l + ?rl) + \ * log \ (?l' + ?l + ?rl) \ \ * log \ (?l' + ?r') + \ * log \ (?l' + ?rl + ?r') + log \ (?l' + ?rl + ?r' + ?l)" (is "?L\?R") proof(rule powr_le_cancel_iff[THEN iffD1, OF a1]) show "\ powr ?L \ \ powr ?R" using A2[of ?r' ?l' ?rl ?l] by(simp add: powr_add add.commute add.left_commute mult.commute[of \] powr_powr[symmetric]) qed thus ?case using 3 sp 1 \rr \ Leaf\ by(auto simp add: Am_simp3' \_def log_divide algebra_simps) qed qed lemma Am_ub3: assumes "bst t" shows "Am t \ log \ (size1 t) + 1" proof cases assume "t = Leaf" thus ?thesis by(simp add: Am_def) next assume "t \ Leaf" have [arith]: "log \ 2 > 0" by simp show ?thesis using Am_ub[OF assms \t \ Leaf\] by(simp add: Am_def log_divide) qed end subsubsection "Optimal Interpretation" lemma mult_root_eq_root: "n>0 \ y \ 0 \ root n x * y = root n (x * (y ^ n))" by(simp add: real_root_mult real_root_pos2) lemma mult_root_eq_root2: "n>0 \ y \ 0 \ y * root n x = root n ((y ^ n) * x)" by(simp add: real_root_mult real_root_pos2) lemma powr_inverse_numeral: "0 < x \ x powr (1 / numeral n) = root (numeral n) x" by (simp add: root_powr_inverse) lemmas root_simps = mult_root_eq_root mult_root_eq_root2 powr_inverse_numeral lemma nl31: "\ (l'::real) \ 1; r' \ 1; lr \ 1; r \ 1 \ \ 4 * (l' + r') * (lr + r) \ (l' + lr + r' + r)^2" by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [r + ~1*l' + lr + ~1*r']^2))))") lemma nl32: assumes "(l'::real) \ 1" "r' \ 1" "lr \ 1" "r \ 1" shows "4 * (l' + r') * (lr + r) * (lr + r' + r) \ (l' + lr + r' + r)^3" proof - have 1: "lr + r' + r \ l' + lr + r' + r" using assms by arith have 2: "0 \ (l' + lr + r' + r)^2" by (rule zero_le_power2) have 3: "0 \ lr + r' + r" using assms by arith from mult_mono[OF nl31[OF assms] 1 2 3] show ?thesis by(simp add: ac_simps numeral_eq_Suc) qed lemma nl3: assumes "(l'::real) \ 1" "r' \ 1" "lr \ 1" "r \ 1" shows "4 * (l' + r')^2 * (lr + r) * (lr + r' + r) \ (l' + lr + r') * (l' + lr + r' + r)^3" proof- have 1: "l' + r' \ l' + lr + r'" using assms by arith have 2: "0 \ (l' + lr + r' + r)^3" using assms by simp have 3: "0 \ l' + r'" using assms by arith from mult_mono[OF nl32[OF assms] 1 2 3] show ?thesis unfolding power2_eq_square by (simp only: ac_simps) qed lemma nl41: assumes "(l'::real) \ 1" "r' \ 1" "ll \ 1" "r \ 1" shows "4 * (l' + ll) * (r' + r) \ (l' + ll + r' + r)^2" using assms by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [r + ~1*l' + ~1*ll + r']^2))))") lemma nl42: assumes "(l'::real) \ 1" "r' \ 1" "ll \ 1" "r \ 1" shows "4 * (l' + r') * (l' + ll) * (r' + r) \ (l' + ll + r' + r)^3" proof - have 1: "l' + r' \ l' + ll + r' + r" using assms by arith have 2: "0 \ (l' + ll + r' + r)^2" by (rule zero_le_power2) have 3: "0 \ l' + r'" using assms by arith from mult_mono[OF nl41[OF assms] 1 2 3] show ?thesis by(simp add: ac_simps numeral_eq_Suc del: distrib_right_numeral) qed lemma nl4: assumes "(l'::real) \ 1" "r' \ 1" "ll \ 1" "r \ 1" shows "4 * (l' + r')^2 * (l' + ll) * (r' + r) \ (l' + ll + r') * (l' + ll + r' + r)^3" proof- have 1: "l' + r' \ l' + ll + r'" using assms by arith have 2: "0 \ (l' + ll + r' + r)^3" using assms by simp have 3: "0 \ l' + r'" using assms by arith from mult_mono[OF nl42[OF assms] 1 2 3] show ?thesis unfolding power2_eq_square by (simp only: ac_simps) qed lemma cancel: "x>(0::real) \ c * x ^ 2 * y * z \ u * v \ c * x ^ 3 * y * z \ x * u * v" by(simp add: power2_eq_square power3_eq_cube) interpretation S34: Splay_Analysis "root 3 4" "1/3" proof (standard, goal_cases) case 2 thus ?case by(simp add: root_simps) (auto simp: numeral_eq_Suc split_mult_pos_le intro!: mult_mono) next case 3 thus ?case by(simp add: root_simps cancel nl3) next case 4 thus ?case by(simp add: root_simps cancel nl4) qed simp lemma log4_log2: "log 4 x = log 2 x / 2" proof - have "log 4 x = log (2^2) x" by simp also have "\ = log 2 x / 2" by(simp only: log_base_pow) finally show ?thesis . qed declare log_base_root[simp] lemma A34_ub: assumes "bst t" shows "S34.A a t \ (3/2) * log 2 (size1 t) + 1" proof - have "S34.A a t \ log (root 3 4) (size1 t) + 1" by(rule S34.A_ub3[OF assms]) also have "\ = (3/2) * log 2 (size t + 1) + 1" by(simp add: log4_log2) finally show ?thesis by simp qed lemma Am34_ub: assumes "bst t" shows "S34.Am t \ (3/2) * log 2 (size1 t) + 1" proof - have "S34.Am t \ log (root 3 4) (size1 t) + 1" by(rule S34.Am_ub3[OF assms]) also have "\ = (3/2) * log 2 (size1 t) + 1" by(simp add: log4_log2) finally show ?thesis by simp qed subsubsection "Overall analysis" fun U where "U Empty [] = 1" | "U (Splay _) [t] = (3/2) * log 2 (size1 t) + 1" | "U (Insert _) [t] = 2 * log 2 (size1 t) + 3/2" | "U (Delete _) [t] = 3 * log 2 (size1 t) + 2" interpretation Amortized where arity = arity and exec = exec and inv = bst and cost = cost and \ = S34.\ and U = U proof (standard, goal_cases) case (1 ss f) show ?case proof (cases f) case Empty thus ?thesis using 1 by auto next case (Splay a) then obtain t where "ss = [t]" "bst t" using 1 by auto with Splay bst_splay[OF \bst t\, of a] show ?thesis by (auto split: tree.split) next case (Insert a) then obtain t where "ss = [t]" "bst t" using 1 by auto with bst_splay[OF \bst t\, of a] Insert show ?thesis by (auto simp: splay_bstL[OF \bst t\] splay_bstR[OF \bst t\] split: tree.split) next case (Delete a) then obtain t where "ss = [t]" "bst t" using 1 by auto with 1 Delete show ?thesis by(simp add: bst_delete) qed next case (2 t) show ?case by(induction t)(simp_all add: S34.\_def) next case (3 ss f) show ?case (is "?l \ ?r") proof(cases f) case Empty thus ?thesis using 3(2) by(simp add: S34.A_def) next case (Splay a) then obtain t where "ss = [t]" "bst t" using 3 by auto thus ?thesis using S34.A_ub3[OF \bst t\] Splay by(simp add: S34.A_def log4_log2) next case [simp]: (Insert a) obtain t where [simp]: "ss = [t]" and "bst t" using 3 by auto show ?thesis proof cases - assume "t = Leaf" thus ?thesis by(simp add: S34.\_def log4_log2 T_insert_def) + assume "t = Leaf" thus ?thesis by(simp add: S34.\_def log4_log2) next assume "t \ Leaf" then obtain l e r where [simp]: "splay a t = Node l e r" by (metis tree.exhaust splay_Leaf_iff) let ?t = "real(T_splay a t)" let ?Plr = "S34.\ l + S34.\ r" let ?Ps = "S34.\ t" let ?slr = "real(size1 l) + real(size1 r)" let ?LR = "log 2 (1 + ?slr)" have opt: "?t + S34.\ (splay a t) - ?Ps \ 3/2 * log 2 (real (size1 t)) + 1" using S34.A_ub3[OF \bst t\, simplified S34.A_def, of a] by (simp add: log4_log2) from less_linear[of e a] show ?thesis proof (elim disjE) assume "e=a" have nneg: "log 2 (1 + real (size t)) \ 0" by simp thus ?thesis using \t \ Leaf\ opt \e=a\ - apply(simp add: field_simps T_insert_def) using nneg by arith + apply(simp add: field_simps) using nneg by arith next let ?L = "log 2 (real(size1 l) + 1)" assume "e a" by simp hence "?l = (?t + ?Plr - ?Ps) + ?L / 2 + ?LR / 2" - using \t \ Leaf\ \e by(simp add: S34.\_def log4_log2 T_insert_def) + using \t \ Leaf\ \e by(simp add: S34.\_def log4_log2) also have "?t + ?Plr - ?Ps \ log 2 ?slr + 1" using opt size_splay[of a t,symmetric] by(simp add: S34.\_def log4_log2) also have "?L/2 \ log 2 ?slr / 2" by(simp) also have "?LR/2 \ log 2 ?slr / 2 + 1/2" proof - have "?LR/2 \ log 2 (2 * ?slr) / 2" by simp also have "\ \ log 2 ?slr / 2 + 1/2" by (simp add: log_mult del:distrib_left_numeral) finally show ?thesis . qed finally show ?thesis using size_splay[of a t,symmetric] by simp next let ?R = "log 2 (2 + real(size r))" assume "a a" by simp hence "?l = (?t + ?Plr - ?Ps) + ?R / 2 + ?LR / 2" - using \t \ Leaf\ \a by(simp add: S34.\_def log4_log2 T_insert_def) + using \t \ Leaf\ \a by(simp add: S34.\_def log4_log2) also have "?t + ?Plr - ?Ps \ log 2 ?slr + 1" using opt size_splay[of a t,symmetric] by(simp add: S34.\_def log4_log2) also have "?R/2 \ log 2 ?slr / 2" by(simp) also have "?LR/2 \ log 2 ?slr / 2 + 1/2" proof - have "?LR/2 \ log 2 (2 * ?slr) / 2" by simp also have "\ \ log 2 ?slr / 2 + 1/2" by (simp add: log_mult del:distrib_left_numeral) finally show ?thesis . qed finally show ?thesis using size_splay[of a t,symmetric] by simp qed qed next case [simp]: (Delete a) obtain t where [simp]: "ss = [t]" and "bst t" using 3 by auto show ?thesis proof (cases t) case Leaf thus ?thesis - by(simp add: Splay_Tree.delete_def T_delete_def S34.\_def log4_log2) + by(simp add: Splay_Tree.delete_def S34.\_def log4_log2) next case [simp]: (Node ls x rs) then obtain l e r where sp[simp]: "splay a (Node ls x rs) = Node l e r" by (metis tree.exhaust splay_Leaf_iff) let ?t = "real(T_splay a t)" let ?Plr = "S34.\ l + S34.\ r" let ?Ps = "S34.\ t" let ?slr = "real(size1 l) + real(size1 r)" let ?LR = "log 2 (1 + ?slr)" let ?lslr = "log 2 (real (size ls) + (real (size rs) + 2))" have "?lslr \ 0" by simp have opt: "?t + S34.\ (splay a t) - ?Ps \ 3/2 * log 2 (real (size1 t)) + 1" using S34.A_ub3[OF \bst t\, simplified S34.A_def, of a] by (simp add: log4_log2 field_simps) show ?thesis proof (cases "e=a") case False thus ?thesis using opt - apply(simp add: Splay_Tree.delete_def T_delete_def field_simps) + apply(simp add: Splay_Tree.delete_def field_simps) using \?lslr \ 0\ by arith next case [simp]: True show ?thesis proof (cases l) case Leaf have "S34.\ Leaf r \ 0" by(simp add: S34.\_def) thus ?thesis using Leaf opt - apply(simp add: Splay_Tree.delete_def T_delete_def field_simps) + apply(simp add: Splay_Tree.delete_def field_simps) using \?lslr \ 0\ by arith next case (Node ll y lr) then obtain l' y' r' where [simp]: "splay_max (Node ll y lr) = Node l' y' r'" using splay_max_Leaf_iff tree.exhaust by blast have "bst l" using bst_splay[OF \bst t\, of a] by simp have "S34.\ r' \ 0" apply (induction r') by (auto simp add: S34.\_def) have optm: "real(T_splay_max l) + S34.\ (splay_max l) - S34.\ l \ 3/2 * log 2 (real (size1 l)) + 1" using S34.Am_ub3[OF \bst l\, simplified S34.Am_def] by (simp add: log4_log2 field_simps Node) have 1: "log 4 (2+(real(size l')+real(size r))) \ log 4 (2+(real(size l)+real(size r)))" using size_splay_max[of l] Node by simp have 2: "log 4 (2 + (real (size l') + real (size r'))) \ 0" by simp have 3: "S34.\ l' r \ S34.\ l' r' + S34.\ l r" apply(simp add: S34.\_def) using 1 2 by arith have 4: "log 2 (real(size ll) + (real(size lr) + 2)) \ ?lslr" using size_if_splay[OF sp] Node by simp show ?thesis using add_mono[OF opt optm] Node 3 - apply(simp add: Splay_Tree.delete_def T_delete_def field_simps) + apply(simp add: Splay_Tree.delete_def field_simps) using 4 \S34.\ r' \ 0\ by arith qed qed qed qed qed end diff --git a/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm.thy b/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm.thy --- a/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm.thy +++ b/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm.thy @@ -1,222 +1,225 @@ subsection \Pure pattern matching rule sets\ theory Rewriting_Pterm imports Rewriting_Pterm_Elim begin type_synonym prule = "name \ pterm" primrec prule :: "prule \ bool" where "prule (_, rhs) \ wellformed rhs \ closed rhs \ is_abs rhs" lemma pruleI[intro!]: "wellformed rhs \ closed rhs \ is_abs rhs \ prule (name, rhs)" by simp locale prules = constants C_info "fst |`| rs" for C_info and rs :: "prule fset" + assumes all_rules: "fBall rs prule" assumes fmap: "is_fmap rs" assumes not_shadows: "fBall rs (\(_, rhs). \ shadows_consts rhs)" assumes welldefined_rs: "fBall rs (\(_, rhs). welldefined rhs)" subsubsection \Rewriting\ inductive prewrite :: "prule fset \ pterm \ pterm \ bool" ("_/ \\<^sub>p/ _ \/ _" [50,0,50] 50) for rs where step: "(name, rhs) |\| rs \ rs \\<^sub>p Pconst name \ rhs" | beta: "c |\| cs \ c \ t \ t' \ rs \\<^sub>p Pabs cs $\<^sub>p t \ t'" | "fun": "rs \\<^sub>p t \ t' \ rs \\<^sub>p t $\<^sub>p u \ t' $\<^sub>p u" | arg: "rs \\<^sub>p u \ u' \ rs \\<^sub>p t $\<^sub>p u \ t $\<^sub>p u'" global_interpretation prewrite: rewriting "prewrite rs" for rs by standard (auto intro: prewrite.intros simp: app_pterm_def)+ abbreviation prewrite_rt :: "prule fset \ pterm \ pterm \ bool" ("_/ \\<^sub>p/ _ \*/ _" [50,0,50] 50) where "prewrite_rt rs \ (prewrite rs)\<^sup>*\<^sup>*" subsubsection \Translation from @{typ irule_set} to @{typ "prule fset"}\ definition finished :: "irule_set \ bool" where "finished rs = fBall rs (\(_, irs). arity irs = 0)" definition translate_rhs :: "irules \ pterm" where "translate_rhs = snd \ fthe_elem" definition compile :: "irule_set \ prule fset" where "compile = fimage (map_prod id translate_rhs)" lemma compile_heads: "fst |`| compile rs = fst |`| rs" unfolding compile_def by simp subsubsection \Correctness of translation\ lemma arity_zero_shape: assumes "arity_compatibles rs" "arity rs = 0" "is_fmap rs" "rs \ {||}" obtains t where "rs = {| ([], t) |}" proof - from assms obtain ppats prhs where "(ppats, prhs) |\| rs" by fast moreover { fix pats rhs assume "(pats, rhs) |\| rs" with assms have "length pats = 0" by (metis arity_compatible_length) hence "pats = []" by simp } note all = this ultimately have proto: "([], prhs) |\| rs" by auto have "fBall rs (\(pats, rhs). pats = [] \ rhs = prhs)" proof safe fix pats rhs assume cur: "(pats, rhs) |\| rs" with all show "pats = []" . with cur have "([], rhs) |\| rs" by auto with proto show "rhs = prhs" using assms by (auto dest: is_fmapD) qed hence "fBall rs (\r. r = ([], prhs))" by blast with assms have "rs = {| ([], prhs) |}" by (simp add: singleton_fset_is) thus thesis by (rule that) qed lemma (in irules) compile_rules: assumes "finished rs" shows "prules C_info (compile rs)" proof show "is_fmap (compile rs)" using fmap unfolding compile_def map_prod_def id_apply by (rule is_fmap_image) next show "fdisjnt (fst |`| compile rs) C" unfolding compile_def using disjnt by simp next have "fBall (compile rs) prule" "fBall (compile rs) (\(_, rhs). \ shadows_consts rhs)" "fBall (compile rs) (\(_, rhs). welldefined rhs)" proof (safe del: fsubsetI) fix name rhs assume "(name, rhs) |\| compile rs" (* FIXME clone of compile_correct *) then obtain irs where "(name, irs) |\| rs" "rhs = translate_rhs irs" unfolding compile_def by force hence "is_fmap irs" "irs \ {||}" "arity irs = 0" using assms inner unfolding finished_def by blast+ moreover have "arity_compatibles irs" - using \(name, irs) |\| rs\ inner by (blast dest: fpairwiseD) + using \(name, irs) |\| rs\ inner + by (metis (no_types, lifting) case_prodD) ultimately obtain u where "irs = {| ([], u) |}" by (metis arity_zero_shape) hence "rhs = u" and u: "([], u) |\| irs" unfolding \rhs = _\ translate_rhs_def by simp+ hence "abs_ish [] u" - using inner \(name, irs) |\| rs\ by blast + using inner \(name, irs) |\| rs\ by fastforce thus "is_abs rhs" unfolding abs_ish_def \rhs = u\ by simp show "wellformed rhs" using u \(name, irs) |\| rs\ inner unfolding \rhs = u\ - by blast + by fastforce have "closed_except u {||}" using u inner \(name, irs) |\| rs\ - by (metis (mono_tags, lifting) case_prod_conv fbspec freess_empty) + by fastforce thus "closed rhs" unfolding \rhs = u\ . { assume "shadows_consts rhs" hence "shadows_consts u" unfolding compile_def \rhs = u\ by simp moreover have "\ shadows_consts u" - using inner \([], u) |\| irs\ \(name, irs) |\| rs\ by blast + using inner \([], u) |\| irs\ \(name, irs) |\| rs\ by fastforce ultimately show False by blast } have "welldefined u" using fbspec[OF inner \(name, irs) |\| rs\, simplified] \([], u) |\| irs\ by blast thus "welldefined rhs" unfolding \rhs = u\ compile_def by simp qed thus "fBall (compile rs) prule" "fBall (compile rs) (\(_, rhs). \ pre_constants.shadows_consts C_info (fst |`| compile rs) rhs)" "fBall (compile rs) (\(_, rhs). pre_constants.welldefined C_info (fst |`| compile rs) rhs)" unfolding compile_heads by auto next show "distinct all_constructors" by (fact distinct_ctr) qed theorem (in irules) compile_correct: assumes "compile rs \\<^sub>p t \ t'" "finished rs" shows "rs \\<^sub>i t \ t'" using assms(1) proof induction case (step name rhs) then obtain irs where "rhs = translate_rhs irs" "(name, irs) |\| rs" unfolding compile_def by force hence "arity_compatibles irs" - using inner by (blast dest: fpairwiseD) + using inner + by (metis (no_types, lifting) case_prodD) have "is_fmap irs" "irs \ {||}" "arity irs = 0" using assms inner \(name, irs) |\| rs\ unfolding finished_def by blast+ then obtain u where "irs = {| ([], u) |}" using \arity_compatibles irs\ by (metis arity_zero_shape) show ?case unfolding \rhs = _\ apply (rule irewrite.step) apply fact unfolding \irs = _\ translate_rhs_def irewrite_step_def by (auto simp: const_term_def) qed (auto intro: irewrite.intros) theorem (in irules) compile_complete: assumes "rs \\<^sub>i t \ t'" "finished rs" shows "compile rs \\<^sub>p t \ t'" using assms(1) proof induction case (step name irs params rhs t t') hence "arity_compatibles irs" - using inner by (blast dest: fpairwiseD) + using inner + by (metis (no_types, lifting) case_prodD) have "is_fmap irs" "irs \ {||}" "arity irs = 0" using assms inner step unfolding finished_def by blast+ then obtain u where "irs = {| ([], u) |}" using \arity_compatibles irs\ by (metis arity_zero_shape) with step have "name, [], u \\<^sub>i t \ t'" by simp hence "t = Pconst name" unfolding irewrite_step_def by (cases t) (auto split: if_splits simp: const_term_def) hence "t' = u" using \name, [], u \\<^sub>i t \ t'\ unfolding irewrite_step_def by (cases t) (auto split: if_splits simp: const_term_def) have "(name, t') |\| compile rs" unfolding compile_def proof show "(name, t') = map_prod id translate_rhs (name, irs)" using \irs = _\ \t' = u\ by (simp add: split_beta translate_rhs_def) qed fact thus ?case unfolding \t = _\ by (rule prewrite.step) qed (auto intro: prewrite.intros) export_code compile finished checking Scala end \ No newline at end of file diff --git a/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy b/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy --- a/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy +++ b/thys/CakeML_Codegen/Rewriting/Rewriting_Pterm_Elim.thy @@ -1,1750 +1,1750 @@ section \Higher-order term rewriting with explicit pattern matching\ theory Rewriting_Pterm_Elim imports Rewriting_Nterm "../Terms/Pterm" begin subsection \Intermediate rule sets\ type_synonym irules = "(term list \ pterm) fset" type_synonym irule_set = "(name \ irules) fset" locale pre_irules = constants C_info "fst |`| rs" for C_info and rs :: "irule_set" locale irules = pre_irules + assumes fmap: "is_fmap rs" assumes nonempty: "rs \ {||}" assumes inner: "fBall rs (\(_, irs). arity_compatibles irs \ is_fmap irs \ patterns_compatibles irs \ irs \ {||} \ fBall irs (\(pats, rhs). linears pats \ abs_ish pats rhs \ closed_except rhs (freess pats) \ fdisjnt (freess pats) all_consts \ wellformed rhs \ \ shadows_consts rhs \ welldefined rhs))" lemma (in pre_irules) irulesI: assumes "\name irs. (name, irs) |\| rs \ arity_compatibles irs" assumes "\name irs. (name, irs) |\| rs \ is_fmap irs" assumes "\name irs. (name, irs) |\| rs \ patterns_compatibles irs" assumes "\name irs. (name, irs) |\| rs \ irs \ {||}" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ linears pats" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ abs_ish pats rhs" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ fdisjnt (freess pats) all_consts" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ closed_except rhs (freess pats)" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ wellformed rhs" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ \ shadows_consts rhs" assumes "\name irs pats rhs. (name, irs) |\| rs \ (pats, rhs) |\| irs \ welldefined rhs" assumes "is_fmap rs" "rs \ {||}" shows "irules C_info rs" proof unfold_locales show "is_fmap rs" using assms(12) . next show "rs \ {||}" using assms(13) . next show "fBall rs (\(_, irs). Rewriting_Nterm.arity_compatibles irs \ is_fmap irs \ patterns_compatibles irs \ irs \ {||} \ fBall irs (\(pats, rhs). linears pats \ abs_ish pats rhs \ closed_except rhs (freess pats) \ fdisjnt (freess pats) all_consts \ pre_strong_term_class.wellformed rhs \ \ shadows_consts rhs \ consts rhs |\| all_consts))" using assms(1-11) by (intro prod_BallI conjI) metis+ qed lemmas irulesI[intro!] = pre_irules.irulesI[unfolded pre_irules_def] subsubsection \Translation from @{typ nterm} to @{typ pterm}\ fun nterm_to_pterm :: "nterm \ pterm" where "nterm_to_pterm (Nvar s) = Pvar s" | "nterm_to_pterm (Nconst s) = Pconst s" | "nterm_to_pterm (t\<^sub>1 $\<^sub>n t\<^sub>2) = nterm_to_pterm t\<^sub>1 $\<^sub>p nterm_to_pterm t\<^sub>2" | "nterm_to_pterm (\\<^sub>n x. t) = (\\<^sub>p x. nterm_to_pterm t)" lemma nterm_to_pterm_inj: "nterm_to_pterm x = nterm_to_pterm y \ x = y" by (induction y arbitrary: x) (auto elim: nterm_to_pterm.elims) lemma nterm_to_pterm: assumes "no_abs t" shows "nterm_to_pterm t = convert_term t" using assms apply induction apply auto by (auto simp: free_nterm_def free_pterm_def const_nterm_def const_pterm_def app_nterm_def app_pterm_def) lemma nterm_to_pterm_frees[simp]: "frees (nterm_to_pterm t) = frees t" by (induct t) auto lemma closed_nterm_to_pterm[intro]: "closed_except (nterm_to_pterm t) (frees t)" unfolding closed_except_def by simp lemma (in constants) shadows_nterm_to_pterm[simp]: "shadows_consts (nterm_to_pterm t) = shadows_consts t" by (induct t) (auto simp: shadows_consts_def fdisjnt_alt_def) lemma wellformed_nterm_to_pterm[intro]: "wellformed (nterm_to_pterm t)" by (induct t) auto lemma consts_nterm_to_pterm[simp]: "consts (nterm_to_pterm t) = consts t" by (induct t) auto subsubsection \Translation from @{typ crule_set} to @{typ irule_set}\ definition translate_crules :: "crules \ irules" where "translate_crules = fimage (map_prod id nterm_to_pterm)" definition compile :: "crule_set \ irule_set" where "compile = fimage (map_prod id translate_crules)" lemma compile_heads: "fst |`| compile rs = fst |`| rs" unfolding compile_def by simp lemma (in crules) compile_rules: "irules C_info (compile rs)" proof have "is_fmap rs" using fmap by simp thus "is_fmap (compile rs)" unfolding compile_def map_prod_def id_apply by (rule is_fmap_image) show "compile rs \ {||}" using nonempty unfolding compile_def by auto show "constants C_info (fst |`| compile rs)" proof show "fdisjnt (fst |`| compile rs) C" using disjnt unfolding compile_def by force next show "distinct all_constructors" by (fact distinct_ctr) qed fix name irs assume irs: "(name, irs) |\| compile rs" then obtain irs' where "(name, irs') |\| rs" "irs = translate_crules irs'" unfolding compile_def by force hence "arity_compatibles irs'" using inner by (metis (no_types, lifting) case_prodD) thus "arity_compatibles irs" unfolding \irs = translate_crules irs'\ translate_crules_def by (force dest: fpairwiseD) have "patterns_compatibles irs'" using \(name, irs') |\| rs\ inner by (blast dest: fpairwiseD) thus "patterns_compatibles irs" unfolding \irs = _\ translate_crules_def by (auto dest: fpairwiseD) have "is_fmap irs'" using \(name, irs') |\| rs\ inner by auto thus "is_fmap irs" unfolding \irs = translate_crules irs'\ translate_crules_def map_prod_def id_apply by (rule is_fmap_image) have "irs' \ {||}" using \(name, irs') |\| rs\ inner by auto thus "irs \ {||}" unfolding \irs = translate_crules irs'\ translate_crules_def by simp fix pats rhs assume "(pats, rhs) |\| irs" then obtain rhs' where "(pats, rhs') |\| irs'" "rhs = nterm_to_pterm rhs'" unfolding \irs = translate_crules irs'\ translate_crules_def by force hence "linears pats" "pats \ []" "frees rhs' |\| freess pats" "\ shadows_consts rhs'" using fbspec[OF inner \(name, irs') |\| rs\] by blast+ show "linears pats" by fact show "closed_except rhs (freess pats)" unfolding \rhs = nterm_to_pterm rhs'\ using \frees rhs' |\| freess pats\ by (metis dual_order.trans closed_nterm_to_pterm closed_except_def) show "wellformed rhs" unfolding \rhs = nterm_to_pterm rhs'\ by auto have "fdisjnt (freess pats) all_consts" using \(pats, rhs') |\| irs'\ \(name, irs') |\| rs\ inner by auto thus "fdisjnt (freess pats) (pre_constants.all_consts C_info (fst |`| compile rs))" unfolding compile_def by simp have "\ shadows_consts rhs" unfolding \rhs = _\ using \\ shadows_consts _\ by simp thus "\ pre_constants.shadows_consts C_info (fst |`| compile rs) rhs" unfolding compile_heads . show "abs_ish pats rhs" using \pats \ []\ unfolding abs_ish_def by simp have "welldefined rhs'" using fbspec[OF inner \(name, irs') |\| rs\, simplified] using \(pats, rhs') |\| irs'\ by blast thus "pre_constants.welldefined C_info (fst |`| compile rs) rhs" unfolding compile_def \rhs = _\ by simp qed sublocale crules \ crules_as_irules: irules C_info "compile rs" by (fact compile_rules) subsubsection \Transformation of @{typ irule_set}\ definition transform_irules :: "irules \ irules" where "transform_irules rs = ( if arity rs = 0 then rs else map_prod id Pabs |`| fgroup_by (\(pats, rhs). (butlast pats, (last pats, rhs))) rs)" lemma arity_compatibles_transform_irules: assumes "arity_compatibles rs" shows "arity_compatibles (transform_irules rs)" proof (cases "arity rs = 0") case True thus ?thesis unfolding transform_irules_def using assms by simp next case False let ?rs' = "transform_irules rs" let ?f = "\(pats, rhs). (butlast pats, (last pats, rhs))" let ?grp = "fgroup_by ?f rs" have rs': "?rs' = map_prod id Pabs |`| ?grp" using False unfolding transform_irules_def by simp show ?thesis proof safe fix pats\<^sub>1 rhs\<^sub>1 pats\<^sub>2 rhs\<^sub>2 assume "(pats\<^sub>1, rhs\<^sub>1) |\| ?rs'" "(pats\<^sub>2, rhs\<^sub>2) |\| ?rs'" then obtain rhs\<^sub>1' rhs\<^sub>2' where "(pats\<^sub>1, rhs\<^sub>1') |\| ?grp" "(pats\<^sub>2, rhs\<^sub>2') |\| ?grp" unfolding rs' by auto then obtain pats\<^sub>1' pats\<^sub>2' x y \ \dummies\ where "fst (?f (pats\<^sub>1', x)) = pats\<^sub>1" "(pats\<^sub>1', x) |\| rs" and "fst (?f (pats\<^sub>2', y)) = pats\<^sub>2" "(pats\<^sub>2', y) |\| rs" by (fastforce simp: split_beta elim: fgroup_byE2) hence "pats\<^sub>1 = butlast pats\<^sub>1'" "pats\<^sub>2 = butlast pats\<^sub>2'" "length pats\<^sub>1' = length pats\<^sub>2'" using assms by (force dest: fpairwiseD)+ thus "length pats\<^sub>1 = length pats\<^sub>2" by auto qed qed lemma arity_transform_irules: assumes "arity_compatibles rs" "rs \ {||}" shows "arity (transform_irules rs) = (if arity rs = 0 then 0 else arity rs - 1)" proof (cases "arity rs = 0") case True thus ?thesis unfolding transform_irules_def by simp next case False let ?f = "\(pats, rhs). (butlast pats, (last pats, rhs))" let ?grp = "fgroup_by ?f rs" let ?rs' = "map_prod id Pabs |`| ?grp" have "arity ?rs' = arity rs - 1" proof (rule arityI) show "fBall ?rs' (\(pats, _). length pats = arity rs - 1)" proof (rule prod_fBallI) fix pats rhs assume "(pats, rhs) |\| ?rs'" then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" by force then obtain pats' x \ \dummy\ where "pats = butlast pats'" "(pats', x) |\| rs" by (fastforce simp: split_beta elim: fgroup_byE2) hence "length pats' = arity rs" using assms by (metis arity_compatible_length) thus "length pats = arity rs - 1" unfolding \pats = butlast pats'\ using False by simp qed next show "?rs' \ {||}" using assms by (simp add: fgroup_by_nonempty) qed with False show ?thesis unfolding transform_irules_def by simp qed definition transform_irule_set :: "irule_set \ irule_set" where "transform_irule_set = fimage (map_prod id transform_irules)" lemma transform_irule_set_heads: "fst |`| transform_irule_set rs = fst |`| rs" unfolding transform_irule_set_def by simp lemma (in irules) rules_transform: "irules C_info (transform_irule_set rs)" proof have "is_fmap rs" using fmap by simp thus "is_fmap (transform_irule_set rs)" unfolding transform_irule_set_def map_prod_def id_apply by (rule is_fmap_image) show "transform_irule_set rs \ {||}" using nonempty unfolding transform_irule_set_def by auto show "constants C_info (fst |`| transform_irule_set rs)" proof show "fdisjnt (fst |`| transform_irule_set rs) C" using disjnt unfolding transform_irule_set_def by force next show "distinct all_constructors" by (fact distinct_ctr) qed fix name irs assume irs: "(name, irs) |\| transform_irule_set rs" then obtain irs' where "(name, irs') |\| rs" "irs = transform_irules irs'" unfolding transform_irule_set_def by force hence "arity_compatibles irs'" using inner by (metis (no_types, lifting) case_prodD) thus "arity_compatibles irs" unfolding \irs = transform_irules irs'\ by (rule arity_compatibles_transform_irules) have "irs' \ {||}" using \(name, irs') |\| rs\ inner by blast thus "irs \ {||}" unfolding \irs = transform_irules irs'\ transform_irules_def by (simp add: fgroup_by_nonempty) let ?f = "\(pats, rhs). (butlast pats, (last pats, rhs))" let ?grp = "fgroup_by ?f irs'" have "patterns_compatibles irs'" using \(name, irs') |\| rs\ inner by (blast dest: fpairwiseD) show "patterns_compatibles irs" proof (cases "arity irs' = 0") case True thus ?thesis unfolding \irs = transform_irules irs'\ transform_irules_def using \patterns_compatibles irs'\ by simp next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp show ?thesis proof safe fix pats\<^sub>1 rhs\<^sub>1 pats\<^sub>2 rhs\<^sub>2 assume "(pats\<^sub>1, rhs\<^sub>1) |\| irs" "(pats\<^sub>2, rhs\<^sub>2) |\| irs" with irs' obtain cs\<^sub>1 cs\<^sub>2 where "(pats\<^sub>1, cs\<^sub>1) |\| ?grp" "(pats\<^sub>2, cs\<^sub>2) |\| ?grp" by force then obtain pats\<^sub>1' pats\<^sub>2' and x y \ \dummies\ where "(pats\<^sub>1', x) |\| irs'" "(pats\<^sub>2', y) |\| irs'" and "pats\<^sub>1 = butlast pats\<^sub>1'" "pats\<^sub>2 = butlast pats\<^sub>2'" unfolding irs' by (fastforce elim: fgroup_byE2) hence "patterns_compatible pats\<^sub>1' pats\<^sub>2'" using \patterns_compatibles irs'\ by (auto dest: fpairwiseD) thus "patterns_compatible pats\<^sub>1 pats\<^sub>2" unfolding \pats\<^sub>1 = _\ \pats\<^sub>2 = _\ by auto qed qed have "is_fmap irs'" using \(name, irs') |\| rs\ inner by blast show "is_fmap irs" proof (cases "arity irs' = 0") case True thus ?thesis unfolding \irs = transform_irules irs'\ transform_irules_def using \is_fmap irs'\ by simp next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp show ?thesis proof fix pats rhs\<^sub>1 rhs\<^sub>2 assume "(pats, rhs\<^sub>1) |\| irs" "(pats, rhs\<^sub>2) |\| irs" with irs' obtain cs\<^sub>1 cs\<^sub>2 where "(pats, cs\<^sub>1) |\| ?grp" "rhs\<^sub>1 = Pabs cs\<^sub>1" and "(pats, cs\<^sub>2) |\| ?grp" "rhs\<^sub>2 = Pabs cs\<^sub>2" by force moreover have "is_fmap ?grp" by auto ultimately show "rhs\<^sub>1 = rhs\<^sub>2" by (auto dest: is_fmapD) qed qed fix pats rhs assume "(pats, rhs) |\| irs" show "linears pats" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" using \(pats, rhs) |\| irs\ by force then obtain pats' x \ \dummy\ where "fst (?f (pats', x)) = pats" "(pats', x) |\| irs'" by (fastforce simp: split_beta elim: fgroup_byE2) hence "pats = butlast pats'" by simp moreover have "linears pats'" using \(pats', x) |\| irs'\ \(name, irs') |\| rs\ inner by auto ultimately show ?thesis by auto qed have "fdisjnt (freess pats) all_consts" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" using \(pats, rhs) |\| irs\ by force then obtain pats' x \ \dummy\ where "fst (?f (pats', x)) = pats" "(pats', x) |\| irs'" by (fastforce simp: split_beta elim: fgroup_byE2) hence "pats = butlast pats'" by simp moreover have "fdisjnt (freess pats') all_consts" using \(pats', x) |\| irs'\ \(name, irs') |\| rs\ inner by auto ultimately show ?thesis by (metis subsetI in_set_butlastD freess_subset fdisjnt_subset_left) qed thus "fdisjnt (freess pats) (pre_constants.all_consts C_info (fst |`| transform_irule_set rs))" unfolding transform_irule_set_def by simp show "closed_except rhs (freess pats)" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force show ?thesis unfolding \rhs = Pabs cs\ closed_except_simps proof safe fix pat t assume "(pat, t) |\| cs" then obtain pats' where "(pats', t) |\| irs'" "?f (pats', t) = (pats, (pat, t))" using \(pats, cs) |\| ?grp\ by auto hence "closed_except t (freess pats')" using \(name, irs') |\| rs\ inner by auto have "pats' \ []" using \arity_compatibles irs'\ \(pats', t) |\| irs'\ False by (metis list.size(3) arity_compatible_length) hence "pats' = pats @ [pat]" using \?f (pats', t) = (pats, (pat, t))\ by (fastforce simp: split_beta snoc_eq_iff_butlast) hence "freess pats |\| frees pat = freess pats'" unfolding freess_def by auto thus "closed_except t (freess pats |\| frees pat)" using \closed_except t (freess pats')\ by simp qed qed show "wellformed rhs" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force show ?thesis unfolding \rhs = Pabs cs\ proof (rule wellformed_PabsI) show "cs \ {||}" using \(pats, cs) |\| ?grp\ \irs' \ {||}\ by (meson femptyE fgroup_by_nonempty_inner) next show "is_fmap cs" proof fix pat t\<^sub>1 t\<^sub>2 assume "(pat, t\<^sub>1) |\| cs" "(pat, t\<^sub>2) |\| cs" then obtain pats\<^sub>1' pats\<^sub>2' where "(pats\<^sub>1', t\<^sub>1) |\| irs'" "?f (pats\<^sub>1', t\<^sub>1) = (pats, (pat, t\<^sub>1))" and "(pats\<^sub>2', t\<^sub>2) |\| irs'" "?f (pats\<^sub>2', t\<^sub>2) = (pats, (pat, t\<^sub>2))" using \(pats, cs) |\| ?grp\ by force moreover hence "pats\<^sub>1' \ []" "pats\<^sub>2' \ []" using \arity_compatibles irs'\ False unfolding prod.case by (metis list.size(3) arity_compatible_length)+ ultimately have "pats\<^sub>1' = pats @ [pat]" "pats\<^sub>2' = pats @ [pat]" unfolding split_beta fst_conv snd_conv by (metis prod.inject snoc_eq_iff_butlast)+ with \is_fmap irs'\ show "t\<^sub>1 = t\<^sub>2" using \(pats\<^sub>1', t\<^sub>1) |\| irs'\ \(pats\<^sub>2', t\<^sub>2) |\| irs'\ by (blast dest: is_fmapD) qed next show "pattern_compatibles cs" proof safe fix pat\<^sub>1 rhs\<^sub>1 pat\<^sub>2 rhs\<^sub>2 assume "(pat\<^sub>1, rhs\<^sub>1) |\| cs" "(pat\<^sub>2, rhs\<^sub>2) |\| cs" then obtain pats\<^sub>1' pats\<^sub>2' where "(pats\<^sub>1', rhs\<^sub>1) |\| irs'" "?f (pats\<^sub>1', rhs\<^sub>1) = (pats, (pat\<^sub>1, rhs\<^sub>1))" and "(pats\<^sub>2', rhs\<^sub>2) |\| irs'" "?f (pats\<^sub>2', rhs\<^sub>2) = (pats, (pat\<^sub>2, rhs\<^sub>2))" using \(pats, cs) |\| ?grp\ by force moreover hence "pats\<^sub>1' \ []" "pats\<^sub>2' \ []" using \arity_compatibles irs'\ False unfolding prod.case by (metis list.size(3) arity_compatible_length)+ ultimately have "pats\<^sub>1' = pats @ [pat\<^sub>1]" "pats\<^sub>2' = pats @ [pat\<^sub>2]" unfolding split_beta fst_conv snd_conv by (metis prod.inject snoc_eq_iff_butlast)+ moreover have "patterns_compatible pats\<^sub>1' pats\<^sub>2'" using \(pats\<^sub>1', rhs\<^sub>1) |\| irs'\ \(pats\<^sub>2', rhs\<^sub>2) |\| irs'\ \patterns_compatibles irs'\ by (auto dest: fpairwiseD) ultimately show "pattern_compatible pat\<^sub>1 pat\<^sub>2" by (auto elim: rev_accum_rel_snoc_eqE) qed next fix pat t assume "(pat, t) |\| cs" then obtain pats' where "(pats', t) |\| irs'" "pat = last pats'" using \(pats, cs) |\| ?grp\ by auto moreover hence "pats' \ []" using \arity_compatibles irs'\ False by (metis list.size(3) arity_compatible_length) ultimately have "pat \ set pats'" by auto moreover have "linears pats'" using \(pats', t) |\| irs'\ \(name, irs') |\| rs\ inner by auto ultimately show "linear pat" by (metis linears_linear) show "wellformed t" using \(pats', t) |\| irs'\ \(name, irs') |\| rs\ inner by auto qed qed have "\ shadows_consts rhs" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force show ?thesis unfolding \rhs = _\ proof assume "shadows_consts (Pabs cs)" then obtain pat t where "(pat, t) |\| cs" "shadows_consts t \ shadows_consts pat" by force then obtain pats' where "(pats', t) |\| irs'" "pat = last pats'" using \(pats, cs) |\| ?grp\ by auto moreover hence "pats' \ []" using \arity_compatibles irs'\ False by (metis list.size(3) arity_compatible_length) ultimately have "pat \ set pats'" by auto show False using \shadows_consts t \ shadows_consts pat\ proof assume "shadows_consts t" thus False using \(name, irs') |\| rs\ \(pats', t) |\| irs'\ inner by auto next assume "shadows_consts pat" have "fdisjnt (freess pats') all_consts" using \(name, irs') |\| rs\ \(pats', t) |\| irs'\ inner by auto have "fdisjnt (frees pat) all_consts" apply (rule fdisjnt_subset_left) apply (subst freess_single[symmetric]) apply (rule freess_subset) apply simp apply fact+ done thus False using \shadows_consts pat\ unfolding shadows_consts_def fdisjnt_alt_def by auto qed qed qed thus "\ pre_constants.shadows_consts C_info (fst |`| transform_irule_set rs) rhs" by (simp add: transform_irule_set_heads) show "abs_ish pats rhs" proof (cases "arity irs' = 0") case True thus ?thesis using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def by auto next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force thus ?thesis unfolding abs_ish_def by (simp add: is_abs_def term_cases_def) qed have "welldefined rhs" proof (cases "arity irs' = 0") case True hence \(pats, rhs) |\| irs'\ using \(pats, rhs) |\| irs\ \(name, irs') |\| rs\ inner unfolding \irs = transform_irules irs'\ transform_irules_def by (smt fBallE split_conv) thus ?thesis unfolding transform_irule_set_def using fbspec[OF inner \(name, irs') |\| rs\, simplified] by force next case False hence irs': "irs = map_prod id Pabs |`| ?grp" unfolding \irs = transform_irules irs'\ transform_irules_def by simp then obtain cs where "(pats, cs) |\| ?grp" "rhs = Pabs cs" using \(pats, rhs) |\| irs\ by force show ?thesis unfolding \rhs = _\ apply simp apply (rule ffUnion_least) unfolding ball_simps apply rule apply (rename_tac x, case_tac x, hypsubst_thin) apply simp subgoal premises prems for pat t proof - from prems obtain pats' where "(pats', t) |\| irs'" using \(pats, cs) |\| ?grp\ by auto hence "welldefined t" using fbspec[OF inner \(name, irs') |\| rs\, simplified] by blast thus ?thesis unfolding transform_irule_set_def by simp qed done qed thus "pre_constants.welldefined C_info (fst |`| transform_irule_set rs) rhs" unfolding transform_irule_set_heads . qed subsubsection \Matching and rewriting\ definition irewrite_step :: "name \ term list \ pterm \ pterm \ pterm option" where "irewrite_step name pats rhs t = map_option (subst rhs) (match (name $$ pats) t)" abbreviation irewrite_step' :: "name \ term list \ pterm \ pterm \ pterm \ bool" ("_, _, _ \\<^sub>i/ _ \/ _" [50,0,50] 50) where "name, pats, rhs \\<^sub>i t \ u \ irewrite_step name pats rhs t = Some u" lemma irewrite_stepI: assumes "match (name $$ pats) t = Some env" "subst rhs env = u" shows "name, pats, rhs \\<^sub>i t \ u" using assms unfolding irewrite_step_def by simp inductive irewrite :: "irule_set \ pterm \ pterm \ bool" ("_/ \\<^sub>i/ _ \/ _" [50,0,50] 50) for irs where step: "\ (name, rs) |\| irs; (pats, rhs) |\| rs; name, pats, rhs \\<^sub>i t \ t' \ \ irs \\<^sub>i t \ t'" | beta: "\ c |\| cs; c \ t \ t' \ \ irs \\<^sub>i Pabs cs $\<^sub>p t \ t'" | "fun": "irs \\<^sub>i t \ t' \ irs \\<^sub>i t $\<^sub>p u \ t' $\<^sub>p u" | arg: "irs \\<^sub>i u \ u' \ irs \\<^sub>i t $\<^sub>p u \ t $\<^sub>p u'" global_interpretation irewrite: rewriting "irewrite rs" for rs by standard (auto intro: irewrite.intros simp: app_pterm_def)+ abbreviation irewrite_rt :: "irule_set \ pterm \ pterm \ bool" ("_/ \\<^sub>i/ _ \*/ _" [50,0,50] 50) where "irewrite_rt rs \ (irewrite rs)\<^sup>*\<^sup>*" lemma (in irules) irewrite_closed: assumes "rs \\<^sub>i t \ u" "closed t" shows "closed u" using assms proof induction case (step name rs pats rhs t t') then obtain env where "match (name $$ pats) t = Some env" "t' = subst rhs env" unfolding irewrite_step_def by auto hence "closed_env env" using step by (auto intro: closed.match) show ?case unfolding \t' = _\ apply (subst closed_except_def) apply (subst subst_frees) apply fact apply (subst match_dom) apply fact apply (subst frees_list_comb) apply simp apply (subst closed_except_def[symmetric]) - using inner step by blast + using inner step by fastforce next case (beta c cs t t') then obtain pat rhs where "c = (pat, rhs)" by (cases c) auto with beta obtain env where "match pat t = Some env" "t' = subst rhs env" by auto moreover have "closed t" using beta unfolding closed_except_def by simp ultimately have "closed_env env" using beta by (auto intro: closed.match) show ?case unfolding \t' = subst rhs env\ apply (subst closed_except_def) apply (subst subst_frees) apply fact apply (subst match_dom) apply fact apply simp apply (subst closed_except_def[symmetric]) using inner beta \c = _\ by (auto simp: closed_except_simps) qed (auto simp: closed_except_def) corollary (in irules) irewrite_rt_closed: assumes "rs \\<^sub>i t \* u" "closed t" shows "closed u" using assms by induction (auto intro: irewrite_closed) subsubsection \Correctness of translation\ abbreviation irelated :: "nterm \ pterm \ bool" ("_ \\<^sub>i _" [0,50] 50) where "n \\<^sub>i p \ nterm_to_pterm n = p" global_interpretation irelated: term_struct_rel_strong irelated by standard (auto simp: app_pterm_def app_nterm_def const_pterm_def const_nterm_def elim: nterm_to_pterm.elims) lemma irelated_vars: "t \\<^sub>i u \ frees t = frees u" by auto lemma irelated_no_abs: assumes "t \\<^sub>i u" shows "no_abs t \ no_abs u" using assms apply (induction arbitrary: t) apply (auto elim!: nterm_to_pterm.elims) apply (fold const_nterm_def const_pterm_def free_nterm_def free_pterm_def app_pterm_def app_nterm_def) by auto lemma irelated_subst: assumes "t \\<^sub>i u" "irelated.P_env nenv penv" shows "subst t nenv \\<^sub>i subst u penv" using assms proof (induction arbitrary: nenv penv u rule: nterm_to_pterm.induct) case (1 s) then show ?case by (auto elim!: fmrel_cases[where x = s]) next case 4 from 4(2)[symmetric] show ?case apply simp apply (rule 4) apply simp using 4(3) by (simp add: fmrel_drop) qed auto lemma related_irewrite_step: assumes "name, pats, nterm_to_pterm rhs \\<^sub>i u \ u'" "t \\<^sub>i u" obtains t' where "unsplit_rule (name, pats, rhs) \ t \ t'" "t' \\<^sub>i u'" proof - let ?rhs' = "nterm_to_pterm rhs" let ?x = "name $$ pats" from assms obtain env where "match ?x u = Some env" "u' = subst ?rhs' env" unfolding irewrite_step_def by blast then obtain nenv where "match ?x t = Some nenv" "irelated.P_env nenv env" using assms by (metis Option.is_none_def not_None_eq option.rel_distinct(1) option.sel rel_option_unfold irelated.match_rel) show thesis proof show "unsplit_rule (name, pats, rhs) \ t \ subst rhs nenv" using \match ?x t = _\ by auto next show "subst rhs nenv \\<^sub>i u'" unfolding \u' = _\ using \irelated.P_env nenv env\ by (auto intro: irelated_subst) qed qed theorem (in nrules) compile_correct: assumes "compile (consts_of rs) \\<^sub>i u \ u'" "t \\<^sub>i u" "closed t" obtains t' where "rs \\<^sub>n t \ t'" "t' \\<^sub>i u'" using assms(1-3) proof (induction arbitrary: t thesis rule: irewrite.induct) case (step name irs pats rhs u u') then obtain crs where "irs = translate_crules crs" "(name, crs) |\| consts_of rs" unfolding compile_def by force moreover with step obtain rhs' where "rhs = nterm_to_pterm rhs'" "(pats, rhs') |\| crs" unfolding translate_crules_def by force ultimately obtain rule where "split_rule rule = (name, (pats, rhs'))" "rule |\| rs" unfolding consts_of_def by blast hence "nrule rule" using all_rules by blast obtain t' where "unsplit_rule (name, pats, rhs') \ t \ t'" "t' \\<^sub>i u'" using \name, pats, rhs \\<^sub>i u \ u'\ \t \\<^sub>i u\ unfolding \rhs = nterm_to_pterm rhs'\ by (elim related_irewrite_step) hence "rule \ t \ t'" using \nrule rule\ \split_rule rule = (name, (pats, rhs'))\ by (metis unsplit_split) show ?case proof (rule step.prems) show "rs \\<^sub>n t \ t'" apply (rule nrewrite.step) apply fact apply fact done next show "t' \\<^sub>i u'" by fact qed next case (beta c cs u u') then obtain pat rhs where "c = (pat, rhs)" "(pat, rhs) |\| cs" by (cases c) auto obtain v w where "t = v $\<^sub>n w" "v \\<^sub>i Pabs cs" "w \\<^sub>i u" using \t \\<^sub>i Pabs cs $\<^sub>p u\ by (auto elim: nterm_to_pterm.elims) obtain x nrhs irhs where "v = (\\<^sub>n x. nrhs)" "cs = {| (Free x, irhs) |}" "nrhs \\<^sub>i irhs" using \v \\<^sub>i Pabs cs\ by (auto elim: nterm_to_pterm.elims) hence "t = (\\<^sub>n x. nrhs) $\<^sub>n w" "\\<^sub>n x. nrhs \\<^sub>i \\<^sub>p x. irhs" unfolding \t = v $\<^sub>n w\ using \v \\<^sub>i Pabs cs\ by auto have "pat = Free x" "rhs = irhs" using \cs = {| (Free x, irhs) |}\ \(pat, rhs) |\| cs\ by auto hence "(Free x, irhs) \ u \ u'" using beta \c = _\ by simp hence "u' = subst irhs (fmap_of_list [(x, u)])" by simp show ?case proof (rule beta.prems) show "rs \\<^sub>n t \ subst nrhs (fmap_of_list [(x, w)])" unfolding \t = (\\<^sub>n x. nrhs) $\<^sub>n w\ by (rule nrewrite.beta) next show "subst nrhs (fmap_of_list [(x, w)]) \\<^sub>i u'" unfolding \u' = subst irhs _\ apply (rule irelated_subst) apply fact apply simp apply rule apply rule apply fact done qed next case ("fun" v v' u) obtain w x where "t = w $\<^sub>n x" "w \\<^sub>i v" "x \\<^sub>i u" using \t \\<^sub>i v $\<^sub>p u\ by (auto elim: nterm_to_pterm.elims) with "fun" obtain w' where "rs \\<^sub>n w \ w'" "w' \\<^sub>i v'" unfolding closed_except_def by auto show ?case proof (rule fun.prems) show "rs \\<^sub>n t \ w' $\<^sub>n x" unfolding \t = w $\<^sub>n x\ by (rule nrewrite.fun) fact next show "w' $\<^sub>n x \\<^sub>i v' $\<^sub>p u" by auto fact+ qed next case (arg u u' v) obtain w x where "t = w $\<^sub>n x" "w \\<^sub>i v" "x \\<^sub>i u" using \ t \\<^sub>i v $\<^sub>p u\ by (auto elim: nterm_to_pterm.elims) with arg obtain x' where "rs \\<^sub>n x \ x'" "x' \\<^sub>i u'" unfolding closed_except_def by auto show ?case proof (rule arg.prems) show "rs \\<^sub>n t \ w $\<^sub>n x'" unfolding \t = w $\<^sub>n x\ by (rule nrewrite.arg) fact next show "w $\<^sub>n x' \\<^sub>i v $\<^sub>p u'" by auto fact+ qed qed corollary (in nrules) compile_correct_rt: assumes "compile (consts_of rs) \\<^sub>i u \* u'" "t \\<^sub>i u" "closed t" obtains t' where "rs \\<^sub>n t \* t'" "t' \\<^sub>i u'" using assms proof (induction arbitrary: thesis t) (* FIXME clone of transform_correct_rt, maybe locale? *) case (step u' u'') obtain t' where "rs \\<^sub>n t \* t'" "t' \\<^sub>i u'" using step by blast obtain t'' where "rs \\<^sub>n t' \* t''" "t'' \\<^sub>i u''" proof (rule compile_correct) show "compile (consts_of rs) \\<^sub>i u' \ u''" "t' \\<^sub>i u'" by fact+ next show "closed t'" using \rs \\<^sub>n t \* t'\ \closed t\ by (rule nrewrite_rt_closed) qed blast show ?case proof (rule step.prems) show "rs \\<^sub>n t \* t''" using \rs \\<^sub>n t \* t'\ \rs \\<^sub>n t' \* t''\ by auto qed fact qed blast subsubsection \Completeness of translation\ lemma (in nrules) compile_complete: assumes "rs \\<^sub>n t \ t'" "closed t" shows "compile (consts_of rs) \\<^sub>i nterm_to_pterm t \ nterm_to_pterm t'" using assms proof induction case (step r t t') then obtain pat rhs' where "r = (pat, rhs')" by force then have "(pat, rhs') |\| rs" "(pat, rhs') \ t \ t'" using step by blast+ then have "nrule (pat, rhs')" using all_rules by blast then obtain name pats where "(name, (pats, rhs')) = split_rule r" "pat = name $$ pats" unfolding split_rule_def \r = _\ apply atomize_elim by (auto simp: split_beta) obtain crs where "(name, crs) |\| consts_of rs" "(pats, rhs') |\| crs" using step \_ = split_rule r\ \r = _\ by (metis consts_of_def fgroup_by_complete fst_conv snd_conv) then obtain irs where "irs = translate_crules crs" by blast then have "(name, irs) |\| compile (consts_of rs)" unfolding compile_def using \(name, _) |\| _\ by (metis fimageI id_def map_prod_simp) obtain rhs where "rhs = nterm_to_pterm rhs'" "(pats, rhs) |\| irs" using \irs = _\ \_ |\| crs\ unfolding translate_crules_def by (metis fimageI id_def map_prod_simp) from step obtain env' where "match pat t = Some env'" "t' = subst rhs' env'" unfolding \r = _\ using rewrite_step.simps by force then obtain env where "match pat (nterm_to_pterm t) = Some env" "irelated.P_env env' env" by (metis irelated.match_rel option_rel_Some1) then have "subst rhs env = nterm_to_pterm t'" unfolding \t' = _\ apply - apply (rule sym) apply (rule irelated_subst) unfolding \rhs = _\ by auto have "name, pats, rhs \\<^sub>i nterm_to_pterm t \ nterm_to_pterm t'" apply (rule irewrite_stepI) using \match _ _ = Some env\ unfolding \pat = _\ apply assumption by fact show ?case by rule fact+ next case (beta x t t') obtain c where "c = (Free x, nterm_to_pterm t)" by blast from beta have "closed (nterm_to_pterm t')" using closed_nterm_to_pterm[where t = t'] unfolding closed_except_def by auto show ?case apply simp apply rule using \c = _\ by (fastforce intro: irelated_subst[THEN sym])+ next case ("fun" t t' u) show ?case apply simp apply rule apply (rule "fun") using "fun" unfolding closed_except_def apply simp done next case (arg u u' t) show ?case apply simp apply rule apply (rule arg) using arg unfolding closed_except_def by simp qed subsubsection \Correctness of transformation\ abbreviation irules_deferred_matches :: "pterm list \ irules \ (term \ pterm) fset" where "irules_deferred_matches args \ fselect (\(pats, rhs). map_option (\env. (last pats, subst rhs env)) (matchs (butlast pats) args))" context irules begin inductive prelated :: "pterm \ pterm \ bool" ("_ \\<^sub>p _" [0,50] 50) where const: "Pconst x \\<^sub>p Pconst x" | var: "Pvar x \\<^sub>p Pvar x" | app: "t\<^sub>1 \\<^sub>p u\<^sub>1 \ t\<^sub>2 \\<^sub>p u\<^sub>2 \ t\<^sub>1 $\<^sub>p t\<^sub>2 \\<^sub>p u\<^sub>1 $\<^sub>p u\<^sub>2" | pat: "rel_fset (rel_prod (=) prelated) cs\<^sub>1 cs\<^sub>2 \ Pabs cs\<^sub>1 \\<^sub>p Pabs cs\<^sub>2" | "defer": "(name, rsi) |\| rs \ 0 < arity rsi \ rel_fset (rel_prod (=) prelated) (irules_deferred_matches args rsi) cs \ list_all closed args \ name $$ args \\<^sub>p Pabs cs" inductive_cases prelated_absE[consumes 1, case_names pat "defer"]: "t \\<^sub>p Pabs cs" lemma prelated_refl[intro!]: "t \\<^sub>p t" proof (induction t) case Pabs thus ?case by (auto simp: snds.simps intro!: prelated.pat rel_fset_refl_strong rel_prod.intros) qed (auto intro: prelated.intros) sublocale prelated: term_struct_rel prelated by standard (auto simp: const_pterm_def app_pterm_def intro: prelated.intros elim: prelated.cases) lemma prelated_pvars: assumes "t \\<^sub>p u" shows "frees t = frees u" using assms proof (induction rule: prelated.induct) case (pat cs\<^sub>1 cs\<^sub>2) show ?case apply simp apply (rule arg_cong[where f = ffUnion]) apply (rule rel_fset_image_eq) apply fact apply auto done next case ("defer" name rsi args cs) { fix pat t assume "(pat, t) |\| cs" with "defer" obtain t' where "(pat, t') |\| irules_deferred_matches args rsi" "frees t = frees t'" by (auto elim: rel_fsetE2) then obtain pats rhs env where "pat = last pats" "(pats, rhs) |\| rsi" and "matchs (butlast pats) args = Some env" "t' = subst rhs env" by auto have "closed_except rhs (freess pats)" "linears pats" using \(pats, rhs) |\| rsi\ \(name, rsi) |\| rs\ inner by auto have "arity_compatibles rsi" using "defer" inner by (metis (no_types, lifting) case_prodD) have "length pats > 0" by (subst arity_compatible_length) fact+ hence "pats = butlast pats @ [last pats]" by simp note \frees t = frees t'\ also have "frees t' = frees rhs - fmdom env" unfolding \t' = _\ apply (rule subst_frees) apply (rule closed.matchs) apply fact+ done also have "\ = frees rhs - freess (butlast pats)" using \matchs _ _ = _\ by (metis matchs_dom) also have "\ |\| freess pats - freess (butlast pats)" using \closed_except _ _\ by (auto simp: closed_except_def) also have "\ = frees (last pats) |-| freess (butlast pats)" by (subst \pats = _\) (simp add: funion_fminus) also have "\ = frees (last pats)" proof (rule fminus_triv) have "fdisjnt (freess (butlast pats)) (freess [last pats])" using \linears pats\ \pats = _\ by (metis linears_appendD) thus "frees (last pats) |\| freess (butlast pats) = {||}" by (fastforce simp: fdisjnt_alt_def) qed also have "\ = frees pat" unfolding \pat = _\ .. finally have "frees t |\| frees pat" . } hence "closed (Pabs cs)" unfolding closed_except_simps by (auto simp: closed_except_def) moreover have "closed (name $$ args)" unfolding closed_list_comb by fact ultimately show ?case unfolding closed_except_def by simp qed auto corollary prelated_closed: "t \\<^sub>p u \ closed t \ closed u" unfolding closed_except_def by (auto simp: prelated_pvars) lemma prelated_no_abs_right: assumes "t \\<^sub>p u" "no_abs u" shows "t = u" using assms apply (induction rule: prelated.induct) apply auto apply (fold app_pterm_def) apply auto done corollary env_prelated_refl[intro!]: "prelated.P_env env env" by (auto intro: fmap.rel_refl) text \ The following, more general statement does not hold: @{prop "t \\<^sub>p u \ rel_option prelated.P_env (match x t) (match x u)"} If @{text t} and @{text u} are related because of the @{thm [source=true] prelated.defer} rule, they have completely different shapes. Establishing @{prop "is_abs t \ is_abs u"} as a precondition would rule out this case, but at the same time be too restrictive. Instead, we use @{thm prelated.related_match}. \ lemma prelated_subst: assumes "t\<^sub>1 \\<^sub>p t\<^sub>2" "prelated.P_env env\<^sub>1 env\<^sub>2" shows "subst t\<^sub>1 env\<^sub>1 \\<^sub>p subst t\<^sub>2 env\<^sub>2" using assms proof (induction arbitrary: env\<^sub>1 env\<^sub>2 rule: prelated.induct) case (var x) thus ?case proof (cases rule: fmrel_cases[where x = x]) case none thus ?thesis by (auto intro: prelated.var) next case (some t u) thus ?thesis by simp qed next case (pat cs\<^sub>1 cs\<^sub>2) let ?drop = "\env. \(pat::term). fmdrop_fset (frees pat) env" from pat have "prelated.P_env (?drop env\<^sub>1 pat) (?drop env\<^sub>2 pat)" for pat by blast with pat show ?case by (auto intro!: prelated.pat rel_fset_image) next case ("defer" name rsi args cs) have "name $$ args \\<^sub>p Pabs cs" apply (rule prelated.defer) apply fact+ apply (rule fset.rel_mono_strong) apply fact apply force apply fact done moreover have "closed (name $$ args)" unfolding closed_list_comb by fact ultimately have "closed (Pabs cs)" by (metis prelated_closed) let ?drop = "\env. \pat. fmdrop_fset (frees pat) env" let ?f = "\env. (\(pat, rhs). (pat, subst rhs (?drop env pat)))" have "name $$ args \\<^sub>p Pabs (?f env\<^sub>2 |`| cs)" proof (rule prelated.defer) show "(name, rsi) |\| rs" "0 < arity rsi" "list_all closed args" using "defer" by auto next { fix pat\<^sub>1 rhs\<^sub>1 fix pat\<^sub>2 rhs\<^sub>2 assume "(pat\<^sub>2, rhs\<^sub>2) |\| cs" assume "pat\<^sub>1 = pat\<^sub>2" "rhs\<^sub>1 \\<^sub>p rhs\<^sub>2" have "rhs\<^sub>1 \\<^sub>p subst rhs\<^sub>2 (fmdrop_fset (frees pat\<^sub>2) env\<^sub>2)" by (subst subst_closed_pabs) fact+ } hence "rel_fset (rel_prod (=) prelated) (id |`| irules_deferred_matches args rsi) (?f env\<^sub>2 |`| cs)" by (force intro!: rel_fset_image[OF \rel_fset _ _ _\]) thus "rel_fset (rel_prod (=) prelated) (irules_deferred_matches args rsi) (?f env\<^sub>2 |`| cs)" by simp qed moreover have "map (\t. subst t env\<^sub>1) args = args" apply (rule map_idI) apply (rule subst_closed_id) using "defer" by (simp add: list_all_iff) ultimately show ?case by (simp add: subst_list_comb) qed (auto intro: prelated.intros) lemma prelated_step: assumes "name, pats, rhs \\<^sub>i u \ u'" "t \\<^sub>p u" obtains t' where "name, pats, rhs \\<^sub>i t \ t'" "t' \\<^sub>p u'" proof - let ?lhs = "name $$ pats" from assms obtain env where "match ?lhs u = Some env" "u' = subst rhs env" unfolding irewrite_step_def by blast then obtain env' where "match ?lhs t = Some env'" "prelated.P_env env' env" using assms by (auto elim: prelated.related_match) hence "subst rhs env' \\<^sub>p subst rhs env" using assms by (auto intro: prelated_subst) show thesis proof show "name, pats, rhs \\<^sub>i t \ subst rhs env'" unfolding irewrite_step_def using \match ?lhs t = Some env'\ by simp next show "subst rhs env' \\<^sub>p u'" unfolding \u' = subst rhs env\ by fact qed qed (* FIXME write using relators *) lemma prelated_beta: \ \same problem as @{thm [source=true] prelated.related_match}\ assumes "(pat, rhs\<^sub>2) \ t\<^sub>2 \ u\<^sub>2" "rhs\<^sub>1 \\<^sub>p rhs\<^sub>2" "t\<^sub>1 \\<^sub>p t\<^sub>2" obtains u\<^sub>1 where "(pat, rhs\<^sub>1) \ t\<^sub>1 \ u\<^sub>1" "u\<^sub>1 \\<^sub>p u\<^sub>2" proof - from assms obtain env\<^sub>2 where "match pat t\<^sub>2 = Some env\<^sub>2" "u\<^sub>2 = subst rhs\<^sub>2 env\<^sub>2" by auto with assms obtain env\<^sub>1 where "match pat t\<^sub>1 = Some env\<^sub>1" "prelated.P_env env\<^sub>1 env\<^sub>2" by (auto elim: prelated.related_match) with assms have "subst rhs\<^sub>1 env\<^sub>1 \\<^sub>p subst rhs\<^sub>2 env\<^sub>2" by (auto intro: prelated_subst) show thesis proof show "(pat, rhs\<^sub>1) \ t\<^sub>1 \ subst rhs\<^sub>1 env\<^sub>1" using \match pat t\<^sub>1 = _\ by simp next show "subst rhs\<^sub>1 env\<^sub>1 \\<^sub>p u\<^sub>2" unfolding \u\<^sub>2 = _\ by fact qed qed theorem transform_correct: assumes "transform_irule_set rs \\<^sub>i u \ u'" "t \\<^sub>p u" "closed t" obtains t' where "rs \\<^sub>i t \* t'" \ \zero or one step\ and "t' \\<^sub>p u'" using assms(1-3) proof (induction arbitrary: t thesis rule: irewrite.induct) case (beta c cs\<^sub>2 u\<^sub>2 x\<^sub>2') obtain v u\<^sub>1 where "t = v $\<^sub>p u\<^sub>1" "v \\<^sub>p Pabs cs\<^sub>2" "u\<^sub>1 \\<^sub>p u\<^sub>2" using \t \\<^sub>p Pabs cs\<^sub>2 $\<^sub>p u\<^sub>2\ by cases with beta have "closed u\<^sub>1" by (simp add: closed_except_def) obtain pat rhs\<^sub>2 where "c = (pat, rhs\<^sub>2)" by (cases c) auto from \v \\<^sub>p Pabs cs\<^sub>2\ show ?case proof (cases rule: prelated_absE) case (pat cs\<^sub>1) with beta \c = _\ obtain rhs\<^sub>1 where "(pat, rhs\<^sub>1) |\| cs\<^sub>1" "rhs\<^sub>1 \\<^sub>p rhs\<^sub>2" by (auto elim: rel_fsetE2) with beta obtain x\<^sub>1' where "(pat, rhs\<^sub>1) \ u\<^sub>1 \ x\<^sub>1'" "x\<^sub>1' \\<^sub>p x\<^sub>2'" using \u\<^sub>1 \\<^sub>p u\<^sub>2\ assms \c = _\ by (auto elim: prelated_beta simp del: rewrite_step.simps) show ?thesis proof (rule beta.prems) show "rs \\<^sub>i t \* x\<^sub>1'" unfolding \t = _\ \v = _\ by (intro r_into_rtranclp irewrite.beta) fact+ next show "x\<^sub>1' \\<^sub>p x\<^sub>2'" by fact qed next case ("defer" name rsi args) with beta \c = _\ obtain rhs\<^sub>1' where "(pat, rhs\<^sub>1') |\| irules_deferred_matches args rsi" "rhs\<^sub>1' \\<^sub>p rhs\<^sub>2" by (auto elim: rel_fsetE2) then obtain env\<^sub>a rhs\<^sub>1 pats where "matchs (butlast pats) args = Some env\<^sub>a" "pat = last pats" "rhs\<^sub>1' = subst rhs\<^sub>1 env\<^sub>a" and "(pats, rhs\<^sub>1) |\| rsi" by auto hence "linears pats" using \(name, rsi) |\| rs\ inner unfolding irules_def by auto have "arity_compatibles rsi" using "defer" inner by (metis (no_types, lifting) case_prodD) have "length pats > 0" by (subst arity_compatible_length) fact+ hence "pats = butlast pats @ [pat]" unfolding \pat = _\ by simp from beta \c = _\ obtain env\<^sub>b where "match pat u\<^sub>2 = Some env\<^sub>b" "x\<^sub>2' = subst rhs\<^sub>2 env\<^sub>b" by fastforce with \u\<^sub>1 \\<^sub>p u\<^sub>2\ obtain env\<^sub>b' where "match pat u\<^sub>1 = Some env\<^sub>b'" "prelated.P_env env\<^sub>b' env\<^sub>b" by (metis prelated.related_match) have "closed_env env\<^sub>a" by (rule closed.matchs) fact+ have "closed_env env\<^sub>b'" apply (rule closed.matchs[where pats = "[pat]" and ts = "[u\<^sub>1]"]) apply simp apply fact apply simp apply fact done have "fmdom env\<^sub>a = freess (butlast pats)" by (rule matchs_dom) fact moreover have "fmdom env\<^sub>b' = frees pat" by (rule match_dom) fact moreover have "fdisjnt (freess (butlast pats)) (frees pat)" using \pats = _\ \linears pats\ by (metis freess_single linears_appendD(3)) ultimately have "fdisjnt (fmdom env\<^sub>a) (fmdom env\<^sub>b')" by simp show ?thesis proof (rule beta.prems) have "rs \\<^sub>i name $$ args $\<^sub>p u\<^sub>1 \ subst rhs\<^sub>1' env\<^sub>b'" proof (rule irewrite.step) show "(name, rsi) |\| rs" "(pats, rhs\<^sub>1) |\| rsi" by fact+ next show "name, pats, rhs\<^sub>1 \\<^sub>i name $$ args $\<^sub>p u\<^sub>1 \ subst rhs\<^sub>1' env\<^sub>b'" apply (rule irewrite_stepI) apply (fold app_pterm_def) apply (subst list_comb_snoc) apply (subst matchs_match_list_comb) apply (subst \pats = _\) apply (rule matchs_appI) apply fact apply simp apply fact unfolding \rhs\<^sub>1' = _\ apply (rule subst_indep') apply fact+ done qed thus "rs \\<^sub>i t \* subst rhs\<^sub>1' env\<^sub>b'" unfolding \t = _\ \v = _\ by (rule r_into_rtranclp) next show "subst rhs\<^sub>1' env\<^sub>b' \\<^sub>p x\<^sub>2'" unfolding \x\<^sub>2' = _\ by (rule prelated_subst) fact+ qed qed next case (step name rs\<^sub>2 pats rhs u u') then obtain rs\<^sub>1 where "rs\<^sub>2 = transform_irules rs\<^sub>1" "(name, rs\<^sub>1) |\| rs" unfolding transform_irule_set_def by force hence "arity_compatibles rs\<^sub>1" using inner by (metis (no_types, lifting) case_prodD) show ?case proof (cases "arity rs\<^sub>1 = 0") case True hence "rs\<^sub>2 = rs\<^sub>1" unfolding \rs\<^sub>2 = _\ transform_irules_def by simp with step have "(pats, rhs) |\| rs\<^sub>1" by simp from step obtain t' where "name, pats, rhs \\<^sub>i t \ t'" "t' \\<^sub>p u'" using assms by (auto elim: prelated_step) show ?thesis proof (rule step.prems) show "rs \\<^sub>i t \* t'" by (intro conjI exI r_into_rtranclp irewrite.step) fact+ qed fact next let ?f = "\(pats, rhs). (butlast pats, last pats, rhs)" let ?grp = "fgroup_by ?f rs\<^sub>1" case False hence "rs\<^sub>2 = map_prod id Pabs |`| ?grp" unfolding \rs\<^sub>2 = _\ transform_irules_def by simp with step obtain cs where "rhs = Pabs cs" "(pats, cs) |\| ?grp" by force from step obtain env\<^sub>2 where "match (name $$ pats) u = Some env\<^sub>2" "u' = subst rhs env\<^sub>2" unfolding irewrite_step_def by auto then obtain args\<^sub>2 where "u = name $$ args\<^sub>2" "matchs pats args\<^sub>2 = Some env\<^sub>2" by (auto elim: match_list_combE) with step obtain args\<^sub>1 where "t = name $$ args\<^sub>1" "list_all2 prelated args\<^sub>1 args\<^sub>2" by (auto elim: prelated.list_combE) then obtain env\<^sub>1 where "matchs pats args\<^sub>1 = Some env\<^sub>1" "prelated.P_env env\<^sub>1 env\<^sub>2" using \matchs pats args\<^sub>2 = _\ by (metis prelated.related_matchs) hence "fmdom env\<^sub>1 = freess pats" by (auto simp: matchs_dom) obtain cs' where "u' = Pabs cs'" unfolding \u' = _\ \rhs = _\ by auto hence "cs' = (\(pat, rhs). (pat, subst rhs (fmdrop_fset (frees pat) env\<^sub>2 ))) |`| cs" using \u' = subst rhs env\<^sub>2\ unfolding \rhs = _\ by simp show ?thesis proof (rule step.prems) show "rs \\<^sub>i t \* t" by (rule rtranclp.rtrancl_refl) next show "t \\<^sub>p u'" unfolding \u' = Pabs cs'\ \t = _\ proof (intro prelated.defer rel_fsetI; safe?) show "(name, rs\<^sub>1) |\| rs" by fact next show "0 < arity rs\<^sub>1" using False by simp next show "list_all closed args\<^sub>1" using \closed t\ unfolding \t = _\ closed_list_comb . next fix pat rhs' assume "(pat, rhs') |\| irules_deferred_matches args\<^sub>1 rs\<^sub>1" then obtain pats' rhs env where "(pats', rhs) |\| rs\<^sub>1" and "matchs (butlast pats') args\<^sub>1 = Some env" "pat = last pats'" "rhs' = subst rhs env" by auto with False have "pats' \ []" using \arity_compatibles rs\<^sub>1\ by (metis list.size(3) arity_compatible_length) hence "butlast pats' @ [last pats'] = pats'" by simp from \(pats, cs) |\| ?grp\ obtain pats\<^sub>e rhs\<^sub>e where "(pats\<^sub>e, rhs\<^sub>e) |\| rs\<^sub>1" "pats = butlast pats\<^sub>e" by (auto elim: fgroup_byE2) have "patterns_compatible (butlast pats') pats" unfolding \pats = _\ apply (rule rev_accum_rel_butlast) using \(pats', rhs) |\| rs\<^sub>1\ \(pats\<^sub>e, rhs\<^sub>e) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| rs\ inner by (blast dest: fpairwiseD) interpret irules': irules C_info "transform_irule_set rs" by (rule rules_transform) have "butlast pats' = pats" "env = env\<^sub>1" apply (rule matchs_compatible_eq) subgoal by fact subgoal apply (rule linears_butlastI) using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| rs\ inner by auto subgoal using \(pats, _) |\| rs\<^sub>2\ \(name, rs\<^sub>2) |\| transform_irule_set rs\ using irules'.inner by auto apply fact+ subgoal apply (rule matchs_compatible_eq) apply fact apply (rule linears_butlastI) using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| rs\ inner apply auto [] using \(pats, _) |\| rs\<^sub>2\ \(name, rs\<^sub>2) |\| transform_irule_set rs\ using irules'.inner apply auto[] by fact+ done let ?rhs_subst = "\env. subst rhs (fmdrop_fset (frees pat) env)" have "fmdom env\<^sub>2 = freess pats" using \match (_ $$ _) _ = Some env\<^sub>2\ by (simp add: match_dom) show "fBex cs' (rel_prod (=) prelated (pat, rhs'))" unfolding \rhs' = _\ proof (rule fBexI, rule rel_prod.intros) have "fdisjnt (freess (butlast pats')) (frees (last pats'))" apply (subst freess_single[symmetric]) apply (rule linears_appendD) apply (subst \butlast pats' @ [last pats'] = pats'\) using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| rs\ inner by auto show "subst rhs env \\<^sub>p ?rhs_subst env\<^sub>2" apply (rule prelated_subst) apply (rule prelated_refl) unfolding fmfilter_alt_defs apply (subst fmfilter_true) subgoal premises prems for x y using fmdomI[OF prems] unfolding \pat = _\ \fmdom env\<^sub>2 = _\ apply (subst (asm) \butlast pats' = pats\[symmetric]) using \fdisjnt (freess (butlast pats')) (frees (last pats'))\ by (auto simp: fdisjnt_alt_def) subgoal unfolding \env = _\ by fact done next have "(pat, rhs) |\| cs" unfolding \pat = _\ apply (rule fgroup_byD[where a = "(x, y)" for x y]) apply fact apply simp apply (intro conjI) apply fact apply (rule refl)+ apply fact done thus "(pat, ?rhs_subst env\<^sub>2) |\| cs'" unfolding \cs' = _\ by force qed simp next fix pat rhs' assume "(pat, rhs') |\| cs'" then obtain rhs where "(pat, rhs) |\| cs" and "rhs' = subst rhs (fmdrop_fset (frees pat) env\<^sub>2 )" unfolding \cs' = _\ by auto with \(pats, cs) |\| ?grp\ obtain pats' where "(pats', rhs) |\| rs\<^sub>1" "pats = butlast pats'" "pat = last pats'" by auto with False have "length pats' \ 0" using \arity_compatibles _\ by (metis arity_compatible_length) hence "pats' = pats @ [pat]" unfolding \pats = _\ \pat = _\ by auto moreover have "linears pats'" using \(pats', rhs) |\| rs\<^sub>1\ \(name, rs\<^sub>1) |\| _\ inner by auto ultimately have "fdisjnt (fmdom env\<^sub>1) (frees pat)" unfolding \fmdom env\<^sub>1 = _\ by (auto dest: linears_appendD) let ?rhs_subst = "\env. subst rhs (fmdrop_fset (frees pat) env)" show "fBex (irules_deferred_matches args\<^sub>1 rs\<^sub>1) (\e. rel_prod (=) prelated e (pat, rhs'))" unfolding \rhs' = _\ proof (rule fBexI, rule rel_prod.intros) show "?rhs_subst env\<^sub>1 \\<^sub>p ?rhs_subst env\<^sub>2" using \prelated.P_env env\<^sub>1 env\<^sub>2\ inner by (auto intro: prelated_subst) next have "matchs (butlast pats') args\<^sub>1 = Some env\<^sub>1" using \matchs pats args\<^sub>1 = _\ \pats = _\ by simp moreover have "subst rhs env\<^sub>1 = ?rhs_subst env\<^sub>1" apply (rule arg_cong[where f = "subst rhs"]) unfolding fmfilter_alt_defs apply (rule fmfilter_true[symmetric]) using \fdisjnt (fmdom env\<^sub>1) _\ by (auto simp: fdisjnt_alt_def intro: fmdomI) ultimately show "(pat, ?rhs_subst env\<^sub>1) |\| irules_deferred_matches args\<^sub>1 rs\<^sub>1" using \(pats', rhs) |\| rs\<^sub>1\ \pat = last pats'\ by auto qed simp qed qed qed next case ("fun" v v' u) obtain w x where "t = w $\<^sub>p x" "w \\<^sub>p v" "x \\<^sub>p u" "closed w" using \t \\<^sub>p v $\<^sub>p u\ \closed t\ by cases (auto simp: closed_except_def) with "fun" obtain w' where "rs \\<^sub>i w \* w'" "w' \\<^sub>p v'" by blast show ?case proof (rule fun.prems) show "rs \\<^sub>i t \* w' $\<^sub>p x" unfolding \t = _\ by (intro irewrite.rt_comb[unfolded app_pterm_def] rtranclp.rtrancl_refl) fact next show "w' $\<^sub>p x \\<^sub>p v' $\<^sub>p u" by (rule prelated.app) fact+ qed next case (arg u u' v) obtain w x where "t = w $\<^sub>p x" "w \\<^sub>p v" "x \\<^sub>p u" "closed x" using \t \\<^sub>p v $\<^sub>p u\ \closed t\ by cases (auto simp: closed_except_def) with arg obtain x' where "rs \\<^sub>i x \* x'" "x' \\<^sub>p u'" by blast show ?case proof (rule arg.prems) show "rs \\<^sub>i t \* w $\<^sub>p x'" unfolding \t = w $\<^sub>p x\ by (intro irewrite.rt_comb[unfolded app_pterm_def] rtranclp.rtrancl_refl) fact next show "w $\<^sub>p x' \\<^sub>p v $\<^sub>p u'" by (rule prelated.app) fact+ qed qed end subsubsection \Completeness of transformation\ lemma (in irules) transform_completeness: assumes "rs \\<^sub>i t \ t'" "closed t" shows "transform_irule_set rs \\<^sub>i t \* t'" using assms proof induction case (step name irs' pats' rhs' t t') then obtain irs where "irs = transform_irules irs'" "(name, irs) |\| transform_irule_set rs" unfolding transform_irule_set_def by (metis fimageI id_apply map_prod_simp) show ?case proof (cases "arity irs' = 0") case True hence "irs = irs'" unfolding \irs = _\ unfolding transform_irules_def by simp with step have "(pats', rhs') |\| irs" "name, pats', rhs' \\<^sub>i t \ t'" by blast+ have "transform_irule_set rs \\<^sub>i t \* t'" apply (rule r_into_rtranclp) apply rule by fact+ show ?thesis by fact next let ?f = "\(pats, rhs). (butlast pats, last pats, rhs)" let ?grp = "fgroup_by ?f irs'" note closed_except_def [simp add] case False then have "irs = map_prod id Pabs |`| ?grp" unfolding \irs = _\ unfolding transform_irules_def by simp with False have "irs = transform_irules irs'" unfolding transform_irules_def by simp obtain pat pats where "pat = last pats'" "pats = butlast pats'" by blast from step False have "length pats' \ 0" using arity_compatible_length inner by (smt (verit, ccfv_threshold) case_prodD) then have "pats' = pats @ [pat]" unfolding \pat = _\ \pats = _\ by simp from step have "linears pats'" using inner by auto then have "fdisjnt (freess pats) (frees pat)" unfolding \pats' = _\ using linears_appendD(3) freess_single by force from step obtain cs where "(pats, cs) |\| ?grp" unfolding \pats = _\ by (metis (no_types, lifting) fgroup_by_complete fst_conv prod.simps(2)) with step have "(pat, rhs') |\| cs" unfolding \pat = _\ \pats = _\ by (meson fgroup_byD old.prod.case) have "(pats, Pabs cs) |\| irs" using \irs = map_prod id Pabs |`| ?grp\ \(pats, cs) |\| _\ by (metis (no_types, lifting) eq_snd_iff fst_conv fst_map_prod id_def rev_fimage_eqI snd_map_prod) from step obtain env' where "match (name $$ pats') t = Some env'" "subst rhs' env' = t'" using irewrite_step_def by auto have "name $$ pats' = (name $$ pats) $ pat" unfolding \pats' = _\ by (simp add: app_term_def) then obtain t\<^sub>0 t\<^sub>1 env\<^sub>0 env\<^sub>1 where "t = t\<^sub>0 $\<^sub>p t\<^sub>1" "match (name $$ pats) t\<^sub>0 = Some env\<^sub>0" "match pat t\<^sub>1 = Some env\<^sub>1" "env' = env\<^sub>0 ++\<^sub>f env\<^sub>1" using match_appE_split[OF \match (name $$ pats') _ = _\[unfolded \name $$ pats' = _\], unfolded app_pterm_def] by blast with step have "closed t\<^sub>0" "closed t\<^sub>1" by auto then have "closed_env env\<^sub>0" "closed_env env\<^sub>1" using match_vars[OF \match _ t\<^sub>0 = _\] match_vars[OF \match _ t\<^sub>1 = _\] unfolding closed_except_def by auto obtain t\<^sub>0' where "subst (Pabs cs) env\<^sub>0 = t\<^sub>0'" by blast then obtain cs' where "t\<^sub>0' = Pabs cs'" "cs' = ((\(pat, rhs). (pat, subst rhs (fmdrop_fset (frees pat) env\<^sub>0))) |`| cs)" using subst_pterm.simps(3) by blast obtain rhs where "subst rhs' (fmdrop_fset (frees pat) env\<^sub>0) = rhs" by blast then have "(pat, rhs) |\| cs'" unfolding \cs' = _\ using \_ |\| cs\ by (metis (mono_tags, lifting) old.prod.case rev_fimage_eqI) have "env\<^sub>0 ++\<^sub>f env\<^sub>1 = (fmdrop_fset (frees pat) env\<^sub>0) ++\<^sub>f env\<^sub>1" apply (subst fmadd_drop_left_dom[symmetric]) using \match pat _ = _\ match_dom by metis have "fdisjnt (fmdom env\<^sub>0) (fmdom env\<^sub>1)" using match_dom using \match pat _ = _\ \match (name $$ pats) _ = _\ using \fdisjnt _ _\ unfolding fdisjnt_alt_def by (metis matchs_dom match_list_combE) have "subst rhs env\<^sub>1 = t'" unfolding \_ = rhs\[symmetric] unfolding \_ = t'\[symmetric] unfolding \env' = _\ unfolding \env\<^sub>0 ++\<^sub>f _ = _\ apply (subst subst_indep') using \closed_env env\<^sub>0\ apply blast using \fdisjnt (fmdom _) _\ unfolding fdisjnt_alt_def by auto show ?thesis unfolding \t = _\ apply rule apply (rule r_into_rtranclp) apply (rule irewrite.intros(3)) apply rule apply fact+ apply (rule irewrite_stepI) apply fact+ unfolding \t\<^sub>0' = _\ apply rule apply fact using \match pat t\<^sub>1 = _\ \subst rhs _ = _\ by force qed qed (auto intro: irewrite.rt_comb[unfolded app_pterm_def] intro!: irewrite.intros simp: closed_except_def) subsubsection \Computability\ export_code compile transform_irules checking Scala SML end \ No newline at end of file diff --git a/thys/ML_Unification/Examples/E_Unification_Examples.thy b/thys/ML_Unification/Examples/E_Unification_Examples.thy --- a/thys/ML_Unification/Examples/E_Unification_Examples.thy +++ b/thys/ML_Unification/Examples/E_Unification_Examples.thy @@ -1,156 +1,163 @@ \<^marker>\creator "Kevin Kappelmann"\ section \E-Unification Examples\ theory E_Unification_Examples imports Main ML_Unification_HOL_Setup + Unify_Assumption_Tactic Unify_Fact_Tactic + Unify_Resolve_Tactics begin paragraph \Summary\ text \Sample applications of e-unifiers, methods, etc. introduced in this session.\ experiment begin subsection \Using The Simplifier For Unification.\ inductive_set even :: "nat set" where zero: "0 \ even" | step: "n \ even \ Suc (Suc n) \ even" text \Premises of the form @{term "SIMPS_TO_UNIF lhs rhs"} are solved by @{ML_structure Simplifier_Unification}. It first normalises @{term lhs} and then unifies the normalisation with @{term rhs}. See also @{theory ML_Unification.ML_Unification_HOL_Setup}.\ lemma [uhint where prio = Prio.LOW]: "n \ 0 \ PROP SIMPS_TO_UNIF (n - 1) m \ n \ Suc m" unfolding SIMPS_TO_UNIF_eq by linarith text \By default, below unification methods use @{ML Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify}, which is a combination of various practical unification algorithms.\ schematic_goal "(\x. x + 4 = n) \ Suc ?x = n" by uassm lemma "6 \ even" apply (urule step) apply (urule step) apply (urule step) apply (urule zero) done lemma "(220 + (80 - 2 * 2)) \ even" apply (urule step) apply (urule step)+ apply (urule zero) done lemma assumes "[a,b,c] = [c,b,a]" shows "[a] @ [b,c] = [c,b,a]" using assms by uassm lemma "x \ ({z, y, x} \ S) \ {x}" by (ufact TrueI) -lemma "(x + (y :: nat))^2 \ x^2 + 2*x*y + y^2 + 4 * y + x - y" +schematic_goal "(x + (y :: nat))^2 \ x^2 + 2*x*y + y^2 + 4 * y + x - y" supply power2_sum[simp] by (ufact TrueI) lemma assumes "\s. P (Suc (Suc 0)) (s(x := (1 :: nat), x := 1 + 1 * 4 - 3))" shows "P 2 (s(x := 2))" - by (ufact assms[of s]) + by (ufact assms) subsection \Providing Canonical Solutions With Unification Hints\ -lemma length_nil_eq_zero [uhint]: "length [] \ 0" by simp - -schematic_goal "length ?xs = 0" - by (ufact refl) - lemma sub_self_eq_zero [uhint]: "(n :: nat) - n \ 0" by simp schematic_goal "n - ?m = (0 :: nat)" by (ufact refl) -text \The following fails because, by default, @{ML Standard_Unification_Hints.try_hints} -uses the higher-order pattern unifier to unify hints against a given disagreement pair, and -@{term 0} cannot be higher-order pattern unified with @{term "length []"}. The unification of the -hint requires the use of yet another hint, namely @{term "length [] = 0"} (cf. above).\ +text \The following example shows a non-trivial interplay of the simplifier and unification hints: +Using just unification, the hint @{thm sub_self_eq_zero} is not applicable in the following example +since @{term 0} cannot be unified with @{term "length []"}. +However, the simplifier can rewrite @{term "length []"} to @{term 0} and the hint can then be applied.\ + +(*uncomment to see the trace*) +declare [[ML_map_context \Logger.set_log_levels Logger.root Logger.TRACE\]] schematic_goal "n - ?m = length []" - \ \by (ufact refl)\ - oops + by (ufact refl) -text \There are two ways to fix this: -\<^enum> We allow the recursive use of unification hints when unifying @{thm sub_self_eq_zero} and our goal. -\<^enum> We use a different unification hint that makes the recursive use of unification hints explicit.\ +text \There are also two ways to solve this using only unification hints: +\<^enum> We allow the recursive use of unification hints when unifying @{thm sub_self_eq_zero} and our goal +and register @{term "length [] = 0"} as an additional hint. +\<^enum> We use an alternative for @{thm sub_self_eq_zero} that makes the recursive use of unification +hints explicit and register @{term "length [] = 0"} as an additional hint.\ + +lemma length_nil_eq [uhint]: "length [] = 0" by simp text \Solution 1: we can use @{attribute rec_uhint} for recursive usages of hints. Warning: recursive hint applications easily loop.\ + schematic_goal "n - ?m = length []" + supply [[ucombine del = \(Standard_Unification_Combine.default_metadata \<^binding>\simp_unif\)\]] + (*doesn't work*) + \ \by (ufact refl)\ supply sub_self_eq_zero[rec_uhint] by (ufact refl) text \Solution 2: make the recursion explicit in the hint.\ lemma [uhint]: "k \ 0 \ (n :: nat) \ m \ n - m \ k" by simp schematic_goal "n - ?m = length []" + supply [[ucombine del = \(Standard_Unification_Combine.default_metadata \<^binding>\simp_unif\)\]] by (ufact refl) - subsection \Strenghten Unification With Unification Hints\ lemma assumes [uhint]: "n = m" shows "n - m = (0 :: nat)" by (ufact refl) lemma assumes "x = y" shows "y = x" supply eq_commute[uhint] by (ufact assms) paragraph \Unfolding definitions.\ definition "mysuc n = Suc n" lemma assumes "\m. Suc n > mysuc m" shows "mysuc n > Suc 3" supply mysuc_def[uhint] by (ufact assms) paragraph \Discharging meta impliciations with object-level implications\ lemma [uhint]: "Trueprop A \ A' \ Trueprop B \ B' \ Trueprop (A \ B) \ (PROP A' \ PROP B')" using atomize_imp[symmetric] by simp lemma assumes "A \ (B \ C) \ D" shows "A \ (B \ C) \ D" using assms by ufact subsection \Better Control Over Meta Variable Instantiations\ text \Consider the following type-inference problem.\ schematic_goal assumes app_typeI: "\f x. (\x. ArgT x \ DomT x (f x)) \ ArgT x \ DomT x (f x)" and f_type: "\x. ArgT x \ DomT x (f x)" and x_type: "ArgT x" shows "?T (f x)" apply (urule app_typeI) \\compare with the following application, creating an (unintuitive) higher-order instantiation\ (* apply (rule app_typeI) *) oops end end diff --git a/thys/ML_Unification/Examples/Unification_Hints_Reification_Examples.thy b/thys/ML_Unification/Examples/Unification_Hints_Reification_Examples.thy --- a/thys/ML_Unification/Examples/Unification_Hints_Reification_Examples.thy +++ b/thys/ML_Unification/Examples/Unification_Hints_Reification_Examples.thy @@ -1,213 +1,216 @@ \<^marker>\creator "Kevin Kappelmann"\ \<^marker>\contributor "Paul Bachmann"\ section \Examples: Reification Via Unification Hints\ theory Unification_Hints_Reification_Examples imports HOL.Rat ML_Unification_HOL_Setup Unify_Fact_Tactic + Unify_Resolve_Tactics begin paragraph \Summary\ text \Reification via unification hints. For an introduction to unification hints refer to \<^cite>\"unif-hints"\. We support a generalisation of unification hints as described in @{theory ML_Unification.ML_Unification_Hints}.\ subsection \Setup\ text \One-time setup to obtain a unifier with unification hints for the purpose of reification. We could also simply use the standard unification hints @{attribute uhint} and @{attribute rec_uhint}, but having separate instances is a cleaner approach.\ ML\ @{functor_instance struct_name = Reification_Unification_Hints and functor_name = Term_Index_Unification_Hints and id = \"reify"\ and more_args = \ structure TI = Discrimination_Tree val init_args = { concl_unifier = NONE, prems_unifier = NONE, normalisers = SOME Higher_Order_Pattern_Unification.norms_unify, retrieval = SOME (Term_Index_Unification_Hints_Args.mk_sym_retrieval TI.norm_term TI.unifiables), hint_preprocessor = SOME (Standard_Unification_Hints.get_hint_preprocessor (Context.the_generic_context ())) }\} val reify_unify = Unification_Combinator.add_fallback_unifier (fn unif_theory => - Higher_Order_Pattern_Unification.e_unify Unification_Util.unify_types unif_theory unif_theory) + Higher_Order_Pattern_Unification.e_unify Unification_Util.unify_types unif_theory unif_theory + |> Type_Unification.e_unify Unification_Util.unify_types) (Reification_Unification_Hints.try_hints |> Unification_Combinator.norm_unifier (#norm_term Higher_Order_Pattern_Unification.norms_unify)) \ local_setup \Reification_Unification_Hints.setup_attribute NONE\ text \Premises of hints should again be unified by the reification unifier.\ declare [[reify_uhint where prems_unifier = reify_unify]] subsection \Formulas with Quantifiers and Environment\ text \The following example is taken from HOL-Library.Reflection\_Examples. It is recommended to compare the approach presented here with the reflection tactic presented in said theory.\ datatype form = TrueF | FalseF | Less nat nat | And form form | Or form form | Neg form | ExQ form primrec interp :: "form \ ('a::ord) list \ bool" where "interp TrueF vs \ True" | "interp FalseF vs \ False" | "interp (Less i j) vs \ vs ! i < vs ! j" | "interp (And f1 f2) vs \ interp f1 vs \ interp f2 vs" | "interp (Or f1 f2) vs \ interp f1 vs \ interp f2 vs" | "interp (Neg f) vs \ \ interp f vs" | "interp (ExQ f) vs \ (\v. interp f (v # vs))" paragraph \Reification with unification and recursive hint unification for conclusion\ text \The following illustrates how to use the equations @{thm interp.simps} directly as unification hints for reification.\ experiment begin text \Hints for list lookup.\ declare List.nth_Cons_Suc[reify_uhint where prio = Prio.LOW] and List.nth_Cons_0[reify_uhint] text \Hints to reify formulas of type @{type bool} into formulas of type @{type form}.\ declare interp.simps[reify_uhint] text \We have to allow the hint unifier to recursively look for hints during unification of the hint's conclusion.\ declare [[reify_uhint where concl_unifier = reify_unify]] (*uncomment the following to see the hint unification process*) (* declare [[ML_map_context \Logger.set_log_levels Unification_Base.logger Logger.DEBUG\]] *) schematic_goal "interp ?f (?vs :: ('a :: ord) list) = (\(x :: 'a). x < y \ \(\(z :: 'a). v < z \ \False))" by (ufact refl where unifier = reify_unify) text \While this all works nicely if set up correctly, it can be rather difficult to understand and debug the recursive unification process for a hint's conclusion. In the next paragraph, we present an alternative that is closer to the examples presented in the original unification hints paper \<^cite>\"unif-hints"\.\ end paragraph \Reification with matching without recursion for conclusion\ text \We disallow the hint unifier to recursively look for hints while unifying the conclusion; instead, we only allow the hint unifier to match the hint's conclusion against the disagreement terms.\ -declare [[reify_uhint where concl_unifier = Higher_Order_Pattern_Unification.match]] +declare [[reify_uhint where concl_unifier = + \Higher_Order_Pattern_Unification.match |> Type_Unification.e_match Unification_Util.match_types\]] text \However, this also means that we now have to write our hints such that the hint's conclusion can successfully be matched against the disagreement terms. In particular, the disagreement terms may still contain meta variables that we want to instantiate with the help of the unification hints. Essentially, a hint then describes a canonical instantiation for these meta variables.\ experiment begin lemma [reify_uhint where prio = Prio.LOW]: "n \ Suc n' \ vs \ v # vs' \ vs' ! n' \ x \ vs ! n \ x" by simp lemma [reify_uhint]: "n \ 0 \ vs \ x # vs' \ vs ! n \ x" by simp lemma [reify_uhint]: "\e \ ExQ f; \v. interp f (v # vs) \ P v\ \ interp e vs \ \v. P v" "\e \ Less i j; x \ vs ! i; y \ vs ! j\ \ interp e vs \ x < y" "\e \ And f1 f2; interp f1 vs \ r1; interp f2 vs \ r2\ \ interp e vs \ r1 \ r2" "\e \ Or f1 f2; interp f1 vs \ r1; interp f2 vs \ r2\ \ interp e vs \ r1 \ r2" "e \ Neg f \ interp f vs \ r \ interp e vs \ \r" "e \ TrueF \ interp e vs \ True" "e \ FalseF \ interp e vs \ False" by simp_all schematic_goal "interp ?f (?vs :: ('a :: ord) list) = (\(x :: 'a). x < y \ \(\(z :: 'a). v < z \ \False))" by (urule refl where unifier = reify_unify) end text \The next examples are modification from \<^cite>\"unif-hints"\.\ subsection \Simple Arithmetic\ datatype add_expr = Var int | Add add_expr add_expr fun eval_add_expr :: "add_expr \ int" where "eval_add_expr (Var i) = i" | "eval_add_expr (Add ex1 ex2) = eval_add_expr ex1 + eval_add_expr ex2" lemma eval_add_expr_Var [reify_uhint where prio = Prio.LOW]: "e \ Var i \ eval_add_expr e \ i" by simp lemma eval_add_expr_add [reify_uhint]: "e \ Add e1 e2 \ eval_add_expr e1 \ m \ eval_add_expr e2 \ n \ eval_add_expr e \ m + n" by simp ML_command\ val t1 = Proof_Context.read_term_pattern @{context} "eval_add_expr ?e" val t2 = Proof_Context.read_term_pattern @{context} "1 + (2 + 7) :: int" val _ = Unification_Util.log_unif_results @{context} (t1, t2) (reify_unify []) \ schematic_goal "eval_add_expr ?e = (1 + (2 + 7) :: int)" by (urule refl where unifier = reify_unify) subsection \Arithmetic with Environment\ datatype mul_expr = Unit | Var nat | Mul mul_expr mul_expr | Inv mul_expr fun eval_mul_expr :: "mul_expr \ rat list \ rat" where "eval_mul_expr (Unit, \) = 1" | "eval_mul_expr (Var i, \) = \ ! i" | "eval_mul_expr (Mul e1 e2, \) = eval_mul_expr (e1, \) * eval_mul_expr (e2, \)" | "eval_mul_expr (Inv e, \) = inverse (eval_mul_expr (e, \))" text \Split @{term e} into an expression and an environment.\ lemma [reify_uhint where prio = Prio.VERY_LOW]: "e \ (e1, \) \ eval_mul_expr (e1, \) \ n \ eval_mul_expr e \ n" by simp text \Hints for environment lookup.\ lemma [reify_uhint where prio = Prio.LOW]: "e \ Var (Suc p) \ \ \ s # \ \ n \ eval_mul_expr (Var p, \) \ eval_mul_expr (e, \) \ n" by simp lemma [reify_uhint]: "e \ Var 0 \ \ \ n # \ \ eval_mul_expr (e, \) \ n" by simp lemma [reify_uhint]: "e1 \ Inv e2 \ n \ eval_mul_expr (e2, \) \ eval_mul_expr (e1, \) \ inverse n" "e \ Mul e1 e2 \ m \ eval_mul_expr (e1, \) \ n \ eval_mul_expr (e2, \) \ eval_mul_expr (e, \) \ m * n" "e \ Unit \ eval_mul_expr (e, \) \ 1" by simp_all ML_command\ val t1 = Proof_Context.read_term_pattern @{context} "eval_mul_expr ?e" val t2 = Proof_Context.read_term_pattern @{context} "1 * inverse 3 * 5 :: rat" val _ = Unification_Util.log_unif_results' 1 @{context} (t2, t1) (reify_unify []) \ schematic_goal "eval_mul_expr ?e = (1 * inverse 3 * 5 :: rat)" by (ufact refl where unifier = reify_unify) end diff --git a/thys/ML_Unification/Logger/ML_Logger_Examples.thy b/thys/ML_Unification/Logger/ML_Logger_Examples.thy --- a/thys/ML_Unification/Logger/ML_Logger_Examples.thy +++ b/thys/ML_Unification/Logger/ML_Logger_Examples.thy @@ -1,111 +1,111 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Examples\ theory ML_Logger_Examples imports ML_Logger Setup_Result_Commands begin text \First some simple, barebone logging: print some information.\ ML_command\ \ \the following two are equivalent\ - val _ = Logger.log Logger.root_logger Logger.INFO @{context} (K "hello root logger") - val _ = @{log Logger.INFO Logger.root_logger} @{context} (K "hello root logger") + val _ = Logger.log Logger.root Logger.INFO @{context} (K "hello root logger") + val _ = @{log Logger.INFO Logger.root} @{context} (K "hello root logger") \ ML_command\ - val logger = Logger.root_logger + val logger = Logger.root val _ = @{log} @{context} (K "hello root logger") \ \\@{log}\ is equivalent to \Logger.log logger Logger.INFO\\ \ text \To guarantee the existence of a "logger" in an ML structure, one should use the \HAS_LOGGER\ signature.\ ML\ structure My_Struct : sig include HAS_LOGGER val get_n : Proof.context -> int end = struct - val logger = Logger.setup_new_logger Logger.root_logger "My_Struct" + val logger = Logger.setup_new_logger Logger.root "My_Struct" fun get_n ctxt = (@{log} ctxt (K "retrieving n..."); 42) end \ ML_command\val n = My_Struct.get_n @{context}\ text\We can set up a hierarchy of loggers\ ML\ - val logger = Logger.root_logger - val parent1 = Logger.setup_new_logger Logger.root_logger "Parent1" + val logger = Logger.root + val parent1 = Logger.setup_new_logger Logger.root "Parent1" val child1 = Logger.setup_new_logger parent1 "Child1" val child2 = Logger.setup_new_logger parent1 "Child2" - val parent2 = Logger.setup_new_logger Logger.root_logger "Parent2" + val parent2 = Logger.setup_new_logger Logger.root "Parent2" \ ML_command\ - (@{log Logger.INFO Logger.root_logger} @{context} (K "Hello root logger"); + (@{log Logger.INFO Logger.root} @{context} (K "Hello root logger"); @{log Logger.INFO parent1} @{context} (K "Hello parent1"); @{log Logger.INFO child1} @{context} (K "Hello child1"); @{log Logger.INFO child2} @{context} (K "Hello child2"); @{log Logger.INFO parent2} @{context} (K "Hello parent2")) \ text \We can use different log levels to show/surpress messages. The log levels are based on Apache's Log4J 2 \<^url>\https://logging.apache.org/log4j/2.x/manual/customloglevels.html\.\ ML_command\@{log Logger.DEBUG parent1} @{context} (K "Hello parent1")\ \\prints nothings\ declare [[ML_map_context \Logger.set_log_level parent1 Logger.DEBUG\]] ML_command\@{log Logger.DEBUG parent1} @{context} (K "Hello parent1")\ \\prints message\ ML_command\Logger.ALL\ \\ctrl+click on the value to see all log levels\ text \We can set options for all loggers below a given logger. Below, we set the log level for all loggers below (and including) \<^ML>\parent1\ to error, thus disabling warning messages.\ ML_command\ (@{log Logger.WARN parent1} @{context} (K "Warning from parent1"); @{log Logger.WARN child1} @{context} (K "Warning from child1")) \ declare [[ML_map_context \Logger.set_log_levels parent1 Logger.ERR\]] ML_command\ (@{log Logger.WARN parent1} @{context} (K "Warning from parent1"); @{log Logger.WARN child1} @{context} (K "Warning from child1")) \ declare [[ML_map_context \Logger.set_log_levels parent1 Logger.INFO\]] text \We can set message filters.\ -declare [[ML_map_context \Logger.set_msg_filters Logger.root_logger (match_string "Third")\]] +declare [[ML_map_context \Logger.set_msg_filters Logger.root (match_string "Third")\]] ML_command\ (@{log Logger.INFO parent1} @{context} (K "First message"); @{log Logger.INFO child1} @{context} (K "Second message"); @{log Logger.INFO child2} @{context} (K "Third message"); @{log Logger.INFO parent2} @{context} (K "Fourth message")) \ -declare [[ML_map_context \Logger.set_msg_filters Logger.root_logger (K true)\]] +declare [[ML_map_context \Logger.set_msg_filters Logger.root (K true)\]] text \One can also use different output channels (e.g. files) and hide/show some additional logging information. Ctrl+click on below values and explore.\ ML_command\Logger.set_output; Logger.set_show_logger; Logging_Antiquotation.show_log_pos\ text \To set up (local) loggers outside ML environments, @{theory ML_Unification.Setup_Result_Commands} contains two commands, @{command setup_result} and @{command local_setup_result}.\ experiment begin -local_setup_result local_logger = \Logger.new_logger Logger.root_logger "Local"\ +local_setup_result local_logger = \Logger.new_logger Logger.root "Local"\ ML_command\@{log Logger.INFO local_logger} @{context} (K "Hello local world")\ end text \\local_logger\ is no longer available. The follow thus does not work:\ \ \ML_command\@{log Logger.INFO local_logger} @{context} (K "Hello local world")\\ text \Let us create another logger in the global context.\ -setup_result some_logger = \Logger.new_logger Logger.root_logger "Some_Logger"\ +setup_result some_logger = \Logger.new_logger Logger.root "Some_Logger"\ ML_command\@{log Logger.INFO some_logger} @{context} (K "Hello world")\ text \Let us delete it again.\ declare [[ML_map_context \Logger.delete_logger some_logger\]] text \The logger can no longer be found in the logger hierarchy\ ML_command\@{log Logger.INFO some_logger} @{context} (K "Hello world")\ end diff --git a/thys/ML_Unification/Logger/logger.ML b/thys/ML_Unification/Logger/logger.ML --- a/thys/ML_Unification/Logger/logger.ML +++ b/thys/ML_Unification/Logger/logger.ML @@ -1,225 +1,228 @@ (* Title: Logger/logger.ML Author: Kevin Kappelmann, Paul Bachmann Hierarchical logger, indexed on bindings. The log levels are based on Apache's Log4J 2 https://logging.apache.org/log4j/2.x/manual/customloglevels.html *) signature LOGGER = sig type log_level = int val OFF : log_level val FATAL : log_level (*error log level*) val ERR : log_level val WARN : log_level val INFO : log_level val DEBUG : log_level val TRACE : log_level val ALL : log_level type logger_name = bstring type logger_binding - val root_logger : logger_binding + val root : logger_binding (*takes parent logger and creates a child logger with the given name*) val create_child_logger : logger_binding -> logger_name -> logger_binding val pretty_binding : logger_binding -> Pretty.T type logger_output = log_level -> string -> unit val default_output : logger_output type msg_filter = string -> bool type logger_config val config : logger_output -> log_level -> msg_filter -> bool -> logger_config val default_config : logger_config val set_config_output : logger_output -> logger_config -> logger_config val set_config_log_level : log_level -> logger_config -> logger_config (*set message filter: only log messages with positive results*) val set_config_msg_filter : msg_filter -> logger_config -> logger_config (*set whether to print the logger's name when logging*) val set_config_show_logger : bool -> logger_config -> logger_config val lookup_logger : logger_binding -> Context.generic -> logger_config option val insert_logger : (logger_binding * logger_config) -> Context.generic -> Context.generic val insert_logger_safe : (logger_binding * logger_config) -> Context.generic -> Context.generic val delete_logger : logger_binding -> Context.generic -> Context.generic val cut_loggers : logger_binding -> Context.generic -> Context.generic val set_logger : logger_binding -> (logger_config -> logger_config) -> Context.generic -> Context.generic val set_loggers : logger_binding -> (logger_config -> logger_config) -> Context.generic -> Context.generic val set_output : logger_binding -> logger_output -> Context.generic -> Context.generic val set_outputs : logger_binding -> logger_output -> Context.generic -> Context.generic val set_log_level : logger_binding -> log_level -> Context.generic -> Context.generic val set_log_levels : logger_binding -> log_level -> Context.generic -> Context.generic val set_msg_filter : logger_binding -> msg_filter -> Context.generic -> Context.generic val set_msg_filters : logger_binding -> msg_filter -> Context.generic -> Context.generic val set_show_logger : logger_binding -> bool -> Context.generic -> Context.generic val set_show_loggers : logger_binding -> bool -> Context.generic -> Context.generic (*creates and inserts child logger with default configuration into context*) val new_logger : logger_binding -> logger_name -> Context.generic -> (logger_binding * Context.generic) (*registers new logger to background context*) val setup_new_logger : logger_binding -> logger_name -> logger_binding (*prints message created by passed function to the logger's output if the logger's log_level is greater or equal than the passed level and the message filter answers positively; uses lazy computation of the message to avoid computations in case the log level blocks the logging.*) val log : logger_binding -> log_level -> Proof.context -> (unit -> string) -> unit (* logging functions for different log levels *) val fatal : logger_binding -> Proof.context -> (unit -> string) -> unit val err : logger_binding -> Proof.context -> (unit -> string) -> unit val warn : logger_binding -> Proof.context -> (unit -> string) -> unit val info : logger_binding -> Proof.context -> (unit -> string) -> unit val debug : logger_binding -> Proof.context -> (unit -> string) -> unit val trace : logger_binding -> Proof.context -> (unit -> string) -> unit end structure Logger : LOGGER = struct type log_level = int (*values for different log levels*) val OFF = 0 val FATAL = 100 val ERR = 200 val WARN = 300 val INFO = 400 val DEBUG = 500 val TRACE = 600 val ALL = 1000 type logger_name = bstring datatype logger_binding = Logger_Binding of binding fun binding_of (Logger_Binding binding) = binding -val root_logger_name = "Root" -val root_logger = Binding.name root_logger_name |> Logger_Binding +val root_name = "Root" +val root = Binding.name root_name |> Logger_Binding fun create_child_logger (Logger_Binding parent) child_name = let val child = Binding.qualify_name true parent child_name in if Symbol_Pos.is_identifier (Binding.name_of child) then Logger_Binding child else error ("Bad identifier for child logger " ^ ML_Syntax.print_string child_name) end val pretty_binding = Binding.pretty o binding_of type logger_output = log_level -> string -> unit fun default_output log_level = if log_level <= WARN then warning else if log_level < DEBUG then writeln else tracing type msg_filter = string -> bool type logger_config = { log_level : log_level, msg_filter : msg_filter, output : logger_output, show_logger : bool } fun config output log_level msg_filter show_logger = {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} val default_config = config default_output INFO (K true) true fun set_config_log_level log_level {output, show_logger, msg_filter,...} = {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} fun set_config_output output {log_level, show_logger, msg_filter,...} = {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} fun set_config_msg_filter msg_filter {output, log_level, show_logger,...} = {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} fun set_config_show_logger show_logger {output, log_level, msg_filter,...} = {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} fun log_config binding {log_level, output, msg_filter, show_logger} level message_f = if level > log_level then () else let val msg = message_f () - |> show_logger ? (fn msg => cat_lines [ - "Logger: " ^ Pretty.string_of (pretty_binding binding), - msg]) + |> show_logger ? (fn msg => Pretty.fbreaks [ + Pretty.block [Pretty.str "Logger: ", pretty_binding binding], + Pretty.str msg + ] |> Pretty.block |> Pretty.string_of) in if msg_filter msg then output level msg else () end structure BT = Binding_Tree -val init_tree = BT.insert (binding_of root_logger, default_config) BT.empty +val init_tree = BT.insert (binding_of root, default_config) BT.empty structure Logger_Data = Generic_Data( type T = logger_config BT.binding_tree val empty = init_tree val merge = BT.merge ) fun lookup_logger (Logger_Binding binding) = Logger_Data.get #> (fn bt => BT.lookup bt binding) fun insert_logger bcp context = (Logger_Data.map (apfst binding_of bcp |> BT.insert) context) handle (exn as BT.INSERT) => (warning (Pretty.block [ Pretty.str "Logger ", pretty_binding (fst bcp), Pretty.str " already added." ] |> Pretty.string_of); Exn.reraise exn) fun insert_logger_safe bcp = Logger_Data.map (apfst binding_of bcp |> BT.insert_safe) fun delete_logger (Logger_Binding binding) = Logger_Data.map (BT.delete_safe binding) fun cut_loggers (Logger_Binding binding) = Logger_Data.map (BT.cut_safe binding) fun set_logger (Logger_Binding binding) f = Logger_Data.map (BT.map binding (Option.map f)) fun set_loggers (Logger_Binding binding) f = Logger_Data.map (BT.map_below binding (Option.map f)) fun set_output binding = set_logger binding o set_config_output fun set_outputs binding = set_loggers binding o set_config_output fun set_log_level binding = set_logger binding o set_config_log_level fun set_log_levels binding = set_loggers binding o set_config_log_level fun set_msg_filter binding = set_logger binding o set_config_msg_filter fun set_msg_filters binding = set_loggers binding o set_config_msg_filter fun set_show_logger binding = set_logger binding o set_config_show_logger fun set_show_loggers binding = set_loggers binding o set_config_show_logger fun new_logger parent child_name context = let val child = create_child_logger parent child_name in (child, insert_logger (child, default_config) context) end fun setup_new_logger parent child_name = Context.>>> (new_logger parent child_name) fun log binding log_level ctxt message_f = case lookup_logger binding (Context.Proof ctxt) of SOME config => log_config binding config log_level message_f | NONE => - let fun warn_msg _ = - "Logger " ^ Pretty.string_of (pretty_binding binding) ^ " not found." + let fun warn_msg _ = Pretty.fbreaks [ + Pretty.block [Pretty.str "Logger ", pretty_binding binding, Pretty.str " not found."], + Pretty.block (map Pretty.str ["Passed message: ", message_f ()]) + ] |> Pretty.block |> Pretty.string_of in - if binding = root_logger + if binding = root then default_output WARN (warn_msg ()) - else log root_logger WARN ctxt warn_msg + else log root WARN ctxt warn_msg end fun fatal binding = log binding FATAL fun err binding = log binding ERR fun warn binding = log binding WARN fun info binding = log binding INFO fun debug binding = log binding DEBUG fun trace binding = log binding TRACE end (*structures that use a logger should implement this signature*) signature HAS_LOGGER = sig val logger : Logger.logger_binding end diff --git a/thys/ML_Unification/ML_Unification_HOL_Setup.thy b/thys/ML_Unification/ML_Unification_HOL_Setup.thy --- a/thys/ML_Unification/ML_Unification_HOL_Setup.thy +++ b/thys/ML_Unification/ML_Unification_HOL_Setup.thy @@ -1,36 +1,34 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Setup for HOL\ theory ML_Unification_HOL_Setup imports HOL.HOL ML_Unification_Hints begin lemma eq_eq_True: "P \ (P \ Trueprop True)" by standard+ simp_all declare [[uhint where hint_preprocessor = \Unification_Hints_Base.obj_logic_hint_preprocessor @{thm atomize_eq[symmetric]} (Conv.rewr_conv @{thm eq_eq_True})\]] and [[rec_uhint where hint_preprocessor = \Unification_Hints_Base.obj_logic_hint_preprocessor @{thm atomize_eq[symmetric]} (Conv.rewr_conv @{thm eq_eq_True})\]] lemma eq_TrueI: "PROP P \ PROP P \ Trueprop True" by (standard) simp declare [[ucombine add = \Standard_Unification_Combine.eunif_data (Simplifier_Unification.SIMPS_TO_unify @{thm eq_TrueI} |> Unification_Combinator.norm_closed_unifier (#norm_term Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify) |> Unification_Combinator.unifier_from_closed_unifier |> K) - (Standard_Unification_Combine.default_metadata \<^binding>\SIMPS_TO_unif\)\]] + (Standard_Unification_Combine.metadata \<^binding>\SIMPS_TO_unif\ Prio.HIGH)\]] declare [[ucombine add = \Standard_Unification_Combine.eunif_data - (Simplifier_Unification.SIMPS_TO_UNIF_unify @{thm eq_TrueI} + (Simplifier_Unification.simp_unify_progress Envir.aeconv + (Simplifier_Unification.SIMPS_TO_UNIF_unify @{thm eq_TrueI}) + (#norm_term Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify) Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify - (Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify - |> Unification_Combinator.norm_unifier Envir_Normalisation.beta_norm_term_unif) - |> Unification_Combinator.norm_closed_unifier - (#norm_term Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify) - |> Unification_Combinator.unifier_from_closed_unifier + Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify |> K) - (Standard_Unification_Combine.default_metadata \<^binding>\SIMPS_TO_UNIF_unif\)\]] + (Standard_Unification_Combine.metadata \<^binding>\SIMPS_TO_UNIF_unif\ Prio.HIGH)\]] end diff --git a/thys/ML_Unification/ML_Utils/Tactics/tactic_util.ML b/thys/ML_Unification/ML_Utils/Tactics/tactic_util.ML --- a/thys/ML_Unification/ML_Utils/Tactics/tactic_util.ML +++ b/thys/ML_Unification/ML_Utils/Tactics/tactic_util.ML @@ -1,260 +1,266 @@ (* Title: ML_Utils/tactic_util.ML Author: Kevin Kappelmann Tactic utilities. *) signature TACTIC_UTIL = sig include HAS_LOGGER (* tactic combinators*) val HEADGOAL : (int -> 'a) -> 'a + val THEN : ('a -> 'b Seq.seq) * ('b -> 'c Seq.seq) -> 'a -> 'c Seq.seq + val THEN' : ('a -> 'b -> 'c Seq.seq) * ('a -> 'c -> 'd Seq.seq) -> 'a -> 'b -> 'd Seq.seq + val TRY' : ('a -> tactic) -> 'a -> tactic val EVERY_ARG : ('a -> tactic) -> 'a list -> tactic val EVERY_ARG' : ('a -> 'b -> tactic) -> 'a list -> 'b -> tactic val CONCAT : tactic list -> tactic val CONCAT' : ('a -> tactic) list -> 'a -> tactic val FOCUS_PARAMS' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic val FOCUS_PARAMS_CTXT : (Proof.context -> tactic) -> Proof.context -> int -> tactic val FOCUS_PARAMS_CTXT' : (Proof.context -> int -> tactic) -> Proof.context -> int -> tactic val FOCUS_PARAMS_FIXED' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic val FOCUS_PREMS' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic val FOCUS' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic val SUBPROOF' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic val CSUBGOAL_DATA : (cterm -> 'a) -> ('a -> int -> tactic) -> int -> tactic val CSUBGOAL_STRIPPED : (cterm Binders.binders * (cterm list * cterm) -> 'a) -> ('a -> int -> tactic) -> int -> tactic val SUBGOAL_DATA : (term -> 'a) -> ('a -> int -> tactic) -> int -> tactic val SUBGOAL_STRIPPED : ((string * typ) list * (term list * term) -> 'a) -> ('a -> int -> tactic) -> int -> tactic (* tactics *) val insert_tac : thm list -> Proof.context -> int -> tactic (*thin_tac n i deletes nth premise of the ith subgoal*) val thin_tac : int -> int -> tactic val thin_tacs : int list -> int -> tactic val set_kernel_ho_unif_bounds : int -> int -> Proof.context -> Proof.context val silence_kernel_ho_bounds_exceeded : Proof.context -> Proof.context val safe_simp_tac : Proof.context -> int -> tactic (*nth_eq_assume_tac n i solves ith subgoal by assumption, without unification, preferring premise n*) val nth_eq_assume_tac : int -> int -> tactic (*resolution without lifting of premises and parameters*) val no_lift_biresolve_tac : bool -> thm -> int -> Proof.context -> int -> tactic val no_lift_resolve_tac : thm -> int -> Proof.context -> int -> tactic val no_lift_eresolve_tac : thm -> int -> Proof.context -> int -> tactic (*rewrite subgoal according to the given equality theorem "lhs \ subgoal"*) val rewrite_subgoal_tac : thm -> Proof.context -> int -> tactic (*creates equality theorem for a subgoal from an equality theorem for the subgoal's conclusion; fails if the equality theorem contains a tpair or implicit Assumption mentioning one of the bound variables*) val eq_subgoal_from_eq_concl : cterm Binders.binders -> cterm list -> thm -> Proof.context -> thm option (*rewrite a subgoal given an equality theorem and environment for the subgoal's conclusion*) val rewrite_subgoal_unif_concl : Envir_Normalisation.term_normaliser -> Envir_Normalisation.thm_normaliser -> Envir_Normalisation.thm_normaliser -> term Binders.binders -> (Envir.env * thm) -> Proof.context -> int -> tactic (*protect premises starting from (and excluding) the given index in the specified subgoal*) val protect_tac : int -> Proof.context -> int -> tactic (*unprotect the subgoal*) val unprotect_tac : Proof.context -> int -> tactic (*protect, then apply tactic, then unprotect*) val protected_tac : int -> (int -> tactic) -> Proof.context -> int -> tactic (*focus on the specified subgoal, introducing only the specified list of premises as theorems in the focus*) val focus_prems_tac : int list -> (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic (*focus_prems_tac and then deletes all focused premises*) val focus_delete_prems_tac : int list -> (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic (*apply tactic to the specified subgoal, where the state is given as a cterm*) val apply_tac : (int -> tactic) -> int -> cterm -> thm Seq.seq end structure Tactic_Util : TACTIC_UTIL = struct -val logger = Logger.setup_new_logger Logger.root_logger "Tactic_Util" +val logger = Logger.setup_new_logger Logger.root "Tactic_Util" (* tactic combinators *) fun HEADGOAL f = f 1 +fun (tac1 THEN tac2) = Seq.THEN (tac1, tac2) +fun (tac1 THEN' tac2) x = tac1 x THEN tac2 x + fun TRY' tac = tac ORELSE' K all_tac fun EVERY_ARG tac = EVERY o map tac fun EVERY_ARG' tac = EVERY' o map tac fun CONCAT tacs = fold_rev (curry op APPEND) tacs no_tac fun CONCAT' tacs = fold_rev (curry op APPEND') tacs (K no_tac) fun FOCUS_PARAMS' tac = Subgoal.FOCUS_PARAMS (HEADGOAL o tac) fun FOCUS_PARAMS_FIXED' tac = Subgoal.FOCUS_PARAMS_FIXED (HEADGOAL o tac) fun FOCUS_PREMS' tac = Subgoal.FOCUS_PREMS (HEADGOAL o tac) fun FOCUS' tac = Subgoal.FOCUS (HEADGOAL o tac) fun SUBPROOF' tac = Subgoal.SUBPROOF (HEADGOAL o tac) fun FOCUS_PARAMS_CTXT tac = Subgoal.FOCUS_PARAMS (#context #> tac) fun FOCUS_PARAMS_CTXT' tac = FOCUS_PARAMS' (#context #> tac) fun CSUBGOAL_DATA f tac = CSUBGOAL (uncurry tac o apfst f) fun CSUBGOAL_STRIPPED f = CSUBGOAL_DATA (f o Term_Util.strip_csubgoal) fun SUBGOAL_DATA f tac = SUBGOAL (uncurry tac o apfst f) fun SUBGOAL_STRIPPED f = SUBGOAL_DATA (f o Term_Util.strip_subgoal) (* tactics *) fun insert_tac thms ctxt = Method.insert_tac ctxt thms fun thin_tac n = if n < 1 then K no_tac else rotate_tac (n - 1) THEN' (DETERM o eresolve0_tac [thin_rl]) THEN' rotate_tac (~n + 1) val thin_tacs = sort int_ord #> map_index ((op -) o swap) #> EVERY_ARG' thin_tac fun set_kernel_ho_unif_bounds trace_bound search_bound = Config.put Unify.unify_trace_bound trace_bound #> Config.put Unify.search_bound search_bound val silence_kernel_ho_bounds_exceeded = Context_Position.set_visible false (*some "safe" solvers create instantiations via flex-flex pairs, which we disallow here*) fun safe_simp_tac ctxt i state = let val eq_flexps = Thm.tpairs_of #> pair (Thm.tpairs_of state) #> eq_list (eq_pair (op aconv) (op aconv)) (*stop higher-order unifier from entering difficult unification problems*) val ctxt = set_kernel_ho_unif_bounds 1 1 ctxt |> silence_kernel_ho_bounds_exceeded in Simplifier.safe_simp_tac ctxt i state |> Seq.filter eq_flexps end fun nth_eq_assume_tac n = if n < 1 then K no_tac else rotate_tac (n - 1) THEN' eq_assume_tac (*resolution*) fun no_lift_biresolve_tac eres brule nsubgoals ctxt = Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = true} (eres, brule, nsubgoals) #> PRIMSEQ val no_lift_resolve_tac = no_lift_biresolve_tac false val no_lift_eresolve_tac = no_lift_biresolve_tac true fun rewrite_subgoal_tac eq_thm = let val (lhs, rhs) = Thm.dest_equals (Thm.cconcl_of eq_thm) val eq_elim = Thm.instantiate' [] [SOME lhs, SOME rhs] Drule.equal_elim_rule1 val thm = Thm.implies_elim eq_elim eq_thm in no_lift_resolve_tac thm 1 end fun eq_subgoal_from_eq_concl cbinders cprems eq_thm ctxt = let fun all_abstract ((x, T), cfree) = let val all_const = Logic.all_const T |> Thm.cterm_of ctxt in Thm_Util.abstract_rule ctxt x cfree #> Option.map (Drule.arg_cong_rule all_const) end in (*introduce the premises*) SOME (fold_rev (Drule.imp_cong_rule o Thm.reflexive) cprems eq_thm) (*introduce abstractions for the fresh Frees*) |> fold (Option.mapPartial o all_abstract) cbinders end fun rewrite_subgoal_unif_concl norm_binders norm_state norm_eq_thm binders (env_eq_thm as (env, _)) ctxt i = let (*updates binders with their normalised type*) fun updated_binders env = Binders.norm_binders (norm_binders env) binders val binders = updated_binders env fun smash_tpairs_if_occurs (_, bvar) = Seq.maps (Unification_Util.smash_tpairs_if_occurs ctxt bvar) - fun rewrite_tac env eq_thm (rparams, (prems, _)) = + fun rewrite_tac env eq_thm prems = let val binders = updated_binders env val cterm_of = Thm.cterm_of ctxt val cprems = map (cterm_of o Binders.replace_binders binders) prems in eq_subgoal_from_eq_concl (map (apsnd cterm_of) binders) cprems (norm_eq_thm ctxt env eq_thm) ctxt |> Option.map (fn eq_thm => rewrite_subgoal_tac eq_thm ctxt) |> the_default (K no_tac) end in (*smash tpairs in case some binder occurs*) fold smash_tpairs_if_occurs binders (Seq.single env_eq_thm) |> Seq.lifts (fn (env, eq_thm) => PRIMITIVE (norm_state ctxt env) - THEN SUBGOAL_STRIPPED I (rewrite_tac env eq_thm) i) + THEN SUBGOAL_STRIPPED (fst o snd) (rewrite_tac env eq_thm) i) end fun protect_tac n ctxt = let fun protect cbinders (cprems, cconcl) i = let val nprems = length cprems in if nprems < n then (@{log Logger.WARN} ctxt (fn _ => Pretty.block [ Pretty.str "Not enough premises to protect. Given number: ", Pretty.str (string_of_int n), Pretty.str ". But there are only ", Pretty.str (string_of_int nprems), Pretty.str " premise(s) in subgoal ", Pretty.str (string_of_int i), Pretty.str ": ", Logic.close_prop (map (apfst fst o apsnd Thm.term_of) cbinders) (map Thm.term_of cprems) (Thm.term_of cconcl) |> Syntax.pretty_term ctxt ] |> Pretty.string_of); no_tac) else let val (cprems_unprotected, cconcl_protected) = chop n cprems ||> Drule.list_implies o rpair cconcl in @{thm Pure.prop_def} |> Thm.instantiate' [] [SOME cconcl_protected] |> (fn eq_thm => eq_subgoal_from_eq_concl cbinders cprems_unprotected eq_thm ctxt) |> Option.map (fn protected_eq_thm => rewrite_subgoal_tac protected_eq_thm ctxt i) |> the_default no_tac end end in if n < 0 then K no_tac else CSUBGOAL_STRIPPED I (uncurry protect) end fun unprotect_tac ctxt = match_tac ctxt [Drule.protectI] fun protected_tac n tac ctxt = protect_tac n ctxt THEN' tac THEN_ALL_NEW (unprotect_tac ctxt) fun focus_prems_tac is tac ctxt = CONVERSION (Conversion_Util.move_prems_to_front_conv is) THEN' protect_tac (length is) ctxt THEN' FOCUS_PREMS' (fn {context=ctxt, params=params, prems=prems, asms=asms, concl=concl, schematics=schematics} => let val concl = Term_Util.unprotect concl in unprotect_tac ctxt THEN' tac {context=ctxt, params=params, prems=prems, asms=asms, concl=concl, schematics=schematics} end ) ctxt fun focus_delete_prems_tac is tac ctxt = focus_prems_tac is tac ctxt THEN' thin_tacs (1 upto length is) fun apply_tac tac i = Goal.init #> tac i #> Seq.map Goal.conclude end diff --git a/thys/ML_Unification/ML_Utils/Theorems/thm_util.ML b/thys/ML_Unification/ML_Utils/Theorems/thm_util.ML --- a/thys/ML_Unification/ML_Utils/Theorems/thm_util.ML +++ b/thys/ML_Unification/ML_Utils/Theorems/thm_util.ML @@ -1,47 +1,47 @@ (* Title: ML_Utils/thm_util.ML Author: Kevin Kappelmann Theorem utilities. *) signature THM_UTIL = sig include HAS_LOGGER val pretty_THM : Proof.context -> string * int * thm list -> Pretty.T (*fails if the equality theorem contains a tpair or implicit Assumption mentioning the bound variables*) val abstract_rule : Proof.context -> string -> cterm -> thm -> thm option (*"match_implies_elim prem thm" matches the first premise of thm against prem and then removes the premise; without lifting*) val match_implies_elim : thm -> thm -> thm val protect : thm -> thm end structure Thm_Util : THM_UTIL = struct -val logger = Logger.setup_new_logger Logger.root_logger "Thm_Util" +val logger = Logger.setup_new_logger Logger.root "Thm_Util" fun pretty_THM ctxt (msg, subgoal, thms) = Pretty.block [ Pretty.str msg, Pretty.fbrk, Pretty.str "Subgoal number ", Pretty.str (string_of_int subgoal), Pretty.fbrk, Pretty.str "Theorems ", Pretty.list "[" "]" (map (Thm.pretty_thm ctxt) thms) ] fun abstract_rule ctxt n ct thm = SOME (Thm.abstract_rule n ct thm) handle THM data => (@{log Logger.DEBUG} ctxt (fn _ => pretty_THM ctxt data |> Pretty.string_of); NONE) fun match_implies_elim prem thm = Thm.instantiate (Thm.first_order_match (Thm.cprem_of thm 1, Thm.cprop_of prem)) thm |> (fn thm => Thm.implies_elim thm prem) fun protect thm = Drule.protectI |> Thm.instantiate (TVars.empty, Vars.make [((("A", 0), propT), Thm.cprop_of thm)]) |> Thm.elim_implies thm end diff --git a/thys/ML_Unification/Normalisations/envir_normalisation.ML b/thys/ML_Unification/Normalisations/envir_normalisation.ML --- a/thys/ML_Unification/Normalisations/envir_normalisation.ML +++ b/thys/ML_Unification/Normalisations/envir_normalisation.ML @@ -1,244 +1,244 @@ (* Title: ML_Unification/envir_normalisation.ML Author: Kevin Kappelmann Normalisations with respect to an environment. Adapted and generalised from envir.ML *) signature ENVIR_NORMALISATION = sig (* types *) type type_normaliser = Type.tyenv -> typ -> typ val norm_type_match : type_normaliser val norm_type_unif : type_normaliser (* terms *) type term_type_normaliser = Type.tyenv -> Term_Normalisation.term_normaliser val norm_term_types : type_normaliser -> term_type_normaliser type term_normaliser = Envir.env -> Term_Normalisation.term_normaliser (** matching **) val norm_term_types_match : term_type_normaliser (*without beta-normalisation*) val norm_term_match : term_normaliser val eta_short_norm_term_match : term_normaliser (*with beta-normalisation*) val beta_norm_term_match : term_normaliser val beta_eta_short_norm_term_match : term_normaliser (** unification **) val norm_term_types_unif : term_type_normaliser (*without beta-normalisation*) val norm_term_unif : term_normaliser val eta_short_norm_term_unif : term_normaliser (*with beta-normalisation*) val beta_norm_term_unif : term_normaliser val beta_eta_short_norm_term_unif : term_normaliser (* theorems *) type thm_normaliser = Proof.context -> Envir.env -> thm -> thm type thm_type_normaliser = thm_normaliser - val norm_thm_types : type_normaliser -> thm_normaliser - val norm_thm_types_match : thm_normaliser - val norm_thm_types_unif : thm_normaliser + val norm_thm_types : type_normaliser -> thm_type_normaliser + val norm_thm_types_match : thm_type_normaliser + val norm_thm_types_unif : thm_type_normaliser val norm_thm : type_normaliser -> term_normaliser -> thm_normaliser val norm_thm_match : thm_normaliser val norm_thm_unif : thm_normaliser val eta_short_norm_thm : type_normaliser -> term_normaliser -> thm_normaliser val eta_short_norm_thm_match : thm_normaliser val eta_short_norm_thm_unif : thm_normaliser (*with beta-normalisation and eta-contraction*) val beta_eta_short_norm_thm : type_normaliser -> term_normaliser -> thm_normaliser val beta_eta_short_norm_thm_match : thm_normaliser val beta_eta_short_norm_thm_unif : thm_normaliser end structure Envir_Normalisation : ENVIR_NORMALISATION = struct (* types *) type type_normaliser = Type.tyenv -> typ -> typ val norm_type_match = Envir.subst_type val norm_type_unif = Envir.norm_type (* terms *) type term_type_normaliser = Type.tyenv -> Term_Normalisation.term_normaliser fun norm_term_types norm_type = map_types o norm_type type term_normaliser = Envir.env -> Term_Normalisation.term_normaliser (** matching **) val norm_term_types_match = norm_term_types norm_type_match fun norm_abs_same2 normT norm (a, T, body) = Abs (a, normT T, Same.commit norm body) handle Same.SAME => Abs (a, T, norm body) fun norm_abs_comb_same beta_norm norm (abs_args as (_, _, body)) arg = if beta_norm then Same.commit norm (subst_bound (arg, body)) else let val f = Abs abs_args in (norm f $ Same.commit norm arg handle Same.SAME => f $ norm arg) end fun norm_comb_same beta_norm norm f t = (case norm f of (nf as Abs (_, _, body)) => if beta_norm then Same.commit norm (subst_bound (t, body)) else nf $ Same.commit norm t | nf => nf $ Same.commit norm t) handle Same.SAME => f $ norm t fun norm_term_match1 beta_norm tenv : term Same.operation = let fun norm (Var v) = (case Envir.lookup1 tenv v of SOME u => u | NONE => raise Same.SAME) | norm (Abs (a, T, body)) = Abs (a, T, norm body) | norm (Abs abs_args $ t) = norm_abs_comb_same beta_norm norm abs_args t | norm (f $ t) = norm_comb_same beta_norm norm f t | norm _ = raise Same.SAME in norm end fun norm_term_match2 beta_norm (Envir.Envir {tenv, tyenv, ...}) : term Same.operation = let val normT = Envir.subst_type_same tyenv fun norm (Const (a, T)) = Const (a, normT T) | norm (Free (a, T)) = Free (a, normT T) | norm (v as Var (xi, T)) = let fun lookup v = (case Envir.lookup1 tenv (dest_Var v) of SOME u => u | NONE => raise Same.SAME) in (normT T |> (fn T => Same.commit lookup (Var (xi, T)))) handle Same.SAME => lookup v end | norm (Abs args) = norm_abs_same2 normT norm args | norm (Abs abs_args $ t) = norm_abs_comb_same beta_norm norm abs_args t | norm (f $ t) = norm_comb_same beta_norm norm f t | norm _ = raise Same.SAME in norm end fun norm_term_match_same beta_norm (envir as Envir.Envir {tenv, tyenv, ...}) = if Vartab.is_empty tyenv then norm_term_match1 beta_norm tenv else norm_term_match2 beta_norm envir fun norm_term_match env = Same.commit (norm_term_match_same false env) fun beta_norm_term_match env = Same.commit (norm_term_match_same true env) fun eta_short_norm_term_match env = norm_term_match env #> Term_Normalisation.eta_short fun beta_eta_short_norm_term_match env = beta_norm_term_match env #> Term_Normalisation.eta_short (** unification **) val norm_term_types_unif = norm_term_types norm_type_unif fun norm_type_unif_same tyenv : typ Same.operation = let fun norm (Type (a, Ts)) = Type (a, Same.map norm Ts) | norm (TFree _) = raise Same.SAME | norm (TVar v) = (case Type.lookup tyenv v of SOME U => Same.commit norm U | NONE => raise Same.SAME) in norm end fun norm_term_unif1 beta_norm tenv : term Same.operation = let fun norm (Var v) = (case Envir.lookup1 tenv v of SOME u => Same.commit norm u | NONE => raise Same.SAME) | norm (Abs (a, T, body)) = Abs (a, T, norm body) | norm (Abs abs_args $ t) = norm_abs_comb_same beta_norm norm abs_args t | norm (f $ t) = norm_comb_same beta_norm norm f t | norm _ = raise Same.SAME in norm end fun norm_term_unif2 beta_norm (envir as Envir.Envir {tyenv, ...}) : term Same.operation = let val normT = norm_type_unif_same tyenv fun norm (Const (a, T)) = Const (a, normT T) | norm (Free (a, T)) = Free (a, normT T) | norm (Var (xi, T)) = (case Envir.lookup envir (xi, T) of SOME u => Same.commit norm u | NONE => Var (xi, normT T)) | norm (Abs args) = norm_abs_same2 normT norm args | norm (Abs abs_args $ t) = norm_abs_comb_same beta_norm norm abs_args t | norm (f $ t) = norm_comb_same beta_norm norm f t | norm _ = raise Same.SAME in norm end fun norm_term_unif_same beta_norm (envir as Envir.Envir {tenv, tyenv, ...}) = if Vartab.is_empty tyenv then norm_term_unif1 beta_norm tenv else norm_term_unif2 beta_norm envir fun norm_term_unif env = Same.commit (norm_term_unif_same false env) fun beta_norm_term_unif env = Same.commit (norm_term_unif_same true env) fun eta_short_norm_term_unif env = norm_term_unif env #> Term_Normalisation.eta_short fun beta_eta_short_norm_term_unif env = beta_norm_term_unif env #> Term_Normalisation.eta_short type thm_normaliser = Proof.context -> Envir.env -> thm -> thm type thm_type_normaliser = thm_normaliser (** theorems **) (*collect and normalise TVars of a term*) fun collect_norm_types norm_type ctxt tyenv t = let val norm_type' = norm_type tyenv fun prep_type_entry x = (x, Thm.ctyp_of ctxt (norm_type' (TVar x))) in fold_types (fold_atyps (fn TVar v => TVars.add (prep_type_entry v) | _ => I)) t |> TVars.build end (*collect and normalise Vars of a term*) fun collect_norm_terms norm_type norm_term ctxt (env as Envir.Envir {tyenv,...}) t = let val norm_type' = norm_type tyenv val norm_term' = norm_term env fun prep_term_entry (x as (n, T)) = ((n, norm_type' T), Thm.cterm_of ctxt (norm_term' (Var x))) in fold_aterms (fn Var v => Vars.add (prep_term_entry v) | _ => I) t |> Vars.build end (*normalise types of a theorem*) fun norm_thm_types norm_types ctxt (Envir.Envir {tyenv, ...}) thm = let val prop = Thm.full_prop_of thm val type_inst = collect_norm_types norm_types ctxt tyenv prop val inst = (type_inst, Vars.empty) in Thm.instantiate inst thm end val norm_thm_types_match = norm_thm_types norm_type_match val norm_thm_types_unif = norm_thm_types norm_type_unif (*normalise a theorem*) fun norm_thm norm_types norm_terms ctxt (env as Envir.Envir {tyenv,...}) thm = let val prop = Thm.full_prop_of thm val type_inst = collect_norm_types norm_types ctxt tyenv prop val term_inst = collect_norm_terms norm_types norm_terms ctxt env prop val inst = (type_inst, term_inst) in Thm.instantiate inst thm end val norm_thm_match = norm_thm norm_type_match norm_term_match val norm_thm_unif = norm_thm norm_type_unif norm_term_unif fun eta_short_norm_thm norm_types norm_terms ctxt env = norm_thm norm_types norm_terms ctxt env #> Conv.fconv_rule Thm.eta_conversion val eta_short_norm_thm_match = eta_short_norm_thm norm_type_match norm_term_match val eta_short_norm_thm_unif = eta_short_norm_thm norm_type_unif norm_term_unif fun beta_eta_short_norm_thm norm_types norm_terms ctxt env = norm_thm norm_types norm_terms ctxt env #> Conv.fconv_rule Drule.beta_eta_conversion val beta_eta_short_norm_thm_match = beta_eta_short_norm_thm norm_type_match norm_term_match val beta_eta_short_norm_thm_unif = beta_eta_short_norm_thm norm_type_unif norm_term_unif end diff --git a/thys/ML_Unification/Simps_To/Simps_To.thy b/thys/ML_Unification/Simps_To/Simps_To.thy --- a/thys/ML_Unification/Simps_To/Simps_To.thy +++ b/thys/ML_Unification/Simps_To/Simps_To.thy @@ -1,72 +1,70 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Simps To\ theory Simps_To imports ML_Tactic_Utils - ML_Theorem_Utils - ML_Unification_Base Setup_Result_Commands begin paragraph \Summary\ text \Simple frameworks to ask for the simp-normal form of a term on the user-level.\ -setup_result simps_to_base_logger = \Logger.new_logger Logger.root_logger "Simps_To_Base"\ +setup_result simps_to_base_logger = \Logger.new_logger Logger.root "Simps_To_Base"\ paragraph \Using Simplification On Left Term\ definition "SIMPS_TO s t \ (s \ t)" lemma SIMPS_TO_eq: "SIMPS_TO s t \ (s \ t)" unfolding SIMPS_TO_def by simp text \Prevent simplification of second/right argument\ lemma SIMPS_TO_cong [cong]: "s \ s' \ SIMPS_TO s t \ SIMPS_TO s' t" by simp lemma SIMPS_TOI: "PROP SIMPS_TO s s" unfolding SIMPS_TO_eq by simp lemma SIMPS_TOD: "PROP SIMPS_TO s t \ s \ t" unfolding SIMPS_TO_eq by simp ML_file\simps_to.ML\ paragraph \Using Simplification On Left Term Followed By Unification\ definition "SIMPS_TO_UNIF s t \ (s \ t)" text \Prevent simplification\ lemma SIMPS_TO_UNIF_cong [cong]: "SIMPS_TO_UNIF s t \ SIMPS_TO_UNIF s t" by simp lemma SIMPS_TO_UNIF_eq: "SIMPS_TO_UNIF s t \ (s \ t)" unfolding SIMPS_TO_UNIF_def by simp lemma SIMPS_TO_UNIFI: "PROP SIMPS_TO s s' \ s' \ t \ PROP SIMPS_TO_UNIF s t" unfolding SIMPS_TO_UNIF_eq SIMPS_TO_eq by simp lemma SIMPS_TO_UNIFD: "PROP SIMPS_TO_UNIF s t \ s \ t" unfolding SIMPS_TO_UNIF_eq by simp ML_file\simps_to_unif.ML\ paragraph \Examples\ experiment begin lemma assumes [simp]: "P \ Q" and [simp]: "Q \ R" shows "PROP SIMPS_TO P Q" apply simp \\Note: only the left-hand side is simplified.\ ML_command\ \\obtaining the normal form theorem for a term in ML\ Simps_To.SIMPS_TO_thm_resultsq (simp_tac @{context}) @{context} @{cterm P} |> Seq.list_of |> map @{print} \ oops schematic_goal assumes [simp]: "P \ Q" and [simp]: "Q \ R" shows "PROP SIMPS_TO P ?Q" - apply simp - by (rule SIMPS_TOI) + by (tactic \Simps_To.SIMPS_TO_tac (Simps_To.simp_inst_tac (simp_tac @{context})) @{context} 1\) + end end diff --git a/thys/ML_Unification/Simps_To/simps_to.ML b/thys/ML_Unification/Simps_To/simps_to.ML --- a/thys/ML_Unification/Simps_To/simps_to.ML +++ b/thys/ML_Unification/Simps_To/simps_to.ML @@ -1,63 +1,72 @@ (* Title: ML_Unification/simps_to.ML Author: Kevin Kappelmann Create SIMPS_TO theorems. *) signature SIMPS_TO = sig include HAS_LOGGER val dest_SIMPS_TO : term -> (term * term) val cdest_SIMPS_TO : cterm -> (cterm * cterm) - val mk_SIMPS_TO_cprop : cterm -> cterm -> cterm - val mk_SIMPS_TO_var_cprop : Proof.context -> cterm -> cterm + val mk_SIMPS_TO_cprop : cterm * cterm -> cterm val finish_SIMPS_TO_tac : Proof.context -> int -> tactic + val SIMPS_TO_tac : (int -> tactic) -> Proof.context -> int -> tactic - val SIMPS_TO_thmsq : (int -> tactic) -> Proof.context -> cterm -> thm Seq.seq - val SIMPS_TO_thm_resultsq : (int -> tactic) -> Proof.context -> cterm -> - (thm * cterm) Seq.seq + val SIMPS_TO_thmsq : (int -> tactic) -> Proof.context -> cterm * cterm -> thm Seq.seq + + val simp_inst_tac : (int -> tactic) -> int -> tactic + + val SIMPS_TO_thm_resultsq : (int -> tactic) -> Proof.context -> cterm -> (thm * cterm) Seq.seq end structure Simps_To : SIMPS_TO = struct val logger = Logger.setup_new_logger simps_to_base_logger "Simps_To" structure Util = Tactic_Util val dest_SIMPS_TO = \<^Const_fn>\SIMPS_TO _ for lhs rhs => \(lhs, rhs)\\ val cdest_SIMPS_TO = Thm.dest_comb #>> Thm.dest_arg -fun mk_SIMPS_TO_cprop clhs crhs = \<^instantiate>\'a = \Thm.ctyp_of_cterm clhs\ and clhs and crhs +fun mk_SIMPS_TO_cprop (clhs, crhs) = \<^instantiate>\'a = \Thm.ctyp_of_cterm clhs\ and clhs and crhs in cprop\PROP (SIMPS_TO clhs crhs)\ for clhs :: 'a\ -fun mk_SIMPS_TO_var_cprop ctxt ct = - Var (("x", Thm.maxidx_of_cterm ct + 1), Thm.typ_of_cterm ct) - |> Thm.cterm_of ctxt - |> mk_SIMPS_TO_cprop ct - fun finish_SIMPS_TO_tac ctxt = match_tac ctxt [@{thm SIMPS_TOI}] +fun SIMPS_TO_tac simp_tac ctxt = simp_tac THEN' finish_SIMPS_TO_tac ctxt -fun SIMPS_TO_thmsq simps_to_tac ctxt ct = - let fun inst_tac cconcl _ = - PRIMITIVE (cdest_SIMPS_TO cconcl |> swap |> Thm.match |> Thm.instantiate) +fun SIMPS_TO_thmsq simp_tac ctxt (ctp as (clhs, crhs)) = + (let val goal = mk_SIMPS_TO_cprop ctp in - (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + (@{log Logger.TRACE} ctxt (fn _ => Pretty.block [ Pretty.str "Creating ", Syntax.pretty_term ctxt @{term SIMPS_TO}, Pretty.str " theorems for ", - Syntax.pretty_term ctxt (Thm.term_of ct) + Unification_Util.pretty_terms ctxt (map Thm.term_of [clhs, crhs]) ] |> Pretty.string_of); - mk_SIMPS_TO_var_cprop ctxt ct - |> Util.HEADGOAL (Util.apply_tac ( - simps_to_tac - THEN' Tactic_Util.CSUBGOAL_STRIPPED (snd o snd) inst_tac - THEN' finish_SIMPS_TO_tac ctxt))) + Util.HEADGOAL (Util.apply_tac (SIMPS_TO_tac simp_tac ctxt)) goal) + end) + handle TYPE _ => (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Types of terms ", + Unification_Util.pretty_terms ctxt (map Thm.term_of [clhs, crhs]), + Pretty.str " not equal" + ] |> Pretty.string_of); + Seq.empty) + +fun simp_inst_tac simp_tac = + let fun inst_tac cconcl = + PRIMITIVE (cdest_SIMPS_TO cconcl |> swap |> Thm.first_order_match |> Thm.instantiate) + in simp_tac THEN' Tactic_Util.CSUBGOAL_STRIPPED (snd o snd) (K o inst_tac) end + +fun SIMPS_TO_thm_resultsq simp_tac ctxt ct = + let val cvar = Term_Util.fresh_var "x" (Thm.typ_of_cterm ct) (Thm.maxidx_of_cterm ct) + |> fst |> Thm.cterm_of ctxt + in + SIMPS_TO_thmsq (simp_inst_tac simp_tac) ctxt (ct, cvar) + |> Seq.map (fn thm => (thm, Thm.cconcl_of thm |> cdest_SIMPS_TO |> snd)) end -val SIMPS_TO_thm_resultsq = - Seq.map (fn thm => (thm, Thm.cconcl_of thm |> Thm.dest_arg)) ooo SIMPS_TO_thmsq - end \ No newline at end of file diff --git a/thys/ML_Unification/Simps_To/simps_to_unif.ML b/thys/ML_Unification/Simps_To/simps_to_unif.ML --- a/thys/ML_Unification/Simps_To/simps_to_unif.ML +++ b/thys/ML_Unification/Simps_To/simps_to_unif.ML @@ -1,71 +1,67 @@ (* Title: ML_Unification/simps_to_unif.ML Author: Kevin Kappelmann Create SIMPS_TO_UNIF theorems. *) signature SIMPS_TO_UNIF = sig include HAS_LOGGER val dest_SIMPS_TO_UNIF : term -> (term * term) val cdest_SIMPS_TO_UNIF : cterm -> (cterm * cterm) - val mk_SIMPS_TO_UNIF_cprop : cterm -> cterm -> cterm + val mk_SIMPS_TO_UNIF_cprop : cterm * cterm -> cterm - val mk_SIMPS_TO_UNIF_thm : thm -> thm -> thm + val SIMPS_TO_UNIF_tac : (int -> tactic) -> (int -> thm -> 'a Seq.seq) -> Proof.context -> int -> + thm -> 'a Seq.seq val SIMPS_TO_UNIF_env_thmsq : (int -> tactic) -> Unification_Base.normalisers -> - Unification_Base.unifier -> term Binders.binders -> Proof.context -> term * term -> Envir.env -> + Unification_Base.closed_unifier -> Proof.context -> term * term -> Envir.env -> (Envir.env * thm) Seq.seq end structure Simps_To_Unif : SIMPS_TO_UNIF = struct val logger = Logger.setup_new_logger simps_to_base_logger "Simps_To_Unif" val dest_SIMPS_TO_UNIF = \<^Const_fn>\SIMPS_TO_UNIF _ for lhs rhs => \(lhs, rhs)\\ val cdest_SIMPS_TO_UNIF = Thm.dest_comb #>> Thm.dest_arg -fun mk_SIMPS_TO_UNIF_cprop clhs crhs = \<^instantiate>\'a = \Thm.ctyp_of_cterm clhs\ and clhs and crhs - in cprop\PROP (SIMPS_TO_UNIF clhs crhs)\ for clhs :: 'a\ - -fun mk_SIMPS_TO_UNIF_thm SIMPS_TO_thm eq_thm = - Drule.incr_indexes2 SIMPS_TO_thm eq_thm @{thm SIMPS_TO_UNIFI} - |> Thm_Util.match_implies_elim SIMPS_TO_thm - |> Thm_Util.match_implies_elim eq_thm +fun mk_SIMPS_TO_UNIF_cprop (clhs, crhs) = \<^instantiate>\'a = \Thm.ctyp_of_cterm clhs\ + and clhs and crhs in cprop\PROP (SIMPS_TO_UNIF clhs crhs)\ for clhs :: 'a\ -fun SIMPS_TO_UNIF_env_thmsq simps_to_tac norms unif binders ctxt (lhs, rhs) env = - let - val SIMPS_TO_res = Binders.replace_binders binders lhs - |> Thm.cterm_of ctxt - |> Simps_To.SIMPS_TO_thm_resultsq simps_to_tac ctxt - fun unif_res (SIMPS_TO_thm, clhs_norm) = - let val tp = (Binders.replace_frees binders (Thm.term_of clhs_norm), rhs) - in - (@{log Logger.TRACE} ctxt (fn _ => Pretty.block [ - Pretty.str "Result: ", - Thm.pretty_thm ctxt SIMPS_TO_thm, - Pretty.str " Now unifying ", - Unification_Util.pretty_unif_problem ctxt tp - ] |> Pretty.string_of); - unif binders ctxt tp env - |> Seq.map (pair SIMPS_TO_thm)) - end - fun SIMPS_TO_UNIF_res (SIMPS_TO_thm, (env, eq_thm)) = - let val (SIMPS_TO_thm, eq_thm) = - (#norm_thm norms ctxt env SIMPS_TO_thm, #norm_unif_thm norms ctxt env eq_thm) - in (env, mk_SIMPS_TO_UNIF_thm SIMPS_TO_thm eq_thm) end +fun SIMPS_TO_UNIF_tac simp_tac eq_tac ctxt = + Tactic_Util.THEN' ( + match_tac ctxt [@{thm SIMPS_TO_UNIFI}] + THEN' Simps_To.SIMPS_TO_tac (Simps_To.simp_inst_tac simp_tac) ctxt, + eq_tac) + +fun SIMPS_TO_UNIF_env_thmsq simp_tac norms unif ctxt (tp as (lhs, rhs)) env = + (let + val goal = apply2 (Thm.cterm_of ctxt) tp |> mk_SIMPS_TO_UNIF_cprop |> Goal.init + fun eq_tac i state = + let + val tp = Thm.cprem_of state i |> Thm.dest_equals |> apply2 Thm.term_of + fun norm_resolve (env, eq_thm) = #norm_unif_thm norms ctxt env eq_thm + |> Thm.implies_elim (#norm_thm norms ctxt env state) + |> pair env + in unif ctxt tp env |> Seq.map norm_resolve end in - (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + (@{log Logger.TRACE} ctxt (fn _ => Pretty.block [ Pretty.str "Creating ", Syntax.pretty_term ctxt @{term SIMPS_TO_UNIF}, Pretty.str " theorems for ", Unification_Util.pretty_terms ctxt [lhs, rhs] ] |> Pretty.string_of); - SIMPS_TO_res - |> Seq.maps unif_res - |> Seq.map SIMPS_TO_UNIF_res) - end + Tactic_Util.HEADGOAL (SIMPS_TO_UNIF_tac simp_tac eq_tac ctxt) goal + |> Seq.map (apsnd Goal.conclude)) + end) + handle TYPE _ => (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Types of terms ", + Unification_Util.pretty_terms ctxt [lhs, rhs], + Pretty.str " not equal" + ] |> Pretty.string_of); + Seq.empty) end diff --git a/thys/ML_Unification/Term_Index/discrimination_tree.ML b/thys/ML_Unification/Term_Index/discrimination_tree.ML --- a/thys/ML_Unification/Term_Index/discrimination_tree.ML +++ b/thys/ML_Unification/Term_Index/discrimination_tree.ML @@ -1,179 +1,181 @@ (* Title: Term_Index/discrimination_tree.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory, Sebastian Willenbrink, Kevin Kappelmann, TU Munich Discrimination trees based on Pure/net.ML. Requires operands to be beta-eta-short normalised. *) structure Discrimination_Tree : TERM_INDEX = struct structure TIB = Term_Index_Base val norm_term = Term_Normalisation.beta_eta_short datatype key = CombK | VarK | AtomK of string (*Keys are preorder lists of symbols -- Combinations, Vars, Atoms. Any term whose head is a Var is regarded entirely as a Var. Abstractions are also regarded as Vars; this covers eta-conversion and "near" eta-conversions such as %x. ?P (?f x).*) fun add_key_of_terms (t, ks) = let fun rands (t, ks) = case t of (*other cases impossible due to beta-eta-short norm*) t1 $ t2 => CombK :: rands (t1, add_key_of_terms (t2, ks)) | Const (n, _) => AtomK n :: ks | Free (n, _) => AtomK n :: ks | Bound i => AtomK (Name.bound i) :: ks in case head_of t of Var _ => VarK :: ks | Abs _ => VarK :: ks | _ => rands (t, ks) end (*convert a term to a list of keys*) fun key_of_term t = add_key_of_terms (t, []) (*Trees indexed by key lists: each arc is labelled by a key. Each node contains a list of values, and arcs to children. The empty key addresses the entire tree. Lookup functions preserve order of values stored at same level.*) datatype 'a term_index = Leaf of 'a list | Tree of {comb: 'a term_index, var: 'a term_index, atoms: 'a term_index Symtab.table} val empty = Leaf [] fun is_empty (Leaf []) = true | is_empty _ = false val empty_tree = Tree {comb=empty, var=empty, atoms=Symtab.empty} (* insert *) (*Adds value x to the list at the node addressed by the keys. Creates node if not already present. is_dup decides whether a value is a duplicate. The empty list of keys generates a Leaf node, others a Tree node.*) fun insert_keys is_dup x ksdtp = case ksdtp of ([], Leaf xs) => if exists is_dup xs then raise TIB.INSERT else Leaf (x :: xs) | (keys, Leaf []) => insert_keys is_dup x (keys, empty_tree) (*expand the tree*) | (CombK :: keys, Tree {comb, var, atoms}) => Tree {comb=insert_keys is_dup x (keys, comb), var=var, atoms=atoms} | (VarK :: keys, Tree {comb, var, atoms}) => Tree {comb=comb, var=insert_keys is_dup x (keys, var), atoms=atoms} | (AtomK a :: keys, Tree {comb, var, atoms}) => let val atoms' = Symtab.map_default (a, empty) (fn dt' => insert_keys is_dup x (keys, dt')) atoms in Tree {comb=comb, var=var, atoms=atoms'} end fun insert is_dup (t, x) dt = insert_keys is_dup x (key_of_term t, dt) fun insert_safe is_dup entry dt = insert is_dup entry dt handle TIB.INSERT => dt (* delete *) (*Create a new Tree node if it would be nonempty*) fun new_tree (args as {comb, var, atoms}) = if is_empty comb andalso is_empty var andalso Symtab.is_empty atoms then empty else Tree args (*Deletes values from the list at the node addressed by the keys. Raises DELETE if absent. Collapses the tree if possible. sel is the selector for values to be deleted.*) fun delete_keys sel ksdtp = case ksdtp of ([], Leaf xs) => if exists sel xs then Leaf (filter_out sel xs) else raise TIB.DELETE | (_, Leaf []) => raise TIB.DELETE | (CombK :: keys, Tree {comb, var, atoms}) => new_tree {comb=delete_keys sel (keys,comb), var=var, atoms=atoms} | (VarK :: keys, Tree {comb, var, atoms}) => new_tree {comb=comb, var=delete_keys sel (keys,var), atoms=atoms} | (AtomK a :: keys, Tree {comb, var, atoms}) => let val atoms' = (case Symtab.lookup atoms a of NONE => raise TIB.DELETE | SOME dt' => (case delete_keys sel (keys, dt') of Leaf [] => Symtab.delete a atoms | dt'' => Symtab.update (a, dt'') atoms)) in new_tree {comb=comb, var=var, atoms=atoms'} end fun delete eq t dt = delete_keys eq (key_of_term t, dt) fun delete_safe eq t dt = delete eq t dt handle TIB.DELETE => dt (* retrievals *) type 'a retrieval = 'a term_index -> term -> 'a list fun variants_keys dt keys = case (dt, keys) of (Leaf xs, []) => xs | (Leaf _, (_ :: _)) => [] (*non-empty keys and empty dt*) | (Tree {comb, ...}, (CombK :: keys)) => variants_keys comb keys | (Tree {var, ...}, (VarK :: keys)) => variants_keys var keys | (Tree {atoms, ...}, (AtomK a :: keys)) => case Symtab.lookup atoms a of SOME dt => variants_keys dt keys | NONE => [] fun variants dt t = variants_keys dt (key_of_term t) (*Skipping a term in a tree. Recursively skip 2 levels if a combination*) fun tree_skip (Leaf _) dts = dts | tree_skip (Tree {comb, var, atoms}) dts = var :: dts |> Symtab.fold (fn (_, dt) => fn dts => dt :: dts) atoms (* tree_skip comb [] only skips first term, so another skip is required *) |> fold_rev tree_skip (tree_skip comb []) (*conses the linked tree, if present, to dts*) fun look1 (atoms, a) dts = case Symtab.lookup atoms a of NONE => dts | SOME dt => dt :: dts datatype query = Instances | Generalisations | Unifiables (*Return the nodes accessible from the term (cons them before trees) Var in tree matches any term. Abs or Var in object: for unification, regarded as a wildcard; otherwise only matches a variable.*) fun query q t dt dts = let fun rands _ (Leaf _) dts = dts | rands t (Tree {comb, atoms, ...}) dts = case t of f $ t => fold_rev (query q t) (rands f comb []) dts | Const (n, _) => look1 (atoms, n) dts | Free (n, _) => look1 (atoms, n) dts | Bound i => look1 (atoms, Name.bound i) dts in case dt of Leaf _ => dts | Tree {var, ...} => case (head_of t, q) of (Var _, Generalisations) => var :: dts (*only matches Var in dt*) | (Var _, _) => tree_skip dt dts | (Abs _, Generalisations) => var :: dts (*only a Var can match*) (*A var instantiation in the abstraction could allow an eta-reduction, so regard the abstraction as a wildcard.*) | (Abs _, _) => tree_skip dt dts | (_, Instances) => rands t dt dts | (_, _) => rands t dt (var :: dts) (*var could also match*) end fun extract_leaves ls = maps (fn Leaf xs => xs) ls fun generalisations dt t = query Generalisations t dt [] |> extract_leaves fun unifiables dt t = query Unifiables t dt [] |> extract_leaves fun instances dt t = query Instances t dt [] |> extract_leaves (* merge *) fun cons_fst x (xs, y) = (x :: xs, y) fun dest (Leaf xs) = map (pair []) xs | dest (Tree {comb, var, atoms}) = flat [ map (cons_fst CombK) (dest comb), map (cons_fst VarK) (dest var), maps (fn (a, net) => map (cons_fst (AtomK a)) (dest net)) (Symtab.dest atoms) ] fun content dt = map #2 (dest dt) fun merge eq dt1 dt2 = if pointer_eq (dt1, dt2) then dt1 else if is_empty dt1 then dt2 - else fold (fn (ks, v) => fn dt => insert_keys (curry eq v) v (ks, dt)) (dest dt2) dt1 + else fold + (fn (ks, v) => fn dt => insert_keys (curry eq v) v (ks, dt) handle TIB.INSERT => dt) + (dest dt2) dt1 end diff --git a/thys/ML_Unification/Term_Index/term_index.ML b/thys/ML_Unification/Term_Index/term_index.ML --- a/thys/ML_Unification/Term_Index/term_index.ML +++ b/thys/ML_Unification/Term_Index/term_index.ML @@ -1,46 +1,47 @@ (* Title: Term_Index/term_index.ML Author: Kevin Kappelmann, Sebastian Willenbrink Signatures for term indexes. For a brief overview, see: https://en.wikipedia.org/wiki/Term_indexing *) signature TERM_INDEX_BASE = sig exception INSERT exception DELETE end structure Term_Index_Base : TERM_INDEX_BASE = struct exception INSERT exception DELETE end signature TERM_INDEX = sig type 'a term_index val empty : 'a term_index val is_empty : 'a term_index -> bool val content : 'a term_index -> 'a list (*puts a term into normal form as required by the index structure*) val norm_term : Term_Normalisation.term_normaliser (*first argument decides which values should be considered as duplicates; raises Term_Index_Base.INSERT if the term, value pair is already inserted*) val insert : ('a -> bool) -> (term * 'a) -> 'a term_index -> 'a term_index val insert_safe : ('a -> bool) -> (term * 'a) -> 'a term_index -> 'a term_index (*first argument selects which values for the given term should be removed; raises Term_Index_Base.DELETE if no value is deleted*) val delete : ('a -> bool) -> term -> 'a term_index -> 'a term_index val delete_safe : ('a -> bool) -> term -> 'a term_index -> 'a term_index type 'a retrieval = 'a term_index -> term -> 'a list val variants : 'a retrieval val generalisations : 'a retrieval val instances : 'a retrieval val unifiables : 'a retrieval - (*first argument is the equality test for values*) + (*first argument is the equality test for values; + prefers values from the first index in case of duplicates*) val merge : (('a * 'a) -> bool) -> 'a term_index -> 'a term_index -> 'a term_index -end \ No newline at end of file +end diff --git a/thys/ML_Unification/Term_Index/term_index_data.ML b/thys/ML_Unification/Term_Index/term_index_data.ML --- a/thys/ML_Unification/Term_Index/term_index_data.ML +++ b/thys/ML_Unification/Term_Index/term_index_data.ML @@ -1,32 +1,32 @@ (* Title: Term_Index/term_index_data.ML Author: Kevin Kappelmann Term index data stored in generic contexts. *) signature TERM_INDEX_DATA_ARGS = sig type data structure TI : TERM_INDEX val data_eq : data * data -> bool end functor Term_Index_Generic_Data_Args(P : TERM_INDEX_DATA_ARGS) : GENERIC_DATA_ARGS = struct type T = P.data P.TI.term_index val empty = P.TI.empty -fun merge (ti1, ti2)= P.TI.merge P.data_eq ti1 ti2 +fun merge (ti1, ti2) = P.TI.merge P.data_eq ti1 ti2 end signature TERM_INDEX_DATA = sig structure TI : TERM_INDEX structure TI_Data : GENERIC_DATA end functor Term_Index_Data(P : TERM_INDEX_DATA_ARGS) : TERM_INDEX_DATA = struct structure TI = P.TI structure TI_Data = Generic_Data(Term_Index_Generic_Data_Args(P)) end diff --git a/thys/ML_Unification/Tests/First_Order_ML_Unification_Tests.thy b/thys/ML_Unification/Tests/First_Order_ML_Unification_Tests.thy --- a/thys/ML_Unification/Tests/First_Order_ML_Unification_Tests.thy +++ b/thys/ML_Unification/Tests/First_Order_ML_Unification_Tests.thy @@ -1,238 +1,240 @@ \<^marker>\creator "Kevin Kappelmann"\ \<^marker>\contributor "Paul Bachmann"\ subsection \First-Order ML Unification Tests\ theory First_Order_ML_Unification_Tests imports ML_Unification_Tests_Base begin paragraph \Summary\ text \Tests for @{ML_structure "First_Order_Unification"}.\ ML\ structure Prop = SpecCheck_Property structure UC = Unification_Combinator open Unification_Tests_Base structure Unif = First_Order_Unification structure Norm = Envir_Normalisation val match = Unif.match [] val match_hints = let fun match binders = UC.add_fallback_matcher (Unif.e_match Unification_Util.match_types) ((fn binders => (Hints.map_retrieval (Hints.mk_retrieval Hints.TI.generalisations |> K) - #> Hints.UH.map_concl_unifier (Higher_Order_Pattern_Unification.match |> K) + #> Hints.UH.map_concl_unifier (Higher_Order_Pattern_Unification.match + |> Type_Unification.e_match Unification_Util.match_types |> K) #> Hints.UH.map_normalisers (Unification_Util.beta_eta_short_norms_match |> K) #> Hints.UH.map_prems_unifier (match |> UC.norm_matcher Norm.beta_norm_term_match |> K)) |> Context.proof_map #> Test_Unification_Hints.try_hints binders) |> UC.norm_matcher (#norm_term Unif.norms_match)) binders in match [] end val unify = Unif.unify [] val unify_hints = let fun unif binders = UC.add_fallback_unifier (Unif.e_unify Unification_Util.unify_types) ((fn binders => - (Hints.UH.map_concl_unifier (Higher_Order_Pattern_Unification.match |> K) + (Hints.UH.map_concl_unifier (Higher_Order_Pattern_Unification.match + |> Type_Unification.e_match Unification_Util.match_types |> K) #> Hints.UH.map_normalisers (Unification_Util.beta_eta_short_norms_unif |> K) #> Hints.UH.map_prems_unifier (unif |> UC.norm_unifier Norm.beta_norm_term_unif |> K)) |> Context.proof_map #> Test_Unification_Hints.try_hints binders) |> UC.norm_unifier (#norm_term Unif.norms_unify)) binders in unif [] end \ subsubsection \Matching\ paragraph \Unit Tests\ ML_command\ let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} |> Variable.declare_term @{term "f :: nat \ bool \ nat"} val tests = map (apply2 (Syntax.read_term ctxt))[ ("(?x :: nat \ ?'Z) 1", "f 1"), ("?x :: nat", "(?x + 1 :: nat)"), ("(g :: nat \ nat \ nat) ?x ?x", "(g :: nat \ nat \ nat) (?x + 1) (?x + 1)"), ("\y. (?x :: nat \ ?'Z) y", "\y. f y"), ("(g :: ?'X \ ?'Y \ ?'Z) (x :: ?'X) (y :: ?'Y)", "(g :: ?'Y \ ?'Z \ ?'Z) (x :: ?'Y) (y :: ?'Z)"), ("(g :: ?'Z \ ?'X)", "\(x :: ?'X). (g :: ?'X \ ?'Y) x"), ("\ (x :: nat). (0 :: nat)", "\ (x :: nat). (0 :: nat)"), ("\ (x :: nat). x", "\ (x :: nat). x"), ("\ (x :: nat) (y :: bool). (x, y)", "\ (x :: nat) (y :: bool). (x, y)"), ("\ (x :: ?'A) (y :: bool). (?x :: ?'A \ bool \ ?'Z) x y", "\ (x :: nat) (y :: bool). f x y"), ("\(x :: ?'X). (g :: ?'X \ ?'X) x", "(g :: ?'X \ ?'X)"), ("?g ?x ?y d", "g ?y ?x d"), ("f 0 True", "(\x y. f y x) True 0"), ("\ (x :: ?'a) y. f y", "\(x :: ?'b). f"), ("\y z. (?x :: nat \ bool \ nat) y z", "f"), ("\x. (?f :: nat \ bool \ nat) 0 x", "\x. f 0 x") ] val check = check_unit_tests_hints_match tests true [] in Lecker.test_group ctxt () [ check "match" match, check "match_hints" match_hints ] end \ subparagraph \Asymmetry\ ML_command\ let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} |> Variable.declare_term @{term "f :: nat \ bool \ nat"} val tests = map (apply2 (Syntax.read_term ctxt))[ ("f 1", "(?x :: nat \ ?'Z) 1") ] val check = check_unit_tests_hints_match tests false [] in Lecker.test_group ctxt () [ check "match" match, check "match_hints" match_hints ] end \ subparagraph \Ground Hints\ ML_command\ let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt) o Syntax.read_term ctxt) [ "?x \ 2 \ ?x + (-?x :: int) \ 0", "?x \ 0 \ ?y \ 0 \ (?x :: int) + ?y \ 0" ] val tests = map (apply2 (Syntax.read_term ctxt))[ ("(2 + -2) + (0 + 0) :: int", "0 :: int") ] val check_hints = check_unit_tests_hints_match tests in Lecker.test_group ctxt () [ check_hints false [] "match" match, check_hints false [] "match_hints without hints" match_hints, check_hints true hints "match_hints with hints" match_hints ] end \ subsubsection \Unification\ paragraph \Generated Tests\ ML_command\ structure Test_Params = struct val unify = unify val unify_hints = unify_hints val params = { nv = 4, ni = 2, max_h = 5, max_args = 4 } end structure First_Order_Tests = First_Order_Unification_Tests(Test_Params) val _ = First_Order_Tests.tests @{context} (SpecCheck_Random.new ()) \ paragraph \Unit Tests\ subparagraph \Occurs-Check\ ML_command\ let val unit_tests = [ ( Var (("x", 0), TVar (("X", 0), [])), Var (("y", 0), TVar (("X", 0), []) --> TFree ("'a", [])) $ Var (("x", 0), TVar (("X", 0), [])) ), ( Var (("y", 0), TVar (("X", 0), []) --> TFree ("'a", [])) $ Var (("x", 0), TVar (("X", 0), [])), Var (("x", 0), TVar (("X", 0), [])) ), ( Free (("f", [TVar (("X", 0), []), TVar (("X", 0), [])] ---> TFree ("'a", []))) $ Var (("x", 0), TVar (("X", 0), [])) $ Var (("y", 0), TVar (("X", 0), [])), Free (("f", [TVar (("X", 0), []), TVar (("X", 0), [])] ---> TFree ("'a", []))) $ Var (("y", 0), TVar (("X", 0), [])) $ ( Free (("g", TVar (("X", 0), []) --> TVar (("X", 0), []))) $ Var (("x", 0), TVar (("X", 0), [])) ) ), ( Free (("f", [TVar (("X", 0), []), TVar (("Y", 0), [])] ---> TFree ("'a", []))) $ Var (("x", 0), TVar (("X", 0), [])) $ Var (("y", 0), TVar (("Y", 0), [])), Free (("f", [TVar (("Y", 0), []), TVar (("X", 0), [])] ---> TFree ("'a", []))) $ Var (("y", 0), TVar (("Y", 0), [])) $ ( Free (("g", TVar (("X", 0), []) --> TVar (("X", 0), []))) $ Var (("x", 0), TVar (("X", 0), [])) ) ) ] fun check name unif ctxt _ = check_list unit_tests ("occurs-check for " ^ name) (Prop.prop (not o terms_unify_thms_correct_unif ctxt unif)) ctxt |> K () in Lecker.test_group @{context} () [ check "unify" unify, check "unify_hints" unify_hints ] end \ subparagraph \Lambda-Abstractions\ ML_command\ let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} |> Variable.declare_term @{term "f :: nat \ bool \ nat"} val tests = map (apply2 (Syntax.read_term ctxt))[ ("\ (x :: nat). (0 :: nat)", "\ (x :: nat). (0 :: nat)"), ("\ (x :: nat). x", "\ (x :: nat). x"), ("\ (x :: nat) (y :: bool). (x, y)", "\ (x :: nat) (y :: bool). (x, y)"), ("\ (x :: ?'a) y. f y", "\(x :: ?'b). f"), ("\ (x :: nat) (y :: bool). f x y", "\ (x :: nat) (y :: bool). (?x :: nat \ bool \ ?'Z) x y"), ("\x. (?f :: nat \ bool \ nat) 0 x", "\x. f ?g x"), ("(?f :: nat \ bool \ nat) (?n :: nat)", "?g :: bool \ nat") ] val check = check_unit_tests_hints_unif tests true [] in Lecker.test_group ctxt () [ check "unify" unify, check "unify_hints" unify_hints ] end \ ML_command\ let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt) o Syntax.read_term ctxt) [ "?x \ (0 :: nat) \ ?y \ ?z \ ?x + ?y \ ?z" ] val tests = map (apply2 (Syntax.read_term ctxt))[ ("\ (x :: nat) (y :: bool). x", "\ (x :: nat) (y :: bool). 0 + x"), ("0 + (\ (x :: nat) (y :: bool). x) 0 ?y", "0 + (\ (x :: nat) (y :: bool). 0 + x) 0 ?z") ] val check_hints = check_unit_tests_hints_unif tests in Lecker.test_group ctxt () [ check_hints false [] "unify" unify, check_hints false [] "unify_hints without hints" unify_hints, check_hints true hints "unify_hints with hints" unify_hints ] end \ end diff --git a/thys/ML_Unification/Tests/Higher_Order_Pattern_ML_Unification_Tests.thy b/thys/ML_Unification/Tests/Higher_Order_Pattern_ML_Unification_Tests.thy --- a/thys/ML_Unification/Tests/Higher_Order_Pattern_ML_Unification_Tests.thy +++ b/thys/ML_Unification/Tests/Higher_Order_Pattern_ML_Unification_Tests.thy @@ -1,186 +1,192 @@ \<^marker>\creator "Kevin Kappelmann"\ \<^marker>\contributor "Paul Bachmann"\ subsection \Higher-Order Pattern ML Unification Tests\ theory Higher_Order_Pattern_ML_Unification_Tests imports ML_Unification_Tests_Base begin paragraph \Summary\ text \Tests for @{ML_structure "Higher_Order_Pattern_Unification"}.\ ML\ structure Prop = SpecCheck_Property structure UC = Unification_Combinator open Unification_Tests_Base structure Unif = Higher_Order_Pattern_Unification - val match = Unif.match [] + val norm_match_types = Type_Unification.e_match Unification_Util.match_types + val match = Unif.match |> norm_match_types + val closed_match = match [] val match_hints = let fun match binders = UC.add_fallback_matcher - (fn match_theory => Unif.e_match Unification_Util.match_types match_theory match_theory) + (fn match_theory => Unif.e_match Unification_Util.match_types match_theory match_theory + |> norm_match_types) ((fn binders => (Hints.map_retrieval (Hints.mk_retrieval Hints.TI.generalisations |> K) - #> Hints.UH.map_concl_unifier (Unif.match |> K) + #> Hints.UH.map_concl_unifier (match |> K) #> Hints.UH.map_normalisers (Unif.norms_match |> K) #> Hints.UH.map_prems_unifier (match |> K)) |> Context.proof_map #> Test_Unification_Hints.try_hints binders) |> UC.norm_matcher (#norm_term Unif.norms_match)) binders in match [] end - val unify = Unif.unify [] + val norm_unif_types = Type_Unification.e_unify Unification_Util.unify_types + val unify = Unif.unify |> norm_unif_types + val closed_unify = unify [] val unify_hints = let fun unif binders = UC.add_fallback_unifier - (fn unif_theory => Unif.e_unify Unification_Util.unify_types unif_theory unif_theory) + (fn unif_theory => Unif.e_unify Unification_Util.unify_types unif_theory unif_theory + |> norm_unif_types) ((fn binders => - (Hints.UH.map_concl_unifier (Unif.match |> K) + (Hints.UH.map_concl_unifier (match |> K) #> Hints.UH.map_normalisers (Unif.norms_unify |> K) #> Hints.UH.map_prems_unifier (unif |> K)) |> Context.proof_map #> Test_Unification_Hints.try_hints binders) |> UC.norm_unifier (#norm_term Unif.norms_unify)) binders in unif [] end \ subsubsection \Matching\ paragraph \Unit Tests\ ML_command\ let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} |> Variable.declare_term @{term "f :: nat \ bool \ nat"} val tests = map (apply2 (Syntax.read_term ctxt))[ ("?x :: nat", "(?x + 1 :: nat)"), ("(g :: nat \ nat \ nat) ?x ?x", "(g :: nat \ nat \ nat) (?x + 1) (?x + 1)"), ("\y. (?x :: nat \ ?'Z) y", "\y. f y"), ("(g :: ?'X \ ?'Y \ ?'Z) (x :: ?'X) (y :: ?'Y)", "(g :: ?'Y \ ?'Z \ ?'Z) (x :: ?'Y) (y :: ?'Z)"), ("(g :: ?'Z \ ?'X)", "\(x :: ?'X). (g :: ?'X \ ?'Y) x"), ("\ (x :: nat). (0 :: nat)", "\ (x :: nat). (0 :: nat)"), ("\ (x :: nat). x", "\ (x :: nat). x"), ("\ (x :: nat) (y :: bool). (x, y)", "\ (x :: nat) (y :: bool). (x, y)"), ("\ (x :: ?'A) (y :: bool). (?x :: ?'A \ bool \ ?'Z) x y", "\ (x :: nat) (y :: bool). f x y"), ("\(x :: ?'X). (g :: ?'X \ ?'X) x", "(g :: ?'X \ ?'X)"), ("\y. (?x :: nat \ bool \ nat) y", "\y. f y"), ("\y z. (?x :: nat \ bool \ nat) y z", "f") ] val check = check_unit_tests_hints_match tests true [] in Lecker.test_group ctxt () [ - check "match" match, + check "match" closed_match, check "match_hints" match_hints ] end \ subparagraph \Asymmetry\ ML_command\ let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} |> Variable.declare_term @{term "f :: nat \ bool \ nat"} val tests = map (apply2 (Syntax.read_term ctxt))[ ("\y. f y", "\y. (?x :: nat \ bool \ nat) y") ] val check = check_unit_tests_hints_match tests false [] in Lecker.test_group ctxt () [ - check "match" match, + check "match" closed_match, check "match_hints" match_hints ] end \ subparagraph \Ground Hints\ ML_command\ let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt) o Syntax.read_term ctxt) [ "?x \ 2 \ ?x + (-?x :: int) \ 0", "?x \ 0 \ ?y \ 0 \ (?x :: int) + ?y \ 0" ] val tests = map (apply2 (Syntax.read_term ctxt))[ ("(2 + -2) + (0 + 0) :: int", "0 :: int") ] val check_hints = check_unit_tests_hints_match tests in Lecker.test_group ctxt () [ - check_hints false [] "match" match, + check_hints false [] "match" closed_match, check_hints false [] "match_hints without hints" match_hints, check_hints true hints "match_hints with hints" match_hints ] end \ subsubsection \Unification\ paragraph \Generated Tests\ subparagraph \First-Order\ ML_command\ structure Test_Params = struct - val unify = unify + val unify = closed_unify val unify_hints = unify_hints val params = { nv = 10, ni = 10, max_h = 5, max_args = 4 } end structure First_Order_Tests = First_Order_Unification_Tests(Test_Params) val _ = First_Order_Tests.tests @{context} (SpecCheck_Random.new ()) \ subparagraph \Higher-Order Pattern\ ML_file\higher_order_pattern_unification_tests.ML\ ML_command\ val ctxt = @{context} val tests = Higher_Order_Pattern_Unification_Tests.unit_tests_unifiable ctxt val check_hints = check_unit_tests_hints_unif tests val _ = Lecker.test_group ctxt () [ - check_hints true [] "unify" unify, + check_hints true [] "unify" closed_unify, check_hints true [] "unify_hints" unify_hints ] \ paragraph \Unit Tests\ subparagraph \With Unification Hints\ ML_command\ let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} |> Variable.declare_term @{term "f :: (nat \ nat) \ nat \ nat"} |> Variable.declare_term @{term "g :: nat \ nat \ nat"} |> Variable.declare_term @{term "h :: nat"} val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt) o Syntax.read_term ctxt) [ "?x \ (0 :: nat) \ ?y \ ?z \ ?x + ?y \ ?z", "(?x :: ?'X) \ (?y :: ?'X) \ id ?x \ ?y", "(?x1 :: nat) \ ?x2 \ (?y1 :: nat) \ ?y2 \ ?x1 + ?y1 \ ?y2 + ?x2" ] val tests = map (apply2 (Syntax.read_term ctxt)) [ ("\ x y. 0 + 1 :: nat", "\ x y. 1 :: nat"), ("\ x. 0 + 0 + x :: nat", "\ x. x :: nat"), ("\ x y. 0 + Suc ?x", "\ x y. Suc 2"), ("\ (x :: nat) (y :: nat). y + (0 + x)", "\ (x :: nat) (y :: nat). x + (0 + y)"), ("f (\ u. g (?x, h), h)", "id (f (\ u. ?y, 0 + h))") ] val check_hints = check_unit_tests_hints_unif tests in Lecker.test_group ctxt () [ - check_hints false [] "unify" unify, + check_hints false [] "unify" closed_unify, check_hints false [] "unify_hints without hints" unify_hints, check_hints true hints "unify_hints with hints" unify_hints ] end \ end diff --git a/thys/ML_Unification/Tests/ML_Unification_Tests_Base.thy b/thys/ML_Unification/Tests/ML_Unification_Tests_Base.thy --- a/thys/ML_Unification/Tests/ML_Unification_Tests_Base.thy +++ b/thys/ML_Unification/Tests/ML_Unification_Tests_Base.thy @@ -1,33 +1,35 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Test Setup\ theory ML_Unification_Tests_Base imports ML_Unification_Hints SpecCheck.SpecCheck Main begin paragraph \Summary\ text \Shared setup for unification tests. We use \<^cite>\speccheck\ to generate tests and create unit tests.\ ML\ @{functor_instance struct_name = Test_Unification_Hints and functor_name = Term_Index_Unification_Hints and id = \"test"\ and more_args = \ structure TI = Discrimination_Tree val init_args = { - concl_unifier = SOME Higher_Order_Pattern_Unification.match, + concl_unifier = SOME (Higher_Order_Pattern_Unification.match + |> Type_Unification.e_match Unification_Util.match_types), normalisers = SOME Unification_Util.beta_eta_short_norms_unif, - prems_unifier = SOME Higher_Order_Pattern_Unification.unify, + prems_unifier = SOME (Higher_Order_Pattern_Unification.unify + |> Type_Unification.e_unify Unification_Util.unify_types), retrieval = SOME (Term_Index_Unification_Hints_Args.mk_sym_retrieval TI.norm_term TI.generalisations), hint_preprocessor = SOME (K I) }\} \ ML_file \tests_base.ML\ ML_file \first_order_unification_tests.ML\ end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Attributes/Unification_Attributes.thy b/thys/ML_Unification/Unification_Attributes/Unification_Attributes.thy --- a/thys/ML_Unification/Unification_Attributes/Unification_Attributes.thy +++ b/thys/ML_Unification/Unification_Attributes/Unification_Attributes.thy @@ -1,72 +1,70 @@ \<^marker>\creator "Kevin Kappelmann"\ -section \Unification Attributes\ theory Unification_Attributes - imports Unify_Resolve_Tactics + imports + Unification_Attributes_Base + ML_Unifiers begin paragraph \Summary\ -text \OF attribute with adjustable unifier.\ - -ML_file\unify_of_base.ML\ -ML_file\unify_of.ML\ +text \Setup of OF attribute with adjustable unifier.\ ML\ @{functor_instance struct_name = Standard_Unify_OF and functor_name = Unify_OF and id = \""\ and more_args = \val init_args = { normalisers = SOME Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify, unifier = SOME Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify, mode = SOME (Unify_OF_Args.PM.key Unify_OF_Args.PM.fact) }\} \ local_setup \Standard_Unify_OF.setup_attribute NONE\ paragraph \Examples\ experiment begin lemma assumes h1: "(PROP A \ PROP D) \ PROP E \ PROP C" assumes h2: "PROP B \ PROP D" and h3: "PROP F \ PROP E" shows "(PROP A \ PROP B) \ PROP F \ PROP C" by (fact h1[uOF h2 h3 where mode = resolve]) \\the line below is equivalent\ (* by (fact h1[OF h2 h3]) *) lemma assumes h1: "(PROP A \ PROP A)" assumes h2: "(PROP A \ PROP A) \ PROP B" shows "PROP B" by (fact h2[uOF h1]) \\the line below is equivalent\ (* by (fact h2[uOF h1 where mode = fact]) *) \\Note: @{attribute OF} would not work in this case:\ (* thm h2[OF h1] *) lemma assumes h1: "\x y z. PROP P x y \ PROP P y y \ (PROP A \ PROP A) \ (PROP A \ PROP B) \ PROP C" and h2: "\x y. PROP P x y" and h3 : "PROP A \ PROP A" and h4 : "PROP D \ PROP B" shows "(PROP A \ PROP D) \ PROP C" by (fact h1[uOF h2 h2 h3, uOF h4 where mode = resolve]) lemma assumes h1: "\P x. PROP P x \ PROP E P x" and h2: "PROP P x" shows "PROP E P x" by (fact h1[uOF h2]) \\the following line does not work (multiple unifiers error)\ (* by (fact h1[OF h2]) *) text\We can also specify the unifier to be used:\ lemma assumes h1: "\P. PROP P \ PROP E" and h2: "\P. PROP P" shows "PROP E" by (fact h1[uOF h2 where unifier = First_Order_Unification.unify]) \\the line below is equivalent\ (* using [[uOF unifier = First_Order_Unification.unify]] by (fact h1[uOF h2]) *) end end diff --git a/thys/ML_Unification/Unification_Attributes/Unification_Attributes_Base.thy b/thys/ML_Unification/Unification_Attributes/Unification_Attributes_Base.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Attributes/Unification_Attributes_Base.thy @@ -0,0 +1,13 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Unification Attributes\ +theory Unification_Attributes_Base + imports Unify_Resolve_Tactics_Base +begin + +paragraph \Summary\ +text \OF attribute with adjustable unifier.\ + +ML_file\unify_of_base.ML\ +ML_file\unify_of.ML\ + +end diff --git a/thys/ML_Unification/Unification_Hints/ML_Unification_Hints.thy b/thys/ML_Unification/Unification_Hints/ML_Unification_Hints.thy --- a/thys/ML_Unification/Unification_Hints/ML_Unification_Hints.thy +++ b/thys/ML_Unification/Unification_Hints/ML_Unification_Hints.thy @@ -1,108 +1,98 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Unification Hints\ theory ML_Unification_Hints imports - ML_Generic_Data_Utils - ML_Term_Index + ML_Unification_Hints_Base ML_Unifiers - ML_Unification_Parsers begin paragraph \Summary\ -text \A generalisation of unification hints, originally introduced in \<^cite>\"unif-hints"\. -We support a generalisation that -\<^enum> allows additional universal variables in premises -\<^enum> allows non-atomic left-hand sides for premises -\<^enum> allows arbitrary functions to perform the matching/unification of a hint with a disagreement pair. - -General shape of a hint: -\\y1...yn. (\x1...xn1. lhs1 \ rhs1) \ ... \ (\x1...xnk. lhsk \ rhsk) \ lhs \ rhs\ -\ - -ML_file\unification_hints_base.ML\ -ML_file\unification_hints.ML\ -ML_file\term_index_unification_hints.ML\ +text \Setup of unification hints.\ text \We now set up two unifiers using unification hints. The first one allows for recursive applications of unification hints when unifying a hint's conclusion \lhs \ rhs\ with a goal \lhs' \ rhs'\. -The second only uses a combination of higher-order pattern and first-order unification. -This particularly implies that recursive applications of unification hints have to be made explicit -in the hint itself (cf. @{dir "../Examples"}). +The second disallows recursive applications of unification hints. Recursive applications have to be +made explicit in the hint itself (cf. @{dir "../Examples"}). While the former can be convenient for local hint registrations and quick developments, it is advisable to use the second for global hints to avoid unexpected looping behaviour.\ ML\ @{functor_instance struct_name = Standard_Unification_Hints_Rec and functor_name = Term_Index_Unification_Hints and id = \"rec"\ and more_args = \ structure TI = Discrimination_Tree val init_args = { concl_unifier = SOME Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify, prems_unifier = SOME (Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify |> Unification_Combinator.norm_unifier Envir_Normalisation.beta_norm_term_unif), normalisers = SOME Standard_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 \Standard_Unification_Hints_Rec.setup_attribute NONE\ text\Standard unification hints using @{ML Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify} when looking for hints are accessible via @{attribute rec_uhint}. \<^emph>\Note:\ when we retrieve a potential unification hint with conclusion \lhs \ rhs\ for a goal \lhs' \ rhs'\, we only consider those hints whose lhs potentially higher-order unifies with lhs' or rhs' \<^emph>\without using hints\. For otherwise, any hint \lhs \ rhs\ applied to a goal \rhs \ lhs\ leads to an immediate loop.\ +declare [[ucombine add = \Standard_Unification_Combine.eunif_data + (Standard_Unification_Hints_Rec.try_hints + |> Unification_Combinator.norm_unifier + (#norm_term Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify) + |> K) + (Standard_Unification_Combine.metadata Standard_Unification_Hints_Rec.binding Prio.LOW)\]] + ML\ @{functor_instance struct_name = Standard_Unification_Hints and functor_name = Term_Index_Unification_Hints and id = \""\ and more_args = \ structure TI = Discrimination_Tree val init_args = { - concl_unifier = SOME (fn binders => - Standard_Unification_Combine.map_eunif_datas (K Prio.Table.empty) - |> Context.proof_map - #> Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify binders), + concl_unifier = NONE, prems_unifier = SOME (Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify |> Unification_Combinator.norm_unifier Envir_Normalisation.beta_norm_term_unif), normalisers = SOME Standard_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 \Standard_Unification_Hints.setup_attribute NONE\ +declare [[uhint where concl_unifier = \fn binders => + Standard_Unification_Combine.delete_eunif_data + (Standard_Unification_Combine.metadata Standard_Unification_Hints.binding (Prio.LOW + 1)) + (*TODO: should we also remove the recursive hint unifier here? time will tell...*) + (*#> Standard_Unification_Combine.delete_eunif_data + (Standard_Unification_Combine.metadata Standard_Unification_Hints_Rec.binding Prio.LOW)*) + |> Context.proof_map + #> Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify binders\]] text\Standard unification hints using @{ML Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify} when looking for hints, without using fallback list of unifiers, are accessible via @{attribute uhint}. \<^emph>\Note:\ there will be no recursive usage of unification hints when searching for potential unification hints in this case. See also @{dir "../Examples"}.\ declare [[ucombine add = \Standard_Unification_Combine.eunif_data - (Standard_Unification_Hints_Rec.try_hints - |> Unification_Combinator.norm_unifier - (#norm_term Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify) - |> K) - (Standard_Unification_Combine.default_metadata Standard_Unification_Hints_Rec.binding)\]] - -declare [[ucombine add = \Standard_Unification_Combine.eunif_data (Standard_Unification_Hints.try_hints |> Unification_Combinator.norm_unifier (#norm_term Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify) |> K) - (Standard_Unification_Combine.default_metadata Standard_Unification_Hints.binding)\]] + (Standard_Unification_Combine.metadata Standard_Unification_Hints.binding (Prio.LOW + 1))\]] text\Examples see @{dir "../Examples"}.\ end diff --git a/thys/ML_Unification/Unification_Hints/ML_Unification_Hints_Base.thy b/thys/ML_Unification/Unification_Hints/ML_Unification_Hints_Base.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Hints/ML_Unification_Hints_Base.thy @@ -0,0 +1,30 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Unification Hints\ +theory ML_Unification_Hints_Base + imports + ML_Conversion_Utils + ML_Functor_Instances + ML_Generic_Data_Utils + ML_Priorities + ML_Term_Index + ML_Term_Utils + ML_Unifiers_Base + ML_Unification_Parsers +begin + +paragraph \Summary\ +text \A generalisation of unification hints, originally introduced in \<^cite>\"unif-hints"\. +We support a generalisation that +\<^enum> allows additional universal variables in premises +\<^enum> allows non-atomic left-hand sides for premises +\<^enum> allows arbitrary functions to perform the matching/unification of a hint with a disagreement pair. + +General shape of a hint: +\\y1...yn. (\x1...xn1. lhs1 \ rhs1) \ ... \ (\x1...xnk. lhsk \ rhsk) \ lhs \ rhs\ +\ + +ML_file\unification_hints_base.ML\ +ML_file\unification_hints.ML\ +ML_file\term_index_unification_hints.ML\ + +end diff --git a/thys/ML_Unification/Unification_Hints/term_index_unification_hints.ML b/thys/ML_Unification/Unification_Hints/term_index_unification_hints.ML --- a/thys/ML_Unification/Unification_Hints/term_index_unification_hints.ML +++ b/thys/ML_Unification/Unification_Hints/term_index_unification_hints.ML @@ -1,384 +1,387 @@ (* Title: ML_Unification/term_index_unification_hints.ML Author: Kevin Kappelmann Unification hints stored in a term index. *) @{parse_entries (sig) PARSE_TERM_INDEX_UNIFICATION_HINTS_MODE [add, del]} @{parse_entries (sig) PARSE_TERM_INDEX_UNIFICATION_HINTS_ARGS [concl_unifier, normalisers, prems_unifier, retrieval, hint_preprocessor, prio]} @{parse_entries (sig) PARSE_TERM_INDEX_UNIFICATION_HINTS_CONTEXT_ARGS [concl_unifier, normalisers, prems_unifier, retrieval, hint_preprocessor]} signature TERM_INDEX_UNIFICATION_HINTS_ARGS = sig structure PA : PARSE_TERM_INDEX_UNIFICATION_HINTS_ARGS structure PCA : PARSE_TERM_INDEX_UNIFICATION_HINTS_CONTEXT_ARGS val PCA_entries_from_PA_entries : ('a, 'b, 'c, 'd, 'e, 'f) PA.entries -> ('a, 'b, 'c, 'd, 'e) PCA.entries val PA_entries_from_PCA_entries : ('a, 'b, 'c, 'd, 'e) PCA.entries -> 'f -> ('a, 'b, 'c, 'd, 'e, 'f) PA.entries val UHA_PA_entries_from_PCA_entries : ('a, 'b, 'c, 'd, 'e) PCA.entries -> ('a, 'b, 'c) Unification_Hints_Args.PA.entries structure PM : PARSE_TERM_INDEX_UNIFICATION_HINTS_MODE type mode = PM.key val parse_mode : mode parser type hint_prio = Unification_Hints_Base.unif_hint * Prio.prio val pretty_hint_prio : Proof.context -> hint_prio -> Pretty.T val sort_hint_prios : hint_prio list -> hint_prio list val mk_retrieval : Term_Normalisation.term_normaliser -> ('ti -> term -> hint_prio list) -> 'ti -> Unification_Hints_Base.hint_retrieval val mk_sym_retrieval : Term_Normalisation.term_normaliser -> ('ti -> term -> hint_prio list) -> 'ti -> Unification_Hints_Base.hint_retrieval type 'ti args = (Unification_Base.unifier, Unification_Base.normalisers, Unification_Base.unifier, 'ti -> Unification_Hints_Base.hint_retrieval, Unification_Hints_Base.hint_preprocessor, Prio.prio) PA.entries type 'ti context_args = (Unification_Base.unifier, Unification_Base.normalisers, Unification_Base.unifier, 'ti -> Unification_Hints_Base.hint_retrieval, Unification_Hints_Base.hint_preprocessor) PCA.entries val arg_parsers : (ML_Code_Util.code parser, ML_Code_Util.code parser, ML_Code_Util.code parser, ML_Code_Util.code parser, ML_Code_Util.code parser, Prio.prio context_parser) PA.entries end structure Term_Index_Unification_Hints_Args : TERM_INDEX_UNIFICATION_HINTS_ARGS = struct structure UB = Unification_Base structure EN = Envir_Normalisation structure UHB = Unification_Hints_Base structure UHA = Unification_Hints_Args structure TUHP = Prio @{parse_entries (struct) PA [concl_unifier, normalisers, prems_unifier, retrieval, hint_preprocessor, prio]} @{parse_entries (struct) PCA [concl_unifier, normalisers, prems_unifier, retrieval, hint_preprocessor]} fun PCA_entries_from_PA_entries {concl_unifier = concl_unifier, normalisers = normalisers, prems_unifier = prems_unifier, retrieval = retrieval, hint_preprocessor = hint_preprocessor,...} = {concl_unifier = concl_unifier, normalisers = normalisers, prems_unifier = prems_unifier, retrieval = retrieval, hint_preprocessor = hint_preprocessor} fun PA_entries_from_PCA_entries {concl_unifier = concl_unifier, normalisers = normalisers, prems_unifier = prems_unifier, retrieval = retrieval, hint_preprocessor = hint_preprocessor} prio = {concl_unifier = concl_unifier, normalisers = normalisers, prems_unifier = prems_unifier, retrieval = retrieval, hint_preprocessor = hint_preprocessor, prio = SOME prio} fun UHA_PA_entries_from_PCA_entries {concl_unifier = concl_unifier, normalisers = normalisers, prems_unifier = prems_unifier,...} = {concl_unifier = concl_unifier, normalisers = normalisers, prems_unifier = prems_unifier} @{parse_entries (struct) PM [add, del]} type mode = PM.key val parse_mode = PM.parse_key type hint_prio = UHB.unif_hint * TUHP.prio fun pretty_hint_prio ctxt (hint, prio) = Pretty.block [ UHB.pretty_hint ctxt hint, Pretty.enclose " (" ")" [Pretty.str "priority: ", Prio.pretty prio] ] (*FIXME: use Prio.Table instead of sorted lists*) val sort_hint_prios = sort (Prio.ord o apply2 snd) val hint_seq_of_hint_prios = sort_hint_prios #> Seq.of_list #> Seq.map fst fun mk_retrieval norm_term retrieve ti _ _ (t, _) _ = norm_term t |> retrieve ti |> hint_seq_of_hint_prios fun interleave [] ys = ys | interleave xs [] = xs | interleave (x :: xs) (y :: ys) = x :: y :: interleave xs ys fun mk_sym_retrieval norm_term retrieve ti _ _ (t1, t2) _ = interleave (norm_term t1 |> retrieve ti) (norm_term t2 |> retrieve ti |> map (apfst Unification_Hints_Base.symmetric_hint)) |> hint_seq_of_hint_prios type 'ti args = (UB.unifier, UB.normalisers, UB.unifier, 'ti -> UHB.hint_retrieval, UHB.hint_preprocessor, TUHP.prio) PA.entries type 'ti context_args = (UB.unifier, UB.normalisers, UB.unifier, 'ti -> UHB.hint_retrieval, UHB.hint_preprocessor) PCA.entries val arg_parsers = { concl_unifier = UHA.PA.get_concl_unifier_safe UHA.arg_parsers, normalisers = UHA.PA.get_normalisers_safe UHA.arg_parsers, prems_unifier = UHA.PA.get_prems_unifier_safe UHA.arg_parsers, retrieval = SOME (Parse_Util.nonempty_code (K "retrieval function must not be empty")), hint_preprocessor = SOME (Parse_Util.nonempty_code (K "hint preprocessor must not be empty")), prio = SOME Prio.parse } end signature TERM_INDEX_UNIFICATION_HINTS = sig include HAS_LOGGER structure UH : UNIFICATION_HINTS (*underlying term index*) structure TI : TERM_INDEX structure Data : GENERIC_DATA val get_retrieval : Context.generic -> Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> Unification_Hints_Base.hint_retrieval val map_retrieval : ( (Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> Unification_Hints_Base.hint_retrieval) -> Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> Unification_Hints_Base.hint_retrieval) -> Context.generic -> Context.generic val get_hint_preprocessor : Context.generic -> Unification_Hints_Base.hint_preprocessor val map_hint_preprocessor : (Unification_Hints_Base.hint_preprocessor -> Unification_Hints_Base.hint_preprocessor) -> Context.generic -> Context.generic val get_index : Context.generic -> Term_Index_Unification_Hints_Args.hint_prio TI.term_index val map_index : (Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> Term_Index_Unification_Hints_Args.hint_prio TI.term_index) -> Context.generic -> Context.generic val pretty_index : Proof.context -> Pretty.T val add_hint_prio : Term_Index_Unification_Hints_Args.hint_prio -> Context.generic -> Context.generic val del_hint : Unification_Hints_Base.unif_hint -> Context.generic -> Context.generic val mk_retrieval : Term_Index_Unification_Hints_Args.hint_prio TI.retrieval -> Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> Unification_Hints_Base.hint_retrieval val mk_sym_retrieval : Term_Index_Unification_Hints_Args.hint_prio TI.retrieval -> Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> Unification_Hints_Base.hint_retrieval val try_hints : Unification_Base.unifier val binding : binding val attribute : (Term_Index_Unification_Hints_Args.mode * ((ML_Code_Util.code, ML_Code_Util.code, ML_Code_Util.code, ML_Code_Util.code, ML_Code_Util.code, Prio.prio) Term_Index_Unification_Hints_Args.PA.entries)) * Position.T -> attribute val parse_attribute : attribute context_parser val setup_attribute : string option -> local_theory -> local_theory end functor Term_Index_Unification_Hints(A : sig structure FIA : FUNCTOR_INSTANCE_ARGS structure TI : TERM_INDEX val init_args : (Term_Index_Unification_Hints_Args.hint_prio TI.term_index) Term_Index_Unification_Hints_Args.context_args end) : TERM_INDEX_UNIFICATION_HINTS = struct structure UHB = Unification_Hints_Base structure TUHA = Term_Index_Unification_Hints_Args structure TUHP = Prio structure PA = TUHA.PA structure PCA = TUHA.PCA structure PM = TUHA.PM structure FIU = Functor_Instance_Util(A.FIA) structure TI = A.TI structure MCU = ML_Code_Util structure PU = Parse_Util val logger = Logger.setup_new_logger Unification_Hints_Base.logger FIU.FIA.full_name @{functor_instance struct_name = UH and functor_name = Unification_Hints and accessor = FIU.accessor and id = FIU.FIA.id and more_args = \val init_args = TUHA.UHA_PA_entries_from_PCA_entries A.init_args\} -fun are_hint_variants ctxt hp = +fun are_hint_variants hp = let - val tp = apply2 Thm.prop_of hp - val env = Unification_Util.empty_envir tp - fun match tp = First_Order_Unification.match [] ctxt tp env - in - match tp - |> Seq.maps (fn _ => match (swap tp)) - |> not o fst o General_Util.seq_is_empty - end + val ctp = apply2 Thm.cprop_of hp + val match = Thm.first_order_match + in can match ctp andalso can match (swap ctp) end (* fun are_hint_prio_variants ctxt ((h1, p1), (h2, p2)) = p1 = p2 andalso are_hint_variants ctxt (h1, h2) *) structure Data = Generic_Data(Pair_Generic_Data_Args( struct structure Data1 = Pair_Generic_Data_Args( struct structure Data1 = struct type T = TUHA.hint_prio TI.term_index -> Unification_Hints_Base.hint_retrieval val empty = PCA.get_retrieval A.init_args val merge = fst end structure Data2 = struct type T = UHB.hint_preprocessor val empty = PCA.get_hint_preprocessor A.init_args val merge = fst end end) structure Data2 = Term_Index_Generic_Data_Args( struct type data = TUHA.hint_prio structure TI = TI - fun data_eq ((h1, _), (h2, _))= are_hint_variants (Context.the_local_context ()) (h1, h2) + fun data_eq ((h1, _), (h2, _)) = are_hint_variants (h1, h2) end) end)) val get_retrieval = fst o fst o Data.get val map_retrieval = Data.map o apfst o apfst val get_hint_preprocessor = snd o fst o Data.get val map_hint_preprocessor = Data.map o apfst o apsnd val get_index = snd o Data.get val map_index = Data.map o apsnd fun pretty_index ctxt = get_index (Context.Proof ctxt) |> TI.content |> TUHA.sort_hint_prios |> map (TUHA.pretty_hint_prio ctxt) |> Pretty.fbreaks |> Pretty.block val term_index_key_from_hint = (*TODO: as of now, hints are only indexed on their lhs term; we could index on both terms to avoid more false positives*) UHB.cdest_hint_concl #> fst #> Thm.term_of #> TI.norm_term fun msg_illegal_hint_format ctxt hint = Pretty.block [ Pretty.str "Illegal hint format for ", UHB.pretty_hint ctxt hint ] |> Pretty.string_of fun add_hint_prio (hint, prio) context = let val ctxt = Context.proof_of context val hint = hint |> get_hint_preprocessor context ctxt - val is_hint_variant = curry (are_hint_variants ctxt) hint o fst + val is_hint_variant = curry are_hint_variants hint o fst in (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "Adding hint ", UHB.pretty_hint ctxt hint ] |> Pretty.string_of); TI.insert is_hint_variant (term_index_key_from_hint hint, (hint, prio)) (get_index context) |> (fn ti => map_index (K ti) context)) handle Term_Index_Base.INSERT => (@{log Logger.WARN} ctxt (fn _ => Pretty.block [ Pretty.str "Hint ", UHB.pretty_hint ctxt hint, Pretty.str " already added." ] |> Pretty.string_of); context) | TERM _ => (@{log Logger.WARN} ctxt (fn _ => msg_illegal_hint_format ctxt hint); context) end fun del_hint hint context = let val ctxt = Context.proof_of context val hint = hint |> get_hint_preprocessor context ctxt - val is_hint_variant = curry (are_hint_variants ctxt) hint o fst + val is_hint_variant = curry are_hint_variants hint o fst in (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "Deleting hint ", UHB.pretty_hint ctxt hint ] |> Pretty.string_of); TI.delete is_hint_variant (term_index_key_from_hint hint) (get_index context) |> (fn ti => map_index (K ti) context)) handle Term_Index_Base.DELETE => (@{log Logger.WARN} ctxt (fn _ => Pretty.block [ Pretty.str "Hint ", UHB.pretty_hint ctxt hint, Pretty.str " not found." ] |> Pretty.string_of); context) | TERM _ => (@{log Logger.WARN} ctxt (fn _ => msg_illegal_hint_format ctxt hint); context) end val mk_retrieval = TUHA.mk_retrieval TI.norm_term val mk_sym_retrieval = TUHA.mk_sym_retrieval TI.norm_term -fun try_hints binders ctxt = - let val context = Context.Proof ctxt - in UH.try_hints (get_retrieval context (get_index context)) binders ctxt end +val binding = FIU.mk_binding_id_prefix "uhint" -val binding = FIU.mk_binding_id_prefix "uhint" +fun try_hints binders ctxt tp = + let + val context = Context.Proof ctxt + val _ = @{log Logger.DEBUG} ctxt (fn _ => + Pretty.block [ + Pretty.str "Trying unification hints ", + Binding.pretty binding, + Pretty.str " for ", + Unification_Util.pretty_unif_problem ctxt tp + ] |> Pretty.string_of) + in UH.try_hints (get_retrieval context (get_index context)) binders ctxt tp end val default_entries = PA.empty_entries () |> fold PA.set_entry [PA.prio TUHP.MEDIUM] fun parse_arg_entries mode = let val parsers = TUHA.arg_parsers val parse_value = PA.parse_entry (PA.get_concl_unifier parsers |> Scan.lift) (PA.get_normalisers parsers |> Scan.lift) (PA.get_prems_unifier parsers |> Scan.lift) (PA.get_retrieval parsers |> Scan.lift) (PA.get_hint_preprocessor parsers |> Scan.lift) (PA.get_prio parsers) val parse_entry = Parse_Key_Value.parse_entry' (Scan.lift PA.parse_key) (Scan.lift PU.eq) parse_value in PA.parse_entries_required' Parse.and_list1' [] parse_entry default_entries |> mode = PM.key PM.del ? PU.filter_cut' (is_none o PA.get_prio_safe) (K o K "Priority may not be specified for deletion.") end fun attribute ((mode, entries), pos) = let fun update_index thm = Term_Util.no_dummy_pattern (Thm.prop_of thm) ? ( case mode of PM.add _ => add_hint_prio (thm, PA.get_prio entries) | PM.del _ => del_hint thm) val PCA_entries = TUHA.PCA_entries_from_PA_entries entries val UHA_PA_entries = TUHA.UHA_PA_entries_from_PCA_entries PCA_entries val run_code = ML_Attribute.run_map_context o rpair pos fun default_code (context, _) = (SOME context, NONE) val map_retrieval = case PCA.get_retrieval_safe PCA_entries of SOME c => FIU.code_struct_op "map_retrieval" @ MCU.atomic (MCU.read "K" @ MCU.atomic c) |> run_code | NONE => default_code val map_hint_preprocessor = case PCA.get_hint_preprocessor_safe PCA_entries of SOME c => FIU.code_struct_op "map_hint_preprocessor" @ MCU.atomic (MCU.read "K" @ MCU.atomic c) |> run_code | NONE => default_code in ML_Attribute_Util.apply_attribute (Thm.declaration_attribute update_index) #> ML_Attribute_Util.apply_attribute (UH.attribute (UHA_PA_entries, pos)) #> ML_Attribute_Util.apply_attribute map_retrieval #> map_hint_preprocessor end val parse_attribute = (Scan.lift (Scan.optional TUHA.parse_mode (PM.key PM.add)) :-- (fn mode => PU.optional' ((Scan.lift Parse.where_ |-- parse_arg_entries mode) |> Parse.!!!!) default_entries) |> PU.position') >> attribute val setup_attribute = Attrib.local_setup binding (Parse.!!!! parse_attribute) o the_default ("set unification hints arguments (" ^ FIU.FIA.full_name ^ ")") end diff --git a/thys/ML_Unification/Unification_Hints/unification_hints_base.ML b/thys/ML_Unification/Unification_Hints/unification_hints_base.ML --- a/thys/ML_Unification/Unification_Hints/unification_hints_base.ML +++ b/thys/ML_Unification/Unification_Hints/unification_hints_base.ML @@ -1,145 +1,145 @@ (* Title: ML_Unification/unification_hints_base.ML Author: Kevin Kappelmann, Paul Bachmann Unification hints (introduced in "Hints in unification" by Asperti et al, 2009). We support a generalisation that 1. allows additional universal variables in premises 2. allows non-atomic left-hand sides for premises 3. allows arbitrary functions to perform the matching/unification of a hint with a disagreement pair. General shape of a hint: \\y1...yn. (\x1...xn1. lhs1 \ rhs1) \ ... \ (\x1...xnk. lhsk \ rhsk) \ lhs \ rhs\ *) signature UNIFICATION_HINTS_BASE = sig include HAS_LOGGER type unif_hint = thm val cdest_hint : unif_hint -> (cterm * cterm) list * (cterm * cterm) val cdest_hint_concl : unif_hint -> cterm * cterm val symmetric_hint : unif_hint -> unif_hint val pretty_hint : Proof.context -> unif_hint -> Pretty.T type hint_preprocessor = Proof.context -> thm -> unif_hint val obj_logic_hint_preprocessor : thm -> conv -> hint_preprocessor val try_hint : Unification_Base.matcher -> Unification_Base.normalisers -> unif_hint -> Unification_Base.e_unifier type hint_retrieval = term Binders.binders -> Proof.context -> term * term -> Envir.env -> unif_hint Seq.seq val try_hints : hint_retrieval -> Unification_Base.matcher -> Unification_Base.normalisers -> Unification_Base.e_unifier end structure Unification_Hints_Base : UNIFICATION_HINTS_BASE = struct val logger = Logger.setup_new_logger Unification_Base.logger "Unification_Hints_Base" structure GUtil = General_Util structure TUtil = Term_Util structure CUtil = Conversion_Util type unif_hint = thm val cdest_hint = Thm.cprop_of #> TUtil.strip_cimp #>> map Thm.dest_equals ##> Thm.dest_equals val cdest_hint_concl = Thm.cconcl_of #> Thm.dest_equals (*only flips the conclusion since unification calls for the premises should be symmetric*) val symmetric_hint = Conv.concl_conv ~1 CUtil.symmetric_conv |> CUtil.thm_conv val pretty_hint = Thm.pretty_thm type hint_preprocessor = Proof.context -> thm -> unif_hint fun obj_logic_hint_preprocessor eq_eq_meta_eq default_conv ctxt = let fun eq_conv ct = (if can Thm.dest_equals ct then Conv.all_conv else Conv.rewr_conv eq_eq_meta_eq else_conv default_conv) ct val forall_eq_conv = Conversion_Util.repeat_forall_conv (K o K eq_conv) in Conversion_Util.imp_conv (forall_eq_conv ctxt) eq_conv |> Conversion_Util.thm_conv end (*Tries to apply a hint to solve E-unification of (t1 \? t2). Vars in hint are lifted wrt. the passed binders. Unifies the hint's conclusion with (t1, t2) using match. Unifies resulting unification problems using unify. Normalises the terms and theorems after unify with norms. Returns a sequence of (env, thm) pairs.*) fun try_hint match norms hint unify binders ctxt (t1, t2) (Envir.Envir {maxidx, tenv, tyenv}) = let val dest_all_equals = TUtil.strip_all ##> Logic.dest_equals val rev_binders = rev binders val _ = @{log Logger.TRACE} ctxt (fn _ => Pretty.block [ Pretty.str "Trying hint: ", pretty_hint ctxt hint ] |> Pretty.string_of) (*lift hint to include bound variables and increase indexes*) val lifted_hint = (*"P" is just some dummy proposition variable with appropriate index*) Logic.list_all (map fst rev_binders, Var (("P", maxidx + 1), propT)) |> Thm.cterm_of ctxt (*lift bound variables to the hint*) |> GUtil.flip Thm.lift_rule hint val (hint_prems, hint_concl) = Thm.prop_of lifted_hint |> TUtil.strip_subgoal |> snd (*match hint with unification pair*) val (P, Q) = dest_all_equals hint_concl |> snd val match = match binders ctxt val (no_hint_match, match_env_concl_thmpq) = Envir.Envir {maxidx = Thm.maxidx_of lifted_hint, tenv = tenv, tyenv = tyenv} |> match (P, t1) |> Seq.maps (fn (env, thm1) => match (Q, t2) env |> Seq.map (apsnd (pair thm1))) |> GUtil.seq_is_empty in if no_hint_match then (@{log Logger.TRACE} ctxt (K "Hint does not match."); Seq.empty) else let val _ = @{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "Hint ", pretty_hint ctxt hint, Pretty.str ((length hint_prems > 0 ? suffix " Unifying premises...") " matched.") ] |> Pretty.string_of) (*unify each hint premise and collect the theorems while iteratively extending the environment*) fun unif_prem prem = let val (params, PQ_prem) = dest_all_equals prem val (binders, ctxt) = Binders.fix_binders params ctxt fun unif_add_result (env, thms) = unify binders ctxt PQ_prem env |> Seq.map (apsnd (fn thm => (binders, thm) :: thms)) in Seq.maps unif_add_result end fun unify_prems (match_env, concl_thmp) = fold unif_prem hint_prems (Seq.single (match_env, [])) |> Seq.map (rpair concl_thmp) fun inst_discharge ((env_res, prem_thms), concl_thmp) = let (*instantiate the theorems*) val lifted_hint_inst = #norm_thm norms ctxt env_res lifted_hint val norm_term = #norm_term norms env_res val mk_norm_cbinders = map (Thm.cterm_of ctxt o norm_term o snd) val norm_unif_thm = #norm_unif_thm norms ctxt env_res fun forall_intr binders = fold Thm.forall_intr (mk_norm_cbinders binders) o norm_unif_thm val prem_thms_inst = map (uncurry forall_intr) prem_thms |> rev val (concl_thmL, concl_thmR) = apply2 norm_unif_thm concl_thmp (*discharge the hint premises*) val thm_res = Drule.implies_elim_list lifted_hint_inst prem_thms_inst |> Drule.forall_elim_list (mk_norm_cbinders rev_binders) |> GUtil.flip Thm.transitive concl_thmR |> Thm.transitive (Unification_Base.symmetric concl_thmL) in (env_res, thm_res) end in Seq.maps unify_prems match_env_concl_thmpq |> Seq.map inst_discharge end end type hint_retrieval = term Binders.binders -> Proof.context -> term * term -> Envir.env -> unif_hint Seq.seq fun try_hints get_hints match norms unif binders ctxt (t1, t2) env = - (@{log Logger.DEBUG} ctxt (fn _ => + (@{log Logger.TRACE} ctxt (fn _ => Pretty.block [ Pretty.str "Trying unification hints for ", Unification_Util.pretty_unif_problem ctxt (t1, t2) ] |> Pretty.string_of); get_hints binders ctxt (t1, t2) env |> Seq.maps (fn hint => try_hint match norms hint unif binders ctxt (t1, t2) env)) end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic.thy b/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic.thy --- a/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic.thy +++ b/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic.thy @@ -1,70 +1,65 @@ \<^marker>\creator "Kevin Kappelmann"\ -subsection \Assumption Tactic\ theory Unify_Assumption_Tactic imports - ML_Functor_Instances + Unify_Assumption_Tactic_Base ML_Unifiers - ML_Unification_Parsers begin paragraph \Summary\ -text \Assumption tactic and method with adjustable unifier.\ - -ML_file\unify_assumption_base.ML\ -ML_file\unify_assumption.ML\ +text \Setup of assumption tactic and examples.\ ML\ @{functor_instance struct_name = Standard_Unify_Assumption and functor_name = Unify_Assumption and id = \""\ and more_args = \val init_args = { normalisers = SOME Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify, unifier = SOME Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify }\} \ local_setup \Standard_Unify_Assumption.setup_attribute NONE\ local_setup \Standard_Unify_Assumption.setup_method NONE\ paragraph \Examples\ experiment begin lemma "PROP P \ PROP P" by uassm lemma assumes h: "\P. PROP P" shows "PROP P x" using h by uassm schematic_goal "\x. PROP P (c :: 'a) \ PROP ?Y (x :: 'a)" by uassm schematic_goal a: "PROP ?P (y :: 'a) \ PROP ?P (?x :: 'a)" by uassm \\compare the result with following call\ (* by assumption *) schematic_goal "PROP ?P (x :: 'a) \ PROP P (?x :: 'a)" by uassm \\compare the result with following call\ (* by assumption *) schematic_goal "\x. PROP D \ (\P y. PROP P y x) \ PROP C \ PROP P x" by (uassm unifier = Higher_Order_Unification.unify) \\the line below is equivalent\ (* using [[uassm unifier = Higher_Order_Unification.unify]] by uassm *) text \Unlike @{method assumption}, @{method uassm} will not close the goal if the order of premises of the assumption and the goal are different. Compare the following two examples:\ lemma "\x. PROP D \ (\y. PROP A y \ PROP B x) \ PROP C \ PROP A x \ PROP B x" by uassm lemma "\x. PROP D \ (\y. PROP A y \ PROP B x) \ PROP A x \ PROP C \ PROP B x" by assumption (* by uassm *) end end diff --git a/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic_Base.thy b/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic_Base.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic_Base.thy @@ -0,0 +1,16 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Assumption Tactic\ +theory Unify_Assumption_Tactic_Base + imports + ML_Functor_Instances + ML_Tactic_Utils + ML_Unification_Parsers +begin + +paragraph \Summary\ +text \Assumption tactic and method with adjustable unifier.\ + +ML_file\unify_assumption_base.ML\ +ML_file\unify_assumption.ML\ + +end diff --git a/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic.thy b/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic.thy --- a/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic.thy +++ b/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic.thy @@ -1,47 +1,45 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Fact Tactic\ theory Unify_Fact_Tactic imports - Unify_Resolve_Tactics + Unify_Fact_Tactic_Base + ML_Unifiers begin paragraph \Summary\ -text \Fact tactic with adjustable unifier.\ - -ML_file\unify_fact_base.ML\ -ML_file\unify_fact.ML\ +text \Setup of fact tactic and examples.\ ML\ @{functor_instance struct_name = Standard_Unify_Fact and functor_name = Unify_Fact and id = \""\ and more_args = \val init_args = { normalisers = SOME Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify, unifier = SOME Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify }\} \ local_setup \Standard_Unify_Fact.setup_attribute NONE\ local_setup \Standard_Unify_Fact.setup_method NONE\ paragraph \Examples\ experiment begin lemma assumes h: "\x y. PROP P x y" shows "PROP P x y" by (ufact h) lemma assumes "\P y. PROP P y x" shows "PROP P x" by (ufact assms where unifier = Higher_Order_Unification.unify) \\the line below is equivalent\ (* using [[ufact unifier = Higher_Order_Unification.unify]] by (ufact assms) *) lemma assumes "\x y. PROP A x \ PROP B x \ PROP P x" shows "\x y. PROP A x \ PROP B x \ PROP P x" using assms by ufact end end diff --git a/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic_Base.thy b/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic_Base.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic_Base.thy @@ -0,0 +1,14 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Fact Tactic\ +theory Unify_Fact_Tactic_Base + imports + Unify_Resolve_Tactics_Base +begin + +paragraph \Summary\ +text \Fact tactic with adjustable unifier.\ + +ML_file\unify_fact_base.ML\ +ML_file\unify_fact.ML\ + +end diff --git a/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics.thy b/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics.thy --- a/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics.thy +++ b/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics.thy @@ -1,91 +1,89 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Resolution Tactics\ theory Unify_Resolve_Tactics imports - Unify_Assumption_Tactic - ML_Method_Utils + Unify_Resolve_Tactics_Base + ML_Unifiers begin paragraph \Summary\ -text \Resolution tactics and methods with adjustable unifier.\ +text \Setup of resolution tactics and examples.\ -ML_file\unify_resolve_base.ML\ -ML_file\unify_resolve.ML\ ML\ @{functor_instance struct_name = Standard_Unify_Resolve and functor_name = Unify_Resolve and id = \""\ and more_args = \val init_args = { normalisers = SOME Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify, unifier = SOME Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify, mode = SOME (Unify_Resolve_Args.PM.key Unify_Resolve_Args.PM.any), chained = SOME (Unify_Resolve_Args.PCM.key Unify_Resolve_Args.PCM.resolve) }\} \ local_setup \Standard_Unify_Resolve.setup_attribute NONE\ local_setup \Standard_Unify_Resolve.setup_method NONE\ paragraph \Examples\ experiment begin lemma assumes h: "\x. PROP D x \ PROP C x" shows "\x. PROP A x \ PROP B x \ PROP C x" apply (urule h) \\the line below is equivalent\ (* apply (rule h) *) oops lemma assumes h: "PROP C x" shows "PROP C x" by (urule h where unifier = First_Order_Unification.unify) \\the line below is equivalent\ (* using [[urule unifier = First_Order_Unification.unify]] by (urule h) *) lemma assumes h: "\x. PROP A x \ PROP D x" shows "\x. PROP A x \ PROP B x \ PROP C x" \\use (r,e,d,f) to specify the resolution mode (resolution, elim, dest, forward)\ apply (urule (d) h) \\the line below is equivalent\ (* apply (drule h) *) oops text\You can specify how chained facts should be used. By default, @{method urule} works like @{method rule}: it uses chained facts to resolve against the premises of the passed rules.\ lemma assumes h1: "\x. (PROP F x \ PROP E x) \ PROP C x" and h2: "\x. PROP F x \ PROP E x" shows "\x. PROP A x \ PROP B x \ PROP C x" \\Compare all of the following calls:\ (* apply (rule h1) *) (* apply (urule h1) *) (* using h2 apply (rule h1) *) (* using h2 apply (urule h1) *) using h2 apply (urule h1 where chained = fact) (* using h2 apply (urule h1 where chained = insert) *) done text\You can specify whether any or every rule must resolve against the goal:\ lemma assumes h1: "\x y. PROP C y \ PROP D x \ PROP C x" and h2: "\x y. PROP C x \ PROP D x" and h3: "\x y. PROP C x" shows "\x. PROP A x \ PROP B x \ PROP C x" using h3 apply (urule h1 h2 where mode = every) (* using h3 apply (urule h1 h2) *) done lemma assumes h1: "\x y. PROP C y \ PROP A x \ PROP C x" and h2: "\x y. PROP C x \ PROP B x \ PROP D x" and h3: "\x y. PROP C x" shows "\x. PROP A x \ PROP B x \ PROP C x" using h3 apply (urule (d) h1 h2 where mode = every) oops end end diff --git a/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics_Base.thy b/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics_Base.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics_Base.thy @@ -0,0 +1,16 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Resolution Tactics\ +theory Unify_Resolve_Tactics_Base + imports + Unify_Assumption_Tactic_Base + ML_Unifiers_Base + ML_Method_Utils +begin + +paragraph \Summary\ +text \Resolution tactics and methods with adjustable unifier.\ + +ML_file\unify_resolve_base.ML\ +ML_file\unify_resolve.ML\ + +end diff --git a/thys/ML_Unification/Unifiers/ML_Unifiers.thy b/thys/ML_Unification/Unifiers/ML_Unifiers.thy --- a/thys/ML_Unification/Unifiers/ML_Unifiers.thy +++ b/thys/ML_Unification/Unifiers/ML_Unifiers.thy @@ -1,73 +1,69 @@ \<^marker>\creator "Kevin Kappelmann"\ -section \ML Unifiers\ theory ML_Unifiers imports - ML_Unification_Base ML_Functor_Instances ML_Priorities + ML_Unifiers_Base Simps_To begin paragraph \Summary\ -text \Unification modulo equations and combinators for unifiers.\ +text \More unifiers.\ -paragraph \Combinators\ +paragraph \Derived Unifiers\ -ML_file\unification_combinator.ML\ +ML_file\higher_order_pattern_decomp_unification.ML\ +ML_file\var_higher_order_pattern_unification.ML\ + +paragraph \Unification via Tactics\ + +ML_file\tactic_unification.ML\ + +subparagraph \Unification via Simplification\ + +lemma eq_if_SIMPS_TO_UNIF_if_SIMPS_TO: + assumes "PROP SIMPS_TO t t'" + and "PROP SIMPS_TO_UNIF s t'" + shows "s \ t" + using assms by (simp add: SIMPS_TO_eq SIMPS_TO_UNIF_eq) + +ML_file\simplifier_unification.ML\ + +paragraph \Combining Unifiers\ ML_file\unification_combine.ML\ ML\ @{functor_instance struct_name = Standard_Unification_Combine and functor_name = Unification_Combine and id = \""\} \ local_setup \Standard_Unification_Combine.setup_attribute NONE\ - -paragraph \Standard Unifiers\ - -ML_file\higher_order_unification.ML\ -ML_file\higher_order_pattern_unification.ML\ -ML_file\first_order_unification.ML\ - -subparagraph \Derived Unifiers\ - -ML_file\higher_order_pattern_decomp_unification.ML\ -ML_file\var_higher_order_pattern_unification.ML\ - - -paragraph \Unification via Tactics\ - -ML_file\tactic_unification.ML\ - -subparagraph \Unification via Simplification\ - -ML_file\simplifier_unification.ML\ - paragraph \Mixture of Unifiers\ ML_file\mixed_unification.ML\ ML\ @{functor_instance struct_name = Standard_Mixed_Unification and functor_name = Mixed_Unification and id = \""\ and more_args = \structure UC = Standard_Unification_Combine\} \ declare [[ucombine add = \Standard_Unification_Combine.eunif_data (Var_Higher_Order_Pattern_Unification.e_unify Unification_Combinator.fail_unify |> Unification_Combinator.norm_unifier (#norm_term Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify) |> K) (Standard_Unification_Combine.metadata \<^binding>\var_hop_unif\ Prio.HIGH)\]] declare [[ucombine add = \Standard_Unification_Combine.eunif_data - (Simplifier_Unification.simp_unify - |> Unification_Combinator.norm_closed_unifier + (Simplifier_Unification.simp_unify_progress Envir.aeconv Simplifier_Unification.simp_unify (#norm_term Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify) - |> Unification_Combinator.unifier_from_closed_unifier + Standard_Mixed_Unification.norms_first_higherp_decomp_comb_higher_unify + Standard_Mixed_Unification.first_higherp_decomp_comb_higher_unify + |> Type_Unification.e_unify Unification_Util.unify_types |> K) - (Standard_Unification_Combine.default_metadata \<^binding>\simp_unif\)\]] - + (Standard_Unification_Combine.default_metadata \<^binding>\simp_unif\) + \]] end diff --git a/thys/ML_Unification/Unifiers/ML_Unifiers_Base.thy b/thys/ML_Unification/Unifiers/ML_Unifiers_Base.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/ML_Unifiers_Base.thy @@ -0,0 +1,25 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Unifiers\ +theory ML_Unifiers_Base + imports + ML_Unification_Base +begin + +paragraph \Summary\ +text \Unification modulo equations and combinators for unifiers.\ + +paragraph \Combinators\ + +ML_file\unification_combinator.ML\ + +paragraph \Type Unifiers\ + +ML_file\type_unification.ML\ + +paragraph \Standard Unifiers\ + +ML_file\higher_order_unification.ML\ +ML_file\higher_order_pattern_unification.ML\ +ML_file\first_order_unification.ML\ + +end diff --git a/thys/ML_Unification/Unifiers/higher_order_pattern_decomp_unification.ML b/thys/ML_Unification/Unifiers/higher_order_pattern_decomp_unification.ML --- a/thys/ML_Unification/Unifiers/higher_order_pattern_decomp_unification.ML +++ b/thys/ML_Unification/Unifiers/higher_order_pattern_decomp_unification.ML @@ -1,141 +1,140 @@ (* Title: ML_Unification/higher_order_pattern_decomp_unification.ML Author: Kevin Kappelmann Tries to decompose terms that are outside the higher-order pattern fragment into a higher-order pattern prefix and a list of remaining arguments. Example: Let 0, 1 be bound variables and f, g, c be constants. The unification problem ?x 0 1 (?y c 0) \\<^sup>? f (g c 0) is outside the HO-pattern fragment, but it can be decomposed into ?x 0 1 \\<^sup>? f and (?y c 0) \\<^sup>? (g c 0). The latter, in turn, can be decomposed into ?y \\<^sup>? g and c \\<^sup>? c and 0 \\<^sup>? 0. *) signature HIGHER_ORDER_PATTERN_DECOMP_UNIFICATION = sig include HAS_LOGGER val e_match : Unification_Base.matcher -> Unification_Base.e_matcher val e_unify : Unification_Base.unifier -> Unification_Base.e_unifier end structure Higher_Order_Pattern_Decomp_Unification : HIGHER_ORDER_PATTERN_DECOMP_UNIFICATION = struct val logger = Logger.setup_new_logger Unification_Base.logger "Higher_Order_Pattern_Decomp_Unification" structure UUtil = Unification_Util -structure HOP = Higher_Order_Pattern_Unification structure UC = Unification_Combinator exception FALLBACK of Pretty.T (*check if sequence is empty or raise FALLBACK exception*) fun seq_try msg sq = General_Util.seq_try (FALLBACK msg) sq (*compute pattern prefixes*) fun bounds_prefix bounds [] = (rev bounds, []) | bounds_prefix bounds (t :: ts) = if is_Bound t andalso not (member (op =) bounds t) then bounds_prefix (t :: bounds) ts else (rev bounds, t :: ts) val bounds_prefix = bounds_prefix [] datatype ('a, 'b) either = Left of 'a | Right of 'b fun bounds_prefix2 (bounds1, _) ([], _) = Left (rev bounds1, []) | bounds_prefix2 (_, bounds2) (_, []) = Right (rev bounds2, []) | bounds_prefix2 (bounds1, bounds2) (s :: ss, t :: ts) = if is_Bound s andalso not (member (op =) bounds1 s) then if is_Bound t andalso not (member (op =) bounds2 t) then bounds_prefix2 (s :: bounds1, t :: bounds2) (ss, ts) else Right (rev bounds2, t :: ts) else Left (rev bounds1, s :: ss) val bounds_prefix2 = bounds_prefix2 ([], []) fun mk_decomp _ (_, []) = NONE (*no decomposition took place*) | mk_decomp h (args, rem) = SOME (list_comb (h, args), rem) fun mk_bounds_decomp h args = bounds_prefix args |> mk_decomp h fun pattern_prefix_match p = case strip_comb p of (v as Var _, args) => mk_bounds_decomp v args | _ => NONE fun pattern_prefix_unif tp = case apply2 strip_comb tp of ((vl as Var _, argsl), (vr as Var _, argsr)) => (case bounds_prefix2 (argsl, argsr) of Left res => mk_decomp vl res |> Option.map Left | Right res => mk_decomp vr res |> Option.map Right) | ((v as Var _, args), _) => mk_bounds_decomp v args |> Option.map Left | (_, (v as Var _, args)) => mk_bounds_decomp v args |> Option.map Right | _ => NONE fun mk_decomp_other ss t = let val (th, ts) = strip_comb t val n = length ts - length ss in if n < 0 then NONE else chop n ts |> mk_decomp th end fun decompose_msg ctxt tp = Pretty.block [ Pretty.str "Decomposing ", UUtil.pretty_unif_problem ctxt tp, Pretty.str " into higher-order pattern prefix and list of remaining arguments." ] |> Pretty.string_of fun decomposed_msg ctxt (ph, th) (ps, ts) = Pretty.block [ Pretty.str "Decomposition result ", map (UUtil.pretty_unif_problem ctxt) ((ph, th) :: (ps ~~ ts)) |> Pretty.separate "," |> Pretty.block ] |> Pretty.string_of val failed_decompose_msg = Pretty.str "Decomposition failed." fun e_match match_pattern match_theory binders ctxt (pt as (p, t)) env = let val failed_sub_patterns_msg = Pretty.str "Matching of sub-patterns failed." in (@{log Logger.DEBUG} ctxt (fn _ => decompose_msg ctxt pt); case pattern_prefix_match p of SOME (ph, ps) => (case mk_decomp_other ps t of SOME (th, ts) => (@{log Logger.TRACE} ctxt (fn _ => decomposed_msg ctxt (ph, th) (ps, ts)); UUtil.strip_comb_strip_comb (K o K I) match_pattern binders ctxt (ph, th) (ps, ts) env |> seq_try failed_sub_patterns_msg) | NONE => raise FALLBACK failed_decompose_msg) | NONE => raise FALLBACK failed_decompose_msg) handle FALLBACK msg => (@{log Logger.DEBUG} ctxt (fn _ => Pretty.breaks [ msg, UUtil.pretty_call_theory_match ctxt pt ] |> Pretty.block |> Pretty.string_of); match_theory binders ctxt pt env) end fun e_unify unif_pattern unif_theory binders ctxt (tp as (s, t)) env = let val unif_patterns = UUtil.strip_comb_strip_comb Envir_Normalisation.norm_thm_types_unif unif_pattern val failed_sub_patterns_msg = Pretty.str "Unification of sub-patterns failed." in (@{log Logger.DEBUG} ctxt (fn _ => decompose_msg ctxt tp); case pattern_prefix_unif tp of SOME (Left (sh, ss)) => (case mk_decomp_other ss t of SOME (th, ts) => (@{log Logger.TRACE} ctxt (fn _ => decomposed_msg ctxt (sh, th) (ss, ts)); unif_patterns binders ctxt (sh, th) (ss, ts) env |> seq_try failed_sub_patterns_msg) | NONE => raise FALLBACK failed_decompose_msg) | SOME (Right (th, ts)) => (case mk_decomp_other ts s of SOME (sh, ss) => (@{log Logger.TRACE} ctxt (fn _ => decomposed_msg ctxt (sh, th) (ss, ts)); unif_patterns binders ctxt (sh, th) (ss, ts) env |> seq_try failed_sub_patterns_msg) | NONE => raise FALLBACK failed_decompose_msg) | NONE => raise FALLBACK failed_decompose_msg) handle FALLBACK msg => (@{log Logger.DEBUG} ctxt (fn _ => Pretty.breaks [ msg, UUtil.pretty_call_theory_unif ctxt tp ] |> Pretty.block |> Pretty.string_of); unif_theory binders ctxt tp env) end end diff --git a/thys/ML_Unification/Unifiers/higher_order_pattern_unification.ML b/thys/ML_Unification/Unifiers/higher_order_pattern_unification.ML --- a/thys/ML_Unification/Unifiers/higher_order_pattern_unification.ML +++ b/thys/ML_Unification/Unifiers/higher_order_pattern_unification.ML @@ -1,476 +1,474 @@ (* Title: ML_Unification/higher_order_pattern_unification.ML Author: Tobias Nipkow, Christine Heinzelmann, Stefan Berghofer, and Kevin Kappelmann TU Muenchen Higher-Order Patterns due to Tobias Nipkow. E-Unification and theorem construction due to Kevin Kappelmann. See also: Tobias Nipkow. Functional Unification of Higher-Order Patterns. In Proceedings of the 8th IEEE Symposium Logic in Computer Science, 1993. TODO: optimize red by special-casing it *) signature HIGHER_ORDER_PATTERN_UNIFICATION = sig include HAS_LOGGER (* e-matching *) + (*precondition: types of terms are matched*) val e_match : Unification_Base.type_matcher -> Unification_Base.matcher -> Unification_Base.e_matcher val match : Unification_Base.matcher val norms_match : Unification_Base.normalisers (* e-unification *) + (*precondition: types of terms are unified*) val e_unify : Unification_Base.type_unifier -> Unification_Base.unifier -> Unification_Base.e_unifier val unify : Unification_Base.unifier val norms_unify : Unification_Base.normalisers end structure Higher_Order_Pattern_Unification : HIGHER_ORDER_PATTERN_UNIFICATION = struct val logger = Logger.setup_new_logger Unification_Base.logger "Higher_Order_Pattern_Unification" (* shared utils *) structure Util = Unification_Util structure Norm = Envir_Normalisation exception FALLBACK (*check if sequence is empty or raise FALLBACK exception*) fun seq_try sq = General_Util.seq_try FALLBACK sq (*predicate: downto0 (is, n) <=> is = [n, n - 1, ..., 0]*) fun downto0 ([], n) = n = ~1 | downto0 (i :: is, n) = i = n andalso downto0 (is, n - 1) fun mkabs (binders, is, t) = let fun abstract i acc = let val ((x, T), _) = nth binders i in Abs (x, T, acc) end in fold_rev abstract is t end exception IDX fun idx [] _ = raise IDX | idx (i :: is) j = if i = j then length is else idx is j fun ints_of bs = (*collects arguments and checks if they are all distinct, bound variables*) let fun dest_check (Bound i) acc = if member (op =) acc i then raise Unification_Base.PATTERN else i :: acc | dest_check _ _ = raise Unification_Base.PATTERN in fold_rev dest_check bs [] end fun app (s, []) = s | app (s, (i :: is)) = app (s $ Bound i, is) (* matching *) fun mapbnd f = let fun mpb d (Bound i) = if i < d then Bound i else Bound (f (i - d) + d) | mpb d (Abs (s, T, t)) = Abs (s, T, mpb (d + 1) t) | mpb d (u1 $ u2) = (mpb d u1) $ (mpb d u2) | mpb _ atom = atom in mpb 0 end fun red (Abs (_, _, s)) (i :: is) js = red s is (i :: js) | red t [] [] = t | red t is jn = app (mapbnd (nth jn) t, is) exception OPEN_TERM fun match_bind (tenv, binders, ixn, T, is, t) = let val js = loose_bnos t in if null is then if null js then Vartab.update_new (ixn, (T, t)) tenv else raise OPEN_TERM else if subset (op =) (js, is) then let val t' = if downto0 (is, length binders - 1) then t else mapbnd (idx is) t in Vartab.update_new (ixn, (T, mkabs (binders, is, t'))) tenv end else raise OPEN_TERM end (** higher-order pattern E-unification **) val norms_match = Util.beta_eta_short_norms_match -fun e_match match_types match_theory_pattern match_theory_fail binders ctxt (pt as (p, t)) env = +fun e_match match_types match_theory_pattern match_theory_fail binders ctxt pt env = let val norm_pt = apfst o (#norm_term norms_match) fun rigid_rigid ctxt p n env = (*normalise the types in rigid-rigid cases*) Util.rigid_rigid Norm.norm_term_types_match match_types ctxt p n env |> seq_try handle Unification_Base.UNIF => Seq.empty (*types do not match*) (*generated theorem is already normalised wrt. the resulting environment*) fun match binders ctxt (pt as (p, t)) (env as Envir.Envir {maxidx, tenv, tyenv}) = let (*call match_theory on failure*) fun handle_failure match_theory failure_msg = (@{log Logger.DEBUG} ctxt (fn _ => Pretty.breaks [ failure_msg (), Util.pretty_call_theory_match ctxt (norm_pt env pt) ] |> Pretty.block |> Pretty.string_of); match_theory binders ctxt pt env) in (case pt of (Abs (np, _, tp), Abs (nt, Tt, tt)) => let val name = if np = "" then nt else np in Util.abstract_abstract (K I) match binders ctxt name Tt (tp, tt) env |> seq_try end (*eta-expand on the fly*) | (Abs (np, Tp, tp), _) => let val Tp' = Envir.subst_type tyenv Tp in Util.abstract_abstract (K I) match binders ctxt np Tp' (tp, incr_boundvars 1 t $ Bound 0) env |> seq_try end | (_, Abs (nt, Tt, tt)) => Util.abstract_abstract (K I) match binders ctxt nt Tt (incr_boundvars 1 p $ Bound 0, tt) env |> seq_try | (Bound i, Bound j) => Util.bound_bound binders ctxt (i, j) |> Seq.map (pair env) | (Free _, Free g) => rigid_rigid ctxt p g env | (Const _, Const d) => rigid_rigid ctxt p d env | _ => case strip_comb p of (Var (x, T), ps) => let val is = ints_of ps val T' = Norm.norm_type_match tyenv T in case Envir.lookup1 tenv (x, T') of NONE => ((let val tenv' = match_bind (tenv, binders, x, T', is, t) val env' = Envir.Envir {maxidx=maxidx, tenv=tenv', tyenv=tyenv} val t' = Binders.replace_binders binders t in Seq.single (env', Unification_Base.reflexive_term ctxt t') end) handle OPEN_TERM => (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "Failed to unify (open term) ", Util.pretty_unif_problem ctxt (norm_pt env pt) ] |> Pretty.string_of); Seq.empty)) | SOME ph' => if Envir.aeconv (red ph' is [], t) then Seq.single (env, Unification_Base.reflexive_term ctxt (Binders.replace_binders binders t)) else raise FALLBACK end | (ph, ps) => let val (th, ts) = strip_comb t in case (ph, th) of (Abs _, _) => raise Unification_Base.PATTERN | (_, Abs _) => raise Unification_Base.PATTERN | _ => if null ps orelse null ts then raise FALLBACK else (*Note: types of theorems are already normalised ==> we pass the identity type normaliser*) Util.strip_comb_strip_comb (K o K I) match binders ctxt (ph, th) (ps, ts) env |> seq_try end) handle FALLBACK => handle_failure match_theory_fail (fn _ => Pretty.str "Higher-order pattern matching failed.") | Unification_Base.PATTERN => handle_failure match_theory_pattern (fn _ => Pretty.str "Terms not in Pattern fragment.") end - val binder_types = Binders.binder_types binders in (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "Higher-order pattern matching ", Util.pretty_unif_problem ctxt (norm_pt env pt) ] |> Pretty.string_of); - match_types ctxt (Envir.fastype env binder_types p, fastype_of1 (binder_types, t)) env - |> match binders ctxt pt) + match binders ctxt pt env) handle Unification_Base.UNIF => Seq.empty (*types do not match*) end - (*standard higher-order pattern matching*) val match = e_match Util.match_types Unification_Combinator.fail_match Unification_Combinator.fail_match (*unification*) fun string_of_term ctxt env binders t = (map (Free o fst) binders, t) |> subst_bounds |> Norm.norm_term_unif env |> Syntax.string_of_term ctxt fun bname binders = fst o fst o nth binders fun bnames binders is = map (bname binders) is |> space_implode " " fun proj_fail ctxt (env, binders, F, _, is, t) reason = @{log Logger.DEBUG} ctxt (fn _ => let val f = Term.string_of_vname F val xs = bnames binders is val u = string_of_term ctxt env binders t in cat_lines [ "Cannot unify variable " ^ f ^ " (depending on bound variables [" ^ xs ^ "])", "with term " ^ u, reason () ] end) fun proj_fail_unif ctxt (params as (_, binders, _, _, is, t)) = let fun reason () = let val ys = bnames binders (subtract (op =) is (loose_bnos t)) in "Term contains additional bound variable(s) " ^ ys end in proj_fail ctxt params reason end fun proj_fail_type_app ctxt (params as (env, binders, _, _, _, _)) var_app = let fun reason () = let val var_app = string_of_term ctxt env binders var_app in "Projection " ^ var_app ^ " is not well-typed." end in proj_fail ctxt params reason end fun ocheck_fail ctxt (F, t, binders, env) = @{log Logger.DEBUG} ctxt (fn _ => cat_lines [ "Variable " ^ Term.string_of_vname F ^ " occurs in term", string_of_term ctxt env binders t, "Cannot unify!" ]) fun occurs (F, t, env) = let fun occ (Var (G, T)) = (case Envir.lookup env (G, T) of SOME t => occ t | NONE => F = G) | occ (t1 $ t2) = occ t1 orelse occ t2 | occ (Abs (_, _, t)) = occ t | occ _ = false in occ t end fun ints_of' env ts = ints_of (map (Envir.head_norm env) ts) (*split_type ([T1,....,Tn]---> T) n = ([Tn,...,T1], T)*) fun split_type t n = let fun split (T, 0, Ts) = (Ts, T) | split (Type ("fun", [T1, T2]), n, Ts) = split (T2, n - 1, T1 :: Ts) | split _ = raise Fail "split_type" in split (t, n, []) end fun type_of_G (Envir.Envir {tyenv,...}) (T, n, is) = let val (Ts, U) = split_type (Norm.norm_type_unif tyenv T) n in map (nth Ts) is ---> U end fun mk_hnf (binders, is, G, js) = mkabs (binders, is, app (G, js)) fun mk_new_hnf (env, binders, is, F as (a, _), T, js) = let val (env', G) = Envir.genvar a (env, type_of_G env (T, length is, js)) in Envir.update ((F, T), mk_hnf (binders, is, G, js)) env' end (*mk_proj_list is = [ |is| - k - 1 | 0 <= k < |is| and is[k] >= 0 ]*) fun mk_proj_list is = let fun mk (SOME _) (acc, j) = (j :: acc, j + 1) | mk NONE (acc, j) = (acc, j + 1) in fold_rev mk is ([], 0) |> fst end fun proj ctxt (s, env, binders, is) = let fun trans d i = if i < d then i else idx is (i - d) + d fun pr (s, env, d, binders) = (case Envir.head_norm env s of Abs (a, T, t) => let val (binder, _) = Binders.fix_binder (a, T) ctxt val (t', env') = pr (t, env, d + 1, binder :: binders) in (Abs (a, T, t'), env') end | t => (case strip_comb t of (c as Const _, ts) => let val (ts', env') = prs (ts, env, d, binders) in (list_comb (c, ts'), env') end | (f as Free _, ts) => let val (ts', env') = prs (ts, env, d, binders) in (list_comb (f, ts'), env') end | (Bound i, ts) => let val j = trans d i val (ts', env') = prs (ts, env, d, binders) in (list_comb (Bound j, ts'), env') end | (Var (F as (a, _), Fty), ts) => let val js = ints_of' env ts val js' = map (try (trans d)) js val ks = mk_proj_list js' val ls = map_filter I js' val Hty = type_of_G env (Fty, length js, ks) val (env', H) = Envir.genvar a (env, Hty) val env'' = Envir.update ((F, Fty), mk_hnf (binders, js, H, ks)) env' in (app (H, ls), env'') end | _ => raise Unification_Base.PATTERN)) and prs (s :: ss, env, d, binders) = let val (s', env1) = pr (s, env, d, binders) val (ss', env2) = prs (ss, env1, d, binders) in (s' :: ss', env2) end | prs ([], env, _, _) = ([], env) in if downto0 (is, length binders - 1) then (s, env) else pr (s, env, 0, binders) end (*mk_ff_list (is, js) = [ |is| - k - 1 | 0 <= k < |is| and is[k] = js[k] ]*) fun mk_ff_list (is,js) = let fun mk ([], [], _) = [] | mk (i :: is, j :: js, k) = if i = j then k :: mk (is, js, k - 1) else mk (is, js, k - 1) | mk _ = raise Fail "mk_ff_list" in mk (is, js, length is - 1) end; fun app_free (Envir.Envir {tyenv,...}) binders n T is = let val norm_type = Norm.norm_type_unif tyenv in list_comb (Var (n, norm_type T), map (map_types norm_type o Binders.nth_binder_data binders) is) end fun flexflex1 unify_types ctxt (env, binders, F, Fty, Fty', is, js) = let val env' = unify_types ctxt (Fty, Fty') env val thm = app_free env binders F Fty is |> Unification_Base.reflexive_term ctxt in if is = js then (env', thm) else let val ks = mk_ff_list (is, js) val env'' = mk_new_hnf (env', binders, is, F, Fty, ks) in (env'', thm) end end fun flexflex2 unify_types ctxt (env, binders, F, Fty, is, G, Gty, js) = let val var_app = app_free env binders F Fty is val binder_types = Binders.binder_types binders val env' = unify_types ctxt (apply2 (Envir.fastype env binder_types) (var_app, app (Var (G, Gty), js))) env val thm = Unification_Base.reflexive_term ctxt var_app fun ff (F, Fty, is, G as (a, _), Gty, js) = if subset (op =) (js, is) then let val t = mkabs (binders, is, app (Var (G, Gty), map (idx is) js)) val env'' = Envir.update ((F, Fty), t) env' in (env'', thm) end else let val ks = inter (op =) js is val Hty = type_of_G env' (Fty, length is, map (idx is) ks) val (env'', H) = Envir.genvar a (env', Hty) fun lam is = mkabs (binders, is, app (H, map (idx is) ks)) val env''' = Envir.update ((F, Fty), lam is) env'' |> Envir.update ((G, Gty), lam js) in (env''', thm) end in if is_less (Term_Ord.indexname_ord (G, F)) then ff (F, Fty, is, G, Gty, js) else ff (G, Gty, js, F, Fty, is) end fun flexrigid unify_types ctxt (params as (env, binders, F, Fty, is, t)) = if occurs (F, t, env) then (ocheck_fail ctxt (F, t, binders, env); raise FALLBACK) else let val var_app = app_free env binders F Fty is in let val binder_types = Binders.binder_types binders val env' = unify_types ctxt (apply2 (Envir.fastype env binder_types) (var_app, t)) env val (u, env'') = proj ctxt (t, env', binders, is) val env''' = Envir.update ((F, Fty), mkabs (binders, is, u)) env'' val thm = Unification_Base.reflexive_term ctxt var_app in (env''', thm) end handle IDX => (proj_fail_unif ctxt params; raise OPEN_TERM) | TYPE _ => (proj_fail_type_app ctxt params var_app; raise Unification_Base.UNIF) end (** higher-order pattern E-unification **) val norms_unify = Util.beta_eta_short_norms_unif fun e_unify unify_types unify_theory_pattern unify_theory_fail binders ctxt st env = let fun unif binders ctxt st env = let val (st' as (s', t')) = apply2 (Envir.head_norm env) st fun rigid_rigid ctxt t n env = (*we do not normalise types in base cases*) Util.rigid_rigid (K I) unify_types ctxt t n env |> seq_try handle Unification_Base.UNIF => Seq.empty (*types do not match*) (*call unify_theory on failure*) fun handle_failure unify_theory failure_msg = (@{log Logger.DEBUG} ctxt (fn _ => Pretty.breaks [ failure_msg (), Util.pretty_call_theory_unif ctxt st' ] |> Pretty.block |> Pretty.string_of); (*Note: st' is already normalised*) unify_theory binders ctxt st' env) in (case st' of (Abs (ns, Ts, ts), Abs (nt, Tt, tt)) => ((let val name = if ns = "" then nt else ns in unify_types ctxt (Ts, Tt) env |> Util.abstract_abstract Norm.norm_term_types_unif unif binders ctxt name Ts (ts, tt) |> seq_try end) handle Unification_Base.UNIF => Seq.empty) (*types do not unify*) (*eta-expand on the fly*) | (Abs (ns, Ts, ts), _) => ((let val st = apply2 (Envir.fastype env (Binders.binder_types binders)) st' in unify_types ctxt st env |> Util.abstract_abstract Norm.norm_term_types_unif unif binders ctxt ns Ts (ts, incr_boundvars 1 t' $ Bound 0) |> seq_try end) handle Unification_Base.UNIF => Seq.empty) (*types do not unify*) | (_, Abs _) => unif binders ctxt (t', s') env |> Seq.map (apsnd Unification_Base.symmetric) | (Bound i, Bound j) => Util.bound_bound binders ctxt (i, j) |> Seq.map (pair env) | (Free _, Free g) => rigid_rigid ctxt s' g env | (Const _, Const d) => rigid_rigid ctxt s' d env (*case distinctions on head term*) | _ => case apply2 strip_comb st' of ((Var (F, Fty), ss), (Var (G, Gty), ts)) => (((if F = G then flexflex1 unify_types ctxt (env, binders, F, Fty, Gty, ints_of' env ss, ints_of' env ts) else flexflex2 unify_types ctxt (env, binders, F, Fty, ints_of' env ss, G, Gty, ints_of' env ts)) |> Seq.single) handle Unification_Base.UNIF => Seq.empty) (*types do not unify*) | ((Var (F, Fty), ss), _) => (flexrigid unify_types ctxt (env, binders, F, Fty, ints_of' env ss, t')|> Seq.single handle OPEN_TERM => Seq.empty | Unification_Base.UNIF => Seq.empty) (*types do not unify*) | (_, (Var (F, Fty), ts)) => (flexrigid unify_types ctxt (env, binders, F, Fty, ints_of' env ts, s') |> Seq.single handle OPEN_TERM => Seq.empty | Unification_Base.UNIF => Seq.empty) (*types do not unify*) | ((sh, ss), (th, ts)) => if null ss orelse null ts then raise FALLBACK else (*but we have to normalise in comb cases*) Util.strip_comb_strip_comb Norm.norm_thm_types_unif unif binders ctxt (sh, th) (ss, ts) env |> seq_try) handle FALLBACK => handle_failure unify_theory_fail (fn _ => Pretty.str "Higher-order pattern unification failed.") | Unification_Base.PATTERN => handle_failure unify_theory_pattern (fn _ => Pretty.str "Terms not in Pattern fragment.") end - val binder_types = Binders.binder_types binders in (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "Higher-order pattern unifying ", Util.pretty_unif_problem ctxt (apply2 (Norm.norm_term_unif env) st) ] |> Pretty.string_of); - unif binders ctxt st (unify_types ctxt (apply2 (Envir.fastype env binder_types) st) env)) + unif binders ctxt st env) handle Unification_Base.UNIF => Seq.empty (*types do not unify*) end (*standard higher-order pattern unification*) val unify = e_unify Util.unify_types Unification_Combinator.fail_unify Unification_Combinator.fail_unify end diff --git a/thys/ML_Unification/Unifiers/higher_order_unification.ML b/thys/ML_Unification/Unifiers/higher_order_unification.ML --- a/thys/ML_Unification/Unifiers/higher_order_unification.ML +++ b/thys/ML_Unification/Unifiers/higher_order_unification.ML @@ -1,60 +1,53 @@ (* Title: ML_Unification/higher_order_unification.ML Author: Kevin Kappelmann TU Muenchen Higher-order unification from the Pure kernel adjusted to fit the Unification_Base.unifier type, i.e. moving flex-flex pairs into the theorem and support of open term unification. *) signature HIGHER_ORDER_UNIFICATION = sig include HAS_LOGGER val unify : Unification_Base.unifier val norms : Unification_Base.normalisers end structure Higher_Order_Unification : HIGHER_ORDER_UNIFICATION = struct val logger = Logger.setup_new_logger Unification_Base.logger "Higher_Order_Unification" structure Util = Unification_Util -fun unify binders ctxt tp env = - let - fun unif env = - let - val init_goal = Logic.mk_equals #> Thm.cterm_of ctxt #> Goal.init - fun create_env_thmq env = - (*replace binders and apply the computed unifier*) - apply2 (Envir_Normalisation.beta_eta_short_norm_term_unif env o Binders.replace_binders binders) tp - (*create a goal such that we can create a theorem including the flex-flex pairs*) - |> init_goal - (*kind of a hack: resolving against the equality theorem will add the - flex-flex pairs to the theorem*) - |> HEADGOAL (match_tac ctxt @{thms Pure.reflexive}) - (*conclude the goal and pair with the environment*) - |> Seq.map (pair env o Goal.conclude) - in - (*find the unifiers*) - Unify.hounifiers (map fst binders) (Context.Proof ctxt, env, [tp]) - (*add the flex-flex pairs to the theorem*) - |> Seq.maps (create_env_thmq o fst) - end - in - ((@{log Logger.DEBUG} ctxt (fn _ => - Pretty.block [ - Pretty.str "Higher-order unifying ", - Util.pretty_unif_problem ctxt (apply2 (Envir.norm_term env) tp) - ] - |> Pretty.string_of); - Util.unify_types ctxt (apply2 (Envir.fastype env (Binders.binder_types binders)) tp) env - |> unif) - handle Unification_Base.UNIF => Seq.empty) (*types do not match*) - end +val unify = + let fun unif binders ctxt tp env = + let + val _ = @{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Higher-order unifying ", + Util.pretty_unif_problem ctxt (apply2 (Envir.norm_term env) tp) + ] |> Pretty.string_of) + val init_goal = Logic.mk_equals #> Thm.cterm_of ctxt #> Goal.init + fun create_env_thmq env = + (*replace binders and apply the computed unifier*) + apply2 (Envir_Normalisation.beta_eta_short_norm_term_unif env o Binders.replace_binders binders) tp + (*create a goal such that we can create a theorem including the flex-flex pairs*) + |> init_goal + (*kind of a hack: resolving against the equality theorem will add the + flex-flex pairs to the theorem*) + |> HEADGOAL (match_tac ctxt @{thms Pure.reflexive}) + (*conclude the goal and pair with the environment*) + |> Seq.map (pair env o Goal.conclude) + in + (*find the unifiers*) + Unify.hounifiers (map fst binders) (Context.Proof ctxt, env, [tp]) + (*add the flex-flex pairs to the theorem*) + |> Seq.maps (create_env_thmq o fst) + end + in Type_Unification.e_unify Util.unify_types unif end val norms = Unification_Util.beta_eta_short_norms_unif end diff --git a/thys/ML_Unification/Unifiers/mixed_unification.ML b/thys/ML_Unification/Unifiers/mixed_unification.ML --- a/thys/ML_Unification/Unifiers/mixed_unification.ML +++ b/thys/ML_Unification/Unifiers/mixed_unification.ML @@ -1,55 +1,55 @@ (* Title: ML_Unification/mixed_unification.ML Author: Kevin Kappelmann Mixture of unification algorithms. *) signature MIXED_UNIFICATION = sig include HAS_LOGGER structure UC : UNIFICATION_COMBINE (*first-order, then higher-order pattern with decomposition, then UC.e_unify, then higher-order unification fallback*) val first_higherp_decomp_comb_higher_unify : Unification_Base.unifier val norms_first_higherp_decomp_comb_higher_unify : Unification_Base.normalisers end functor Mixed_Unification (A : sig structure FIA : FUNCTOR_INSTANCE_ARGS structure UC : UNIFICATION_COMBINE end) : MIXED_UNIFICATION = struct val logger = Logger.setup_new_logger Unification_Base.logger A.FIA.full_name structure UUtil = Unification_Util structure UCO = Unification_Combinator structure UC = A.UC fun first_higherp_decomp_comb_higher_unify binders ctxt tp env = let val unify_types = UUtil.unify_types val comb_higher = UCO.add_fallback_unifier UC.e_unify Higher_Order_Unification.unify val decomp_comb_higher = UCO.add_fallback_unifier (Higher_Order_Pattern_Decomp_Unification.e_unify first_higherp_decomp_comb_higher_unify) comb_higher val higherp_decomp_comb_higher = UCO.add_fallback_unifier - (Higher_Order_Pattern_Unification.e_unify unify_types decomp_comb_higher) - comb_higher + (Higher_Order_Pattern_Unification.e_unify unify_types decomp_comb_higher) + comb_higher + |> Type_Unification.e_unify unify_types val fo_higherp_decomp_comb_higher = UCO.add_fallback_unifier - (First_Order_Unification.e_unify unify_types) - higherp_decomp_comb_higher + (First_Order_Unification.e_unify unify_types) higherp_decomp_comb_higher in (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "First-order with higher-order pattern with decomposition with ", Binding.pretty UC.binding, Pretty.str " with higher-order fallback unifying ", UUtil.pretty_unif_problem ctxt (apply2 (Envir_Normalisation.norm_term_unif env) tp)] |> Pretty.string_of); fo_higherp_decomp_comb_higher binders ctxt tp env) end val norms_first_higherp_decomp_comb_higher_unify = UUtil.beta_eta_short_norms_unif end diff --git a/thys/ML_Unification/Unifiers/simplifier_unification.ML b/thys/ML_Unification/Unifiers/simplifier_unification.ML --- a/thys/ML_Unification/Unifiers/simplifier_unification.ML +++ b/thys/ML_Unification/Unifiers/simplifier_unification.ML @@ -1,75 +1,107 @@ (* Title: ML_Unification/simplifier_unification.ML Author: Kevin Kappelmann Solving equations for unification problems with the simplifier. *) signature SIMPLIFIER_UNIFICATION = sig include HAS_LOGGER - (*solves "s \\<^sup>? t" via simplification*) - val simp_unify : Unification_Base.closed_unifier - (*solves "SIMPS_TO s t \ rhs" via simplification of s when given a matching theorem + (*solves "SIMPS_TO s t \ rhs" via simplification of s when given a theorem "SIMPS_TO s t \ SIMPS_TO s t \ rhs"*) val SIMPS_TO_unify : thm -> Unification_Base.closed_unifier (*solves "SIMPS_TO_UNIF s t \ rhs" via simplification of s to s', followed by unification of - "s' \\<^sup>? t", when given a matching theorem "SIMPS_TO_UNIF s t \ SIMPS_TO_UNIF s t \ rhs"*) - val SIMPS_TO_UNIF_unify : thm -> Unification_Base.normalisers -> Unification_Base.unifier -> + "s' \\<^sup>? t", when given a theorem "SIMPS_TO_UNIF s t \ SIMPS_TO_UNIF s t \ rhs"*) + val SIMPS_TO_UNIF_unify : thm -> Unification_Base.normalisers -> + Unification_Base.closed_unifier -> Unification_Base.closed_unifier + (*solves "s \\<^sup>? t" via simplification followed by unification*) + val simp_unify : Unification_Base.normalisers -> Unification_Base.closed_unifier -> Unification_Base.closed_unifier + (*solves "s \\<^sup>? t" via simplification followed by unification; aborts if no progress was made + during simplification*) + val simp_unify_progress : (term * term -> bool) -> (Unification_Base.normalisers -> + Unification_Base.closed_unifier -> Unification_Base.closed_unifier) -> + Envir_Normalisation.term_normaliser -> Unification_Base.normalisers -> + Unification_Base.e_unifier end structure Simplifier_Unification : SIMPLIFIER_UNIFICATION = struct val logger = Logger.setup_new_logger Unification_Base.logger "Simplifier_Unification" -(*some "safe" solvers create instantiations via flex-flex pairs, which we disallow*) -val safe_simp_tac = Tactic_Util.safe_simp_tac +structure Util = Tactic_Util -fun simp_unify ctxt tp = - (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ - Pretty.str "Solving unification problem via simplification ", - Unification_Util.pretty_unif_problem ctxt tp - ] |> Pretty.string_of); - Tactic_Unification.unify (Tactic_Unification.env_tac_from_no_inst_tac (safe_simp_tac ctxt) |> K) - ctxt tp) +(*some "safe" solvers create instantiations via flex-flex pairs, which we disallow*) +val safe_simp_tac = Util.safe_simp_tac -fun preprocess_tac ctxt = match_tac ctxt o single +fun match_tac ctxt = Tactic.match_tac ctxt o single fun SIMPS_TO_unify preprocess_rule ctxt = let fun tac (tp as (lhs, _)) = if can Simps_To.dest_SIMPS_TO lhs then (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "Solving ", Syntax.pretty_term ctxt @{term SIMPS_TO}, Pretty.str " unification problem ", Unification_Util.pretty_unif_problem ctxt tp ] |> Pretty.string_of); - preprocess_tac ctxt preprocess_rule - THEN' safe_simp_tac ctxt - THEN' Simps_To.finish_SIMPS_TO_tac ctxt) + match_tac ctxt preprocess_rule + THEN' Simps_To.SIMPS_TO_tac (safe_simp_tac ctxt) ctxt) else K no_tac in Tactic_Unification.unify (Tactic_Unification.env_tac_from_no_inst_tac o tac) ctxt end +fun SIMPS_TO_UNIF_env_thm_tac norms unif ctxt env tSIMPS_TO_UNIF state = + let + val tp = Simps_To_Unif.dest_SIMPS_TO_UNIF tSIMPS_TO_UNIF + fun norm_resolve (env, thm) = #norm_thm norms ctxt env state |> Thm.elim_implies thm |> pair env + in + Simps_To_Unif.SIMPS_TO_UNIF_env_thmsq (safe_simp_tac ctxt) norms unif ctxt tp env + |> Seq.map norm_resolve + end + fun SIMPS_TO_UNIF_unify preprocess_rule norms unif ctxt = - let fun tac ctxt (tp as (lhs, _)) i (env, state) = - (let - val simps_to_tp = Simps_To_Unif.dest_SIMPS_TO_UNIF lhs - fun solve_tac state = - Simps_To_Unif.SIMPS_TO_UNIF_env_thmsq (safe_simp_tac ctxt) norms unif [] ctxt - simps_to_tp env - |> Seq.map (fn (env, thm) => (env, #norm_thm norms ctxt env state |> Thm.elim_implies thm)) - in + let fun tac (tp as (lhs, _)) i (env, state) = + if can Simps_To_Unif.dest_SIMPS_TO_UNIF lhs + then (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "Solving ", Syntax.pretty_term ctxt @{term SIMPS_TO_UNIF}, Pretty.str " unification problem ", Unification_Util.pretty_unif_problem ctxt tp ] |> Pretty.string_of); - preprocess_tac ctxt preprocess_rule i state - |> Seq.maps solve_tac) - end) - handle TERM _ => Seq.empty - in Tactic_Unification.unify (tac ctxt) ctxt end + match_tac ctxt preprocess_rule i state + |> Seq.maps (SIMPS_TO_UNIF_env_thm_tac norms unif ctxt env lhs)) + else Seq.empty + in Tactic_Unification.unify tac ctxt end + +fun simp_unify norms unif ctxt = + let + val simp_tac = safe_simp_tac ctxt + fun SIMPS_TO_UNIF_env_thm_tac' env i state = SIMPS_TO_UNIF_env_thm_tac norms unif ctxt env + (Thm.cprem_of state i |> Thm.term_of) state + fun eq_tac tp i (env, state) = + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Solving unification problem via simplification followed by unification ", + Unification_Util.pretty_unif_problem ctxt tp + ] |> Pretty.string_of); + Util.THEN' ( + match_tac ctxt @{thm eq_if_SIMPS_TO_UNIF_if_SIMPS_TO} + THEN' Simps_To.SIMPS_TO_tac (Simps_To.simp_inst_tac simp_tac) ctxt, + SIMPS_TO_UNIF_env_thm_tac' env) i state) + in Tactic_Unification.unify eq_tac ctxt end + +fun simp_unify_progress teq simp_unify norm_term norms unif binders ctxt tp env = + let + val tp as (lhs, rhs) = apply2 (Binders.replace_binders binders #> norm_term env) tp + fun unify ctxt (tp' as (lhs', rhs')) = if teq (lhs, lhs') andalso teq (rhs, rhs') + then (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Simplification of ", + Unification_Util.pretty_unif_problem ctxt tp, + Pretty.str " failed (no progress)" + ] |> Pretty.string_of); + K Seq.empty) + else unif binders ctxt (apply2 (Binders.replace_frees binders) tp') + in simp_unify norms unify ctxt tp env end end diff --git a/thys/ML_Unification/Unifiers/tactic_unification.ML b/thys/ML_Unification/Unifiers/tactic_unification.ML --- a/thys/ML_Unification/Unifiers/tactic_unification.ML +++ b/thys/ML_Unification/Unifiers/tactic_unification.ML @@ -1,29 +1,38 @@ (* Title: ML_Unification/tactic_unification.ML Author: Kevin Kappelmann Solving equations for unification problems with environment-aware tactics. *) signature TACTIC_UNIFICATION = sig + include HAS_LOGGER + type env_tactic = Envir.env * thm -> (Envir.env * thm) Seq.seq (*create environment-aware tactic from a tactic that *DOES NOT* instantiate meta variables*) val env_tac_from_no_inst_tac : (int -> tactic) -> int -> env_tactic val unify : (term * term -> int -> env_tactic) -> Unification_Base.closed_unifier end structure Tactic_Unification : TACTIC_UNIFICATION = struct +val logger = Logger.setup_new_logger Unification_Base.logger "Tactic_Unification" + type env_tactic = Envir.env * thm -> (Envir.env * thm) Seq.seq fun env_tac_from_no_inst_tac tac i (env, state) = tac i state |> Seq.map (pair env) -fun unify tac ctxt tp env = +fun unify tac ctxt (tp as (lhs, rhs)) env = (Logic.mk_equals tp |> Thm.cterm_of ctxt |> Goal.init |> Tactic_Util.HEADGOAL (tac tp) o pair env |> Seq.map_filter (try (apsnd (Goal.finish ctxt)))) - handle TYPE _ => Seq.empty + handle TYPE _ => (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Types of terms ", + Unification_Util.pretty_terms ctxt [lhs, rhs], + Pretty.str " not equal" + ] |> Pretty.string_of); + Seq.empty) end diff --git a/thys/ML_Unification/Unifiers/type_unification.ML b/thys/ML_Unification/Unifiers/type_unification.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/type_unification.ML @@ -0,0 +1,48 @@ +(* Title: ML_Unification/type_unification.ML + Author: Kevin Kappelmann +*) +signature TYPE_UNIFICATION = +sig + include HAS_LOGGER + + (*matches/unifies types of given terms and passes unification problem with updated environment + to the fallback matcher/unifier*) + val e_match : Unification_Base.type_matcher -> Unification_Base.e_matcher + val e_unify : Unification_Base.type_unifier -> Unification_Base.e_unifier +end + +structure Type_Unification : TYPE_UNIFICATION = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger "Type_Unification" + +structure Norm = Envir_Normalisation +structure Util = Unification_Util + +fun e_match match_types match_theory binders ctxt (pt as (p, t)) env = + let val binder_types = Binders.binder_types binders + in + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Matching types of ", + Util.pretty_unif_problem ctxt pt + ] |> Pretty.string_of); + match_types ctxt (Envir.fastype env binder_types p, fastype_of1 (binder_types, t)) env + |> match_theory binders ctxt pt) + handle Unification_Base.UNIF => Seq.empty + end + +fun e_unify unif_types unif_theory binders ctxt st env = + let val binder_types = Binders.binder_types binders + in + (@{log Logger.DEBUG} ctxt (fn _ => + Pretty.block [ + Pretty.str "Unifying types of ", + Util.pretty_unif_problem ctxt (apply2 (Norm.norm_term_unif env) st) + ] + |> Pretty.string_of); + unif_types ctxt (apply2 (Envir.fastype env binder_types) st) env + |> unif_theory binders ctxt st) + handle Unification_Base.UNIF => Seq.empty + end + +end diff --git a/thys/ML_Unification/Unifiers/unification_combinator.ML b/thys/ML_Unification/Unifiers/unification_combinator.ML --- a/thys/ML_Unification/Unifiers/unification_combinator.ML +++ b/thys/ML_Unification/Unifiers/unification_combinator.ML @@ -1,108 +1,132 @@ (* Title: ML_Unification/unification_combinator.ML Author: Kevin Kappelmann Unification combinators. *) signature UNIFICATION_COMBINATOR = sig (* failures *) val fail_closed_unify : Unification_Base.closed_unifier val fail_unify : Unification_Base.unifier val fail_e_unify : Unification_Base.e_unifier val fail_closed_match : Unification_Base.closed_matcher val fail_match : Unification_Base.matcher val fail_e_match : Unification_Base.e_matcher (* fallbacks *) (*"add_fallback_e_unifier u1 u2 u" creates a unifier that calls u1 with fallback unifier "u2 u"*) val add_fallback_e_unifier : Unification_Base.e_unifier -> Unification_Base.e_unifier -> Unification_Base.e_unifier (*"add_fallback_unifier u1 u2" creates a unifier that calls u1 with fallback unifier u2*) val add_fallback_unifier : Unification_Base.e_unifier -> Unification_Base.unifier -> Unification_Base.unifier val add_fallback_e_matcher : Unification_Base.e_matcher -> Unification_Base.e_matcher -> Unification_Base.e_matcher val add_fallback_matcher : Unification_Base.e_matcher -> Unification_Base.matcher -> Unification_Base.matcher (* sequential *) (*try all unifiers in sequence*) val concat_closed_unifiers : Unification_Base.closed_unifier list -> Unification_Base.closed_unifier val concat_closed_matchers : Unification_Base.closed_matcher list -> Unification_Base.closed_matcher val concat_unifiers : Unification_Base.unifier list -> Unification_Base.unifier val concat_matchers : Unification_Base.matcher list -> Unification_Base.matcher val concat_e_unifiers : Unification_Base.e_unifier list -> Unification_Base.e_unifier val concat_e_matchers : Unification_Base.e_matcher list -> Unification_Base.e_matcher (* normalisation *) (*apply normaliser before calling the matcher/unifier*) val norm_closed_unifier : Envir_Normalisation.term_normaliser -> Unification_Base.closed_unifier -> Unification_Base.closed_unifier val norm_closed_matcher : Envir_Normalisation.term_normaliser -> Unification_Base.closed_matcher -> Unification_Base.closed_matcher val norm_unifier : Envir_Normalisation.term_normaliser -> Unification_Base.unifier -> Unification_Base.unifier val norm_matcher : Envir_Normalisation.term_normaliser -> Unification_Base.matcher -> Unification_Base.matcher - (* miscellaneous *) + (* binder replacements *) + val replace_binders_closed_unifier : Unification_Base.closed_unifier -> Unification_Base.unifier + val replace_binders_closed_matcher : Unification_Base.closed_matcher -> Unification_Base.matcher + val replace_binders_unifier : Unification_Base.unifier -> Unification_Base.unifier + val replace_binders_matcher : Unification_Base.matcher -> Unification_Base.matcher + + val replace_frees_closed_unifier : Unification_Base.closed_unifier -> Unification_Base.unifier + val replace_frees_closed_matcher : Unification_Base.closed_matcher -> Unification_Base.matcher + val replace_frees_unifier : Unification_Base.unifier -> Unification_Base.unifier + val replace_frees_matcher : Unification_Base.matcher -> Unification_Base.matcher + (*approximates a unifier (working with bound variables) from a closed unifier by replacing all bound variables with their associated free variables*) val unifier_from_closed_unifier : Unification_Base.closed_unifier -> Unification_Base.unifier val matcher_from_closed_matcher : Unification_Base.closed_matcher -> Unification_Base.matcher end structure Unification_Combinator : UNIFICATION_COMBINATOR = struct structure UB = Unification_Base (* failures *) val fail_closed_unify = K o K o K Seq.empty val fail_unify = K fail_closed_unify val fail_e_unify = K fail_unify val fail_closed_match = fail_closed_unify val fail_match = fail_unify val fail_e_match = fail_e_unify (* fallbacks *) fun add_fallback_e_unifier e_unif1 e_unif2 = e_unif1 o e_unif2 fun add_fallback_unifier e_unif1 e_unif2 = add_fallback_e_unifier e_unif1 (K e_unif2) fail_unify val add_fallback_e_matcher = add_fallback_e_unifier fun add_fallback_matcher e_match1 e_match2 = add_fallback_e_matcher e_match1 (K e_match2) fail_match (* sequential *) fun concat_closed_unifiers unifiers ctxt tp = fold_rev (fn unify => curry Seq.APPEND (unify ctxt tp)) unifiers Seq.fail val concat_closed_matchers = concat_closed_unifiers fun concat_unifiers unifiers binders ctxt tp = fold_rev (fn unify => curry Seq.APPEND (unify binders ctxt tp)) unifiers Seq.fail val concat_matchers = concat_unifiers fun concat_e_unifiers unifiers unif binders ctxt tp = fold_rev (fn unify => curry Seq.APPEND (unify unif binders ctxt tp)) unifiers Seq.fail val concat_e_matchers = concat_e_unifiers (* normalisation *) fun norm_closed_unifier norm_t unif ctxt tp env = unif ctxt (apply2 (norm_t env) tp) env fun norm_closed_matcher norm_p match ctxt (p, t) env = match ctxt (norm_p env p, t) env fun norm_unifier norm_t unif binders ctxt tp env = norm_closed_unifier norm_t (unif (Binders.norm_binders (norm_t env) binders)) ctxt tp env fun norm_matcher norm_p match binders ctxt pt env = norm_closed_matcher norm_p (match (Binders.norm_binders (norm_p env) binders)) ctxt pt env +(* binder replacements *) +fun replace_binders_closed_unifier unif binders ctxt = + unif ctxt o apply2 (Binders.replace_binders binders) +val replace_binders_closed_matcher = replace_binders_closed_unifier + +fun replace_binders_unifier unif binders = replace_binders_closed_unifier (unif binders) binders +fun replace_binders_matcher match binders = replace_binders_closed_matcher (match binders) binders + +fun replace_frees_closed_unifier unif binders ctxt = + unif ctxt o apply2 (Binders.replace_frees binders) +val replace_frees_closed_matcher = replace_frees_closed_unifier + +fun replace_frees_unifier unif binders = replace_frees_closed_unifier (unif binders) binders +fun replace_frees_matcher match binders = replace_frees_closed_matcher (match binders) binders + (* miscellaneous *) -fun unifier_from_closed_unifier unify binders ctxt = - unify ctxt o apply2 (Binders.replace_binders binders) -val matcher_from_closed_matcher = unifier_from_closed_unifier +val unifier_from_closed_unifier = replace_binders_closed_unifier +val matcher_from_closed_matcher = replace_binders_closed_matcher end diff --git a/thys/ML_Unification/Unifiers/var_higher_order_pattern_unification.ML b/thys/ML_Unification/Unifiers/var_higher_order_pattern_unification.ML --- a/thys/ML_Unification/Unifiers/var_higher_order_pattern_unification.ML +++ b/thys/ML_Unification/Unifiers/var_higher_order_pattern_unification.ML @@ -1,70 +1,73 @@ (* Title: ML_Unification/var_higher_order_pattern_unification.ML Author: Kevin Kappelmann Unifies terms of the form ?H b1 ... bn \\<^sup>? t where (1) b1,...,bn are bound variables and (2) ?H does not occur in t with the higher-order pattern matcher. This pattern occurs frequently when unifying theorems with a goal containing bound variables. Example: If we want to solve the goal "\x. ?P x 0" with a lemma ?Q by resolution, the bound variables of the goal are first lifted to lemma, resulting in "?Q x". Next, the terms "?Q x \\<^sup>? ?P x 0" are unified. Unfortunately, this problem falls outside the higher-order pattern fragment since 0 is a constant. However, the problem seems essentially first-order before lifting the bound variables: we should be able to use ?Q to solve any goal G by setting ?Q to G. The unifier of this file returns the expected substitution "?Q := \x. ?P x 0" in this case. *) signature VAR_HIGHER_ORDER_PATTERN_UNIFICATION = sig include HAS_LOGGER val can_var_pattern_unif : term * term -> bool * term list val e_unify : Unification_Base.e_unifier end structure Var_Higher_Order_Pattern_Unification : VAR_HIGHER_ORDER_PATTERN_UNIFICATION = struct val logger = Logger.setup_new_logger Unification_Base.logger "Var_Higher_Order_Pattern_Unification" structure Util = Unification_Util fun can_var_pattern_unif (p, t) = let val (head, args) = strip_comb p in if is_Var head andalso forall is_Bound args andalso not (has_duplicates (op aconv) args) andalso not (Term.exists_subterm (curry (op =) head) t) then (true, args) else (false, []) end fun e_unify unif_theory binders ctxt (tp as (t, p)) env = let exception FALLBACK of Pretty.T - val hop_match_type_unif = Higher_Order_Pattern_Unification.e_match Util.unify_types - Unification_Combinator.fail_match Unification_Combinator.fail_match binders ctxt + val unify_types = Util.unify_types + val hop_match_type_unif = Higher_Order_Pattern_Unification.e_match unify_types + Unification_Combinator.fail_match Unification_Combinator.fail_match + |> Type_Unification.e_unify unify_types + |> (fn unif => unif binders ctxt) val hop_match_type_unif_tp = hop_match_type_unif tp val hop_match_type_unif_pt = hop_match_type_unif (p, t) #> Seq.map (apsnd Unification_Base.symmetric) val seq_try = General_Util.seq_try (FALLBACK (Pretty.str "Higher-order pattern matching failed.")) in (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ Pretty.str "Variable head higher-order pattern unifying ", Util.pretty_unif_problem ctxt tp ] |> Pretty.string_of); case apply2 can_var_pattern_unif (tp, (p, t)) of ((true, boundst), (true, boundsp)) => if length boundst >= length boundsp then hop_match_type_unif_tp env |> seq_try else hop_match_type_unif_pt env |> seq_try | ((true, _), _) => hop_match_type_unif_tp env |> seq_try | (_, (true, _)) => hop_match_type_unif_pt env |> seq_try | _ => raise FALLBACK (Pretty.str "Terms not of the form ?H b1 ... bn \\<^sup>? t where ?H does not occur in t.")) handle FALLBACK msg => (@{log Logger.DEBUG} ctxt (fn _ => Pretty.breaks [ msg, Util.pretty_call_theory_unif ctxt tp ] |> Pretty.block |> Pretty.string_of); unif_theory binders ctxt tp env) end end diff --git a/thys/ML_Unification/unification_base.ML b/thys/ML_Unification/unification_base.ML --- a/thys/ML_Unification/unification_base.ML +++ b/thys/ML_Unification/unification_base.ML @@ -1,74 +1,77 @@ (* Title: ML_Unification/unification_base.ML Author: Kevin Kappelmann Basic definitions for unifiers. *) signature UNIFICATION_BASE = sig include HAS_LOGGER val reflexive : cterm -> thm val combination : thm -> thm -> thm val symmetric : thm -> thm val abstract_rule : Proof.context -> string -> cterm -> thm -> thm option val reflexive_term : Proof.context -> term -> thm (*raised on unsupported input*) exception PATTERN (*raised on unification failure for non-sequence outputs (e.g. type unification)*) exception UNIF type type_unifier = Proof.context -> typ * typ -> Envir.env -> Envir.env type type_matcher = type_unifier type closed_unifier = Proof.context -> term * term -> Envir.env -> (Envir.env * thm) Seq.seq (*term binders stores fresh free variables associated to each loose bound variable*) type unifier = term Binders.binders -> closed_unifier type e_unifier = unifier -> unifier type closed_matcher = closed_unifier type matcher = unifier type e_matcher = matcher -> matcher (* normalisers for matchers/unifiers *) type normalisers = { + (*normaliser for result of unifier*) norm_unif_thm : Envir_Normalisation.thm_normaliser, + (*normaliser for other theorems*) norm_thm : Envir_Normalisation.thm_normaliser, + (*normaliser for terms*) norm_term : Envir_Normalisation.term_normaliser } end structure Unification_Base : UNIFICATION_BASE = struct -val logger = Logger.setup_new_logger Logger.root_logger "Unification_Base" +val logger = Logger.setup_new_logger Logger.root "Unification_Base" val reflexive = Thm.reflexive val combination = Thm.combination val symmetric = Thm.symmetric val abstract_rule = Thm_Util.abstract_rule val reflexive_term = reflexive oo Thm.cterm_of exception PATTERN = Pattern.Pattern exception UNIF = Pattern.Unif type type_unifier = Proof.context -> typ * typ -> Envir.env -> Envir.env type type_matcher = type_unifier type closed_unifier = Proof.context -> term * term -> Envir.env -> (Envir.env * thm) Seq.seq type unifier = term Binders.binders -> closed_unifier type e_unifier = unifier -> unifier type closed_matcher = closed_unifier type matcher = unifier type e_matcher = matcher -> matcher type normalisers = { norm_unif_thm : Envir_Normalisation.thm_normaliser, norm_thm : Envir_Normalisation.thm_normaliser, norm_term : Envir_Normalisation.term_normaliser } end diff --git a/thys/Standard_Borel_Spaces/Abstract_Metrizable_Topology.thy b/thys/Standard_Borel_Spaces/Abstract_Metrizable_Topology.thy --- a/thys/Standard_Borel_Spaces/Abstract_Metrizable_Topology.thy +++ b/thys/Standard_Borel_Spaces/Abstract_Metrizable_Topology.thy @@ -1,1008 +1,1004 @@ (* Title: Abstract_Metrizable_Topology.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) section \Abstract Polish Spaces \ theory Abstract_Metrizable_Topology imports "Set_Based_Metric_Product" begin subsection \ Polish Spaces \ definition "polish_space X \ completely_metrizable_space X \ separable_space X" lemma(in Metric_space) polish_space_mtopology: assumes mcomplete "separable_space mtopology" shows "polish_space mtopology" by (simp add: assms completely_metrizable_space_mtopology polish_space_def) lemma assumes "polish_space X" shows polish_space_imp_completely_metrizable_space: "completely_metrizable_space X" and polish_space_imp_metrizable_space: "metrizable_space X" and polish_space_imp_second_countable: "second_countable X" and polish_space_imp_separable_space: "separable_space X" using assms by(auto simp: completely_metrizable_imp_metrizable_space polish_space_def metrizable_space_separable_iff_second_countable) lemma polish_space_closedin: assumes "polish_space X" "closedin X A" shows "polish_space (subtopology X A)" using assms by(auto simp: completely_metrizable_imp_metrizable_space polish_space_def completely_metrizable_space_closedin second_countable_subtopology metrizable_space_separable_iff_second_countable) lemma polish_space_gdelta_in: assumes "polish_space X" "gdelta_in X A" shows "polish_space (subtopology X A)" using assms by(auto simp: completely_metrizable_imp_metrizable_space polish_space_def completely_metrizable_space_gdelta_in second_countable_subtopology metrizable_space_separable_iff_second_countable) corollary polish_space_openin: assumes "polish_space X" "openin X A" shows "polish_space (subtopology X A)" by (simp add: open_imp_gdelta_in assms polish_space_gdelta_in) lemma homeomorphic_polish_space_aux: assumes "polish_space X" "X homeomorphic_space Y" shows "polish_space Y" using assms by(simp add: homeomorphic_completely_metrizable_space_aux homeomorphic_separable_space polish_space_def) corollary homeomorphic_polish_space: assumes "X homeomorphic_space Y" shows "polish_space X \ polish_space Y" by (meson assms homeomorphic_polish_space_aux homeomorphic_space_sym) lemma polish_space_euclidean[simp]: "polish_space (euclidean :: ('a :: polish_space) topology)" by (simp add: completely_metrizable_space_euclidean polish_space_def second_countable_imp_separable_space) lemma polish_space_countable[simp]: "polish_space (euclidean :: 'a :: {countable,discrete_topology} topology)" proof - interpret discrete_metric "UNIV :: 'a set" . have [simp]:"euclidean = disc.mtopology" by (metis discrete_topology_class.open_discrete discrete_topology_unique istopology_open mtopology_discrete_metric topology_inverse' topspace_euclidean) show ?thesis by(auto simp: polish_space_def intro!: disc.completely_metrizable_space_mtopology mcomplete_discrete_metric countable_space_separable_space) qed lemma polish_space_discrete_topology: "polish_space (discrete_topology I) \ countable I" by (simp add: completely_metrizable_space_discrete_topology polish_space_def separable_space_discrete_topology) lemma polish_space_prod: assumes "polish_space X" and "polish_space Y" shows "polish_space (prod_topology X Y)" using assms by (simp add: completely_metrizable_space_prod_topology polish_space_def separable_space_prod) lemma polish_space_product: assumes "countable I" and "\i. i \ I \ polish_space (S i)" shows "polish_space (product_topology S I)" using assms by(auto simp: separable_countable_product polish_space_def completely_metrizable_space_product_topology) lemma(in Product_metric) polish_spaceI: assumes "\i. i \ I \ separable_space (Metric_space.mtopology (Mi i) (di i))" and "\i. i \ I \ Metric_space.mcomplete (Mi i) (di i)" shows "polish_space Product_metric.mtopology" using assms I by(auto simp: polish_space_def Product_metric_mtopology_eq[symmetric] completely_metrizable_space_product_topology intro!: separable_countable_product Metric_space.completely_metrizable_space_mtopology Md_metric) lemma(in Sum_metric) polish_spaceI: assumes "countable I" and "\i. i \ I \ separable_space (Metric_space.mtopology (Mi i) (di i))" and "\i. i \ I \ Metric_space.mcomplete (Mi i) (di i)" shows "polish_space Sum_metric.mtopology" by(auto simp: polish_space_def intro!: separable_Mi_separable_M assms mcomplete_Mi_mcomplete_M Sum_metric.completely_metrizable_space_mtopology) lemma compact_metrizable_imp_polish_space: assumes "metrizable_space X" "compact_space X" shows "polish_space X" proof - obtain d where "Metric_space (topspace X) d" "Metric_space.mtopology (topspace X) d = X" by (metis assms(1) Metric_space.topspace_mtopology metrizable_space_def) thus ?thesis by (metis Metric_space.compact_space_imp_separable assms(1) assms(2) compact_imp_locally_compact_space locally_compact_imp_completely_metrizable_space polish_space_def) qed subsection \Continuous Embddings\ abbreviation Hilbert_cube_topology :: "(nat \ real) topology" where "Hilbert_cube_topology \ (product_topology (\n. top_of_set {0..1}) UNIV)" lemma topspace_Hilbert_cube: "topspace Hilbert_cube_topology = (\\<^sub>E x\UNIV. {0..1})" by simp lemma polish_space_Hilbert_cube: "polish_space Hilbert_cube_topology" by(auto intro!: polish_space_closedin polish_space_product) abbreviation Cantor_space_topology :: "(nat \ real) topology" where "Cantor_space_topology \ (product_topology (\n. top_of_set {0,1}) UNIV)" lemma topspace_Cantor_space: "topspace Cantor_space_topology = (\\<^sub>E x\UNIV. {0,1})" by simp lemma polish_space_Cantor_space: "polish_space Cantor_space_topology" by(auto intro!: polish_space_closedin polish_space_product) corollary completely_metrizable_space_homeo_image_gdelta_in: assumes "completely_metrizable_space X" "completely_metrizable_space Y" "B \ topspace Y" "X homeomorphic_space subtopology Y B" shows "gdelta_in Y B" using assms completely_metrizable_space_eq_gdelta_in homeomorphic_completely_metrizable_space by blast subsubsection \ Embedding into Hilbert Cube\ lemma embedding_into_Hilbert_cube: assumes "metrizable_space X" "separable_space X" shows "\A \ topspace Hilbert_cube_topology. X homeomorphic_space (subtopology Hilbert_cube_topology A)" proof - consider "X = trivial_topology" | "topspace X \ {}" by auto then show ?thesis proof cases case 1 then show ?thesis by(auto intro!: exI[where x="{}"] simp: homeomorphic_empty_space_eq) next case S_ne:2 then obtain U where U:"countable U" "dense_in X U" "U \ {}" using assms(2) by(auto simp: separable_space_def2 dense_in_nonempty) obtain xn where xn:"\n::nat. xn n \ U" "U = range xn" by (metis U(1) U(3) from_nat_into range_from_nat_into) then have xns:"xn n \ topspace X" for n using dense_in_subset[OF U(2)] by auto obtain d' where d':"Metric_space (topspace X) d'" "Metric_space.mtopology (topspace X) d' = X" by (metis Metric_space.topspace_mtopology assms(1) metrizable_space_def) interpret ms': Metric_space "topspace X" d' by fact define d where "d = ms'.capped_dist (1/2)" have d: "Metric_space.mtopology (topspace X) d = X" "\x y. d x y < 1" by(simp add: d_def ms'.mtopology_capped_metric d') (simp add: d_def ms'.capped_dist_def) interpret ms: Metric_space "topspace X" d by (simp add: d_def ms'.capped_dist) define f where "f \ (\x n. d x (xn n))" have f_inj:"inj_on f (topspace X)" proof fix x y assume xy:"x \ topspace X" "y \ topspace X" "f x = f y" then have "\n. d x (xn n) = d y (xn n)" by(auto simp: f_def dest: fun_cong) hence d2:"d x y \ 2 * d x (xn n)" for n using ms.triangle[OF xy(1) _ xy(2),of "xn n",simplified ms.commute[of "xn n" y]] dense_in_subset[OF U(2)] xn(1)[of n] by auto have "d x y < \" if "\ > 0" for \ proof - have "0 < \ / 2" using that by simp then obtain n where "d x (xn n) < \ / 2" using ms.mdense_def2[of U,simplified d(1)] U(2) xy(1) xn(2) by blast with d2[of n] show ?thesis by simp qed hence "d x y = 0" by (metis ms.nonneg[of x y] dual_order.irrefl order_neq_le_trans) thus "x = y" using xy by simp qed have f_img: "f ` topspace X \ topspace Hilbert_cube_topology" using d(2) ms.nonneg by(auto simp: topspace_Hilbert_cube f_def less_le_not_le) have f_cont: "continuous_map X Hilbert_cube_topology f" unfolding continuous_map_componentwise_UNIV f_def continuous_map_in_subtopology proof safe show "continuous_map X euclideanreal (\x. d x (xn k))" for k proof(rule continuous_map_eq[of _ _ "mdist_set ms.Self {xn k}"]) show "continuous_map X euclideanreal (mdist_set ms.Self {xn k})" by (metis d(1) mdist_set_uniformly_continuous ms.mdist_Self ms.mspace_Self mtopology_of_def mtopology_of_euclidean uniformly_continuous_imp_continuous_map) next fix x assume "x \ topspace X" then show "mdist_set ms.Self {xn k} x = d x (xn k)" by(auto simp: ms.mdist_set_Self xns) qed next show "d x (xn k) \ {0..1}" for x k using d(2) ms.nonneg by(auto simp: less_le_not_le) qed hence f_cont': "continuous_map X (subtopology Hilbert_cube_topology (f ` topspace X)) f" using continuous_map_into_subtopology by blast obtain g where g: "g ` (f ` topspace X) = topspace X" "\x. x \ topspace X \ g (f x) = x" "\x. x \ f ` topspace X \ f (g x) = x" by (meson f_inj f_the_inv_into_f the_inv_into_f_eq the_inv_into_onto) have g_cont: "continuous_map (subtopology Hilbert_cube_topology (f ` topspace X)) X g" proof - interpret m01: Submetric UNIV dist "{0..1::real}" by(simp add: Submetric_def Submetric_axioms_def Met_TC.Metric_space_axioms) have m01_eq: "m01.sub.mtopology = top_of_set {0..1}" using m01.mtopology_submetric by auto have m01_polish: "polish_space m01.sub.mtopology" by(auto simp: m01_eq intro!: polish_space_closedin) interpret m01': Metric_space "{0..1::real}" "\x y. if 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 then dist x y else 0" by(auto intro!: Metric_space_eq[OF m01.sub.Metric_space_axioms]) metric have m01'_eq: "m01'.mtopology = top_of_set {0..1}" by(auto intro!: Metric_space_eq_mtopology[OF m01.sub.Metric_space_axioms,simplified m01_eq,symmetric]) metric have "dist x y \ 1" "dist x y \ 0" if "x \ 0" "x \ 1" "y \ 0" "y \ 1" for x y :: real using dist_real_def that by auto then interpret ppm: Product_metric "1/2" "UNIV :: nat set" id id "\_. {0..1::real}" "\_ x y. if 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 then dist x y else 0" 1 by(auto intro!: product_metric_natI Metric_space_eq[OF m01.sub.Metric_space_axioms] simp: m01.sub.commute) have Hilbert_cube_eq: "ppm.Product_metric.mtopology = Hilbert_cube_topology" by(simp add: ppm.Product_metric_mtopology_eq[symmetric] m01'_eq) interpret f_S: Submetric "\\<^sub>E x\UNIV. {0..1}" ppm.product_dist "f ` topspace X" by(auto simp: Submetric_def ppm.Product_metric.Metric_space_axioms Submetric_axioms_def f_def order.strict_implies_order[OF d(2)]) have 1:"subtopology Hilbert_cube_topology (f ` topspace X) = f_S.sub.mtopology" using Hilbert_cube_eq f_S.mtopology_submetric by auto - have "continuous_map (mtopology_of f_S.sub.Self) (mtopology_of ms.Self) g" - unfolding continuous_map_metric_limitin f_S.sub.mspace_Self ms.mspace_Self - unfolding mtopology_of_def f_S.sub.mspace_Self ms.mspace_Self f_S.sub.mdist_Self ms.mdist_Self + have "continuous_map f_S.sub.mtopology ms.mtopology g" + unfolding continuous_map_iff_limit_seq[OF f_S.sub.first_countable_mtopology] proof safe - show "x \ topspace X \ g (f x) \ topspace X" for x - by(simp add: g(2)) - next fix yn y assume h:" limitin f_S.sub.mtopology yn y sequentially" have h':"limitin ppm.Product_metric.mtopology yn y sequentially" using f_S.limitin_submetric_iff h by blast hence m01_conv:"\n. limitin m01'.mtopology (\i. yn i n) (y n) sequentially" "y \ UNIV \\<^sub>E {0..1}" by(auto simp: ppm.limitin_M_iff_limitin_Mi) have "\N. \n\N. \zn. yn n = f zn \ zn \ topspace X" using h g by(simp only: f_S.sub.limit_metric_sequentially) (meson imageE ppm.K_pos) then obtain N' zn where zn:"\n. n \ N' \ f (zn n) = yn n" "\n. n \ N' \ zn n \ topspace X" by metis obtain z where z:"f z = y" "z \ topspace X" using h f_S.sub.limitin_mspace by blast show "limitin ms.mtopology (\n. g (yn n)) (g y) sequentially" unfolding ms.limit_metric_sequentially proof safe fix \ :: real assume he: "0 < \" then have "0 < \ / 3" by simp then obtain m where m:"d z (xn m) < \ / 3" using ms.mdense_def2[of U,simplified d(1)] U(2) z(2) xn(2) by blast have "\e. e>0 \ \N. \n\N. yn n m \ {0..1} \ dist (yn n m) (y m) < e" using m01_conv(1)[of m,simplified m01'.limit_metric_sequentially] by fastforce from this[OF \0 < \ / 3\] obtain N where "\n. n \ N \ \yn n m - y m\ < \ / 3" "\n. n \ N \ yn n m \ {0..1}" by(auto simp: dist_real_def) hence N:"\n. n \ N \ yn n m < \ / 3 + y m" by (metis abs_diff_less_iff add.commute) have "\N. \n\N. zn n \ topspace X \ d (zn n) z < \" proof(safe intro!: exI[where x="max N N'"]) fix n assume "max N N' \ n" then have "N \ n" "N' \ n" by auto then have "d (zn n) z \ f (zn n) m + d z (xn m)" using ms.triangle[OF zn(2)[of n] xns[of m] z(2),simplified ms.commute[of "xn m" z]] by(auto simp: f_def) also have "... < \ / 3 + y m + d z (xn m)" using N[OF \N\n\] zn(1)[of n] \N' \ n\ by simp also have "... = \ / 3 + d z (xn m) + d z (xn m)" by(simp add: z(1)[symmetric] f_def) also have "... < \" using m by auto finally show "d (zn n) z < \" . qed(use zn in auto) thus "\N. \n\N. g (yn n) \ topspace X \ d (g (yn n)) (g y) < \" by (metis dual_order.trans nle_le zn(1) z(1) g(2)[OF z(2)] g(2)[OF zn(2)]) qed(use g z in auto) qed hence "continuous_map f_S.sub.mtopology ms.mtopology g" by(auto simp: mtopology_of_def) thus ?thesis by(simp add: d(1) 1) qed show ?thesis using f_img g(2,3) f_cont' g_cont by(auto intro!: exI[where x="f ` topspace X"] homeomorphic_maps_imp_homeomorphic_space[where f=f and g=g] simp: homeomorphic_maps_def) qed qed corollary embedding_into_Hilbert_cube_gdelta_in: assumes "polish_space X" shows "\A. gdelta_in Hilbert_cube_topology A \ X homeomorphic_space (subtopology Hilbert_cube_topology A)" proof - obtain A where h:"A \ topspace Hilbert_cube_topology" "X homeomorphic_space subtopology Hilbert_cube_topology A" using embedding_into_Hilbert_cube polish_space_imp_metrizable_space polish_space_imp_separable_space assms by blast with completely_metrizable_space_homeo_image_gdelta_in[OF polish_space_imp_completely_metrizable_space[OF assms] polish_space_imp_completely_metrizable_space[OF polish_space_Hilbert_cube] h(1,2)] show ?thesis by blast qed subsubsection \ Embedding from Cantor Space \ lemma embedding_from_Cantor_space: assumes "polish_space X" "uncountable (topspace X)" shows "\A. gdelta_in X A \ Cantor_space_topology homeomorphic_space (subtopology X A)" proof - obtain U P where up: "countable U" "openin X U" "perfect_set X P""U \ P = topspace X" "U \ P = {}" "\a. a \ {} \ openin (subtopology X P) a \ uncountable a" using Cantor_Bendixon[OF polish_space_imp_second_countable[OF assms(1)]] by auto have P: "closedin X P" "P \ topspace X" "uncountable P" using countable_Un_iff[of U P] up(1) assms up(4) by(simp_all add: perfect_setD[OF up(3)]) then have pp: "polish_space (subtopology X P)" by (simp add: assms(1) polish_space_closedin) have Ptop: "topspace (subtopology X P) = P" using P(2) by auto obtain U where U: "countable U" "dense_in (subtopology X P) U" using polish_space_def pp separable_space_def2 by auto with uncountable_infinite[OF P(3)] P(2) have "infinite U" by (metis Metric_space.t1_space_mtopology Ptop assms(1) completely_metrizable_space_def dense_in_infinite polish_space_def t1_space_subtopology) obtain d where "Metric_space P d" and d:"Metric_space.mtopology P d = subtopology X P" and mdc: "Metric_space.mcomplete P d" by (metis Metric_space.topspace_mtopology Ptop completely_metrizable_space_def polish_space_def pp) interpret md: Metric_space P d by fact define xn where "xn \ from_nat_into U" have xn: "bij_betw xn UNIV U" "\n m. n \ m \ xn n \ xn m" "\n. xn n \ U" "\n. xn n \ P" "md.mdense (range xn)" using bij_betw_from_nat_into[OF U(1) \infinite U\] dense_in_subset[OF U(2)] d U(2) range_from_nat_into[OF infinite_imp_nonempty[OF \infinite U\] U(1)] by(auto simp add: xn_def U(1) \infinite U\ from_nat_into[OF infinite_imp_nonempty[OF \infinite U\]]) have perfect:"perfect_space md.mtopology" using d perfect_set_subtopology up(3) by simp define jn where "jn \ (\n. LEAST i. i > n \ md.mcball (xn i) ((1/2)^i) \ md.mball (xn n) ((1/2)^n) - md.mball (xn n) ((1/2)^i))" define kn where "kn \ (\n. LEAST k. k > jn n \ md.mcball (xn k) ((1/2)^k) \ md.mball (xn n) ((1/2)^jn n))" have dxmxn: "\n n'. \m. m > n \ m > n' \ (1/2)^(m-1) < d (xn n) (xn m) \ d (xn n) (xn m) < (1/2)^(Suc n')" proof safe fix n n' have hinfin':"infinite (md.mball x e \ (range xn))" if "x \ P" "e > 0" for x e proof assume h_fin:"finite (md.mball x e \ range xn)" have h_nen:"md.mball x e \ range xn \ {}" using xn(5) that by(auto simp: md.mdense_def) have infin: "infinite (md.mball x e)" using md.perfect_set_mball_infinite[OF perfect] that by simp then obtain y where y:"y \ md.mball x e" "y \ range xn" using h_fin by(metis inf.absorb_iff2 inf_commute subsetI) define e' where "e' = Min {d y xk |xk. xk \ md.mball x e \ range xn}" have fin: "finite {d y xk |xk. xk \ md.mball x e \ range xn}" using finite_imageI[OF h_fin,of "d y"] by (metis Setcompr_eq_image) have nen: "{d y xk |xk. xk \ md.mball x e \ range xn} \ {}" using h_nen by auto have "e' > 0" unfolding e'_def Min_gr_iff[OF fin nen] proof safe fix l assume "xn l \ md.mball x e" with y show "0 < d y (xn l)" by auto qed obtain e'' where e'': "e'' > 0" "md.mball y e'' \ md.mball x e" "y \ md.mball y e''" by (meson md.centre_in_mball_iff md.in_mball md.openin_mball md.openin_mtopology y(1)) define \ where "\ \ min e' e''" have "\ > 0" using e''(1) \e' > 0\ by(simp add: \_def) then obtain m where m: "d y (xn m) < \" using md.mdense_def2[of "range xn"] xn(5) y(1) by fastforce consider "xn m \ md.mball x e" | "xn m \ P - md.mball x e" using xn(4) by auto then show False proof cases case 1 then have "e' \ d y (xn m)" using Min_le_iff[OF fin nen] by(auto simp: e'_def) thus ?thesis using m by(simp add: \_def) next case 2 then have "xn m \ md.mball y e''" using e''(2) by auto hence "e'' \ d y (xn m)" using y e'' xn by auto thus ?thesis using m by(simp add: \_def) qed qed have hinfin:"infinite (md.mball x e \ (xn ` {l<..}))" if "x \ P" "e > 0" for x e l proof assume "finite (md.mball x e \ xn ` {l<..})" moreover have "finite (md.mball x e \ xn ` {..l})" by simp moreover have "(md.mball x e \ (range xn)) = (md.mball x e \ xn ` {l<..}) \ (md.mball x e \ xn ` {..l})" by fastforce ultimately have "finite (md.mball x e \ (range xn))" by auto with hinfin'[OF that] show False .. qed have "infinite (md.mball (xn n) ((1/2)^Suc n'))" using md.perfect_set_mball_infinite[OF perfect] xn(4)[of n] by simp then obtain x where x: "x \ md.mball (xn n) ((1/2)^Suc n')" "x \ xn n" by (metis finite_insert finite_subset infinite_imp_nonempty singletonI subsetI) then obtain e where e: "e > 0" "md.mball x e \ md.mball (xn n) ((1/2)^Suc n')" "x \ md.mball x e" by (meson md.centre_in_mball_iff md.in_mball md.openin_mball md.openin_mtopology) have "d (xn n) x > 0" using xn x by simp then obtain m' where m': "m' - 1 > 0" "(1/2)^(m' - 1) < d (xn n) x" by (metis One_nat_def diff_Suc_Suc diff_zero one_less_numeral_iff reals_power_lt_ex semiring_norm(76)) define m where "m \ max m' (max n' (Suc n))" then have "m \ m'" "m \ n'" "m \ Suc n" by simp_all hence m: "m - 1 > 0" "(1/2)^(m - 1) < d (xn n) x" "m > n" using m' less_trans[OF _ m'(2),of "(1 / 2) ^ (m - 1)"] by auto (metis diff_less_mono le_eq_less_or_eq) define \ where "\ \ min e (d (xn n) x - (1/2)^(m - 1))" have "\ > 0" using e m by(simp add: \_def) have ball_le:"md.mball x \ \ md.mball (xn n) ((1 / 2) ^ Suc n')" using e(2) by(auto simp add: \_def) obtain k where k: "xn k \ md.mball x \" "k > m" using \\ > 0\ infinite_imp_nonempty[OF hinfin,of _ \] x(1) by fastforce show "\m>n. n' < m \ (1 / 2) ^ (m - 1) < d (xn n) (xn m) \ d (xn n) (xn m) < (1 / 2) ^ Suc n'" proof(intro exI[where x=k] conjI) have "(1 / 2) ^ (k - 1) < (1 / (2 :: real)) ^ (m - 1)" using k(2) m(3) by simp also have "... = d (xn n) x + ((1/2)^ (m - 1) - d (xn n) x)" by simp also have "... < d (xn n) x - d (xn k) x" using k by(auto simp: \_def md.commute) also have "... \ d (xn n) (xn k)" using xn x md.mdist_reverse_triangle[of "xn n" x "xn k"] by(auto simp: md.commute) finally show "(1 / 2) ^ (k - 1) < d (xn n) (xn k)" . qed(use \m \ n'\ k ball_le m(3) in auto) qed have "jn n > n \ md.mcball (xn (jn n)) ((1/2)^(jn n)) \ md.mball (xn n) ((1/2)^n) - md.mball (xn n) ((1/2)^(jn n))" for n unfolding jn_def proof(rule LeastI_ex) obtain m where m:"m > n" "(1 / 2) ^ (m - 1) < d (xn n) (xn m)" "d (xn n) (xn m) < (1 / 2) ^ Suc n" using dxmxn by auto show "\x>n. md.mcball (xn x) ((1 / 2) ^ x) \ md.mball (xn n) ((1 / 2) ^ n) - md.mball (xn n) ((1 / 2) ^ x)" proof(safe intro!: exI[where x=m] m(1)) fix x assume h:"x \ md.mcball (xn m) ((1 / 2) ^ m)" have 1:"d (xn n) x < (1 / 2) ^ n" proof - have "d (xn n) x < (1 / 2) ^ Suc n + (1 / 2) ^ m" using m(3) md.triangle[OF xn(4)[of n] xn(4)[of m],of x] h by auto also have "... \ (1 / 2) ^ Suc n + (1 / 2) ^ Suc n" by (metis Suc_lessI add_mono divide_less_eq_1_pos divide_pos_pos less_eq_real_def m(1) one_less_numeral_iff power_strict_decreasing_iff semiring_norm(76) zero_less_numeral zero_less_one) finally show ?thesis by simp qed have 2:"(1 / 2) ^ m \ d (xn n) x" proof - have "(1 / 2) ^ (m - 1) < d (xn n) x + (1 / 2) ^ m" using order.strict_trans2[OF m(2) md.triangle[OF xn(4)[of n] _ xn(4)[of m]]] h md.commute by fastforce hence "(1 / 2) ^ (m - 1) - (1 / 2) ^ m \ d (xn n) x" by simp thus ?thesis using not0_implies_Suc[OF gr_implies_not0[OF m(1)]] by auto qed show "x \ md.mball (xn n) ((1 / 2) ^ n)" "x \ md.mball (xn n) ((1 / 2) ^ m) \ False" using xn h 1 2 by auto qed qed hence jn: "\n. jn n > n" "\n. md.mcball (xn (jn n)) ((1/2)^(jn n)) \ md.mball (xn n) ((1/2)^n) - md.mball (xn n) ((1/2)^(jn n))" by simp_all have "kn n > jn n \ md.mcball (xn (kn n)) ((1/2)^(kn n)) \ md.mball (xn n) ((1/2)^jn n)" for n unfolding kn_def proof(rule LeastI_ex) obtain m where m:"m > jn n" "d (xn n) (xn m) < (1 / 2) ^ Suc (jn n)" using dxmxn by blast show "\x>jn n. md.mcball (xn x) ((1 / 2) ^ x) \ md.mball (xn n) ((1 / 2) ^ jn n)" proof(intro exI[where x=m] conjI) show "md.mcball (xn m) ((1 / 2) ^ m) \ md.mball (xn n) ((1 / 2) ^ jn n)" proof fix x assume h:"x \ md.mcball (xn m) ((1 / 2) ^ m)" have "d (xn n) x < (1 / 2)^ Suc (jn n) + (1 / 2) ^ m" using md.triangle[OF xn(4)[of n] xn(4)[of m]] h m(2) by fastforce also have "... \ (1 / 2)^ Suc (jn n) + (1 / 2)^ Suc (jn n)" by (metis Suc_le_eq add_mono dual_order.refl less_divide_eq_1_pos linorder_not_less m(1) not_numeral_less_one power_decreasing zero_le_divide_1_iff zero_le_numeral zero_less_numeral) finally show "x \ md.mball (xn n) ((1 / 2) ^ jn n)" using xn(4) h by auto qed qed(use m(1) in auto) qed hence kn: "\n. kn n > jn n" "\n. md.mcball (xn (kn n)) ((1/2)^(kn n)) \ md.mball (xn n) ((1/2)^(jn n))" by simp_all have jnkn_pos: "jn n > 0" "kn n > 0" for n using not0_implies_Suc[OF gr_implies_not0[OF jn(1)[of n]]] kn(1)[of n] by auto define bn :: "real list \ nat" where "bn \ rec_list 1 (\a l t. if a = 0 then jn t else kn t)" have bn_simp: "bn [] = 1" "bn (a # l) = (if a = 0 then jn (bn l) else kn (bn l))" for a l by(simp_all add: bn_def) define to_listn :: "(nat \ real) \ nat \ real list" where "to_listn \ (\x . rec_nat [] (\n t. x n # t))" have to_listn_simp: "to_listn x 0 = []" "to_listn x (Suc n) = x n # to_listn x n" for x n by(simp_all add: to_listn_def) have to_listn_eq: "(\m. m < n \ x m = y m) \ to_listn x n = to_listn y n" for x y n by(induction n) (auto simp: to_listn_simp) have bn_gtn: "bn (to_listn x n) > n" for x n apply(induction n arbitrary: x) using jn(1) kn(1) by(auto simp: bn_simp to_listn_simp) (meson Suc_le_eq le_less less_trans_Suc)+ define rn where "rn \ (\n. Min (range (\x. (1 / 2 :: real) ^ bn (to_listn x n))))" have rn_fin: "finite (range (\x. (1 / 2 :: real) ^ bn (to_listn x n)))" for n proof - have "finite (range (\x. bn (to_listn x n)))" proof(induction n) case ih:(Suc n) have "(range (\x. bn (to_listn x (Suc n)))) \ (range (\x. jn (bn (to_listn x n)))) \ (range (\x. kn (bn (to_listn x n))))" by(auto simp: to_listn_simp bn_simp) moreover have "finite ..." using ih finite_range_imageI by auto ultimately show ?case by(rule finite_subset) qed(simp add: to_listn_simp) thus ?thesis using finite_range_imageI by blast qed have rn_nen: "(range (\x. (1 / 2 :: real) ^ bn (to_listn x n))) \ {}" for n by simp have rn_pos: "0 < rn n" for n by(simp add: Min_gr_iff[OF rn_fin rn_nen] rn_def) have rn_less: "rn n < (1/2)^n" for n using bn_gtn[of n] by(auto simp: rn_def Min_less_iff[OF rn_fin rn_nen]) have cball_le_ball:"md.mcball (xn (bn (a#l))) ((1/2)^(bn (a#l))) \ md.mball (xn (bn l)) ((1/2) ^ (bn l))" for a l using kn(2)[of "bn l"] less_imp_le[OF jn(1)] jn(2) md.mball_subset_concentric[of "(1 / 2) ^ jn (bn l)" "(1 / 2) ^ bn l" "xn (bn l)"] by(auto simp: bn_simp) hence cball_le:"md.mcball (xn (bn (a#l))) ((1/2)^(bn (a#l))) \ md.mcball (xn (bn l)) ((1/2) ^ (bn l))" for a l using md.mball_subset_mcball by blast have cball_disj: "md.mcball (xn (bn (0#l))) ((1/2)^(bn (0#l))) \ md.mcball (xn (bn (1#l))) ((1/2)^(bn (1#l))) = {}" for l using jn(2) kn(2) by(auto simp: bn_simp) have "\x. \l. l \ P \ (\n. md.mcball (xn (bn (to_listn x n))) ((1 / 2) ^ bn (to_listn x n))) = {l}" proof fix x show "\l. l \ P \ (\n. md.mcball (xn (bn (to_listn x n))) ((1 / 2) ^ bn (to_listn x n))) = {l}" proof(safe intro!: md.mcomplete_nest_sing[THEN iffD1,OF mdc,rule_format]) show "md.mcball (xn (bn (to_listn x n))) ((1 / 2) ^ bn (to_listn x n)) = {} \ False" for n using md.mcball_eq_empty xn(4) by auto next show "decseq (\n. md.mcball (xn (bn (to_listn x n))) ((1 / 2) ^ bn (to_listn x n)))" by(intro decseq_SucI,simp add: to_listn_simp cball_le) next fix e :: real assume "0 < e" then obtain N where N: "(1 / 2) ^ N < e" by (meson reals_power_lt_ex rel_simps(49) rel_simps(9)) show "\n a. md.mcball (xn (bn (to_listn x n))) ((1 / 2) ^ bn (to_listn x n)) \ md.mcball a e" proof(safe intro!: exI[where x=N] exI[where x="xn (bn (to_listn x N))"]) fix y assume "y \ md.mcball (xn (bn (to_listn x N))) ((1 / 2) ^ bn (to_listn x N))" then have "y \ md.mcball (xn (bn (to_listn x N))) ((1 / 2) ^ N)" using md.mcball_subset_concentric[OF power_decreasing[OF less_imp_le[OF bn_gtn[of N x]],of "1/2"]] by fastforce thus "y \ md.mcball (xn (bn (to_listn x N))) e" using N \0 < e\ by auto qed qed qed then obtain f where f:"\x. f x \ P" "\x. (\n. md.mcball (xn (bn (to_listn x n))) ((1 / 2) ^ bn (to_listn x n))) = {f x}" by metis hence f': "\n x. f x \ md.mcball (xn (bn (to_listn x n))) ((1 / 2) ^ bn (to_listn x n))" by blast have f'': "f x \ md.mball (xn (bn (to_listn x n))) ((1 / 2) ^ bn (to_listn x n))" for n x using f'[of x "Suc n"] cball_le_ball[of _ "to_listn x n"] by(fastforce simp: to_listn_simp) interpret bdmd: Submetric P d "f ` (\\<^sub>E i\UNIV. {0,1})" by standard (use f in auto) have bdmd_sub: "bdmd.sub.mtopology = subtopology X (f ` (\\<^sub>E i\UNIV. {0,1}))" using f(1) Int_absorb1[of "f ` (UNIV \\<^sub>E {0, 1})" P] by(fastforce simp: bdmd.mtopology_submetric d subtopology_subtopology) let ?d = "\x y. if (x = 0 \ x = 1) \ (y = 0 \ y = 1) then dist x y else 0" interpret d01: Metric_space "{0,1::real}" ?d by(auto simp: Metric_space_def) have d01: "d01.mtopology = top_of_set {0,1}" proof - have "d01.mtopology = Metric_space.mtopology {0,1} dist" by(auto intro!: Metric_space_eq_mtopology simp: Metric_space_def metric_space_class.dist_commute) also have "... = top_of_set {0,1}" by(auto intro!: Submetric.mtopology_submetric[of UNIV dist "{0,1::real}",simplified] simp: Submetric_def Metric_space_def Submetric_axioms_def dist_real_def) finally show ?thesis . qed interpret pd: Product_metric "1/2" UNIV id id "\_. {0,1::real}" "\_. ?d" 1 by(auto intro!: product_metric_natI d01.Metric_space_axioms) have mpd_top: "pd.Product_metric.mtopology = Cantor_space_topology" by(auto simp: pd.Product_metric_mtopology_eq[symmetric] d01 intro!: product_topology_cong) define def_at where "def_at x y \ LEAST n. x n \ y n" for x y :: "nat \ real" have def_atxy: "\n. n < def_at x y \ x n = y n" "x (def_at x y) \ y (def_at x y)" if "x \ y" for x y proof - have "\n. x n \ y n" using that by auto from LeastI_ex[OF this] show "\n. n < def_at x y \ x n = y n" "x (def_at x y) \ y (def_at x y)" using not_less_Least by(auto simp: def_at_def) qed have def_at_le_if: "pd.product_dist x y \ (1/2)^n \ n \ def_at x y" if assm:"x \ y" "x \ (\\<^sub>E i\UNIV. {0,1})" "y \ (\\<^sub>E i\UNIV. {0,1})" for x y n proof - assume h:"pd.product_dist x y \ (1 / 2) ^ n" have "x m = y m" if m_less_n: "m < n" for m proof(rule ccontr) assume nen: "x m \ y m" then have "?d (x m) (y m) = 1" using assm(2,3) by(auto simp: submetric_def) hence "1 \ 2 ^ m * pd.product_dist x y" using pd.product_dist_geq[of m m,simplified,OF assm(2,3)] by simp hence "(1/2)^m \ 2^m * (1/2)^m * pd.product_dist x y" by simp hence "(1/2)^m \ pd.product_dist x y" by (simp add: power_one_over) also have "... \ (1 / 2) ^ n" by(simp add: h) finally show False using that by auto qed thus "n \ def_at x y" by (meson def_atxy(2) linorder_not_le that(1)) qed have def_at_le_then: "pd.product_dist x y \ 2 * (1/2)^n" if assm:"x \ y" "x \ (\\<^sub>E i\UNIV. {0,1})" "y \ (\\<^sub>E i\UNIV. {0,1})" "n \ def_at x y" for x y n proof - have "\m. m < n \ x m = y m" by (metis def_atxy(1) order_less_le_trans that(4)) hence 1:"\m. m < n \ ?d (x m) (y m) = 0" by (simp add: submetric_def) have "pd.product_dist x y = (\i. (1/2)^(i + n) * (?d (x (i + n)) (y (i + n)))) + (\ii. (1/2)^(i + n) * (?d (x (i + n)) (y (i + n))))" by(simp add: 1) also have "... \ (\i. (1/2)^(i + n))" using pd.product_dist_summable' unfolding id_apply by(auto intro!: suminf_le summable_ignore_initial_segment) finally show ?thesis using pd.nsum_of_rK[of n] by simp qed have d_le_def: "d (f x) (f y) \ (1/2)^(def_at x y)" if assm:"x \ y" "x \ (\\<^sub>E i\UNIV. {0,1})" "y \ (\\<^sub>E i\UNIV. {0,1})" for x y proof - have 1:"to_listn x n = to_listn y n" if "n \ def_at x y" for n proof - have "\m. m < n \ x m = y m" by (metis def_atxy(1) order_less_le_trans that) then show ?thesis by(auto intro!: to_listn_eq) qed have "f x \ md.mcball (xn (bn (to_listn x (def_at x y)))) ((1 / 2) ^ bn (to_listn x (def_at x y)))" "f y \ md.mcball (xn (bn (to_listn x (def_at x y)))) ((1 / 2) ^ bn (to_listn x (def_at x y)))" using f'[of x "def_at x y"] f'[of y "def_at x y"] by(auto simp: 1[OF order_refl]) hence "d (f x) (f y) \ 2 * (1 / 2) ^ bn (to_listn x (def_at x y))" using f(1) by(auto intro!: md.mdiameter_is_sup'[OF _ _ md.mdiameter_cball_leq]) also have "... \ (1/2)^(def_at x y)" proof - have "Suc (def_at x y) \ bn (to_listn x (def_at x y))" using bn_gtn[of "def_at x y" x] by simp hence "(1 / 2) ^ bn (to_listn x (def_at x y)) \ (1 / 2 :: real) ^ Suc (def_at x y)" using power_decreasing_iff[OF pd.r] by blast thus ?thesis by simp qed finally show "d (f x) (f y) \ (1/2)^(def_at x y)" . qed have fy_in:"f y \ md.mcball (xn (bn (to_listn x m))) ((1/2)^bn (to_listn x m)) \ \l (\\<^sub>E i\UNIV. {0,1})" "y \ (\\<^sub>E i\UNIV. {0,1})" for x y m proof(induction m) case ih:(Suc m) have "f y \ md.mcball (xn (bn (to_listn x m))) ((1 / 2) ^ bn (to_listn x m))" using ih(2) cball_le by(fastforce simp: to_listn_simp) with ih(1) have k:"k < m \ x k = y k" for k by simp show ?case proof safe fix l assume "l < Suc m" then consider "l < m" | "l = m" using \l < Suc m\ by fastforce thus "x l = y l" proof cases case 2 have 3:"f y \ md.mcball (xn (bn (y l # to_listn y l))) ((1 / 2) ^ bn (y l # to_listn y l))" using f'[of y "Suc l"] by(simp add: to_listn_simp) have 4:"f y \ md.mcball (xn (bn (x l # to_listn y l))) ((1 / 2) ^ bn (x l # to_listn y l))" using ih(2) to_listn_eq[of m x y,OF k] by(simp add: to_listn_simp 2) show ?thesis proof(rule ccontr) assume "x l \ y l" then consider "x l = 0" "y l = 1" | "x l = 1" "y l = 0" using assm(1,2) by(auto simp: PiE_def Pi_def) metis thus False by cases (use cball_disj[of "to_listn y l"] 3 4 in auto) qed qed(simp add: k) qed qed simp have d_le_rn_then: "\e>0. \y \ (\\<^sub>E i\UNIV. {0,1}). x \ y \ d (f x) (f y) < e \ n \ def_at x y" if assm: "x \ (\\<^sub>E i\UNIV. {0,1})" for x n proof(safe intro!: exI[where x="(1/2)^bn (to_listn x n) - d (xn (bn (to_listn x n))) (f x)"]) show "0 < (1 / 2) ^ bn (to_listn x n) - d (xn (bn (to_listn x n))) (f x)" using f'' by auto next fix y assume h:"y \ (\\<^sub>E i\UNIV. {0,1})" "d (f x) (f y) < (1 / 2) ^ bn (to_listn x n) - d (xn (bn (to_listn x n))) (f x)" "x \ y" then have "f y \ md.mcball (xn (bn (to_listn x n))) ((1/2)^bn (to_listn x n))" using md.triangle[OF xn(4)[of "bn (to_listn x n)"] f(1)[of x] f(1)[of y]] by(simp add: xn(4)[of "bn (to_listn x n)"] f(1)[of y] md.mcball_def) with fy_in[OF assm h(1)] have "\m < n. x m = y m" by auto thus "n \ def_at x y" by (meson def_atxy(2) linorder_not_le h(3)) qed have 0: "f ` (\\<^sub>E i\UNIV. {0,1}) \ topspace X" using f(1) P(2) by auto have 1: "continuous_map pd.Product_metric.mtopology bdmd.sub.mtopology f" unfolding pd.Product_metric.metric_continuous_map[OF bdmd.sub.Metric_space_axioms] proof safe fix x :: "nat \ real" and \ :: real assume h:"x \ (\\<^sub>E i\UNIV. {0,1})" "0 < \" then obtain n where n:"(1/2)^n < \" using real_arch_pow_inv[OF _ pd.r(2)] by auto show "\\>0. \y. y\UNIV \\<^sub>E {0, 1} \ pd.product_dist x y < \ \ d (f x) (f y) < \" proof(safe intro!: exI[where x="(1/2)^n"]) fix y assume y:"y \ (\\<^sub>E i\UNIV. {0,1})" "pd.product_dist x y < (1 / 2) ^ n" consider "x = y" | "x \ y" by auto thus "d (f x) (f y) < \" proof cases case 1 with y(1) h md.zero[OF f(1)[of y] f(1)[of y]] show ?thesis by simp next case 2 then have "n \ def_at x y" using h(1) y by(auto intro!: def_at_le_if) have "d (f x) (f y) \ (1/2)^(def_at x y)" using h(1) y(1) by(auto simp: d_le_def[OF 2 h(1) y(1)]) also have "... \ (1/2)^n" using \n \ def_at x y\ by simp finally show ?thesis using n by simp qed qed simp qed have 2: "open_map pd.Product_metric.mtopology bdmd.sub.mtopology f" proof - have "open_map (mtopology_of pd.Product_metric.Self) (subtopology (mtopology_of md.Self) (f ` mspace pd.Product_metric.Self)) f" proof(safe intro!: Metric_space_open_map_from_dist) fix x :: "nat \ real" and \ :: real assume h:"x \ mspace pd.Product_metric.Self" "0 < \" then have x:"x \ (\\<^sub>E i\UNIV. {0,1})" by simp from h obtain n where n: "(1/2)^n < \" using real_arch_pow_inv[OF _ pd.r(2)] by auto obtain e where e: "e > 0" "\y. y \ (\\<^sub>E i\UNIV. {0,1}) \ x \ y \ d (f x) (f y) < e \ Suc n \ def_at x y" using d_le_rn_then[OF x,of "Suc n"] by auto show "\\>0. \y\mspace pd.Product_metric.Self. mdist md.Self (f x) (f y) < \ \ mdist pd.Product_metric.Self x y < \" unfolding md.mdist_Self pd.Product_metric.mspace_Self pd.Product_metric.mdist_Self proof(safe intro!: exI[where x=e]) fix y assume y:"y \ (\\<^sub>E i\UNIV. {0,1})" and "d (f x) (f y) < e" then have d':"d (f x) (f y) < e" using h(1) by simp consider "x = y" | "x \ y" by auto thus "pd.product_dist x y < \" by cases (use pd.Product_metric.zero[OF y y] h(2) def_at_le_then[OF _ x y e(2)[OF y _ d']] n in auto) qed(use e(1) in auto) qed(use f in auto) thus ?thesis by (simp add: bdmd.mtopology_submetric mtopology_of_def) qed have 3: "f ` (topspace pd.Product_metric.mtopology) = topspace bdmd.sub.mtopology" by simp have 4: "inj_on f (topspace pd.Product_metric.mtopology)" unfolding pd.Product_metric.topspace_mtopology proof fix x y assume h:"x \ (\\<^sub>E i\UNIV. {0,1})" "y \ (\\<^sub>E i\UNIV. {0,1})" "f x = f y" show "x = y" proof fix n have "f y \ md.mcball (xn (bn (to_listn x (Suc n)))) ((1/2)^bn (to_listn x (Suc n)))" using f'[of x "Suc n"] by(simp add: h) thus "x n = y n" using fy_in[OF h(1,2),of "Suc n"] by simp qed qed show ?thesis using homeomorphic_map_imp_homeomorphic_space[OF bijective_open_imp_homeomorphic_map[OF 1 2 3 4]] 0 by (metis (no_types, lifting) assms(1) bdmd_sub completely_metrizable_space_homeo_image_gdelta_in mpd_top polish_space_Cantor_space polish_space_def) qed subsection \Borel Spaces generated from Polish Spaces\ lemma closedin_clopen_topology: assumes "polish_space X" "closedin X a" shows "\X'. polish_space X' \ (\u. openin X u \ openin X' u) \ topspace X = topspace X' \ sets (borel_of X) = sets (borel_of X') \ openin X' a \ closedin X' a" proof - have p1:"polish_space (subtopology X a)" by (simp add: assms polish_space_closedin) then obtain da' where da': "Metric_space a da'" "subtopology X a = Metric_space.mtopology a da'" "Metric_space.mcomplete a da'" by (metis Metric_space.topspace_mtopology assms(2) closedin_subset completely_metrizable_space_def polish_space_imp_completely_metrizable_space topspace_subtopology_subset) define da where "da = Metric_space.capped_dist da' (1/2)" have da: "subtopology X a = Metric_space.mtopology a da" "Metric_space.mcomplete a da" using da' by(auto simp: da_def Metric_space.mtopology_capped_metric Metric_space.mcomplete_capped_metric) interpret pa: Metric_space a da using da' by(simp add: Metric_space.capped_dist da_def) have da_bounded: "\x y. da x y < 1" using da' by(auto simp: da_def Metric_space.capped_dist_def) have p2:"polish_space (subtopology X (topspace X - a))" by (meson assms(1) assms(2) closedin_def polish_space_openin) then obtain db' where db': "Metric_space (topspace X - a) db'" "subtopology X (topspace X - a) = Metric_space.mtopology (topspace X - a) db'" "Metric_space.mcomplete (topspace X - a) db'" by (metis Diff_subset Metric_space.topspace_mtopology completely_metrizable_space_def polish_space_imp_completely_metrizable_space topspace_subtopology_subset) define db where "db = Metric_space.capped_dist db' (1/2)" have db: "subtopology X (topspace X - a) = Metric_space.mtopology (topspace X - a) db" "Metric_space.mcomplete (topspace X - a) db" using db' by(auto simp: db_def Metric_space.mtopology_capped_metric Metric_space.mcomplete_capped_metric) interpret pb: Metric_space "topspace X - a" db using db' by(simp add: Metric_space.capped_dist db_def) have db_bounded: "\x y. db x y < 1" using db' by(auto simp: db_def Metric_space.capped_dist_def) interpret p: Sum_metric UNIV "\b. if b then a else topspace X - a" "\b. if b then da else db" using da db da_bounded db_bounded by(auto intro!: sum_metricI simp: disjoint_family_on_def pa.Metric_space_axioms pb.Metric_space_axioms) have 0: "(\i. if i then a else topspace X - a) = topspace X" using closedin_subset assms by auto have 1: "sets (borel_of X) = sets (borel_of p.Sum_metric.mtopology)" proof - have "sigma_sets (topspace X) (Collect (openin X)) = sigma_sets (topspace X) (Collect (openin p.Sum_metric.mtopology))" proof(rule sigma_sets_eqI) fix a assume "a \ Collect (openin X)" then have "openin p.Sum_metric.mtopology a" by(simp only: p.openin_mtopology_iff) (auto simp: 0 da(1)[symmetric] db(1)[symmetric] openin_subtopology dest: openin_subset) thus "a \ sigma_sets (topspace X) (Collect (openin p.Sum_metric.mtopology))" by auto next interpret s: sigma_algebra "topspace X" "sigma_sets (topspace X) (Collect (openin X))" by(auto intro!: sigma_algebra_sigma_sets openin_subset) fix b assume "b \ Collect (openin p.Sum_metric.mtopology)" then have "openin p.Sum_metric.mtopology b" by auto then have b:"b \ topspace X" "openin (subtopology X a) (b \ a)" "openin (subtopology X (topspace X - a)) (b \ (topspace X - a))" by(simp_all only: p.openin_mtopology_iff,insert 0 da(1) db(1)) (auto simp: all_bool_eq) have [simp]: "(b \ a) \ (b \ (topspace X - a)) = b" using Diff_partition b(1) by blast have "(b \ a) \ (b \ (topspace X - a)) \ sigma_sets (topspace X) (Collect (openin X))" proof(rule sigma_sets_Un) have [simp]:"a \ sigma_sets (topspace X) (Collect (openin X))" proof - have "topspace X - (topspace X - a) \ sigma_sets (topspace X) (Collect (openin X))" by(rule sigma_sets.Compl) (use assms in auto) thus ?thesis using double_diff[OF closedin_subset[OF assms(2)]] by simp qed from b(2,3) obtain T T' where T:"openin X T" "openin X T'" and [simp]:"b \ a = T \ a" "b \ (topspace X - a) = T' \ (topspace X - a)" by(auto simp: openin_subtopology) show "b \ a \ sigma_sets (topspace X) (Collect (openin X))" "b \ (topspace X - a) \ sigma_sets (topspace X) (Collect (openin X))" using T assms by auto qed thus "b \ sigma_sets (topspace X) (Collect (openin X))" by simp qed thus ?thesis by(simp only: sets_borel_of p.Sum_metric.topspace_mtopology) (use 0 in auto) qed have 2:"\u. openin X u \ openin p.Sum_metric.mtopology u" by(simp only: p.openin_mtopology_iff) (auto simp: all_bool_eq da(1)[symmetric] db(1)[symmetric] openin_subtopology dest: openin_subset) have 3:"openin p.Sum_metric.mtopology a" by(simp only: p.openin_mtopology_iff) (auto simp: all_bool_eq) have 4:"closedin p.Sum_metric.mtopology a" by (metis 0 2 assms(2) closedin_def p.Sum_metric.topspace_mtopology) have 5: "topspace X = topspace p.Sum_metric.mtopology" by(simp only: p.Sum_metric.topspace_mtopology) (simp only: 0) have 6: "polish_space p.Sum_metric.mtopology" by(rule p.polish_spaceI,insert da(2) db(2) p1 p2) (auto simp: da(1) db(1) polish_space_def) show ?thesis by(rule exI[where x=p.Sum_metric.mtopology]) (insert 5 2 6, simp only: 1 3 4 ,auto) qed lemma polish_space_union_polish: fixes X :: "nat \ 'a topology" assumes "\n. polish_space (X n)" "\n. topspace (X n) = Xt" "\x y. x \ Xt \ y \ Xt \ x \ y \ \Ox Oy. (\n. openin (X n) Ox) \ (\n. openin (X n) Oy) \ x \ Ox \ y \ Oy \ disjnt Ox Oy" defines "Xun \ topology_generated_by (\n. {u. openin (X n) u})" shows "polish_space Xun" proof - have topsXun:"topspace Xun = Xt" using assms(2) by(auto simp: Xun_def dest:openin_subset) define f :: "'a \ nat \ 'a" where "f \ (\x n. x)" have "continuous_map Xun (product_topology X UNIV) f" by(auto simp: assms(2) topsXun f_def continuous_map_componentwise, auto simp: Xun_def openin_topology_generated_by_iff continuous_map_def assms(2) dest:openin_subset[of "X _",simplified assms(2)] ) (insert openin_subopen, fastforce intro!: generate_topology_on.Basis) hence 1: "continuous_map Xun (subtopology (product_topology X UNIV) (f ` (topspace Xun))) f" by(auto simp: continuous_map_in_subtopology) have 2: "inj_on f (topspace Xun)" by(auto simp: inj_on_def f_def dest:fun_cong) have 3: "f ` (topspace Xun) = topspace (subtopology (product_topology X UNIV) (f ` (topspace Xun)))" by(auto simp: topsXun assms(2) f_def) have 4: "open_map Xun (subtopology (product_topology X UNIV) (f ` (topspace Xun))) f" proof(safe intro!: open_map_generated_topo[OF _ 2[simplified Xun_def],simplified Xun_def[symmetric]]) fix u n assume u:"openin (X n) u" show "openin (subtopology (product_topology X UNIV) (f ` topspace Xun)) (f ` u)" unfolding openin_subtopology proof(safe intro!: exI[where x="{ \i. if i = n then a else b i |a b. a \u \ b \ UNIV \ Xt}"]) show "openin (product_topology X UNIV) {\i. if i = n then a else b i |a b. a \u \ b \ UNIV \ Xt}" by(auto simp: openin_product_topology_alt u assms(2) openin_topspace[of "X _",simplified assms(2)] intro!: exI[where x="\i. if i = n then u else Xt"]) (auto simp: PiE_def Pi_def, metis openin_subset[OF u,simplified assms(2)] in_mono) next show "\y. y \ u \ \a b. f y = (\i. if i = n then a else b i) \ a \ u \ b \ UNIV \ Xt" using assms(2) f_def openin_subset u by fastforce next show "\y. y \ u \ f y \ f ` topspace Xun" using openin_subset[OF u] by(auto simp: assms(2) topsXun) next show "\x xa a b. xa \ topspace Xun \ f xa = (\i. if i = n then a else b i) \ a \ u \ b \ UNIV \ Xt \ f xa \ f ` u" using openin_subset[OF u] by(auto simp: topsXun assms(2)) (metis f_def imageI) qed qed have 5:"(subtopology (product_topology X UNIV) (f ` topspace Xun)) homeomorphic_space Xun" using homeomorphic_map_imp_homeomorphic_space[OF bijective_open_imp_homeomorphic_map[OF 1 4 3 2]] by(simp add: homeomorphic_space_sym[of Xun]) show ?thesis proof(safe intro!: homeomorphic_polish_space_aux[OF polish_space_closedin[OF polish_space_product] 5] assms) show "closedin (product_topology X UNIV) (f ` topspace Xun)" proof - have 1: "openin (product_topology X UNIV) ((UNIV \\<^sub>E Xt) - f ` Xt)" proof(rule openin_subopen[THEN iffD2]) show "\x\(UNIV \\<^sub>E Xt) - f ` Xt. \T. openin (product_topology X UNIV) T \ x \ T \ T \ (UNIV \\<^sub>E Xt) - f ` Xt" proof safe fix x assume x:"x \ UNIV \\<^sub>E Xt" "x \ f ` Xt" have "\n. x n \ x 0" proof(rule ccontr) assume " \n. x n \ x 0" then have "\n. x n = x 0" by auto hence "x = (\_. x 0)" by auto thus False using x by(auto simp: f_def topsXun assms(2)) qed then obtain n where n: "n \ 0" "x n \ x 0" by metis from assms(3)[OF _ _ this(2)] x obtain On O0 where h:"\n. openin (X n) On" "\n. openin (X n) O0" "x n \ On" "x 0 \ O0" "disjnt On O0" by fastforce have "openin (product_topology X UNIV) ((\x. x 0) -` O0 \ topspace (product_topology X UNIV))" using continuous_map_product_coordinates[of 0 UNIV X] h(2)[of 0] by blast moreover have "openin (product_topology X UNIV) ((\x. x n) -` On \ topspace (product_topology X UNIV))" using continuous_map_product_coordinates[of n UNIV X] h(1)[of n] by blast ultimately have op: "openin (product_topology X UNIV) ((\T. T 0) -` O0 \ topspace (product_topology X UNIV) \ ((\T. T n) -` On \ topspace (product_topology X UNIV)))" by auto have xin:"x \ (\T. T 0) -` O0 \ topspace (product_topology X UNIV) \ ((\T. T n) -` On \ topspace (product_topology X UNIV))" using x h(3,4) by(auto simp: assms(2)) have subset:"(\T. T 0) -` O0 \ topspace (product_topology X UNIV) \ ((\T. T n) -` On \ topspace (product_topology X UNIV)) \ (UNIV \\<^sub>E Xt) - f ` Xt" using h(5) by(auto simp: assms(2) disjnt_def f_def) show "\T. openin (product_topology X UNIV) T \ x \ T \ T \ (UNIV \\<^sub>E Xt) - f ` Xt" by(rule exI[where x="((\x. x 0) -` O0 \ topspace (product_topology X UNIV)) \ ((\x. x n) -` On \ topspace (product_topology X UNIV))"]) (use op xin subset in auto) qed qed thus ?thesis by(auto simp: closedin_def assms(2) topsXun f_def) qed qed(simp add: f_def) qed lemma sets_clopen_topology: assumes "polish_space X" "a \ sets (borel_of X)" shows "\X'. polish_space X' \ (\u. openin X u \ openin X' u) \ topspace X = topspace X' \ sets (borel_of X) = sets (borel_of X') \ openin X' a \ closedin X' a" proof - have "a \ sigma_sets (topspace X) {U. closedin X U}" using assms by(simp add: sets_borel_of_closed) thus ?thesis proof induction case (Basic a) then show ?case by(simp add: assms closedin_clopen_topology) next case Empty with polish_space_axioms assms show ?case by auto next case (Compl a) then obtain X' where S':"polish_space X'" "(\u. openin X u \ openin X' u)" "topspace X = topspace X'" "sets (borel_of X) = sets (borel_of X')" "openin X' a" "closedin X' a" by auto from closedin_clopen_topology[OF S'(1) S'(6)] S' show ?case by auto next case ih:(Union a) then obtain Si where Si: "\i. polish_space (Si i)" "\u i. openin X u \ openin (Si i) u" "\i::nat. topspace X = topspace (Si i)" "\i. sets (borel_of X) = sets (borel_of (Si i))" "\i. openin (Si i) (a i)" "\i. closedin (Si i) (a i)" by metis define Sun where "Sun \ topology_generated_by (\n. {u. openin (Si n) u})" have Sun1: "polish_space Sun" unfolding Sun_def proof(safe intro!: polish_space_union_polish[where Xt="topspace X"]) fix x y assume xy:"x \ topspace X" "y \ topspace X" "x \ y" then obtain Ox Oy where Oxy: "x \ Ox" "y \ Oy" "openin X Ox" "openin X Oy" "disjnt Ox Oy" using metrizable_imp_Hausdorff_space[OF polish_space_imp_metrizable_space[OF assms(1)]] by(simp only: Hausdorff_space_def) metis show "\Ox Oy. (\x. openin (Si x) Ox) \ (\x. openin (Si x) Oy) \ x \ Ox \ y \ Oy \ disjnt Ox Oy" by(rule exI[where x=Ox],insert Si(2) Oxy, auto intro!: exI[where x=Oy]) qed (use Si in auto) have Suntop:"topspace X = topspace Sun" using Si(3) by(auto simp: Sun_def dest: openin_subset) have Sunsets: "sets (borel_of X) = sets (borel_of Sun)" (is "?lhs = ?rhs") proof - have "?lhs = sigma_sets (topspace X) (\n. {u. openin (Si n) u})" proof show "sets (borel_of X) \ sigma_sets (topspace X) (\n. {u. openin (Si n) u})" using Si(2) by(auto simp: sets_borel_of intro!: sigma_sets_mono') next show "sigma_sets (topspace X) (\n. {u. openin (Si n) u}) \ sets (borel_of X)" by(simp add: sigma_sets_le_sets_iff[of "borel_of X" "\n. {u. openin (Si n) u}",simplified space_borel_of]) (use Si(4) sets_borel_of in fastforce) qed also have "... = ?rhs" using borel_of_second_countable'[OF polish_space_imp_second_countable[OF Sun1],of "\n. {u. openin (Si n) u}"] by (simp add: Sun_def Suntop subbase_in_def subset_Pow_Union) finally show ?thesis . qed have Sun_open: "\u i. openin (Si i) u \ openin Sun u" by(auto simp: Sun_def openin_topology_generated_by_iff intro!: generate_topology_on.Basis) have Sun_opena: "openin Sun (\i. a i)" using Sun_open[OF Si(5),simplified Sun_def] by(auto simp: Sun_def openin_topology_generated_by_iff intro!: generate_topology_on.UN) hence "closedin Sun (topspace Sun - (\i. a i))" by auto from closedin_clopen_topology[OF Sun1 this] show ?case using Suntop Sunsets Sun_open[OF Si(2)] Sun_opena by (metis closedin_def openin_closedin_eq) qed qed end \ No newline at end of file diff --git a/thys/Standard_Borel_Spaces/Lemmas_StandardBorel.thy b/thys/Standard_Borel_Spaces/Lemmas_StandardBorel.thy --- a/thys/Standard_Borel_Spaces/Lemmas_StandardBorel.thy +++ b/thys/Standard_Borel_Spaces/Lemmas_StandardBorel.thy @@ -1,2576 +1,2717 @@ (* Title: Lemmas_StandardBorel.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) text \ We refer to the HOL-Analysis library, the textbooks by Matsuzaka~\cite{topology} and Srivastava~\cite{borelsets}, and the lecture note by Biskup~\cite{standardborel}.\ section \Lemmas\ theory Lemmas_StandardBorel imports "HOL-Probability.Probability" begin subsection \Lemmas for Abstract Topology\ subsubsection \ Generated By \ lemma topology_generated_by_sub: assumes "\U. U \ \ \ (openin X U)" and "openin (topology_generated_by \) U" shows "openin X U" proof - have "generate_topology_on \ U" by (simp add: assms(2) openin_topology_generated_by) then show ?thesis by induction (use assms(1) in auto) qed lemma topology_generated_by_open: "S = topology_generated_by {U | U . openin S U}" unfolding topology_eq proof standard+ fix U assume "openin (topology_generated_by {U |U. openin S U}) U" note this[simplified openin_topology_generated_by_iff] then show "openin S U" by induction auto qed(simp add: openin_topology_generated_by_iff generate_topology_on.Basis) lemma topology_generated_by_eq: assumes "\U. U \ \ \ (openin (topology_generated_by \) U)" and "\U. U \ \ \ (openin (topology_generated_by \) U)" shows "topology_generated_by \ = topology_generated_by \" using topology_generated_by_sub[of \, OF assms(1)] topology_generated_by_sub[of \,OF assms(2)] by(auto simp: topology_eq) lemma topology_generated_by_homeomorphic_spaces: assumes "homeomorphic_map X Y f" "X = topology_generated_by \" shows "Y = topology_generated_by ((`) f ` \)" unfolding topology_eq proof have f:"open_map X Y f" "inj_on f (topspace X)" using assms(1) by (simp_all add: homeomorphic_imp_open_map perfect_injective_eq_homeomorphic_map[symmetric]) obtain g where g: "\x. x \ topspace X \ g (f x) = x" "\y. y \ topspace Y \ f (g y) = y" "open_map Y X g" "inj_on g (topspace Y)" using homeomorphic_map_maps[of X Y f,simplified assms(1)] homeomorphic_imp_open_map homeomorphic_maps_map[of X Y f] homeomorphic_imp_injective_map[of Y X] by blast show "\S. openin Y S = openin (topology_generated_by ((`) f ` \)) S" proof safe fix S assume "openin Y S" then have "openin X (g ` S)" using g(3) by (simp add: open_map_def) hence h:"generate_topology_on \ (g ` S)" by(simp add: assms(2) openin_topology_generated_by_iff) have "S = f ` (g ` S)" using openin_subset[OF \openin Y S\] g(2) by(fastforce simp: image_def) also have "openin (topology_generated_by ((`) f ` \)) ..." using h proof induction case Empty then show ?case by simp next case (Int a b) with inj_on_image_Int[OF f(2),of a b] show ?case by (metis assms(2) openin_Int openin_subset openin_topology_generated_by_iff) next case (UN K) then show ?case by(auto simp: image_Union) next case (Basis s) then show ?case by(auto intro!: generate_topology_on.Basis simp: openin_topology_generated_by_iff) qed finally show "openin (topology_generated_by ((`) f ` \)) S" . next fix S assume "openin (topology_generated_by ((`) f ` \)) S" then have "generate_topology_on ((`) f ` \) S" by(simp add: openin_topology_generated_by_iff) thus "openin Y S" proof induction case (Basis s) then obtain U where u:"U \ \" "s = f ` U" by auto then show ?case using assms(1) assms(2) homeomorphic_map_openness_eq topology_generated_by_Basis by blast qed auto qed qed lemma open_map_generated_topo: assumes "\u. u \ U \ openin S (f ` u)" "inj_on f (topspace (topology_generated_by U))" shows "open_map (topology_generated_by U) S f" unfolding open_map_def proof safe fix u assume "openin (topology_generated_by U) u" then have "generate_topology_on U u" by(simp add: openin_topology_generated_by_iff) thus "openin S (f ` u)" proof induction case (Int a b) then have [simp]:"f ` (a \ b) = f ` a \ f ` b" by (meson assms(2) inj_on_image_Int openin_subset openin_topology_generated_by_iff) from Int show ?case by auto qed (simp_all add: image_Union openin_clauses(3) assms) qed lemma subtopology_generated_by: "subtopology (topology_generated_by \) T = topology_generated_by {T \ U | U. U \ \}" unfolding topology_eq openin_subtopology openin_topology_generated_by_iff proof safe fix A assume "generate_topology_on \ A" then show "generate_topology_on {T \ U |U. U \ \} (A \ T)" proof induction case Empty then show ?case by (simp add: generate_topology_on.Empty) next case (Int a b) moreover have "a \ b \ T = (a \ T) \ (b \ T)" by auto ultimately show ?case by(auto intro!: generate_topology_on.Int) next case (UN K) moreover have "(\ K \ T) = (\ { k \ T | k. k \ K})" by auto ultimately show ?case by(auto intro!: generate_topology_on.UN) next case (Basis s) then show ?case by(auto intro!: generate_topology_on.Basis) qed next fix A assume "generate_topology_on {T \ U |U. U \ \} A" then show "\L. generate_topology_on \ L \ A = L \ T" proof induction case Empty show ?case by(auto intro!: exI[where x="{}"] generate_topology_on.Empty) next case ih:(Int a b) then obtain La Lb where "generate_topology_on \ La" "a = La \ T" "generate_topology_on \ Lb" "b = Lb \ T" by auto thus ?case using ih by(auto intro!: exI[where x="La \ Lb"] generate_topology_on.Int) next case ih:(UN K) then obtain L where "\k. k \ K \ generate_topology_on \ (L k) " "\k. k \ K \ k = (L k) \ T" by metis thus ?case using ih by(auto intro!: exI[where x="\k\K. L k"] generate_topology_on.UN) next case (Basis s) then show ?case using generate_topology_on.Basis by fastforce qed qed lemma prod_topology_generated_by: "topology_generated_by { U \ V | U V. U \ \ \ V \ \} = prod_topology (topology_generated_by \) (topology_generated_by \)" unfolding topology_eq proof safe fix U assume h:"openin (topology_generated_by {U \ V |U V. U \ \ \ V \ \}) U" show "openin (prod_topology (topology_generated_by \) (topology_generated_by \)) U" by(auto simp: openin_prod_Times_iff[of "topology_generated_by \" "topology_generated_by \"] intro!: topology_generated_by_Basis topology_generated_by_sub[OF _ h]) next fix U assume "openin (prod_topology (topology_generated_by \) (topology_generated_by \)) U" then have "\z\U. \V1 V2. openin (topology_generated_by \) V1 \ openin (topology_generated_by \) V2 \ fst z \ V1 \ snd z \ V2 \ V1 \ V2 \ U" by(auto simp: openin_prod_topology_alt) hence "\V1. \z\U. \V2. openin (topology_generated_by \) (V1 z) \ openin (topology_generated_by \) V2 \ fst z \ (V1 z) \ snd z \ V2 \ (V1 z) \ V2 \ U" by(rule bchoice) then obtain V1 where "\z\U. \V2. openin (topology_generated_by \) (V1 z) \ openin (topology_generated_by \) V2 \ fst z \ (V1 z) \ snd z \ V2 \ (V1 z) \ V2 \ U" by auto hence "\V2. \z\U. openin (topology_generated_by \) (V1 z) \ openin (topology_generated_by \) (V2 z) \ fst z \ (V1 z) \ snd z \ (V2 z) \ (V1 z) \ (V2 z) \ U" by(rule bchoice) then obtain V2 where hv12:"\z. z\U \ openin (topology_generated_by \) (V1 z) \ openin (topology_generated_by \) (V2 z) \ fst z \ (V1 z) \ snd z \ (V2 z) \ (V1 z) \ (V2 z) \ U" by auto hence 1:"U = (\z\U. (V1 z) \ (V2 z))" by auto have "openin (topology_generated_by {U \ V |U V. U \ \ \ V \ \}) (\z\U. (V1 z) \ (V2 z))" proof(rule openin_Union) show "\S. S \ (\z. V1 z \ V2 z) ` U \ openin (topology_generated_by {U \ V |U V. U \ \ \ V \ \}) S" proof safe fix x y assume h:"(x,y) \ U" then have "generate_topology_on \ (V1 (x,y))" using hv12 by(auto simp: openin_topology_generated_by_iff) thus "openin (topology_generated_by {U \ V |U V. U \ \ \ V \ \}) (V1 (x, y) \ V2 (x, y))" proof induction case Empty then show ?case by auto next case (Int a b) thus ?case by (auto simp: Sigma_Int_distrib1) next case (UN K) then have "openin (topology_generated_by {U \ V |U V. U \ \ \ V \ \}) (\{ k \ V2 (x, y) | k. k \ K})" by auto moreover have "(\ {k \ V2 (x, y) |k. k \ K}) = (\ K \ V2 (x, y))" by blast ultimately show ?case by simp next case ho:(Basis s) have "generate_topology_on \ (V2 (x,y))" using h hv12 by(auto simp: openin_topology_generated_by_iff) thus ?case proof induction case Empty then show ?case by auto next case (Int a b) then show ?case by (auto simp: Sigma_Int_distrib2) next case (UN K) then have "openin (topology_generated_by {U \ V |U V. U \ \ \ V \ \}) (\ { s \ k | k. k \K})" by auto moreover have "(\ { s \ k | k. k \K}) = s \ \K" by blast ultimately show ?case by simp next case (Basis s') then show ?case using ho by(auto intro!: topology_generated_by_Basis) qed qed qed qed thus "openin (topology_generated_by {U \ V |U V. U \ \ \ V \ \}) U" using 1 by auto qed lemma prod_topology_generated_by_open: "prod_topology S S' = topology_generated_by {U \ V | U V. openin S U \ openin S' V}" using prod_topology_generated_by[of " {U |U. openin S U}" "{U |U. openin S' U}"] topology_generated_by_open[of S,symmetric] topology_generated_by_open[of S'] by auto lemma product_topology_cong: assumes "\i. i \ I \ S i = K i" shows "product_topology S I = product_topology K I" proof - have 1:"{\\<^sub>E i\I. X i |X. (\i. openin (S i) (X i)) \ finite {i. X i \ topspace (S i)}} \ {\\<^sub>E i\I. X i |X. (\i. openin (K i) (X i)) \ finite {i. X i \ topspace (K i)}}" if "\i. i \ I \ S i = K i" for S K :: "_ \ 'b topology" proof fix x assume hx:"x \ {\\<^sub>E i\I. X i |X. (\i. openin (S i) (X i)) \ finite {i. X i \ topspace (S i)}}" then obtain X where hX: "x = (\\<^sub>E i\I. X i)" "\i. openin (S i) (X i)" "finite {i. X i \ topspace (S i)}" by auto define X' where "X' \ (\i. if i \ I then X i else topspace (K i))" have "x = (\\<^sub>E i\I. X' i)" by(auto simp: hX(1) X'_def PiE_def Pi_def) moreover have "finite {i. X' i \ topspace (K i)}" using that by(auto intro!: finite_subset[OF _ hX(3)] simp: X'_def) moreover have "openin (K i) (X' i)" for i using hX(2)[of i] that[of i] by(auto simp: X'_def) ultimately show "x \ {\\<^sub>E i\I. X i |X. (\i. openin (K i) (X i)) \ finite {i. X i \ topspace (K i)}}" by(auto intro!: exI[where x="X'"]) qed have "{\\<^sub>E i\I. X i |X. (\i. openin (S i) (X i)) \ finite {i. X i \ topspace (S i)}} = {\\<^sub>E i\I. X i |X. (\i. openin (K i) (X i)) \ finite {i. X i \ topspace (K i)}}" using 1[of S K] 1[of K S] assms by auto thus ?thesis by(simp add: product_topology_def) qed lemma topology_generated_by_without_empty: "topology_generated_by \ = topology_generated_by { U \ \. U \ {}}" proof(rule topology_generated_by_eq) fix U show "U \ \ \ openin (topology_generated_by { U \ \. U \ {}}) U" by(cases "U = {}") (simp_all add: topology_generated_by_Basis) qed (simp add: topology_generated_by_Basis) lemma topology_from_bij: assumes "bij_betw f A (topspace S)" shows "homeomorphic_map (pullback_topology A f S) S f" "topspace (pullback_topology A f S) = A" proof - note h = bij_betw_imp_surj_on[OF assms] bij_betw_inv_into_left[OF assms] bij_betw_inv_into_right[OF assms] then show [simp]:"topspace (pullback_topology A f S) = A" by(auto simp: topspace_pullback_topology) show "homeomorphic_map (pullback_topology A f S) S f" by(auto simp: homeomorphic_map_maps homeomorphic_maps_def h continuous_map_pullback[OF continuous_map_id,simplified] inv_into_into intro!: exI[where x="inv_into A f"] continuous_map_pullback'[where f=f]) (metis (mono_tags, opaque_lifting) comp_apply continuous_map_eq continuous_map_id h(3) id_apply) qed lemma openin_pullback_topology': assumes "bij_betw f A (topspace S)" shows "openin (pullback_topology A f S) u \ (openin S (f ` u)) \ u \ A" unfolding openin_pullback_topology proof safe fix U assume h:"openin S U" "u = f -` U \ A" from openin_subset[OF this(1)] assms have [simp]:"f ` (f -` U \ A) = U" by(auto simp: image_def vimage_def bij_betw_def) show "openin S (f ` (f -` U \ A))" by(simp add: h) next assume "openin S (f ` u)" "u \ A" with assms show "\U. openin S U \ u = f -` U \ A" by(auto intro!: exI[where x="f ` u"] simp: bij_betw_def inj_on_def) qed subsubsection \ Isolated Point \ definition isolated_points_of :: "'a topology \ 'a set \ 'a set" (infixr "isolated'_points'_of" 80) where "X isolated_points_of A \ {x\topspace X \ A. x \ X derived_set_of A}" lemma isolated_points_of_eq: "X isolated_points_of A = {x\topspace X \ A. \U. x \ U \ openin X U \ U \ (A - {x}) = {}}" unfolding isolated_points_of_def by(auto simp: in_derived_set_of) lemma in_isolated_points_of: "x \ X isolated_points_of A \ x \ topspace X \ x \ A \ (\U. x \ U \ openin X U \ U \ (A - {x}) = {})" by(simp add: isolated_points_of_eq) lemma derived_set_of_eq: "x \ X derived_set_of A \ x \ X closure_of (A - {x})" by(auto simp: in_derived_set_of in_closure_of) subsubsection \ Perfect Set \ definition perfect_set :: "'a topology \ 'a set \ bool" where "perfect_set X A \ closedin X A \ X isolated_points_of A = {}" abbreviation "perfect_space X \ perfect_set X (topspace X)" lemma perfect_space_euclidean: "perfect_space (euclidean :: 'a :: perfect_space topology)" by(auto simp: isolated_points_of_def perfect_set_def derived_set_of_eq closure_interior) lemma perfect_setI: assumes "closedin X A" and "\x T. \x \ A; x \ T; openin X T\ \ \y\x. y \ T \ y \ A" shows "perfect_set X A" using assms by(simp add: perfect_set_def isolated_points_of_def in_derived_set_of) blast lemma perfect_spaceI: assumes "\x T. \x \ T; openin X T\ \ \y\x. y \ T" shows "perfect_space X" using assms by(auto intro!: perfect_setI) (meson in_mono openin_subset) lemma perfect_setD: assumes "perfect_set X A" shows "closedin X A" "A \ topspace X" "\x T. \x \ A; x \ T; openin X T\ \ \y\x. y \ T \ y \ A" using assms closedin_subset[of X A] by(simp_all add: perfect_set_def isolated_points_of_def in_derived_set_of) blast lemma perfect_space_perfect: "perfect_set euclidean (UNIV :: 'a :: perfect_space set)" by(auto simp: perfect_set_def in_isolated_points_of) (metis Int_Diff inf_top.right_neutral insert_Diff not_open_singleton) lemma perfect_set_subtopology: assumes "perfect_set X A" shows "perfect_space (subtopology X A)" using perfect_setD[OF assms] by(auto intro!: perfect_setI simp: inf.absorb_iff2 openin_subtopology) subsubsection \ Bases and Sub-Bases in Abstract Topology\ definition subbase_in :: "['a topology, 'a set set] \ bool" where "subbase_in S \ \ S = topology_generated_by \" definition base_in :: "['a topology, 'a set set] \ bool" where "base_in S \ \ (\U. openin S U \ (\\. U = \\ \ \ \ \))" lemma second_countable_base_in: "second_countable S \ (\\. countable \ \ base_in S \)" proof - have [simp]:"\\. (openin S = arbitrary union_of (\x. x \ \)) \ (\U. openin S U \ (\\. U = \\ \ \ \ \))" by(simp add: arbitrary_def union_of_def fun_eq_iff) metis show ?thesis by(auto simp: second_countable base_in_def) qed definition zero_dimensional :: "'a topology \ bool" where "zero_dimensional S \ (\\. base_in S \ \ (\u\\. openin S u \ closedin S u))" lemma openin_base: assumes "base_in S \ " "U = \\" and "\ \ \" shows "openin S U" using assms by(auto simp: base_in_def) lemma base_is_subbase: assumes "base_in S \" shows "subbase_in S \" unfolding subbase_in_def topology_eq openin_topology_generated_by_iff proof safe fix U assume "openin S U" then obtain \ where hu:"U = \\" "\ \ \" using assms by(auto simp: base_in_def) thus "generate_topology_on \ U" by(auto intro!: generate_topology_on.UN) (auto intro!: generate_topology_on.Basis) next fix U assume "generate_topology_on \ U" then show "openin S U" proof induction case (Basis s) then show ?case using openin_base[OF assms,of s "{s}"] by auto qed auto qed lemma subbase_in_subset: assumes "subbase_in S \" and "U \ \" shows "U \ topspace S" using assms(1)[simplified subbase_in_def] topology_generated_by_topspace assms by auto lemma subbase_in_openin: assumes "subbase_in S \" and "U \ \" shows "openin S U" using assms by(simp add: subbase_in_def openin_topology_generated_by_iff generate_topology_on.Basis) lemma base_in_subset: assumes "base_in S \" and "U \ \" shows "U \ topspace S" using subbase_in_subset[OF base_is_subbase[OF assms(1)] assms(2)] . lemma base_in_openin: assumes "base_in S \" and "U \ \" shows "openin S U" using subbase_in_openin[OF base_is_subbase[OF assms(1)] assms(2)] . lemma base_in_def2: assumes "\U. U \ \ \ openin S U" shows "base_in S \ \ (\U. openin S U \ (\x\U. \W\\. x \ W \ W \ U))" proof assume h:"base_in S \" show "\U. openin S U \ (\x\U. \W\\. x \ W \ W \ U)" proof safe fix U x assume h':"openin S U" "x \ U" then obtain \ where hu: "U = \\" "\ \ \" using h by(auto simp: base_in_def) then obtain W where "x \ W" "W \ \" using h'(2) by blast thus "\W\\. x \ W \ W \ U" using hu by(auto intro!: bexI[where x=W]) qed next assume h:"\U. openin S U \ (\x\U. \W\\. x \ W \ W \ U)" show "base_in S \" unfolding base_in_def proof safe fix U assume "openin S U" then have "\x\U. \W. W\\ \ x \ W \ W \ U" using h by blast hence "\W. \x\U. W x \ \ \ x \ W x \ W x \ U" by(rule bchoice) then obtain W where hw: "\x\U. W x \ \ \ x \ W x \ W x \ U" by auto thus "\\. U = \ \ \ \ \ \" by(auto intro!: exI[where x="W ` U"]) next fix U \ show "\ \ \ \ openin S (\ \)" using assms by auto qed qed lemma base_in_def2': "base_in S \ \ (\b\\. openin S b) \ (\x. openin S x \ (\B'\\. \ B' = x))" proof assume h:"base_in S \" show "(\b\\. openin S b) \ (\x. openin S x \ (\B'\\. \ B' = x))" proof(rule conjI) show "\b\\. openin S b" using openin_base[OF h,of _ "{_}"] by auto next show "\x. openin S x \ (\B'\\. \ B' = x)" using h by(auto simp: base_in_def) qed next assume h:"(\b\\. openin S b) \ (\x. openin S x \ (\B'\\. \ B' = x))" show "base_in S \" unfolding base_in_def proof safe fix U assume "openin S U" then obtain B' where "B'\\" "\ B' = U" using h by blast thus "\\. U = \ \ \ \ \ \" by(auto intro!: exI[where x=B']) next fix U \ show "\ \ \ \ openin S (\ \)" using h by auto qed qed corollary base_in_in_subset: assumes "base_in S \" "openin S u" "x \ u" shows "\v\\. x \ v \ v \ u" using assms base_in_def2 base_in_def2' by fastforce lemma base_in_without_empty: assumes "base_in S \" shows "base_in S {U \ \. U \ {}}" unfolding base_in_def2' proof safe fix x assume "x \ \" " \ openin S x" thus "\y. y \ {}" using base_in_openin[OF assms \x \ \\] by simp next fix x assume "openin S x" then obtain B' where "B' \\" "\ B' = x" using assms by(simp add: base_in_def2') metis thus "\B'\{U \ \. U \ {}}. \ B' = x" by(auto intro!: exI[where x="{y \ B'. y \ {}}"]) qed lemma second_countable_ex_without_empty: assumes "second_countable S" shows "\\. countable \ \ base_in S \ \ (\U\\. U \ {})" proof - obtain \ where "countable \" "base_in S \" using assms second_countable_base_in by blast thus ?thesis by(auto intro!: exI[where x="{U \ \. U \ {}}"] base_in_without_empty) qed lemma subtopology_subbase_in: assumes "subbase_in S \" shows "subbase_in (subtopology S T) {T \ U | U. U \ \}" using assms subtopology_generated_by by(auto simp: subbase_in_def) lemma subtopology_base_in: assumes "base_in S \" shows "base_in (subtopology S T) {T \ U | U. U \ \}" unfolding base_in_def proof fix L show "openin (subtopology S T) L = (\\. L = \ \ \ \ \ {T \ U |U. U \ \})" proof assume "openin (subtopology S T) L " then obtain T' where ht: "openin S T'" "L = T' \ T" by(auto simp: openin_subtopology) then obtain \ where hu: "T' = (\ \)" "\ \ \" using assms by(auto simp: base_in_def) show "\\. L = \ \ \ \ \ {T \ U |U. U \ \}" using hu ht by(auto intro!: exI[where x="{T \ U | U. U \ \}"]) next assume "\\. L = \ \ \ \ \ {T \ U |U. U \ \}" then obtain \ where hu: "L = \ \" "\ \ {T \ U |U. U \ \}" by auto hence "\U\\. \U'\\. U = T \ U'" by blast then obtain k where hk:"\U. U \ \ \ k U \ \" "\U. U \ \ \ U = T \ k U" by metis hence "L = \ {T \ k U |U. U \ \}" using hu by auto also have "... = \ {k U |U. U \ \} \ T" by auto finally have 1:"L = \ {k U |U. U \ \} \ T" . moreover have "openin S (\ {k U |U. U \ \})" using hu hk assms by(auto simp: base_in_def) ultimately show "openin (subtopology S T) L" by(auto intro!: exI[where x="\ {k U |U. U \ \}"] simp: openin_subtopology) qed qed lemma second_countable_subtopology: assumes "second_countable S" shows "second_countable (subtopology S T)" proof - obtain \ where "countable \" "base_in S \" using assms second_countable_base_in by blast thus ?thesis by(auto intro!: exI[where x="{T \ U | U. U \ \}"] simp: second_countable_base_in Setcompr_eq_image dest: subtopology_base_in) qed lemma open_map_with_base: assumes "base_in S \" "\A. A \ \ \ openin S' (f ` A)" shows "open_map S S' f" unfolding open_map_def proof safe fix U assume "openin S U" then obtain \ where "U = \\" "\ \ \" using assms(1) by(auto simp: base_in_def) hence "f ` U = \{ f ` A | A. A \ \}" by blast also have "openin S' ..." using assms(2) \\ \ \\ by auto finally show "openin S' (f ` U)" . qed text \ Construct a base from a subbase.\ lemma finite'_intersection_of_idempot [simp]: "finite' intersection_of finite' intersection_of P = finite' intersection_of P" proof fix A show "(finite' intersection_of finite' intersection_of P) A = (finite' intersection_of P) A" proof assume "(finite' intersection_of finite' intersection_of P) A" then obtain \ where \:"finite' \ \ \ \ Collect (finite' intersection_of P) \ \\ = A" by(auto simp: intersection_of_def) hence "\U\\. \\'. finite' \' \ \' \ Collect P \ \\' = U" by(auto simp: intersection_of_def) then obtain \' where \': "\U. U \ \ \ finite' (\' U)" "\U. U \ \ \ \' U \ Collect P" "\U. U \ \ \ \(\' U) = U" by metis have 1: "\ (\ (\' ` \)) = A" using \ \'(3) by blast show "(finite' intersection_of P) A" unfolding intersection_of_def using \ \'(1,2) 1 by(auto intro!: exI[where x="\U\\. \' U"]) qed(rule finite'_intersection_of_inc) qed lemma finite'_intersection_of_countable: assumes "countable \" shows "countable (Collect (finite' intersection_of (\x. x \ \)))" proof - have "Collect (finite' intersection_of (\x. x \ \)) = (\i\{\'. \' \ {} \ finite \' \ \' \ \}. {\ i})" by(auto simp: intersection_of_def) also have "countable ..." using countable_Collect_finite_subset[OF assms] by(auto intro!: countable_UN[of "{ \'. \' \ {} \ finite \' \ \' \ \}" "\\'. {\\'}"]) (auto intro!: countable_subset[of "{\'. \' \ {} \ finite \' \ \' \ \}" "{A. finite A \ A \ \}"]) finally show ?thesis . qed lemma finite'_intersection_of_openin: assumes "(finite' intersection_of (\x. x \ \)) U" shows "openin (topology_generated_by \) U" unfolding openin_topology_generated_by_iff using assms by(auto simp: generate_topology_on_eq arbitrary_union_of_inc) lemma topology_generated_by_finite_intersections: "topology_generated_by \ = topology_generated_by (Collect (finite' intersection_of (\x. x \ \)))" unfolding topology_eq openin_topology_generated_by_iff by(simp add: generate_topology_on_eq) lemma base_from_subbase: assumes "subbase_in S \" shows "base_in S (Collect (finite' intersection_of (\x. x \ \)))" unfolding subbase_in_def base_in_def assms[simplified subbase_in_def] openin_topology_generated_by_iff by(auto simp: arbitrary_def union_of_def generate_topology_on_eq) lemma countable_base_from_countable_subbase: assumes "countable \" and "subbase_in S \" shows "second_countable S" using finite'_intersection_of_countable[OF assms(1)] base_from_subbase[OF assms(2)] by(auto simp: second_countable_base_in) lemma prod_topology_second_countable: assumes "second_countable S" and "second_countable S'" shows "second_countable (prod_topology S S')" proof - obtain \ \' where ho: "countable \" "base_in S \" "countable \'" "base_in S' \'" using assms by(auto simp: second_countable_base_in) show ?thesis proof(rule countable_base_from_countable_subbase[where \="{ U \ V | U V. U \ \ \ V \ \'}"]) have "{U \ V |U V. U \ \ \ V \ \'} = (\(U,V). U \ V) ` (\ \ \')" by auto also have "countable ..." using ho(1,3) by auto finally show "countable {U \ V |U V. U \ \ \ V \ \'}" . next show "subbase_in (prod_topology S S') {U \ V |U V. U \ \ \ V \ \'}" using base_is_subbase[OF ho(2)] base_is_subbase[OF ho(4)] by(simp add: subbase_in_def prod_topology_generated_by) qed qed text \ Abstract version of the theorem @{thm product_topology_countable_basis}.\ lemma product_topology_countable_base_in: assumes "countable I" and "\i. i \ I \ second_countable (S i)" shows "\\'. countable \' \ base_in (product_topology S I) \' \ (\k \ \'. \X. k = (\\<^sub>E i\I. X i) \ (\i. openin (S i) (X i)) \ finite {i. X i \ topspace (S i)} \ {i. X i \ topspace (S i)} \ I)" proof - obtain \ where ho: "\i. i \ I \ countable (\ i)" "\i. i \ I \ base_in (S i) (\ i)" using assms(2)[simplified second_countable_base_in] by metis show ?thesis unfolding second_countable_base_in proof(intro exI[where x="{\\<^sub>E i\I. U i | U. finite {i\I. U i \ topspace (S i)} \ (\i\{i\I. U i \ topspace (S i)}. U i \ \ i)}"] conjI) show "countable {\\<^sub>E i\I. U i | U. finite {i\I. U i \ topspace (S i)} \ (\i\{i\I. U i \ topspace (S i)}. U i \ \ i)}" (is "countable ?X") proof - have "?X = {\\<^sub>E i\I. U i | U. finite {i\I. U i \ topspace (S i)} \ (\i\{i\I. U i \ topspace (S i)}. U i \ \ i) \ (\i \(UNIV- I). U i = {undefined})}" (is "_ = ?Y") proof (rule set_eqI) show "\x. x \ ?X \ x \ ?Y" proof fix x assume "x \ ?X" then obtain U where hu: "x = (\\<^sub>E i\I. U i)" "finite {i\I. U i \ topspace (S i)}" "(\i\{i\I. U i \ topspace (S i)}. U i \ \ i)" by auto define U' where "U' i \ (if i \ I then U i else {undefined})" for i have "x = (\\<^sub>E i\I. U' i)" using hu(1) by(auto simp: U'_def PiE_def extensional_def Pi_def) moreover have "finite {i\I. U' i \ topspace (S i)}" "(\i\{i\I. U' i \ topspace (S i)}. U' i \ \ i)" "\i \(UNIV- I). U' i = {undefined}" using hu(2,3) by(auto simp: U'_def) (metis (mono_tags, lifting) Collect_cong) ultimately show "x \ ?Y" by auto qed auto qed also have "... = (\U. \\<^sub>E i\I. U i) ` {U. finite {i\I. U i \ topspace (S i)} \ (\i\{i\I. U i \ topspace (S i)}. U i \ \ i) \ (\i \(UNIV- I). U i = {undefined})}" by auto also have "countable ..." proof(rule countable_image) have "{U. finite {i \ I. U i \ topspace (S i)} \ (\i\{i \ I. U i \ topspace (S i)}. U i \ \ i) \ (\i\UNIV - I. U i = {undefined})} = {U. \I'. finite I' \ I' \ I \ (\i\I'. U i \ \ i) \ (\i\(I - I'). U i = topspace (S i)) \ (\i\UNIV - I. U i = {undefined})}" (is "?A = ?B") proof (rule set_eqI) show "\x. x \ ?A \ x \ ?B" proof fix U assume "U \ {U. finite {i \ I. U i \ topspace (S i)} \ (\i\{i \ I. U i \ topspace (S i)}. U i \ \ i) \ (\i\UNIV - I. U i = {undefined})}" then show "U \ {U. \I'. finite I' \ I' \ I \ (\i\I'. U i \ \ i) \ (\i\I - I'. U i = topspace (S i)) \ (\i\UNIV - I. U i = {undefined})}" by auto next fix U assume assm:"U \ {U. \I'. finite I' \ I' \ I \ (\i\I'. U i \ \ i) \ (\i\I - I'. U i = topspace (S i)) \ (\i\UNIV - I. U i = {undefined})}" then obtain I' where hi': "finite I'" "I' \ I" "\i\I'. U i \ \ i" "\i\I - I'. U i = topspace (S i)" "\i\UNIV - I. U i = {undefined}" by auto then have "\i. i \ I \ U i \ topspace (S i) \ i \ I'" by auto hence "{i \ I. U i \ topspace (S i)} \ I'" by auto hence "finite {i \ I. U i \ topspace (S i)}" using hi'(1) by (simp add: rev_finite_subset) thus "U \ {U. finite {i \ I. U i \ topspace (S i)} \ (\i\{i \ I. U i \ topspace (S i)}. U i \ \ i) \ (\i\UNIV - I. U i = {undefined})}" using hi' by auto qed qed also have "... = (\I'\{I'. finite I' \ I' \ I}. {U. (\i\I'. U i \ \ i) \ (\i\I - I'. U i = topspace (S i)) \ (\i\UNIV - I. U i = {undefined})})" by auto also have "countable ..." proof(rule countable_UN[OF countable_Collect_finite_subset[OF assms(1)]]) fix I' assume "I' \ {I'. finite I' \ I' \ I}" hence hi':"finite I'" "I' \ I" by auto have "(\U i. if i \ I' then U i else undefined) ` {U. (\i\I'. U i \ \ i) \ (\i\I - I'. U i = topspace (S i)) \ (\i\UNIV - I. U i = {undefined})} \ (\\<^sub>E i\I'. \ i)" by auto moreover have "countable ..." using hi' by(auto intro!: countable_PiE ho) ultimately have "countable ((\U i. if i \ I' then U i else undefined) ` {U. (\i\I'. U i \ \ i) \ (\i\I - I'. U i = topspace (S i)) \ (\i\UNIV - I. U i = {undefined})})" by(simp add: countable_subset) moreover have "inj_on (\U i. if i \ I' then U i else undefined) {U. (\i\I'. U i \ \ i) \ (\i\I - I'. U i = topspace (S i)) \ (\i\UNIV - I. U i = {undefined})}" (is "inj_on ?f ?X") proof fix x y assume hxy: "x \ ?X" "y \ ?X" "?f x = ?f y" show "x = y" proof fix i consider "i \ I'" | "i \ I - I'" | "i \ UNIV - I" using hi'(2) by blast then show "x i = y i" proof cases case i:1 then show ?thesis using fun_cong[OF hxy(3),of i] by auto next case i:2 then show ?thesis using hxy(1,2) by auto next case i:3 then show ?thesis using hxy(1,2) by auto qed qed qed ultimately show "countable {U. (\i\I'. U i \ \ i) \ (\i\I - I'. U i = topspace (S i)) \ (\i\UNIV - I. U i = {undefined})}" using countable_image_inj_on by auto qed finally show "countable {U. finite {i \ I. U i \ topspace (S i)} \ (\i\{i \ I. U i \ topspace (S i)}. U i \ \ i) \ (\i\UNIV - I. U i = {undefined})}" . qed finally show ?thesis . qed next show "base_in (product_topology S I) {\\<^sub>E i\I. U i |U. finite {i \ I. U i \ topspace (S i)} \ (\i\{i \ I. U i \ topspace (S i)}. U i \ \ i)}" (is "base_in (product_topology S I) ?X") unfolding base_in_def proof safe fix U assume "openin (product_topology S I) U" then have "\x\U. \Ux. finite {i \ I. Ux i \ topspace (S i)} \ (\i\I. openin (S i) (Ux i)) \ x \ Pi\<^sub>E I Ux \ Pi\<^sub>E I Ux \ U" by(simp add: openin_product_topology_alt) hence "\Ux. \x\U. finite {i \ I. Ux x i \ topspace (S i)} \ (\i\I. openin (S i) (Ux x i)) \ x \ Pi\<^sub>E I (Ux x) \ Pi\<^sub>E I (Ux x) \ U" by(rule bchoice) then obtain Ux where hui: "\x. x \ U \ finite {i \ I. Ux x i \ topspace (S i)}" "\x i. x \ U \ i \ I \ openin (S i) (Ux x i)" "\x. x \ U \ x \ Pi\<^sub>E I (Ux x)" "\x. x \ U \ Pi\<^sub>E I (Ux x) \ U" by fastforce then have 1:"\x\U. \i\{i \ I. Ux x i \ topspace (S i)}. \\xj. \xj \ \ i \ Ux x i = \ \xj" using ho[simplified base_in_def] by (metis (no_types, lifting) mem_Collect_eq) have "\x\U. \\xj. \i\{i \ I. Ux x i \ topspace (S i)}. \xj i \ \ i \ Ux x i = \ (\xj i)" by(standard, rule bchoice) (use 1 in simp) hence "\\xj. \x\U. \i\{i \ I. Ux x i \ topspace (S i)}. \xj x i \ \ i \ Ux x i = \ (\xj x i)" by(rule bchoice) then obtain \xj where "\x\U. \i\{i \ I. Ux x i \ topspace (S i)}. \xj x i \ \ i \ Ux x i = \ (\xj x i)" by auto hence huxj: "\x i. x \ U \ i \ {i \ I. Ux x i \ topspace (S i)} \ \xj x i \ \ i" "\x i. x \ U \ i \ {i \ I. Ux x i \ topspace (S i)} \ Ux x i = \ (\xj x i)" by blast+ show "\\. U = \ \ \ \ \ ?X" proof(intro exI[where x="{\\<^sub>E i\I. K i | K. \x\U. finite {i \ I. Ux x i \ topspace (S i)} \ (\i\{i \ I. Ux x i \ topspace (S i)}. K i \ \xj x i) \ (\i\UNIV -{i \ I. Ux x i \ topspace (S i)}. K i = topspace (S i))}"] conjI) show "U = \ {\\<^sub>E i\I. K i | K. \x\U. finite {i \ I. Ux x i \ topspace (S i)} \ (\i\{i \ I. Ux x i \ topspace (S i)}. K i \ \xj x i) \ (\i\UNIV -{i \ I. Ux x i \ topspace (S i)}. K i = topspace (S i))}" proof safe fix x assume hxu:"x \ U" have "\i\{i \ I. Ux x i \ topspace (S i)}. Ux x i = \ (\xj x i)" using huxj[OF hxu] by blast hence "\i\{i \ I. Ux x i \ topspace (S i)}. \Uxj. Uxj \ \xj x i \ x i \ Uxj" using hui(3)[OF hxu] by auto hence "\Uxj. \i\{i \ I. Ux x i \ topspace (S i)}. Uxj i \ \xj x i \ x i \ Uxj i" by(rule bchoice) then obtain Uxj where huxj': "\i. i \ {i \ I. Ux x i \ topspace (S i)} \ Uxj i \ \xj x i" "\i. i \ {i \ I. Ux x i \ topspace (S i)} \ x i \ Uxj i" by auto define K where "K \ (\i. if i \ {i \ I. Ux x i \ topspace (S i)} then Uxj i else topspace (S i))" have "x \ (\\<^sub>E i\I. K i)" using huxj'(2) hui(3,4)[OF hxu] openin_subset[OF hui(2)[OF hxu]] by(auto simp: K_def PiE_def Pi_def) moreover have "\x\U. finite {i \ I. Ux x i \ topspace (S i)} \ (\i\{i \ I. Ux x i \ topspace (S i)}. K i \ \xj x i) \ (\i\UNIV -{i \ I. Ux x i \ topspace (S i)}. K i = topspace (S i))" by(rule bexI[OF _ hxu], rule conjI,simp add: hui(1)[OF hxu]) (use hui(2) hxu openin_subset huxj'(1) K_def in auto) ultimately show "x \ \ {\\<^sub>E i\I. K i | K. \x\U. finite {i \ I. Ux x i \ topspace (S i)} \ (\i\{i \ I. Ux x i \ topspace (S i)}. K i \ \xj x i) \ (\i\UNIV -{i \ I. Ux x i \ topspace (S i)}. K i = topspace (S i))}" by auto next fix x X K u assume hu: "x \ (\\<^sub>E i\I. K i)" "u \ U" "finite {i \ I. Ux u i \ topspace (S i)}" "\i\{i \ I. Ux u i \ topspace (S i)}. K i \ \xj u i" "\i\UNIV -{i \ I. Ux u i \ topspace (S i)}. K i = topspace (S i)" have "\i. i \ {i \ I. Ux u i \ topspace (S i)} \ K i \ Ux u i" using huxj[OF hu(2)] hu(4) by blast moreover have "\i. i \ I - {i \ I. Ux u i \ topspace (S i)} \ K i = Ux u i" using hu(5) by auto ultimately have "\i. i \ I \ K i \ Ux u i" by blast thus "x \ U" using hui(4)[OF hu(2)] hu(1) by blast qed next show "{\\<^sub>E i\I. K i | K. \x\U. finite {i \ I. Ux x i \ topspace (S i)} \ (\i\{i \ I. Ux x i \ topspace (S i)}. K i \ \xj x i) \ (\i\UNIV -{i \ I. Ux x i \ topspace (S i)}. K i = topspace (S i))} \ ?X" proof fix x assume "x \ {\\<^sub>E i\I. K i | K. \x\U. finite {i \ I. Ux x i \ topspace (S i)} \ (\i\{i \ I. Ux x i \ topspace (S i)}. K i \ \xj x i) \ (\i\UNIV -{i \ I. Ux x i \ topspace (S i)}. K i = topspace (S i))}" then obtain u K where hu: "x = (\\<^sub>E i\I. K i)" "u \ U" "finite {i \ I. Ux u i \ topspace (S i)}" "\i\{i \ I. Ux u i \ topspace (S i)}. K i \ \xj u i" "\i\UNIV -{i \ I. Ux u i \ topspace (S i)}. K i = topspace (S i)" by auto have hksubst:"{i \ I. K i \ topspace (S i)} \ {i \ I. Ux u i \ topspace (S i)}" using hu(5) by fastforce hence "finite {i \ I. K i \ topspace (S i)}" using hu(3) by (simp add: finite_subset) moreover have "\i\{i \ I. K i \ topspace (S i)}. K i \ \ i" using huxj(1)[OF hu(2)] hu(4) hksubst by (meson subsetD) ultimately show "x \ ?X" using hu(1) by auto qed qed next fix \ assume "\ \ ?X" have "openin (product_topology S I) u" if hu:"u \ \" for u proof - have hu': "u \ ?X" using \\ \ ?X\ hu by auto then obtain U where hU: "u = (\\<^sub>E i\I. U i)" "finite {i \ I. U i \ topspace (S i)}" "\i\{i \ I. U i \ topspace (S i)}. U i \ \ i" by auto define U' where "U' \ (\i. if i \ {i \ I. U i \ topspace (S i)} then U i else topspace (S i))" have hU': "u = (\\<^sub>E i\I. U' i)" by(auto simp: hU(1) U'_def PiE_def Pi_def) have hUfinite : "finite {i. U' i \ topspace (S i)}" using hU(2) by(auto simp: U'_def) have hUoi: "\i\{i. U' i \ topspace (S i)}. U' i \ \ i" using hU(3) by(auto simp: U'_def) have hUi: "\i\{i. U' i \ topspace (S i)}. i \ I" using hU(2) by(auto simp: U'_def) have hallopen:"openin (S i) (U' i)" for i proof - consider "i \ {i. U' i \ topspace (S i)}" | "i \ {i. U' i \ topspace (S i)}" by auto then show ?thesis proof cases case 1 then show ?thesis using hUoi ho(2)[of i] base_in_openin[of "S i" "\ i" "U' i"] hUi by auto next case 2 then have "U' i = topspace (S i)" by auto thus ?thesis by auto qed qed show "openin (product_topology S I) u" using hallopen hUfinite by(auto intro!: product_topology_basis simp: hU') qed thus "openin (product_topology S I) (\ \)" by auto qed next show "\k\{Pi\<^sub>E I U |U. finite {i \ I. U i \ topspace (S i)} \ (\i\{i \ I. U i \ topspace (S i)}. U i \ \ i)}. \X. k = Pi\<^sub>E I X \ (\i. openin (S i) (X i)) \ finite {i. X i \ topspace (S i)} \ {i. X i \ topspace (S i)} \ I" proof fix k assume "k \ {Pi\<^sub>E I U |U. finite {i \ I. U i \ topspace (S i)} \ (\i\{i \ I. U i \ topspace (S i)}. U i \ \ i)}" then obtain U where hu: "k = (\\<^sub>E i\I. U i)" "finite {i \ I. U i \ topspace (S i)}" "\i\{i \ I. U i \ topspace (S i)}. U i \ \ i" by auto define X where "X \ (\i. if i \ {i \ I. U i \ topspace (S i)} then U i else topspace (S i))" have hX1: "k = (\\<^sub>E i\I. X i)" using hu(1) by(auto simp: X_def PiE_def Pi_def) have hX2: "openin (S i) (X i)" for i using hu(3) base_in_openin[of "S i" _ "U i",OF ho(2)] by(auto simp: X_def) have hX3: "finite {i. X i \ topspace (S i)}" using hu(2) by(auto simp: X_def) have hX4: "{i. X i \ topspace (S i)} \ I" by(auto simp: X_def) show "\X. k = (\\<^sub>E i\I. X i) \ (\i. openin (S i) (X i)) \ finite {i. X i \ topspace (S i)} \ {i. X i \ topspace (S i)} \ I" using hX1 hX2 hX3 hX4 by(auto intro!: exI[where x=X]) qed qed qed lemma product_topology_second_countable: assumes "countable I" and "\i. i \ I \ second_countable (S i)" shows "second_countable (product_topology S I)" using product_topology_countable_base_in[OF assms(1)] assms(2) by(fastforce simp: second_countable_base_in) lemma second_countable_euclidean[simp]: "second_countable (euclidean :: 'a :: second_countable_topology topology)" using ex_countable_basis second_countable_def topological_basis_def by fastforce lemma Cantor_Bendixon: assumes "second_countable X" shows "\U P. countable U \ openin X U \ perfect_set X P \ U \ P = topspace X \ U \ P = {} \ (\a\{}. openin (subtopology X P) a \ uncountable a)" proof - obtain \ where o: "countable \" "base_in X \" using assms by(auto simp: second_countable_base_in) define U where "U \ \ {u\\. countable u}" define P where "P \ topspace X - U" have 1: "countable U" using o(1) by(auto simp: U_def intro!: countable_UN[of _ id,simplified]) have 2: "openin X U" using base_in_openin[OF o(2)] by(auto simp: U_def) have openin_c:"countable v \ v \ U" if "openin X v" for v proof assume "countable v" obtain \ where "v = \\" "\ \ \" using \openin X v\ o(2) by(auto simp: base_in_def) with \countable v\ have "\u. u \ \ \ countable u" by (meson Sup_upper countable_subset) thus "v \ U" using \\ \ \\ by(auto simp: \v = \\\ U_def) qed(rule countable_subset[OF _ 1]) have 3: "perfect_set X P" proof(rule perfect_setI) fix x T assume h:"x \ P" "x \ T" "openin X T" have T_unc:"uncountable T" using openin_c[OF h(3)] h(1,2) by(auto simp: P_def) obtain \ where U:"T = \\" "\ \ \" using h(3) o(2) by(auto simp: base_in_def) then obtain u where u:"u \ \" "uncountable u" using T_unc U_def h(3) openin_c by auto hence "uncountable (u - {x})" by simp hence "\ (u - {x} \ U)" using 1 by (metis countable_subset) then obtain y where "y \ u - {x}" "y \ U" by blast thus "\y. y \ x \ y \ T \ y \ P" using U u base_in_subset[OF o(2),of u] by(auto intro!: exI[where x=y] simp:P_def) qed(use 2 P_def in auto) have 4 : "uncountable a" if "openin (subtopology X P) a" "a \ {}" for a proof assume contable:"countable a" obtain b where b: "openin X b" "a = P \ b" using \openin (subtopology X P) a\ by(auto simp: openin_subtopology) hence "uncountable b" using P_def openin_c that(2) by auto thus False by (metis 1 Diff_Int_distrib2 Int_absorb1 P_def b(1) b(2) contable countable_Int1 openin_subset uncountable_minus_countable) qed show ?thesis using 1 2 3 4 by(auto simp: P_def) qed subsubsection \ Separable Spaces \ definition dense_in :: "['a topology, 'a set] \ bool" where "dense_in S U \ (U \ topspace S \ (\V. openin S V \ V \ {} \ U \ V \ {}))" lemma dense_in_def2: "dense_in S U \ (U \ topspace S \ (S closure_of U) = topspace S)" using dense_intersects_open by(auto simp: dense_in_def closure_of_subset_topspace in_closure_of) auto lemma dense_in_topspace[simp]: "dense_in S (topspace S)" by(auto simp: dense_in_def2) lemma dense_in_subset: assumes "dense_in S U" shows "U \ topspace S" using assms by(simp add: dense_in_def) lemma dense_in_nonempty: assumes "topspace S \ {}" "dense_in S U" shows "U \ {}" using assms by(auto simp: dense_in_def) lemma dense_inI: assumes "U \ topspace S" and "\V. openin S V \ V \ {} \ U \ V \ {}" shows "dense_in S U" using assms by(auto simp: dense_in_def) lemma dense_in_infinite: assumes "t1_space X" "infinite (topspace X)" "dense_in X U" shows "infinite U" proof assume fin: "finite U" then have "closedin X U" by (metis assms(1,3) dense_in_def t1_space_closedin_finite) hence "X closure_of U = U" by (simp add: closure_of_eq) thus False by (metis assms(2) assms(3) dense_in_def2 fin) qed lemma dense_in_prod: assumes "dense_in S U" and "dense_in S' U'" shows "dense_in (prod_topology S S') (U \ U')" proof(rule dense_inI) fix V assume h:"openin (prod_topology S S') V" "V \ {}" then obtain x y where hxy:"(x,y) \ V" by auto then obtain V1 V2 where hv12: "openin S V1" "openin S' V2" "x \ V1" "y \ V2" "V1 \ V2 \ V" using h(1) openin_prod_topology_alt[of S S' V] by blast hence "V1 \ {}" "V2 \ {}" by auto hence "U \ V1 \ {}" "U' \ V2 \ {}" using assms hv12 by(auto simp: dense_in_def) thus "U \ U' \ V \ {}" using hv12 by auto next show "U \ U' \ topspace (prod_topology S S')" using assms by(auto simp add: dense_in_def) qed lemma separable_space_def2:"separable_space S \ (\U. countable U \ dense_in S U)" by(auto simp: separable_space_def dense_in_def2) lemma countable_space_separable_space: assumes "countable (topspace S)" shows "separable_space S" using assms by(auto simp: separable_space_def2 intro!: exI[where x="topspace S"]) lemma separable_space_prod: assumes "separable_space S" and "separable_space S'" shows "separable_space (prod_topology S S')" proof - obtain U U' where "countable U" "dense_in S U" "countable U'" "dense_in S' U'" using assms by(auto simp: separable_space_def2) thus ?thesis by(auto intro!: exI[where x="U\U'"] dense_in_prod simp: separable_space_def2) qed lemma dense_in_product: assumes "\i. i \ I \ dense_in (T i) (U i)" shows "dense_in (product_topology T I) (\\<^sub>E i\I. U i)" proof(rule dense_inI) fix V assume h:"openin (product_topology T I) V" "V \ {}" then obtain x where hx:"x \ V" by auto then obtain K where hk: "finite {i \ I. K i \ topspace (T i)}" "\i\I. openin (T i) (K i)" "x \ (\\<^sub>E i\I. K i)" "(\\<^sub>E i\I. K i) \ V" using h(1) openin_product_topology_alt[of T I V] by auto hence "\i. i \ I \ K i \ {}" by auto hence "\i. i \ I \ U i \ K i \ {}" using assms hk by(auto simp: dense_in_def) hence "(\\<^sub>E i\I. U i) \ (\\<^sub>E i\I. K i) \ {}" by (simp add: PiE_Int PiE_eq_empty_iff) thus "(\\<^sub>E i\I. U i) \ V \ {}" using hk by auto next show "(\\<^sub>E i\I. U i) \ topspace (product_topology T I)" using assms by(auto simp: dense_in_def) qed lemma separable_countable_product: assumes "countable I" and "\i. i \ I \ separable_space (T i)" shows "separable_space (product_topology T I)" proof - consider "\i\I. T i = trivial_topology" | "\i. i \ I \ T i \ trivial_topology" by auto thus ?thesis proof cases case 1 then obtain i where i:"i \ I" "topspace (T i) = {}" by auto show ?thesis unfolding separable_space_def2 dense_in_def proof(intro exI[where x="{}"] conjI) show "\V. openin (product_topology T I) V \ V \ {} \ {} \ V \ {}" proof safe fix V x assume h: "openin (product_topology T I) V" "x \ V" from i have "(product_topology T I) = trivial_topology" using product_topology_trivial_iff by auto with h(1) have "V = {}" by simp thus "x \ {}" using h(2) by auto qed qed auto next case 2 then have "\x. \i\I. x i \ topspace (T i)" "\U. \i\I. countable (U i) \ dense_in (T i) (U i)" using assms(2) by(auto intro!: bchoice simp: separable_space_def2 ex_in_conv) then obtain x U where hxu: "\i. i \ I \ x i \ topspace (T i)" "\i. i \ I \ countable (U i)" "\i. i \ I \ dense_in (T i) (U i)" by auto define U' where "U' \ (\J i. if i \ J then U i else {x i})" show ?thesis unfolding separable_space_def2 proof(intro exI[where x="\{\\<^sub>E i\I. U' J i | J. finite J \ J \ I}"] conjI) have "(\{ \\<^sub>E i\I. U' J i | J. finite J \ J \ I}) = (\ ((\J. \\<^sub>E i\I. U' J i) ` {J. finite J \ J \ I}))" by auto also have "countable ..." proof(rule countable_UN) fix J assume hj:"J \ {J. finite J \ J \ I}" have "inj_on (\f. (\i\J. f i, \i\(I-J). f i)) (\\<^sub>E i\I. U' J i)" proof(rule inj_onI) fix f g assume h:"f \ Pi\<^sub>E I (U' J)" "g \ Pi\<^sub>E I (U' J)" "(restrict f J, restrict f (I - J)) = (restrict g J, restrict g (I - J))" then have "\i. i \ J \ f i = g i" "\i. i \(I-J) \ f i = g i" by(auto simp: restrict_def) meson+ thus "f = g" using h(1,2) by(auto simp: U'_def) (meson PiE_ext) qed moreover have "countable ((\f. (\i\J. f i, \i\(I-J). f i)) ` (\\<^sub>E i\I. U' J i))" (is "countable ?K") proof - have 1:"?K \ (\\<^sub>E i\J. U i) \ (\\<^sub>E i\(I-J). {x i})" using hj by(auto simp: U'_def PiE_def Pi_def) have 2:"countable ..." proof(rule countable_SIGMA) show "countable (Pi\<^sub>E J U)" using hj hxu(2) by(auto intro!: countable_PiE) next have "(\\<^sub>E i\I - J. {x i}) = { \i\I-J. x i }" by(auto simp: PiE_def extensional_def restrict_def Pi_def) thus "countable (\\<^sub>E i\I - J. {x i})" by simp qed show ?thesis by(rule countable_subset[OF 1 2]) qed ultimately show "countable (\\<^sub>E i\I. U' J i)" by(simp add: countable_image_inj_eq) qed(rule countable_Collect_finite_subset[OF assms(1)]) finally show "countable (\{ \\<^sub>E i\I. U' J i | J. finite J \ J \ I})" . next show "dense_in (product_topology T I) (\ {\\<^sub>E i\I. U' J i |J. finite J \ J \ I})" proof(rule dense_inI) fix V assume h:"openin (product_topology T I) V" "V \ {}" then obtain y where hx:"y \ V" by auto then obtain K where hk: "finite {i \ I. K i \ topspace (T i)}" "\i. i\I \ openin (T i) (K i)" "y \ (\\<^sub>E i\I. K i)" "(\\<^sub>E i\I. K i) \ V" using h(1) openin_product_topology_alt[of T I V] by auto hence 3:"\i. i \ I \ K i \ {}" by auto hence 4:"i \ {i \ I. K i \ topspace (T i)} \ K i \ U' {i \ I. K i \ topspace (T i)} i \ {}" for i using hxu(3)[of i] hk(2)[of i] by(auto simp: U'_def dense_in_def) have "\z. \i\{i \ I. K i \ topspace (T i)}. z i \ K i \ U' {i \ I. K i \ topspace (T i)} i" by(rule bchoice) (use 4 in auto) then obtain z where hz: "\i\{i \ I. K i \ topspace (T i)}. z i \ K i \ U' {i \ I. K i \ topspace (T i)} i" by auto have 5: "i \ {i \ I. K i \ topspace (T i)} \ i \ I \ x i \ K i" for i using hxu(1)[of i] by auto have "(\i. if i \ {i \ I. K i \ topspace (T i)} then z i else if i \ I then x i else undefined) \ (\\<^sub>E i\I. U' {i \ I. K i \ topspace (T i)} i) \ (\\<^sub>E i\I. K i)" using 4 5 hz by(auto simp: U'_def) thus "\ {Pi\<^sub>E I (U' J) |J. finite J \ J \ I} \ V \ {}" using hk(1,4) by blast next have "\J. J \ I \ (\\<^sub>E i\I. U' J i) \ topspace (product_topology T I)" using hxu by(auto simp: dense_in_def U'_def PiE_def Pi_def) (metis subsetD) thus "(\ {\\<^sub>E i\I. U' J i |J. finite J \ J \ I}) \ topspace (product_topology T I)" by auto qed qed qed qed lemma separable_finite_product: assumes "finite I" and "\i. i \ I \ separable_space (T i)" shows "separable_space (product_topology T I)" using separable_countable_product[OF countable_finite[OF assms(1)]] assms by auto subsubsection \ $G_{\delta}$ Set \ lemma gdelta_inD: assumes "gdelta_in S A" shows "\\. \ \ {} \ countable \ \ (\b\\. openin S b) \ A = \ \" using assms unfolding gdelta_in_def relative_to_def intersection_of_def by (metis IntD1 Int_insert_right_if1 complete_lattice_class.Inf_insert countable_insert empty_not_insert inf.absorb_iff2 mem_Collect_eq openin_topspace) lemma gdelta_inD': assumes "gdelta_in S A" shows "\U. (\n::nat. openin S (U n)) \ A = \ (range U)" proof- obtain \ where h:"\ \ {}" "countable \" "\b. b\\ \ openin S b" "A = \ \" using gdelta_inD[OF assms] by metis show ?thesis using range_from_nat_into[OF h(1,2)] h(3,4) by(auto intro!: exI[where x="from_nat_into \"]) qed lemma gdelta_in_continuous_map: assumes "continuous_map X Y f" "gdelta_in Y a" shows "gdelta_in X (f -` a \ topspace X)" proof - obtain Ua where u: "Ua \ {}" "countable Ua" "\b. b \ Ua \ openin Y b" "a = \ Ua" using gdelta_inD[OF assms(2)] by metis then have 0:"f -` a \ topspace X = \ {f -` b \ topspace X|b. b \ Ua}" by auto have 1: "{f -` b \ topspace X |b. b \ Ua} \ {}" using u(1) by simp have 2:"countable {f -` b \ topspace X|b. b \ Ua}" using u by (simp add: Setcompr_eq_image) have 3:"\c. c \ {f -` b \ topspace X|b. b \ Ua} \ openin X c" using assms u(3) by blast show ?thesis by (metis (mono_tags, lifting) "0" "1" "2" "3" gdelta_in_Inter open_imp_gdelta_in) qed lemma g_delta_of_inj_open_map: assumes "open_map X Y f" "inj_on f (topspace X)" "gdelta_in X a" shows "gdelta_in Y (f ` a)" proof - obtain Ua where u: "Ua \ {}" "countable Ua" "\b. b \ Ua \ openin X b" "a = \ Ua" using gdelta_inD[OF assms(3)] by metis then obtain j where "j \ Ua" by auto have "f ` a = f ` \ Ua" by(simp add: u(4)) also have "... = \ ((`) f ` Ua)" using u openin_subset by(auto intro!: image_INT[OF assms(2) _ \j \ Ua\,of id,simplified]) also have "... = \ {f ` u|u. u \ Ua}" by auto finally have 0: "f ` a = \ {f ` u |u. u \ Ua}" . have 1:"{f ` u |u. u \ Ua} \ {}" using u(1) by auto have 2:"countable {f ` u |u. u \ Ua}" using u(2) by (simp add: Setcompr_eq_image) have 3: "\c. c \ {f ` u |u. u \ Ua} \ openin Y c" using assms(1) u(3) by(auto simp: open_map_def) show ?thesis by (metis (no_types, lifting) "0" "1" "2" "3" gdelta_in_Inter open_imp_gdelta_in) qed lemma gdelta_in_prod: assumes "gdelta_in X A" "gdelta_in Y B" shows "gdelta_in (prod_topology X Y) (A \ B)" proof - obtain Ua Ub where hu: "Ua \ {}" "countable Ua" "\b. b \ Ua \ openin X b" "A = \ Ua" "Ub \ {}" "countable Ub" "\b. b \ Ub \ openin Y b" "B = \ Ub" by (meson gdelta_inD assms) then have 0:"A \ B = \ {a \ b | a b. a \ Ua \ b \ Ub}" by blast have 1: "{a \ b | a b. a \ Ua \ b \ Ub} \ {}" using hu(1,5) by auto have 2: "countable {a \ b | a b. a \ Ua \ b \ Ub}" proof - have "countable ((\(x, y). x \ y) ` (Ua \ Ub))" using hu(2,6) by(auto intro!: countable_image[of "Ua \ Ub" "\(x,y). x \ y"]) moreover have "... = {a \ b | a b. a \ Ua \ b \ Ub}" by auto ultimately show ?thesis by simp qed have 3: "\c. c \ {a \ b | a b. a \ Ua \ b \ Ub} \ openin (prod_topology X Y) c" using hu(3,7) by(auto simp: openin_prod_Times_iff) show ?thesis by (metis (no_types, lifting) gdelta_in_Inter open_imp_gdelta_in 0 1 2 3) qed corollary gdelta_in_prod1: assumes "gdelta_in X A" shows "gdelta_in (prod_topology X Y) (A \ topspace Y)" by(auto intro!: gdelta_in_prod assms) corollary gdelta_in_prod2: assumes "gdelta_in Y B" shows "gdelta_in (prod_topology X Y) (topspace X \ B)" by(auto intro!: gdelta_in_prod assms) lemma continuous_map_imp_closed_graph': assumes "continuous_map X Y f" "Hausdorff_space Y" shows "closedin (prod_topology Y X) ((\x. (f x,x)) ` topspace X)" using assms closed_map_def closed_map_paired_continuous_map_left by blast +subsubsection \ Continuous Maps on First Countable Topology\ +text \ Generalized version of @{thm Metric_space.eventually_atin_sequentially}\ +lemma eventually_atin_sequentially: + assumes "first_countable X" + shows "eventually P (atin X a) \ (\\. range \ \ topspace X - {a} \ limitin X \ a sequentially \ eventually (\n. P (\ n)) sequentially)" +proof safe + fix an + assume h:"eventually P (atin X a)" "range an \ topspace X - {a}" "limitin X an a sequentially" + then obtain U where U: "openin X U" "a \ U" "\x\U - {a}. P x" + by (auto simp: eventually_atin limitin_topspace) + with h(3) obtain N where "\n\N. an n \ U" + by (meson limitin_sequentially) + with U(3) h(2) show "\\<^sub>F n in sequentially. P (an n)" + unfolding eventually_sequentially by blast +next + assume h:"\an. range an \ topspace X - {a} \ limitin X an a sequentially \ (\\<^sub>F n in sequentially. P (an n))" + consider "a \ topspace X" | "a \ topspace X" + by blast + then show "eventually P (atin X a)" + proof cases + assume a:"a \ topspace X" + from a assms obtain B' where B': "countable B'" "\V. V \ B' \ openin X V" "\U. openin X U \ a \ U \ (\V \ B'. a \ V \ V \ U)" + by(fastforce simp: first_countable_def) + define B where "B \ {V\B'. a \ V}" + have B:"\V. V \ B \ openin X V" "countable B" "B \ {}" "\U. openin X U \ a \ U \ (\V \ B. a \ V \ V \ U)" + using B' B'(3)[OF _ a] by(fastforce simp: B_def)+ + define An where "An \ (\n. \i\n. from_nat_into B i)" + have a_in_An:"a \ An n" for n + by (metis (no_types, lifting) An_def B_def B(3) INT_I from_nat_into mem_Collect_eq) + have openAn:"\n. openin X (An n)" + using B by(auto simp: An_def from_nat_into[OF B(3)] openin_Inter) + have deqseq_An:"decseq An" + by(fastforce simp: decseq_def An_def) + have "\U. openin X U \ a \ U \ Ball (U - {a}) P" + proof(rule ccontr) + assume "\U. openin X U \ a \ U \ Ball (U - {a}) P" + then have "\U. openin X U \ a \ U \ \x \ U - {a}. \ P x" + by blast + hence "\b\An n - {a}. \ P b" for n + using openAn a_in_An by auto + then obtain an where an: "\n. an n \ An n" "\n. an n \ a" "\n. \ P (an n)" + by (metis Diff_iff singletonI) + have "limitin X an a sequentially" + unfolding limitin_sequentially + proof safe + fix U + assume "openin X U" "a \ U" + then obtain V where V:"a \ V" "V \ U" "V \ B" + using B by meson + then obtain N where "V = from_nat_into B N" + by (metis B(2) from_nat_into_surj) + hence "\n. n \ N \ an n \ V" + using an(1) An_def by blast + thus "\N. \n\N. an n \ U" + using V by blast + qed fact + hence 1:"\\<^sub>F n in sequentially. P (an n)" + using an(2) h an(1) openin_subset[OF openAn] by blast + thus False + using an(3) by simp + qed + thus ?thesis + by(simp add: eventually_atin) + qed(auto simp: eventually_atin) +qed + +lemma continuous_map_iff_limit_seq: + assumes "first_countable X" + shows "continuous_map X Y f \ (\xn x. limitin X xn x sequentially \ limitin Y (\n. f (xn n)) (f x) sequentially)" + unfolding continuous_map_atin +proof safe + fix xn x + assume h:"\x\topspace X. limitin Y f (f x) (atin X x)" "limitin X xn x sequentially" + then have limfx: "limitin Y f (f x) (atin X x)" + by(simp add: limitin_topspace) + show "limitin Y (\n. f (xn n)) (f x) sequentially" + unfolding limitin_sequentially + proof safe + fix U + assume U:"openin Y U" "f x \ U" + then have h':"\\. range \ \ topspace X - {x} \ x \ topspace X \ limitin X \ x sequentially \ (\N. \n\N. f (\ n) \ U)" + using limfx by(auto simp: limitin_def eventually_atin_sequentially[OF assms(1)] eventually_sequentially) + show "\N. \n\N. f (xn n) \ U" + proof(cases "finite {n. xn n \ x}") + assume "finite {n. xn n \ x}" + then obtain N where "\n. n \ N \ xn n = x" + using infinite_nat_iff_unbounded_le by blast + then show ?thesis + using U by auto + next + assume inf:"infinite {n. xn n \ x}" + obtain n0 where n0:"\n. n \ n0 \ xn n \ topspace X" + by (meson h(2) limitin_sequentially openin_topspace) + have inf':"infinite ({n. xn n \ x} \ {n0..})" + proof + have 1:"({n. xn n \ x} \ {n0..}) \ ({n. xn n \ x} \ {.. x}" + by auto + assume "finite ({n. xn n \ x} \ {n0..})" + then have "finite (({n. xn n \ x} \ {n0..}) \ ({n. xn n \ x} \ {.. enumerate ({n. xn n \ x} \ {n0..})" + have a: "strict_mono a" "range a = ({n. xn n \ x} \ {n0..})" + using range_enumerate[OF inf'] strict_mono_enumerate[OF inf'] + by(auto simp: a_def) + have "\N. \n\N. f (xn (a n)) \ U" + using limitin_subsequence[OF a(1) h(2)] a(2) n0 + by(auto intro!: h' limitin_topspace[OF h(2)] simp: comp_def) + then obtain N where N:"\n. n \ N \ f (xn (a n)) \ U" + by blast + show "\N. \n\N. f (xn n) \ U" + proof(auto intro!: exI[where x="a N"]) + fix n + assume n:"n \ a N" + show "f (xn n) \ U" + proof (cases "xn n = x") + assume "xn n \ x" + moreover have "n0 \ n" + using seq_suble[OF a(1),of N] n a(2) + by (metis Int_Collect atLeast_def dual_order.trans rangeI) + ultimately obtain n1 where n1:"n = a n1" + by (metis (mono_tags, lifting) Int_Collect atLeast_def imageE mem_Collect_eq a(2)) + have "n1 \ N" + using strict_mono_less_eq[OF a(1),of N n1] n by(simp add: n1) + thus ?thesis + by(auto intro!: N simp: n1) + qed(auto simp: U) + qed + qed + qed(auto intro!: limitin_topspace limfx) +next + fix x + assume h:"\xn x. limitin X xn x sequentially \ limitin Y (\n. f (xn n)) (f x) sequentially" "x \ topspace X" + then have "f x \ topspace Y" + by (meson Abstract_Limits.limitin_const_iff limitin_topspace) + thus "limitin Y f (f x) (atin X x)" + using h by(auto simp: eventually_atin_sequentially[OF assms(1)] limitin_def ) +qed + subsubsection \ Upper-Semicontinuous Functions \ definition upper_semicontinuous_map :: "['a topology, 'a \ 'b :: linorder_topology] \ bool" where "upper_semicontinuous_map X f \ (\a. openin X {x\topspace X. f x < a})" lemma continuous_upper_semicontinuous: assumes "continuous_map X (euclidean :: ('b :: linorder_topology) topology) f" shows "upper_semicontinuous_map X f" unfolding upper_semicontinuous_map_def proof safe fix a :: 'b have *:"openin euclidean U \ openin X {x \ topspace X. f x \ U}" for U using assms by(simp add: continuous_map) have "openin euclidean {.. topspace X. f x < a}" by auto qed lemma upper_semicontinuous_map_iff_closed: "upper_semicontinuous_map X f \ (\a. closedin X {x\topspace X. f x \ a})" proof - have "{x \ topspace X. f x < a} = topspace X - {x \ topspace X. f x \ a}" for a by auto thus ?thesis by (simp add: closedin_def upper_semicontinuous_map_def) qed lemma upper_semicontinuous_map_real_iff: fixes f :: "'a \ real" shows "upper_semicontinuous_map X f \ upper_semicontinuous_map X (\x. ereal (f x))" unfolding upper_semicontinuous_map_def proof safe fix a :: ereal assume h:"\a::real. openin X {x \ topspace X. f x < a}" consider "a = - \" | "a = \" | "a \ - \ \ a \ \" by auto then show "openin X {x \ topspace X. ereal (f x) < a}" proof cases case 3 then have "ereal (f x) < a \ f x < real_of_ereal a" for x by (metis ereal_less_eq(3) linorder_not_less real_of_ereal.elims) thus ?thesis using h by simp qed simp_all next fix a :: real assume h:"\a::ereal. openin X {x \ topspace X. ereal (f x) < a}" then have "openin X {x \ topspace X. ereal (f x) < ereal a}" by blast moreover have"ereal (f x) < real_of_ereal a \ f x < a" for x by auto ultimately show "openin X {x \ topspace X. f x < a}" by auto qed subsubsection \ Lower-Semicontinuous Functions \ definition lower_semicontinuous_map :: "['a topology, 'a \ 'b :: linorder_topology] \ bool" where "lower_semicontinuous_map X f \ (\a. openin X {x\topspace X. a < f x})" lemma continuous_lower_semicontinuous: assumes "continuous_map X (euclidean :: ('b :: linorder_topology) topology) f" shows "lower_semicontinuous_map X f" unfolding lower_semicontinuous_map_def proof safe fix a :: 'b have *:"openin euclidean U \ openin X {x \ topspace X. f x \ U}" for U using assms by(simp add: continuous_map) have "openin euclidean {a<..}" by auto with *[of "{a<..}"] show "openin X {x \ topspace X. a < f x}" by auto qed lemma lower_semicontinuous_map_iff_closed: "lower_semicontinuous_map X f \ (\a. closedin X {x\topspace X. f x \ a})" proof - have "{x \ topspace X. a < f x} = topspace X - {x \ topspace X. f x \ a}" for a by auto thus ?thesis by (simp add: closedin_def lower_semicontinuous_map_def) qed lemma lower_semicontinuous_map_real_iff: fixes f :: "'a \ real" shows "lower_semicontinuous_map X f \ lower_semicontinuous_map X (\x. ereal (f x))" unfolding lower_semicontinuous_map_def proof safe fix a :: ereal assume h:"\a::real. openin X {x \ topspace X. a < f x}" consider "a = - \" | "a = \" | "a \ - \ \ a \ \" by auto then show "openin X {x \ topspace X. a < ereal (f x)}" proof cases case 3 then have "a < ereal (f x) \ real_of_ereal a < f x" for x by (metis ereal_less_eq(3) linorder_not_less real_of_ereal.elims) thus ?thesis using h by simp qed simp_all next fix a :: real assume h:"\a::ereal. openin X {x \ topspace X. a < ereal (f x)}" then have "openin X {x \ topspace X. ereal (f x) > ereal a}" by blast moreover have"ereal (f x) > real_of_ereal a \ a < f x" for x by auto ultimately show "openin X {x \ topspace X. f x > a}" by auto qed subsection \Lemmas for Measure Theory\ subsubsection \ Lemmas for Measurable Sets\ lemma measurable_preserve_sigma_sets: assumes "sets M = sigma_sets \ S" "S \ Pow \" "\a. a \ S \ f ` a \ sets N" "inj_on f (space M)" "f ` space M \ sets N" and "b \ sets M" shows "f ` b \ sets N" proof - have "b \ sigma_sets \ S" using assms(1,6) by simp thus ?thesis proof induction case (Basic a) then show ?case by(rule assms(3)) next case Empty then show ?case by simp next case (Compl a) moreover have " \ = space M" by (metis assms(1) assms(2) sets.sets_into_space sets.top sigma_sets_into_sp sigma_sets_top subset_antisym) ultimately show ?case by (metis Diff_subset assms(2) assms(4) assms(5) inj_on_image_set_diff sets.Diff sigma_sets_into_sp) next case (Union a) then show ?case by (simp add: image_UN) qed qed inductive_set sigma_sets_cinter :: "'a set \ 'a set set \ 'a set set" for sp :: "'a set" and A :: "'a set set" where Basic_c[intro, simp]: "a \ A \ a \ sigma_sets_cinter sp A" | Top_c[simp]: "sp \ sigma_sets_cinter sp A" | Inter_c: "(\i::nat. a i \ sigma_sets_cinter sp A) \ (\i. a i) \ sigma_sets_cinter sp A" | Union_c: "(\i::nat. a i \ sigma_sets_cinter sp A) \ (\i. a i) \ sigma_sets_cinter sp A" inductive_set sigma_sets_cinter_dunion :: "'a set \ 'a set set \ 'a set set" for sp :: "'a set" and A :: "'a set set" where Basic_cd[intro, simp]: "a \ A \ a \ sigma_sets_cinter_dunion sp A" | Top_cd[simp]: "sp \ sigma_sets_cinter_dunion sp A" | Inter_cd: "(\i::nat. a i \ sigma_sets_cinter_dunion sp A) \ (\i. a i) \ sigma_sets_cinter_dunion sp A" | Union_cd: "(\i::nat. a i \ sigma_sets_cinter_dunion sp A) \ disjoint_family a \ (\i. a i) \ sigma_sets_cinter_dunion sp A" lemma sigma_sets_cinter_dunion_subset: "sigma_sets_cinter_dunion sp A \ sigma_sets_cinter sp A" proof safe fix x assume "x \ sigma_sets_cinter_dunion sp A" then show "x \ sigma_sets_cinter sp A" by induction (auto intro!: Union_c Inter_c) qed lemma sigma_sets_cinter_into_sp: assumes "A \ Pow sp" "x \ sigma_sets_cinter sp A" shows "x \ sp" using assms(2) by induction (use assms(1) subsetD in blast)+ lemma sigma_sets_cinter_dunion_into_sp: assumes "A \ Pow sp" "x \ sigma_sets_cinter_dunion sp A" shows "x \ sp" using assms(2) by induction (use assms(1) subsetD in blast)+ lemma sigma_sets_cinter_int: assumes "a \ sigma_sets_cinter sp A" "b \ sigma_sets_cinter sp A" shows "a \ b \ sigma_sets_cinter sp A" proof - have 1:"a \ b = (\i::nat. if i = 0 then a else b)" by auto show ?thesis unfolding 1 by(rule Inter_c,use assms in auto) qed lemma sigma_sets_cinter_dunion_int: assumes "a \ sigma_sets_cinter_dunion sp A" "b \ sigma_sets_cinter_dunion sp A" shows "a \ b \ sigma_sets_cinter_dunion sp A" proof - have 1:"a \ b = (\i::nat. if i = 0 then a else b)" by auto show ?thesis unfolding 1 by(rule Inter_cd,use assms in auto) qed lemma sigma_sets_cinter_un: assumes "a \ sigma_sets_cinter sp A" "b \ sigma_sets_cinter sp A" shows "a \ b \ sigma_sets_cinter sp A" proof - have 1:"a \ b = (\i::nat. if i = 0 then a else b)" by auto show ?thesis unfolding 1 by(rule Union_c,use assms in auto) qed lemma sigma_sets_eq_cinter_dunion: assumes "metrizable_space X" shows "sigma_sets (topspace X) {U. openin X U} = sigma_sets_cinter_dunion (topspace X) {U. openin X U}" proof safe fix a interpret sa: sigma_algebra "topspace X" "sigma_sets (topspace X) {U. openin X U}" by(auto intro!: sigma_algebra_sigma_sets openin_subset) assume "a \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}" then show "a \ sigma_sets (topspace X) {U. openin X U}" by induction auto next have c:"sigma_sets_cinter_dunion (topspace X) {U. openin X U} \ {U\sigma_sets_cinter_dunion (topspace X) {U. openin X U}. topspace X - U \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}}" proof fix a assume a: "a \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}" then show "a \ {U \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}. topspace X - U \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}}" proof induction case a:(Basic_cd a) then have "gdelta_in X (topspace X - a)" by(auto intro!: closed_imp_gdelta_in assms) from gdelta_inD'[OF this] obtain U where U: "\n :: nat. openin X (U n)" "topspace X - a = \ (range U)" by auto show ?case using a U(1) by(auto simp: U(2) intro!: Inter_cd) next case Top_cd then show ?case by auto next case ca:(Inter_cd a) define b where "b \ (\n. (topspace X - a n) \ (\i. if i < n then a i else topspace X))" have bd:"disjoint_family b" using nat_neq_iff by(fastforce simp: disjoint_family_on_def b_def) have bin:"b i \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}" for i unfolding b_def apply(rule sigma_sets_cinter_dunion_int) using ca(2)[of i] apply auto[1] apply(rule Inter_cd) using ca by auto have bun:"topspace X - (\ (range a)) = (\i. b i)" (is "?lhs = ?rhs") proof - { fix x have "x \ ?lhs \ x \ topspace X \ x \ (\i. topspace X - a i)" by auto also have "... \ x \ topspace X \ (\n. x \ topspace X - a n)" by auto also have "... \ x \ topspace X \ (\n. x \ topspace X - a n \ (\i a i))" proof safe fix n assume 1:"x \ a n" "x \ topspace X" define N where "N \ Min {m. m \ n \ x \ a m}" have N:"x \ a N" "N \ n" using linorder_class.Min_in[of "{m. m \ n \ x \ a m}"] 1 by(auto simp: N_def) have N':"x \ a i" if "i < N" for i proof(rule ccontr) assume "x \ a i" then have "N \ i" using linorder_class.Min_le[of "{m. m \ n \ x \ a m}" i] that N(2) by(auto simp: N_def) with that show False by auto qed show "\n. x \ topspace X - a n \ (\i a i)" using N N' by(auto intro!: exI[where x=N] 1) qed auto also have "... \ x \ ?rhs" by(auto simp: b_def) finally have "x \ ?lhs \ x \ ?rhs" . } thus ?thesis by auto qed have "... \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}" by(rule Union_cd) (use bin bd in auto) thus ?case using Inter_cd[of a,OF ca(1)] by(auto simp: bun) next case ca:(Union_cd a) have "topspace X - (\ (range a)) = (\i. (topspace X - a i))" by simp have "... \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}" by(rule Inter_cd) (use ca in auto) then show ?case using Union_cd[of a,OF ca(1,2)] by auto qed qed fix a assume "a \ sigma_sets (topspace X) {U. openin X U}" then show "a \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}" proof induction case a:(Union a) define b where "b \ (\n. a n \ (\i. if i < n then topspace X - a i else topspace X))" have bd:"disjoint_family b" by(auto simp: disjoint_family_on_def b_def) (metis Diff_iff UnCI image_eqI linorder_neqE_nat mem_Collect_eq) have bin:"b i \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}" for i unfolding b_def apply(rule sigma_sets_cinter_dunion_int) using a(2)[of i] apply auto[1] apply(rule Inter_cd) using c a by auto have bun:"(\i. a i) = (\i. b i)" (is "?lhs = ?rhs") proof - { fix x have "x \ ?lhs \ x \ topspace X \ x \ ?lhs" using sigma_sets_cinter_dunion_into_sp[OF _ a(2)] by (metis UN_iff subsetD subset_Pow_Union topspace_def) also have "... \ x \ topspace X \ (\n. x \ a n)" by auto also have "... \ x \ topspace X \ (\n. x \ a n \ (\i topspace X - a i))" proof safe fix n assume 1:"x \ topspace X" "x \ a n" define N where "N \ Min {m. m \ n \ x \ a m}" have N:"x \ a N" "N \ n" using linorder_class.Min_in[of "{m. m \ n \ x \ a m}"] 1 by(auto simp: N_def) have N':"x \ a i" if "i < N" for i proof(rule ccontr) assume "\ x \ a i" then have "N \ i" using linorder_class.Min_le[of "{m. m \ n \ x \ a m}" i] that N(2) by(auto simp: N_def) with that show False by auto qed show "\n. x \ a n \ (\i topspace X - a i)" using N N' 1 by(auto intro!: exI[where x=N]) qed auto also have "... \ x \ ?rhs" proof safe fix m assume "x \ b m" then show "x \ topspace X" "\n. x \ a n \ (\i topspace X - a i)" by(auto intro!: exI[where x=m] simp: b_def) qed(auto simp: b_def) finally have "x \ ?lhs \ x \ ?rhs" . } thus ?thesis by auto qed have "... \ sigma_sets_cinter_dunion (topspace X) {U. openin X U}" by(rule Union_cd) (use bin bd in auto) thus ?case by(auto simp: bun) qed(use c in auto) qed lemma sigma_sets_eq_cinter: assumes "metrizable_space X" shows "sigma_sets (topspace X) {U. openin X U} = sigma_sets_cinter (topspace X) {U. openin X U}" proof safe fix a interpret sa: sigma_algebra "topspace X" "sigma_sets (topspace X) {U. openin X U}" by(auto intro!: sigma_algebra_sigma_sets openin_subset) assume "a \ sigma_sets_cinter (topspace X) {U. openin X U}" then show "a \ sigma_sets (topspace X) {U. openin X U}" by induction auto qed (use sigma_sets_cinter_dunion_subset sigma_sets_eq_cinter_dunion[OF assms] in auto) subsubsection \ Measurable Isomorphisms \ definition measurable_isomorphic_map::"['a measure, 'b measure, 'a \ 'b] \ bool" where "measurable_isomorphic_map M N f \ bij_betw f (space M) (space N) \ f \ M \\<^sub>M N \ the_inv_into (space M) f \ N \\<^sub>M M" lemma measurable_isomorphic_map_sets_cong: assumes "sets M = sets M'" "sets N = sets N'" shows "measurable_isomorphic_map M N f \ measurable_isomorphic_map M' N' f" by(simp add: measurable_isomorphic_map_def sets_eq_imp_space_eq[OF assms(1)] sets_eq_imp_space_eq[OF assms(2)] measurable_cong_sets[OF assms] measurable_cong_sets[OF assms(2,1)]) lemma measurable_isomorphic_map_surj: assumes "measurable_isomorphic_map M N f" shows "f ` space M = space N" using assms by(auto simp: measurable_isomorphic_map_def bij_betw_def) lemma measurable_isomorphic_mapI: assumes "bij_betw f (space M) (space N)" "f \ M \\<^sub>M N" "the_inv_into (space M) f \ N \\<^sub>M M" shows "measurable_isomorphic_map M N f" using assms by(simp add: measurable_isomorphic_map_def) lemma measurable_isomorphic_map_byWitness: assumes "f \ M \\<^sub>M N" "g \ N \\<^sub>M M" "\x. x \ space M \ g (f x) = x" "\x. x \ space N \ f (g x) = x" shows "measurable_isomorphic_map M N f" proof - have *:"bij_betw f (space M) (space N)" using assms by(auto intro!: bij_betw_byWitness[where f'=g] dest:measurable_space) show ?thesis proof(rule measurable_isomorphic_mapI) have "the_inv_into (space M) f x = g x" if "x \ space N" for x by (metis * assms(2) assms(4) bij_betw_imp_inj_on measurable_space that the_inv_into_f_f) thus "the_inv_into (space M) f \ N \\<^sub>M M" using measurable_cong assms(2) by blast qed (simp_all add: * assms(1)) qed lemma measurable_isomorphic_map_restrict_space: assumes "f \ M \\<^sub>M N" "\A. A \ sets M \ f ` A \ sets N" "inj_on f (space M)" shows "measurable_isomorphic_map M (restrict_space N (f ` space M)) f" proof(rule measurable_isomorphic_mapI) show "bij_betw f (space M) (space (restrict_space N (f ` space M)))" by (simp add: assms(2,3) inj_on_imp_bij_betw) next show "f \ M \\<^sub>M restrict_space N (f ` space M)" by (simp add: assms(1) measurable_restrict_space2) next show "the_inv_into (space M) f \ restrict_space N (f ` space M) \\<^sub>M M" proof(rule measurableI) show "x \ space (restrict_space N (f ` space M)) \ the_inv_into (space M) f x \ space M" for x by (simp add: assms(2,3) the_inv_into_into) next fix A assume "A \ sets M" have "the_inv_into (space M) f -` A \ space (restrict_space N (f ` space M)) = f ` A" by (simp add: \A \ sets M\ assms(2,3) sets.sets_into_space the_inv_into_vimage) also note assms(2)[OF \A \ sets M\] finally show "the_inv_into (space M) f -` A \ space (restrict_space N (f ` space M)) \ sets (restrict_space N (f ` space M))" by (simp add: assms(2) sets_restrict_space_iff) qed qed lemma measurable_isomorphic_mapD': assumes "measurable_isomorphic_map M N f" shows "\A. A \ sets M \ f ` A \ sets N" "f \ M \\<^sub>M N" "\g. bij_betw g (space N) (space M) \ g \ N \\<^sub>M M \ (\x \ space M. g (f x) = x) \ (\x\ space N. f (g x) = x) \ (\A\sets N. g ` A \ sets M)" proof - have h:"bij_betw f (space M) (space N)" "f \ M \\<^sub>M N" "the_inv_into (space M) f \ N \\<^sub>M M" using assms by(simp_all add: measurable_isomorphic_map_def) show "f ` A \ sets N" if "A \ sets M" for A proof - have "f ` A = the_inv_into (space M) f -` A \ space N" using the_inv_into_vimage[OF bij_betw_imp_inj_on[OF h(1)] sets.sets_into_space[OF that]] by(simp add: bij_betw_imp_surj_on[OF h(1)]) also have "... \ sets N" using that h(3) by auto finally show ?thesis . qed show "f \ M \\<^sub>M N" using assms by(simp add: measurable_isomorphic_map_def) show "\g. bij_betw g (space N) (space M) \ g \ N \\<^sub>M M \ (\x \ space M. g (f x) = x) \ (\x\ space N. f (g x) = x) \ (\A\sets N. g ` A \ sets M)" proof(rule exI[where x="the_inv_into (space M) f"]) have *:"the_inv_into (space M) f ` A \ sets M" if "A \ sets N" for A proof - have "\x. x \ space M \ the_inv_into (space N) (the_inv_into (space M) f) x = f x" by (metis bij_betw_imp_inj_on bij_betw_the_inv_into h(1) h(2) measurable_space the_inv_into_f_f) from vimage_inter_cong[of "space M" _ f A,OF this] the_inv_into_vimage[OF bij_betw_imp_inj_on[OF bij_betw_the_inv_into[OF h(1)]] sets.sets_into_space[OF that]] bij_betw_imp_surj_on[OF bij_betw_the_inv_into[OF h(1)]] measurable_sets[OF h(2) that] show ?thesis by fastforce qed show "bij_betw (the_inv_into (space M) f) (space N) (space M) \ the_inv_into (space M) f \ N \\<^sub>M M \ (\x\space M. the_inv_into (space M) f (f x) = x) \ (\x\space N. f (the_inv_into (space M) f x) = x) \ (\A\sets N. the_inv_into (space M) f ` A \ sets M)" using bij_betw_the_inv_into[OF h(1)] by (meson * bij_betw_imp_inj_on f_the_inv_into_f_bij_betw h(1) h(3) the_inv_into_f_f) qed qed lemma measurable_isomorphic_map_inv: assumes "measurable_isomorphic_map M N f" shows "measurable_isomorphic_map N M (the_inv_into (space M) f)" using assms[simplified measurable_isomorphic_map_def] by(auto intro!: measurable_isomorphic_map_byWitness[where g=f] bij_betw_the_inv_into f_the_inv_into_f_bij_betw[of f] bij_betw_imp_inj_on the_inv_into_f_f) lemma measurable_isomorphic_map_comp: assumes "measurable_isomorphic_map M N f" and "measurable_isomorphic_map N L g" shows "measurable_isomorphic_map M L (g \ f)" proof - obtain f' g' where [measurable]: "f' \ N \\<^sub>M M" and hf:"\x. x\space M \ f' (f x) = x" "\x. x\space N \ f (f' x) = x" and [measurable]: "g' \ L \\<^sub>M N" and hg:"\x. x\space N \ g' (g x) = x" "\x. x\space L \ g (g' x) = x" using measurable_isomorphic_mapD'[OF assms(1)] measurable_isomorphic_mapD'[OF assms(2)] by metis have [measurable]: "f \ M \\<^sub>M N" "g \ N \\<^sub>M L" using assms by(auto simp: measurable_isomorphic_map_def) from hf hg measurable_space[OF \f \ M \\<^sub>M N\] measurable_space[OF \g' \ L \\<^sub>M N\] show ?thesis by(auto intro!: measurable_isomorphic_map_byWitness[where g="f'\g'"]) qed definition measurable_isomorphic::"['a measure, 'b measure] \ bool" (infixr "measurable'_isomorphic" 50) where "M measurable_isomorphic N \ (\f. measurable_isomorphic_map M N f)" lemma measurable_isomorphic_sets_cong: assumes "sets M = sets M'" "sets N = sets N'" shows "M measurable_isomorphic N \ M' measurable_isomorphic N'" using measurable_isomorphic_map_sets_cong[OF assms] by(auto simp: measurable_isomorphic_def) lemma measurable_isomorphicD: assumes "M measurable_isomorphic N" shows "\f g. f \ M \\<^sub>M N \ g \ N \\<^sub>M M \ (\x\space M. g (f x) = x) \ (\y\space N. f (g y) = y) \ (\A\sets M. f ` A \ sets N) \ (\A\sets N. g ` A \ sets M)" using assms measurable_isomorphic_mapD'[of M N] by (metis (mono_tags, lifting) measurable_isomorphic_def) lemma measurable_isomorphic_cardinality_eq: assumes "M measurable_isomorphic N" shows "space M \ space N" by (meson assms eqpoll_def measurable_isomorphic_def measurable_isomorphic_map_def) lemma measurable_isomorphic_count_spaces: "count_space A measurable_isomorphic count_space B \ A \ B" proof assume "A \ B" then obtain f where f:"bij_betw f A B" by(auto simp: eqpoll_def) then show "count_space A measurable_isomorphic count_space B" by(auto simp: measurable_isomorphic_def measurable_isomorphic_map_def bij_betw_def the_inv_into_into intro!: exI[where x=f]) qed(use measurable_isomorphic_cardinality_eq in fastforce) lemma measurable_isomorphic_byWitness: assumes "f \ M \\<^sub>M N" "\x. x\space M \ g (f x) = x" and "g \ N \\<^sub>M M" "\y. y\space N \ f (g y) = y" shows "M measurable_isomorphic N" by(auto simp: measurable_isomorphic_def assms intro!: exI[where x = f] measurable_isomorphic_map_byWitness[where g=g]) lemma measurable_isomorphic_refl: "M measurable_isomorphic M" by(auto intro!: measurable_isomorphic_byWitness[where f=id and g=id]) lemma measurable_isomorphic_sym: assumes "M measurable_isomorphic N" shows "N measurable_isomorphic M" using assms measurable_isomorphic_map_inv[of M N] by(auto simp: measurable_isomorphic_def) lemma measurable_isomorphic_trans: assumes "M measurable_isomorphic N" and "N measurable_isomorphic L" shows "M measurable_isomorphic L" using assms measurable_isomorphic_map_comp[of M N _ L] by(auto simp: measurable_isomorphic_def) lemma measurable_isomorphic_empty: assumes "space M = {}" "space N = {}" shows "M measurable_isomorphic N" using assms by(auto intro!: measurable_isomorphic_byWitness[where f=undefined and g=undefined] simp: measurable_empty_iff) lemma measurable_isomorphic_empty1: assumes "space M = {}" "M measurable_isomorphic N" shows "space N = {}" using measurable_isomorphicD[OF assms(2)] by(auto simp: measurable_empty_iff[OF assms(1)]) lemma measurable_ismorphic_empty2: assumes "space N = {}" "M measurable_isomorphic N" shows "space M = {}" using measurable_isomorphic_sym[OF assms(2)] assms(1) by(simp add: measurable_isomorphic_empty1) lemma measurable_lift_product: assumes "\i. i \ I \ f i \ (M i) \\<^sub>M (N i)" shows "(\x i. if i \ I then f i (x i) else undefined) \ (\\<^sub>M i\I. M i) \\<^sub>M (\\<^sub>M i\I. N i)" using measurable_space[OF assms] by(auto intro!: measurable_PiM_single' simp: assms measurable_PiM_component_rev space_PiM PiE_iff) lemma measurable_isomorphic_map_lift_product: assumes "\i. i \ I \ measurable_isomorphic_map (M i) (N i) (h i)" shows "measurable_isomorphic_map (\\<^sub>M i\I. M i) (\\<^sub>M i\I. N i) (\x i. if i \ I then h i (x i) else undefined)" proof - obtain h' where "\i. i \ I \ h' i \ (N i) \\<^sub>M (M i)" "\i x. i \ I \ x\space (M i) \ h' i (h i x) = x" "\i x. i \ I \ x\space (N i) \ h i (h' i x) = x" using measurable_isomorphic_mapD'(3)[OF assms] by metis thus ?thesis by(auto intro!: measurable_isomorphic_map_byWitness[OF measurable_lift_product[of I h M N,OF measurable_isomorphic_mapD'(2)[OF assms]] measurable_lift_product[of I h' N M,OF \\i. i \ I \ h' i \ (N i) \\<^sub>M (M i)\]] simp: space_PiM PiE_iff extensional_def) qed lemma measurable_isomorphic_lift_product: assumes "\i. i \ I \ (M i) measurable_isomorphic (N i)" shows "(\\<^sub>M i\I. M i) measurable_isomorphic (\\<^sub>M i\I. N i)" proof - obtain h where "\i. i \ I \ measurable_isomorphic_map (M i) (N i) (h i)" using assms by(auto simp: measurable_isomorphic_def) metis thus ?thesis by(auto intro!: measurable_isomorphic_map_lift_product exI[where x="\x i. if i \ I then h i (x i) else undefined"] simp: measurable_isomorphic_def) qed text \\<^url>\https://math24.net/cantor-schroder-bernstein-theorem.html\\ lemma Schroeder_Bernstein_measurable': assumes "f ` (space M) \ sets N" "g ` (space N) \ sets M" and "measurable_isomorphic_map M (restrict_space N (f ` (space M))) f" and "measurable_isomorphic_map N (restrict_space M (g ` (space N))) g" shows "\h. measurable_isomorphic_map M N h" proof - have hset:"\A. A \ sets M \ f ` A \ sets (restrict_space N (f ` space M))" "\A. A \ sets N \ g ` A \ sets (restrict_space M (g ` space N))" and hfg[measurable]:"f \ M \\<^sub>M restrict_space N (f ` space M)" "g \ N \\<^sub>M restrict_space M (g ` space N)" using measurable_isomorphic_mapD'(1,2)[OF assms(3)] measurable_isomorphic_mapD'(1,2)[OF assms(4)] assms(1,2) by auto have hset2:"\A. A \ sets M \ f ` A \ sets N" "\A. A \ sets N \ g ` A \ sets M" and hfg2[measurable]: "f \ M \\<^sub>M N" "g \ N \\<^sub>M M" using sets.Int_space_eq2[OF assms(1)] sets.Int_space_eq2[OF assms(2)] sets_restrict_space_iff[of "f ` space M" N] sets_restrict_space_iff[of "g ` space N" M] hset measurable_restrict_space2_iff[of f M N] measurable_restrict_space2_iff[of g N M] hfg assms(1,2) by auto have bij1:"bij_betw f (space M) (f ` (space M))" "bij_betw g (space N) (g ` (space N))" using assms(3,4) by(auto simp: measurable_isomorphic_map_def space_restrict_space sets.Int_space_eq2[OF assms(1)] sets.Int_space_eq2[OF assms(2)]) obtain f' g' where hfg1'[measurable]: "f' \ restrict_space N (f ` (space M)) \\<^sub>M M" "g' \ restrict_space M (g ` (space N)) \\<^sub>M N" and hfg':"\x. x\space M \ f' (f x) = x" "\x. x\f ` space M \ f (f' x) = x" "\x. x\space N \ g' (g x) = x" "\x. x\g ` space N \ g (g' x) = x" "bij_betw f' (f ` space M) (space M)" "bij_betw g' (g ` space N) (space N)" using measurable_isomorphic_mapD'(3)[OF assms(3)] measurable_isomorphic_mapD'(3)[OF assms(4)] sets.Int_space_eq2[OF assms(1)] sets.Int_space_eq2[OF assms(2)] by (metis space_restrict_space) have hgfA:"(g \ f) ` A \ sets M" if "A \ sets M" for A using hset2(2)[OF hset2(1)[OF that]] by(simp add: image_comp) define An where "An \ (\n. ((g \ f)^^n) ` (space M - g ` (space N)))" define A where "A \ (\n\UNIV. An n)" have "An n \ sets M" for n proof(induction n) case 0 thus ?case using hset2[OF sets.top] by(simp add: An_def) next case ih:(Suc n) have "An (Suc n) = (g \ f) ` (An n)" by(auto simp add: An_def) thus ?case using hgfA[OF ih] by simp qed hence Asets:"A \ sets M" by(simp add: A_def) have Acompl:"space M - A \ g ` space N" proof - have "space M - A \ space M - An 0" by(auto simp: A_def) also have "... \ g ` space N" by(auto simp: An_def) finally show ?thesis . qed define h where "h \ (\x. if x \ A \ (- space M) then f x else g' x)" define h' where "h' \ (\x. if x \ f ` A then f' x else g x)" have xinA_iff:"x \ A \ h x \ f ` A" if "x \ space M" for x proof assume "h x \ f ` A" show "x \ A" proof(rule ccontr) assume "x \ A" then have "\n. x \ An n" by(auto simp: A_def) from this[of 0] have "x \ g ` (space N)" using that by(auto simp: An_def) have "g' x \ f ` A " using \h x \ f ` A\ \x \ A\ by (simp add: h_def that) hence "g (g' x) \ (g \ f) ` A" by auto hence "x \ (g \ f) ` A" using \x \ g ` (space N)\ by (simp add: hfg'(4)) then obtain n where "x \ (g \ f) ` (An n)" by(auto simp: A_def) hence "x \ An (Suc n)" by(auto simp: An_def) thus False using \\n. x \ An n\ by simp qed qed(simp add: h_def) show ?thesis proof(intro exI[where x=h] measurable_isomorphic_map_byWitness[where g=h']) have "{x \ space M. x \ A \ (- space M)} \ sets M" using sets.Int_space_eq2[OF Asets] Asets by simp moreover have "f \ restrict_space M {x. x \ A \ - space M} \\<^sub>M N" by (simp add: measurable_restrict_space1) moreover have "g' \ restrict_space M {x. x \ A \ (- space M)} \\<^sub>M N" proof - have "sets (restrict_space (restrict_space M (g ` space N)) {x. x \ A \ - space M}) = sets (restrict_space M (g ` space N \ {x. x \ A \ - space M}))" by(simp add: sets_restrict_restrict_space) also have "... = sets (restrict_space M (g ` space N \ {x. x \ space M - A}))" by (metis Compl_iff DiffE DiffI Un_iff) also have "... = sets (restrict_space M {x. x \ space M - A})" by (metis Acompl le_inf_iff mem_Collect_eq subsetI subset_antisym) also have "... = sets (restrict_space M {x. x \ A \ (- space M)})" by (metis Compl_iff DiffE DiffI Un_iff) finally have "sets (restrict_space (restrict_space M (g ` space N)) {x. x \ A \ - space M}) = sets (restrict_space M {x. x \ A \ - space M})" . from measurable_cong_sets[OF this refl] measurable_restrict_space1[OF hfg1'(2),of " {x. x \ A \ - space M}"] show ?thesis by auto qed ultimately show "h \ M \\<^sub>M N" by(simp add: h_def measurable_If_restrict_space_iff) next have "{x \ space N. x \ f ` A} \ sets N" using sets.Int_space_eq2[OF hset2(1)[OF Asets]] hset2(1)[OF Asets] by simp moreover have "f' \ restrict_space N {x. x \ f ` A} \\<^sub>M M" proof - have "sets (restrict_space (restrict_space N (f ` space M)) {x. x \ f ` A}) = sets (restrict_space N (f ` space M \ {x. x \ f ` A}))" by(simp add: sets_restrict_restrict_space) also have "... = sets (restrict_space N {x. x \ f ` A})" proof - have "f ` space M \ {x. x \ f ` A} = {x. x \ f ` A}" using sets.sets_into_space[OF Asets] by auto thus ?thesis by simp qed finally have "sets (restrict_space (restrict_space N (f ` space M)) {x. x \ f ` A}) = sets (restrict_space N {x. x \ f ` A})" . from measurable_cong_sets[OF this refl] measurable_restrict_space1[OF hfg1'(1),of "{x. x \ f ` A}"] show ?thesis by auto qed moreover have "g \ restrict_space N {x. x \ f ` A} \\<^sub>M M" by (simp add: measurable_restrict_space1) ultimately show "h' \ N \\<^sub>M M" by(simp add: h'_def measurable_If_restrict_space_iff) next fix x assume "x \ space M" then consider "x \ A" | "x \ space M - A" by auto thus "h' (h x) = x" proof cases case xa:2 hence "h x \ f ` A" using \x \ space M\ xinA_iff by blast thus ?thesis using Acompl hfg'(4) xa by(auto simp add: h_def h'_def) qed(simp add: h_def h'_def \x \ space M\ hfg'(1)) next fix x assume "x \ space N" then consider "x \ f ` A" | "x \ space N - f ` A" by auto thus "h (h' x) = x" proof cases case hx:1 hence "x \ f ` (space M)" using image_mono[OF sets.sets_into_space[OF Asets],of f] by auto have "h' x = f' x" using hx by(simp add: h'_def) also have "... \ A" using hx sets.sets_into_space[OF Asets] hfg'(1) by auto finally show ?thesis using hfg'(2)[OF \x \ f ` (space M)\] hx by(auto simp: h_def h'_def) next case hx:2 then have "h' x = g x" by(simp add: h'_def) also have "... \ A" proof(rule ccontr) assume "\ g x \ A" then have "g x \ A" by simp then obtain n where hg:"g x \ An n" by(auto simp: A_def) hence "0 < n" using hx by(auto simp: An_def) then obtain n' where [simp]:"n = Suc n'" using not0_implies_Suc by blast then have "g x \ g ` f ` An n'" using hg by(auto simp: An_def) hence "x \ f ` An n'" using inj_on_image_mem_iff[OF bij_betw_imp_inj_on[OF bij1(2)] \x \ space N\,of "f ` An n'"] sets.sets_into_space[OF \An n' \ sets M\] measurable_space[OF hfg2(1)] by auto also have "... \ f ` A" by(auto simp: A_def) finally show False using hx by simp qed finally show ?thesis using hx hfg'(3)[OF \x \ space N\] measurable_space[OF hfg2(2) \x \ space N\] by(auto simp: h_def h'_def) qed qed qed lemma Schroeder_Bernstein_measurable: assumes "f \ M \\<^sub>M N" "\A. A \ sets M \ f ` A \ sets N" "inj_on f (space M)" and "g \ N \\<^sub>M M" "\A. A \ sets N \ g ` A \ sets M" "inj_on g (space N)" shows "\h. measurable_isomorphic_map M N h" using Schroeder_Bernstein_measurable'[OF assms(2)[OF sets.top] assms(5)[OF sets.top] measurable_isomorphic_map_restrict_space[OF assms(1-3)] measurable_isomorphic_map_restrict_space[OF assms(4-6)]] by simp lemma measurable_isomorphic_from_embeddings: assumes "M measurable_isomorphic (restrict_space N B)" "N measurable_isomorphic (restrict_space M A)" and "A \ sets M" "B \ sets N" shows "M measurable_isomorphic N" proof - obtain f g where fg:"measurable_isomorphic_map M (restrict_space N B) f" "measurable_isomorphic_map N (restrict_space M A) g" using assms(1,2) by(auto simp: measurable_isomorphic_def) have [simp]:"f ` space M = B" "g ` space N = A" using measurable_isomorphic_map_surj[OF fg(1)] measurable_isomorphic_map_surj[OF fg(2)] sets.sets_into_space[OF assms(3)] sets.sets_into_space[OF assms(4)] by(auto simp: space_restrict_space) obtain h where "measurable_isomorphic_map M N h" using Schroeder_Bernstein_measurable'[of f M N g] assms(3,4) fg by auto thus ?thesis by(auto simp: measurable_isomorphic_def) qed lemma measurable_isomorphic_antisym: assumes "B measurable_isomorphic (restrict_space C c)" "A measurable_isomorphic (restrict_space B b)" and "c \ sets C" "b \ sets B" "C measurable_isomorphic A" shows "C measurable_isomorphic B" by(rule measurable_isomorphic_from_embeddings[OF measurable_isomorphic_trans[OF assms(5,2)] assms(1) assms(3,4)]) lemma countable_infinite_isomorphisc_to_nat_index: assumes "countable I" and "infinite I" shows "(\\<^sub>M x\I. M) measurable_isomorphic (\\<^sub>M (x::nat)\UNIV. M)" proof(rule measurable_isomorphic_byWitness[where f="\x n. x (from_nat_into I n)" and g="\x. \i\I. x (to_nat_on I i)"]) show "(\x n. x (from_nat_into I n)) \ (\\<^sub>M x\I. M) \\<^sub>M (\\<^sub>M (x::nat)\UNIV. M)" by(auto intro!: measurable_PiM_single' measurable_component_singleton[OF from_nat_into[OF infinite_imp_nonempty[OF assms(2)]]]) (simp add: PiE_iff infinite_imp_nonempty space_PiM from_nat_into[OF infinite_imp_nonempty[OF assms(2)]]) next show "(\x. \i\I. x (to_nat_on I i)) \ (\\<^sub>M (x::nat)\UNIV. M) \\<^sub>M (\\<^sub>M x\I. M)" by(auto intro!: measurable_PiM_single') next show "x \ space (\\<^sub>M x\I. M) \ (\i\I. x (from_nat_into I (to_nat_on I i))) = x" for x by (simp add: assms(1) restrict_ext space_PiM) next show "y \ space (Pi\<^sub>M UNIV (\x. M)) \ (\n. (\i\I. y (to_nat_on I i)) (from_nat_into I n)) = y" for y by (simp add: assms(1) assms(2) from_nat_into infinite_imp_nonempty) qed lemma PiM_PiM_isomorphic_to_PiM: "(\\<^sub>M i\I. \\<^sub>M j\J. M i j) measurable_isomorphic (\\<^sub>M (i,j)\I\J. M i j)" proof(rule measurable_isomorphic_byWitness[where f="\x (i,j). if (i,j) \ I \ J then x i j else undefined" and g="\x i j. if i \ I then undefined j else if j \ J then undefined else x (i,j)"]) have [simp]: "(\\. \ a b) \ (\\<^sub>M i\I. \\<^sub>M j\J. M i j) \\<^sub>M M a b" if "a \ I" "b \ J" for a b using measurable_component_singleton[OF that(1),of "\i. \\<^sub>M j\J. M i j"] measurable_component_singleton[OF that(2),of "M a"] by auto show "(\x (i, j). if (i, j) \ I \ J then x i j else undefined) \ (\\<^sub>M i\I. \\<^sub>M j\J. M i j) \\<^sub>M (\\<^sub>M (i,j)\I\J. M i j)" apply(rule measurable_PiM_single') apply auto[1] apply(auto simp: PiE_def Pi_def space_PiM extensional_def;meson) done next have [simp]: "(\\. \ (i, j)) \ Pi\<^sub>M (I \ J) (\(i, j). M i j) \\<^sub>M M i j" if "i \ I" "j \ J" for i j using measurable_component_singleton[of "(i,j)" "I \ J" "\(i, j). M i j"] that by auto show "(\x i j. if i \ I then undefined j else if j \ J then undefined else x (i, j)) \ (\\<^sub>M (i,j)\I\J. M i j) \\<^sub>M (\\<^sub>M i\I. \\<^sub>M j\J. M i j)" by(auto intro!: measurable_PiM_single') (simp_all add: PiE_iff space_PiM extensional_def) next show "x \ space (\\<^sub>M i\I. \\<^sub>M j\J. M i j) \ (\i j. if i \ I then undefined j else if j \ J then undefined else case (i, j) of (i, j) \ if (i, j) \ I \ J then x i j else undefined) = x" for x by standard+ (auto simp: space_PiM PiE_def Pi_def extensional_def) next show "y \ space (\\<^sub>M (i,j)\I\J. M i j) \ (\(i, j). if (i, j) \ I \ J then if i \ I then undefined j else if j \ J then undefined else y (i, j) else undefined) = y" for y by standard+ (auto simp: space_PiM PiE_def Pi_def extensional_def) qed lemma measurable_isomorphic_map_sigma_sets: assumes "sets M = sigma_sets (space M) U" "measurable_isomorphic_map M N f" shows "sets N = sigma_sets (space N) ((`) f ` U)" proof - from measurable_isomorphic_mapD'[OF assms(2)] obtain g where h: "\A. A \ sets M \ f ` A \ sets N" "f \ M \\<^sub>M N" "bij_betw g (space N) (space M)" "g \ N \\<^sub>M M" "\x. x\space M \ g (f x) = x" "\x. x\space N \ f (g x) = x" "\A. A\sets N \ g ` A \ sets M" by metis interpret s: sigma_algebra "space N" "sigma_sets (space N) ((`) f ` U)" by(auto intro!: sigma_algebra_sigma_sets) (metis assms(1) h(2) measurable_space sets.sets_into_space sigma_sets_superset_generator subsetD) show ?thesis proof safe fix x assume "x \ sets N" from h(7)[OF this] assms(1) have "g ` x \ sigma_sets (space M) U" by simp hence "f ` (g ` x) \ sigma_sets (space N) ((`) f ` U)" proof induction case h:(Compl a) have "f ` (space M - a) = f ` (space M) - f ` a" by(rule inj_on_image_set_diff[where C="space M"], insert assms h) (auto simp: measurable_isomorphic_map_def bij_betw_def sets.sets_into_space) with h show ?case by (metis assms(2) measurable_isomorphic_map_surj s.Diff s.top) qed (auto simp: image_UN) moreover have "f ` (g ` x) = x" using sets.sets_into_space[OF \x \ sets N\] h(6) by(fastforce simp: image_def) ultimately show "x \ sigma_sets (space N) ((`) f ` U)" by simp next interpret s': sigma_algebra "space M" "sigma_sets (space M) U" by(simp add: assms(1)[symmetric] sets.sigma_algebra_axioms) have 1:"\x. x \ U \ x \ space M" by (simp add: s'.sets_into_space) fix x assume assm:"x \ sigma_sets (space N) ((`) f ` U)" then show "x \ sets N" by induction (auto simp: assms(1) h(1)) qed qed subsubsection \Borel Spaces Genereted from Abstract Topologies\ definition borel_of :: "'a topology \ 'a measure" where "borel_of X \ sigma (topspace X) {U. openin X U}" lemma emeasure_borel_of: "emeasure (borel_of X) A = 0" by (simp add: borel_of_def emeasure_sigma) lemma borel_of_euclidean: "borel_of euclidean = borel" by(simp add: borel_of_def borel_def) lemma space_borel_of: "space (borel_of X) = topspace X" by(simp add: space_measure_of_conv borel_of_def) lemma sets_borel_of: "sets (borel_of X) = sigma_sets (topspace X) {U. openin X U}" by (simp add: subset_Pow_Union topspace_def borel_of_def) lemma sets_borel_of_closed: "sets (borel_of X) = sigma_sets (topspace X) {U. closedin X U}" unfolding sets_borel_of proof(safe intro!: sigma_sets_eqI) fix a assume a:"openin X a" have "topspace X - (topspace X - a) \ sigma_sets (topspace X) {U. closedin X U}" by(rule sigma_sets.Compl) (use a in auto) thus "a \ sigma_sets (topspace X) {U. closedin X U}" using openin_subset[OF a] by (simp add: Diff_Diff_Int inf.absorb_iff2) next fix b assume b:"closedin X b" have "topspace X - (topspace X - b) \ sigma_sets (topspace X) {U. openin X U}" by(rule sigma_sets.Compl) (use b in auto) thus "b \ sigma_sets (topspace X) {U. openin X U}" using closedin_subset[OF b] by (simp add: Diff_Diff_Int inf.absorb_iff2) qed lemma borel_of_open: assumes "openin X U" shows "U \ sets (borel_of X)" using assms by (simp add: subset_Pow_Union topspace_def borel_of_def) lemma borel_of_closed: assumes "closedin X U" shows "U \ sets (borel_of X)" using assms sigma_sets.Compl[of "topspace X - U" "topspace X"] by (simp add: closedin_def double_diff sets_borel_of) lemma(in Metric_space) nbh_sets[measurable]: "(\a\A. mball a e) \ sets (borel_of mtopology)" by(auto intro!: borel_of_open openin_clauses(3)) lemma borel_of_gdelta_in: assumes "gdelta_in X U" shows "U \ sets (borel_of X)" using gdelta_inD[OF assms] borel_of_open by(auto intro!: sets.countable_INT'[of _ id,simplified]) lemma borel_of_subtopology: "borel_of (subtopology X U) = restrict_space (borel_of X) U" proof(rule measure_eqI) show "sets (borel_of (subtopology X U)) = sets (restrict_space (borel_of X) U)" unfolding restrict_space_eq_vimage_algebra' sets_vimage_algebra sets_borel_of topspace_subtopology space_borel_of Int_commute[of U] proof(rule sigma_sets_eqI) fix a assume "a \ Collect (openin (subtopology X U))" then obtain T where "openin X T" "a = T \ U" by(auto simp: openin_subtopology) show "a \ sigma_sets (topspace X \ U) {(\x. x) -` A \ (topspace X \ U) |A. A \ sigma_sets (topspace X) (Collect (openin X))}" using openin_subset[OF \openin X T\] \a = T \ U\ by(auto intro!: exI[where x=T] \openin X T\) next fix b assume "b \ {(\x. x) -` A \ (topspace X \ U) |A. A \ sigma_sets (topspace X) (Collect (openin X))}" then obtain T where ht:"b = T \ (topspace X \ U)" "T \ sigma_sets (topspace X) (Collect (openin X))" by auto hence "b = T \ U" proof - have "T \ topspace X" by(rule sigma_sets_into_sp[OF _ ht(2)]) (simp add: subset_Pow_Union topspace_def) thus ?thesis by(auto simp: ht(1)) qed with ht(2) show "b \ sigma_sets (topspace X \ U) (Collect (openin (subtopology X U)))" proof(induction arbitrary: b U) case (Basic a) then show ?case by(auto simp: openin_subtopology) next case Empty then show ?case by simp next case ih:(Compl a) then show ?case by (simp add: Diff_Int_distrib2 sigma_sets.Compl) next case (Union a) then show ?case by (metis UN_extend_simps(4) sigma_sets.Union) qed qed qed(simp add: emeasure_borel_of restrict_space_def emeasure_measure_of_conv) lemma sets_borel_of_discrete_topology: "sets (borel_of (discrete_topology I)) = sets (count_space I)" by (metis Pow_UNIV UNIV_eq_I borel_of_open borel_of_subtopology inf.absorb_iff2 openin_discrete_topology sets_count_space sets_restrict_space sets_restrict_space_count_space subtopology_discrete_topology top_greatest) lemma continuous_map_measurable: assumes "continuous_map X Y f" shows "f \ borel_of X \\<^sub>M borel_of Y" proof(rule measurable_sigma_sets[OF sets_borel_of[of Y]]) show "{U. openin Y U} \ Pow (topspace Y)" by (simp add: subset_Pow_Union topspace_def) next show "f \ space (borel_of X) \ topspace Y" using continuous_map_image_subset_topspace[OF assms] by(auto simp: space_borel_of) next fix U assume "U \ {U. openin Y U}" then have "openin X (f -` U \ topspace X)" using continuous_map[of X Y f] assms by auto thus "f -` U \ space (borel_of X) \ sets (borel_of X)" by(simp add: space_borel_of sets_borel_of) qed lemma upper_semicontinuous_map_measurable: fixes f :: "'a \ 'b :: {linorder_topology, second_countable_topology}" assumes "upper_semicontinuous_map X f" shows "f \ borel_measurable (borel_of X)" using assms by(auto intro!: borel_measurableI_less borel_of_open simp: space_borel_of upper_semicontinuous_map_def) lemma lower_semicontinuous_map_measurable: fixes f :: "'a \ 'b :: {linorder_topology, second_countable_topology}" assumes "lower_semicontinuous_map X f" shows "f \ borel_measurable (borel_of X)" using assms by(auto intro!: borel_measurableI_greater borel_of_open simp: space_borel_of lower_semicontinuous_map_def) lemma open_map_preserves_sets: assumes "open_map S T f" "inj_on f (topspace S)" "A \ sets (borel_of S)" shows "f ` A \ sets (borel_of T)" using assms(3)[simplified sets_borel_of] proof(induction) case (Basic a) with assms(1) show ?case by(auto simp: sets_borel_of open_map_def) next case Empty show ?case by simp next case (Compl a) moreover have "f ` (topspace S - a) = f ` (topspace S) - f ` a" by (metis Diff_subset assms(2) calculation(1) inj_on_image_set_diff sigma_sets_into_sp subset_Pow_Union topspace_def) moreover have "f ` (topspace S) \ sets (borel_of T)" by (meson assms(1) borel_of_open open_map_def openin_topspace) ultimately show ?case by auto next case (Union a) then show ?case by (simp add: image_UN) qed lemma open_map_preserves_sets': assumes "open_map S (subtopology T (f ` (topspace S))) f" "inj_on f (topspace S)" "f ` (topspace S) \ sets (borel_of T)" "A \ sets (borel_of S)" shows "f ` A \ sets (borel_of T)" using assms(4)[simplified sets_borel_of] proof(induction) case (Basic a) then have "openin (subtopology T (f ` (topspace S))) (f ` a)" using assms(1) by(auto simp: open_map_def) hence "f ` a \ sets (borel_of (subtopology T (f ` (topspace S))))" by(simp add: sets_borel_of) hence "f ` a \ sets (restrict_space (borel_of T) (f ` (topspace S)))" by(simp add: borel_of_subtopology) thus ?case by (metis sets_restrict_space_iff assms(3) sets.Int_space_eq2) next case Empty show ?case by simp next case (Compl a) moreover have "f ` (topspace S - a) = f ` (topspace S) - f ` a" by (metis Diff_subset assms(2) calculation(1) inj_on_image_set_diff sigma_sets_into_sp subset_Pow_Union topspace_def) ultimately show ?case using assms(3) by auto next case (Union a) then show ?case by (simp add: image_UN) qed text \ Abstract topology version of @{thm second_countable_borel_measurable}. \ lemma borel_of_second_countable': assumes "second_countable S" and "subbase_in S \" shows "borel_of S = sigma (topspace S) \" unfolding borel_of_def proof(rule sigma_eqI) show "{U. openin S U} \ Pow (topspace S)" by (simp add: subset_Pow_Union topspace_def) next show "\ \ Pow (topspace S)" using subbase_in_subset[OF assms(2)] by auto next interpret s: sigma_algebra "topspace S" "sigma_sets (topspace S) \" using subbase_in_subset[OF assms(2)] by(auto intro!: sigma_algebra_sigma_sets) obtain \ where ho: "countable \" "base_in S \" using assms(1) by(auto simp: second_countable_base_in) show "sigma_sets (topspace S) {U. openin S U} = sigma_sets (topspace S) \" proof(rule sigma_sets_eqI) fix U assume "U \ {U. openin S U}" then have "generate_topology_on \ U" using assms(2) by(simp add: subbase_in_def openin_topology_generated_by_iff) thus "U \ sigma_sets (topspace S) \" proof induction case (UN K) with ho(2) obtain V where hv: "\k. k \ K \ V k \ \" "\k. k \ K \ \ (V k) = k" by(simp add: base_in_def openin_topology_generated_by_iff[symmetric] assms(2)[simplified subbase_in_def,symmetric]) metis define \k where "\k = (\k\K. V k)" have 0:"countable \k" using hv by(auto intro!: countable_subset[OF _ ho(1)] simp: \k_def) have "\ \k = (\A\\k. A)" by auto also have "... = \ K" unfolding \k_def UN_simps by(simp add: hv(2)) finally have 1:"\ \k = \ K" . have "\b\\k. \k\K. b \ k" using hv by (auto simp: \k_def) then obtain V' where hv': "\b. b \ \k \ V' b \ K" and "\b. b \ \k \ b \ V' b" by metis then have "(\b\\k. V' b) \ \K" "\\k \ (\b\\k. V' b)" by auto then have "\K = (\b\\k. V' b)" unfolding 1 by auto also have "\ \ sigma_sets (topspace S) \" using hv' UN by(auto intro!: s.countable_UN' simp: 0) finally show "\K \ sigma_sets (topspace S) \" . qed auto next fix U assume "U \ \" from assms(2)[simplified subbase_in_def] openin_topology_generated_by_iff generate_topology_on.Basis[OF this] show "U \ sigma_sets (topspace S) {U. openin S U}" by auto qed qed text \ Abstract topology version @{thm borel_prod}.\ lemma borel_of_prod: assumes "second_countable S" and "second_countable S'" shows "borel_of S \\<^sub>M borel_of S' = borel_of (prod_topology S S')" proof - have "borel_of S \\<^sub>M borel_of S' = sigma (topspace S \ topspace S') {a \ b |a b. a \ {a. openin S a} \ b \ {b. openin S' b}}" proof - obtain \ \' where ho: "countable \" "base_in S \" "countable \'" "base_in S' \'" using assms by(auto simp: second_countable_base_in) show ?thesis unfolding borel_of_def apply(rule sigma_prod) using topology_generated_by_topspace[of \,simplified base_is_subbase[OF ho(2),simplified subbase_in_def,symmetric]] topology_generated_by_topspace[of \',simplified base_is_subbase[OF ho(4),simplified subbase_in_def,symmetric]] base_in_openin[OF ho(2)] base_in_openin[OF ho(4)] by(auto intro!: exI[where x=\] exI[where x=\'] simp: ho subset_Pow_Union topspace_def) qed also have "... = borel_of (prod_topology S S')" using borel_of_second_countable'[OF prod_topology_second_countable[OF assms],simplified subbase_in_def,OF prod_topology_generated_by_open] by simp finally show ?thesis . qed lemma product_borel_of_measurable: assumes "i \ I" shows "(\x. x i) \ (borel_of (product_topology S I)) \\<^sub>M borel_of (S i)" by(auto intro!: continuous_map_measurable simp: assms) text \ Abstract topology version of @{thm sets_PiM_subset_borel} \ lemma sets_PiM_subset_borel_of: "sets (\\<^sub>M i\I. borel_of (S i)) \ sets (borel_of (product_topology S I))" proof - have *: "(\\<^sub>E i\I. X i) \ sets (borel_of (product_topology S I))" if [measurable]:"\i. X i \ sets (borel_of (S i))" "finite {i. X i \ topspace (S i)}" for X proof - note [measurable] = product_borel_of_measurable define I' where "I' = {i. X i \ topspace (S i)} \ I" have "finite I'" unfolding I'_def using that by simp have "(\\<^sub>E i\I. X i) = (\i\I'. (\x. x i)-`(X i) \ space (borel_of (product_topology S I))) \ space (borel_of (product_topology S I))" proof(standard;standard) fix x assume "x \ Pi\<^sub>E I X" then show "x \ (\i\I'. (\x. x i) -` X i \ space (borel_of (product_topology S I))) \ space (borel_of (product_topology S I))" using sets.sets_into_space[OF that(1)] by(auto simp: PiE_def I'_def Pi_def space_borel_of) next fix x assume 1:"x \ (\i\I'. (\x. x i) -` X i \ space (borel_of (product_topology S I))) \ space (borel_of (product_topology S I))" have "x i \ X i" if hi:"i \ I" for i proof - consider "i \ I' \ I' \ {}" | "i \ I' \ I' = {}" | "i \ I' \ I' \ {}" by auto then show ?thesis apply cases using sets.sets_into_space[OF \\i. X i \ sets (borel_of (S i))\] 1 that by(auto simp: space_borel_of I'_def) qed then show "x \ Pi\<^sub>E I X" using 1 by(auto simp: space_borel_of) qed also have "... \ sets (borel_of (product_topology S I))" using that \finite I'\ by(auto simp: I'_def) finally show ?thesis . qed then have "{Pi\<^sub>E I X |X. (\i. X i \ sets (borel_of (S i))) \ finite {i. X i \ space (borel_of (S i))}} \ sets (borel_of (product_topology S I))" by(auto simp: space_borel_of) show ?thesis unfolding sets_PiM_finite by(rule sets.sigma_sets_subset',fact) (simp add: borel_of_open[OF openin_topspace, of "product_topology S I",simplified] space_borel_of) qed text \ Abstract topology version of @{thm sets_PiM_equal_borel}.\ lemma sets_PiM_equal_borel_of: assumes "countable I" and "\i. i \ I \ second_countable (S i)" shows "sets (\\<^sub>M i\I. borel_of (S i)) = sets (borel_of (product_topology S I))" proof obtain K where hk: "countable K" "base_in (product_topology S I) K" "\k. k \ K \ \X. (k = (\\<^sub>E i\I. X i)) \ (\i. openin (S i) (X i)) \ finite {i. X i \ topspace (S i)} \ {i. X i \ topspace (S i)} \ I" using product_topology_countable_base_in[OF assms(1)] assms(2) by force have *:"k \ sets (\\<^sub>M i\I. borel_of (S i))" if "k \ K" for k proof - obtain X where H: "k = (\\<^sub>E i\I. X i)" "\i. openin (S i) (X i)" "finite {i. X i \ topspace (S i)}" "{i. X i \ topspace (S i)} \ I" using hk(3)[OF \k \ K\] by blast show ?thesis unfolding H(1) sets_PiM_finite using borel_of_open[OF H(2)] H(3) by(auto simp: space_borel_of) qed have **: "U \ sets (\\<^sub>M i\I. borel_of (S i))" if "openin (product_topology S I) U" for U proof - obtain B where "B \ K" "U = (\B)" using \openin (product_topology S I) U\ \base_in (product_topology S I) K\ by (metis base_in_def) have "countable B" using \B \ K\ \countable K\ countable_subset by blast moreover have "k \ sets (\\<^sub>M i\I. borel_of (S i))" if "k \ B" for k using \B \ K\ * that by auto ultimately show ?thesis unfolding \U = (\B)\ by auto qed have "sigma_sets (topspace (product_topology S I)) {U. openin (product_topology S I) U} \ sets (\\<^sub>M i\I. borel_of (S i))" apply (rule sets.sigma_sets_subset') using ** by(auto intro!: sets_PiM_I_countable[OF assms(1)] simp: borel_of_open[OF openin_topspace]) thus " sets (borel_of (product_topology S I)) \ sets (\\<^sub>M i\I. borel_of (S i))" by (simp add: subset_Pow_Union topspace_def borel_of_def) qed(rule sets_PiM_subset_borel_of) lemma homeomorphic_map_borel_isomorphic: assumes "homeomorphic_map X Y f" shows "measurable_isomorphic_map (borel_of X) (borel_of Y) f" proof - obtain g where "homeomorphic_maps X Y f g" using assms by(auto simp: homeomorphic_map_maps) hence "continuous_map X Y f" "continuous_map Y X g" "\x. x \ topspace X \ g (f x) = x" "\y. y \ topspace Y \ f (g y) = y" by(auto simp: homeomorphic_maps_def) thus ?thesis by(auto intro!: measurable_isomorphic_map_byWitness dest: continuous_map_measurable simp: space_borel_of) qed lemma homeomorphic_space_measurable_isomorphic: assumes "S homeomorphic_space T" shows "borel_of S measurable_isomorphic borel_of T" using homeomorphic_map_borel_isomorphic[of S T] assms by(auto simp: measurable_isomorphic_def homeomorphic_space) lemma measurable_isomorphic_borel_map: assumes "sets M = sets (borel_of S)" and f: "measurable_isomorphic_map M N f" shows "\S'. homeomorphic_map S S' f \ sets N = sets (borel_of S')" proof - obtain g where fg:"f \ M \\<^sub>M N" "g \ N \\<^sub>M M" "\x. x\space M \ g (f x) = x" "\y. y\space N \ f (g y) = y" "\A. A\sets M \ f ` A \ sets N" "\A. A\sets N \ g ` A \ sets M" "bij_betw g (space N) (space M)" using measurable_isomorphic_mapD'[OF f] by metis have g:"measurable_isomorphic_map N M g" by(auto intro!: measurable_isomorphic_map_byWitness fg) have g':"bij_betw g (space N) (topspace S)" using fg(7) sets_eq_imp_space_eq[OF assms(1)] by(auto simp: space_borel_of) show ?thesis proof(intro exI[where x="pullback_topology (space N) g S"] conjI) have [simp]: "{U. openin (pullback_topology (space N) g S) U} = (`) f ` {U. openin S U}" unfolding openin_pullback_topology'[OF g'] proof safe fix u assume u:"openin S u" then have 1:"u \ space M" by(simp add: sets_eq_imp_space_eq[OF assms(1)] space_borel_of openin_subset) with fg(3) have "g ` f ` u = u" by(fastforce simp: image_def) with u show "openin S (g ` f ` u)" by simp fix x assume "x \ u" with 1 fg(1) show "f x \ space N" by(auto simp: measurable_space) next fix u assume "openin S (g ` u)" "u \ space N" with fg(4) show "u \ (`) f ` {U. openin S U}" by(auto simp: image_def intro!: exI[where x="g ` u"]) (metis in_mono) qed have [simp]:"g -` topspace S \ space N = space N" using bij_betw_imp_surj_on g' by blast show "sets N = sets (borel_of (pullback_topology (space N) g S))" by(auto simp: sets_borel_of topspace_pullback_topology intro!: measurable_isomorphic_map_sigma_sets[OF assms(1)[simplified sets_borel_of space_borel_of[symmetric] sets_eq_imp_space_eq[OF assms(1),symmetric]] f]) next show "homeomorphic_map S (pullback_topology (space N) g S) f" proof(rule homeomorphic_maps_imp_map[where g=g]) obtain f' where f':"homeomorphic_maps (pullback_topology (space N) g S) S g f'" using topology_from_bij(1)[OF g'] homeomorphic_map_maps by blast have f'2:"f' y = f y" if y:"y \ topspace S" for y proof - have [simp]:"g -` topspace S \ space N = space N" using bij_betw_imp_surj_on g' by blast obtain x where "x \ space N" "y = g x" using g' y by(auto simp: bij_betw_def image_def) thus ?thesis using fg(4) f' by(auto simp: homeomorphic_maps_def topspace_pullback_topology) qed thus "homeomorphic_maps S (pullback_topology (space N) g S) f g" by(auto intro!: homeomorphic_maps_eq[OF f'] simp: homeomorphic_maps_sym[of S]) qed qed qed lemma measurable_isomorphic_borels: assumes "sets M = sets (borel_of S)" "M measurable_isomorphic N" shows "\S'. S homeomorphic_space S' \ sets N = sets (borel_of S')" using measurable_isomorphic_borel_map[OF assms(1)] assms(2) homeomorphic_map_maps by(fastforce simp: measurable_isomorphic_def homeomorphic_space_def ) end \ No newline at end of file diff --git a/thys/Standard_Borel_Spaces/Set_Based_Metric_Space.thy b/thys/Standard_Borel_Spaces/Set_Based_Metric_Space.thy --- a/thys/Standard_Borel_Spaces/Set_Based_Metric_Space.thy +++ b/thys/Standard_Borel_Spaces/Set_Based_Metric_Space.thy @@ -1,2255 +1,2050 @@ (* Title: Set_Based_Metric_Space.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) subsection \Lemmas for Abstract Metric Spaces\ theory Set_Based_Metric_Space imports Lemmas_StandardBorel begin text \ We prove additional lemmas related to set-based metric spaces. \ subsubsection \ Basic Lemmas \ lemma assumes "Metric_space M d" "\x y. x \ M \ y \ M \ d x y = d' x y" and "\x y. d' x y = d' y x" "\x y. d' x y \ 0" shows Metric_space_eq: "Metric_space M d'" and Metric_space_eq_mtopology: "Metric_space.mtopology M d = Metric_space.mtopology M d'" and Metric_space_eq_mcomplete: "Metric_space.mcomplete M d \ Metric_space.mcomplete M d'" proof - interpret m1: Metric_space M d by fact show "Metric_space M d'" using assms by(auto simp: Metric_space_def) then interpret m2:Metric_space M d' . have [simp]: "m1.mball x e = m2.mball x e" for x e using assms by(auto simp: m1.mball_def m2.mball_def) show 1:"m1.mtopology = m2.mtopology" by(auto simp: topology_eq m1.openin_mtopology m2.openin_mtopology) show "m1.mcomplete = m2.mcomplete" by(auto simp: 1 m1.mcomplete_def m2.mcomplete_def m1.MCauchy_def m2.MCauchy_def assms(2) in_mono) qed context Metric_space begin lemma mtopology_base_in_balls: "base_in mtopology {mball a \ | a \. a \ M \ \ > 0}" proof - have 1:"\x. x \ {mball a \ | a \. a \ M \ \ > 0} \ openin mtopology x" by auto show ?thesis unfolding base_in_def2[of "{mball a \ | a \. a \ M \ \ > 0}",OF 1,simplified] by (metis centre_in_mball_iff in_mono openin_mtopology) qed lemma closedin_metric2: "closedin mtopology C \ C \ M \ (\x. x \ C \ (\\>0. mball x \ \ C \ {}))" proof assume h:"closedin mtopology C" have 1: "C \ M" using Metric_space.closedin_metric Metric_space_axioms h by blast show "C \ M \ (\x. x \ C \ (\\>0. mball x \ \ C \ {}))" proof safe fix \ x assume "x \ C" "(0 :: real) < \" "mball x \ \ C = {}" with 1 show False by blast next fix x assume "\\>0. mball x \ \ C \ {}" hence "\xn. xn \ mball x (1 / real (Suc n)) \ C" for n by (meson all_not_in_conv divide_pos_pos of_nat_0_less_iff zero_less_Suc zero_less_one) then obtain xn where xn:"\n. xn n \ mball x (1 / real (Suc n)) \ C" by metis hence xxn:"x \ M" "range xn \ C" using xn by auto have "limitin mtopology xn x sequentially" unfolding limitin_metric eventually_sequentially proof safe fix \ assume "(0 :: real) < \" then obtain N where hN: "1 / real (Suc N) < \" using nat_approx_posE by blast show "\N. \n\N. xn n \ M \ d (xn n) x < \" proof(safe intro!: exI[where x="N"]) fix n assume n[arith]: "N \ n" then have "1 / real (Suc n) < \" by (metis Suc_le_mono hN inverse_of_nat_le nat.distinct(1) order_le_less_trans) with xn[of n] show "d (xn n) x < \" by (simp add: commute) qed(use xxn 1 in auto) qed fact with h 1 xxn show "x \ C" by(auto simp: metric_closedin_iff_sequentially_closed) qed(use 1 in auto) next assume"C \ M \ (\x. (x \ C) \ (\\>0. mball x \ \ C \ {}))" hence h:"C \ M" "\x. (x \ C) \ (\\>0. mball x \ \ C \ {})" by simp_all show "closedin mtopology C" unfolding metric_closedin_iff_sequentially_closed proof safe fix xn x assume h':"range xn \ C" "limitin mtopology xn x sequentially" hence "x \ M" by (simp add: limitin_mspace) have "mball x \ \ C \ {}" if "\ > 0" for \ proof - obtain N where hN:"\n. n \ N \ d (xn n) x < \" using h'(2) \\ > 0\ limit_metric_sequentially by blast have "xn N \ mball x \ \ C" using h'(1) hN[of N] \x \ M\ commute h(1) by fastforce thus "mball x \ \ C \ {}" by auto qed with h(2)[of x] show "x \ C" by simp qed(use h(1) in auto) qed lemma openin_mtopology2: "openin mtopology U \ U \ M \ (\xn x. limitin mtopology xn x sequentially \ x \ U \ (\N. \n\N. xn n \ U))" unfolding openin_mtopology proof safe fix xn x assume h: "\x. x \ U \ (\r>0. mball x r \ U)" "limitin mtopology xn x sequentially" "x \ U" "U \ M" then obtain r where r: "r > 0" "mball x r \ U" by auto with h(2) obtain N where N: "\n. n \ N \ xn n \ M" "\n. n \ N \ d (xn n) x < r" by (metis limit_metric_sequentially) with h have "\N. \n\N. xn n \ mball x r" by(auto intro!: exI[where x=N] simp:commute) with r show "\N. \n\N. xn n \ U" by blast next fix x assume h:"U \ M" "\xn x. limitin mtopology xn x sequentially \ x \ U \ (\N. \n\N. xn n \ U)" "x \ U" show "\r>0. mball x r \ U" proof(rule ccontr) assume "\ (\r>0. mball x r \ U)" then have "\n. \xn\mball x (1 / Suc n). xn \ U" by (meson of_nat_0_less_iff subsetI zero_less_Suc zero_less_divide_1_iff) then obtain xn where xn: "\n. xn n \ mball x (1 / Suc n)" "\n. xn n \ U" by metis have "limitin mtopology xn x sequentially" unfolding limit_metric_sequentially proof safe fix e :: real assume e:"0 < e" then obtain N where N: "1 / real (Suc N) < e" using nat_approx_posE by blast show "\N. \n\N. xn n \ M \ d (xn n) x < e" proof(safe intro!: exI[where x=N]) fix n assume n: "n \ N" then have "1 / Suc n < e" by (metis N Suc_le_mono inverse_of_nat_le nat.distinct(1) order_le_less_trans) thus "d (xn n) x < e" using xn(1)[of n] by(auto simp: commute) qed(use xn in auto) qed(use h in auto) with h(2,3) xn(2) show False by auto qed qed lemma closure_of_mball: "mtopology closure_of mball a e \ mcball a e" by (simp add: closure_of_minimal mball_subset_mcball) lemma interior_of_mcball: "mball a e \ mtopology interior_of mcball a e" by (simp add: interior_of_maximal_eq mball_subset_mcball) lemma isolated_points_of_mtopology: "mtopology isolated_points_of A = {x\M\A. \xn. range xn \ A \ limitin mtopology xn x sequentially \ (\no. \n\no. xn n = x)}" proof safe fix x xn assume h:"x \ mtopology isolated_points_of A" "limitin mtopology xn x sequentially" "range xn \ A" then have ha:"x \ topspace mtopology" "x \ A" "\U. x \ U \ openin mtopology U \ U \ (A - {x}) = {}" by(simp_all add: in_isolated_points_of) then obtain U where u:"x \ U" "openin mtopology U" "U \ (A - {x}) = {}" by auto then obtain e where e: "e > 0" "mball x e \ U" by (meson openin_mtopology) then obtain N where "\n. n \ N \ xn n \ mball x e" using h(2) commute limit_metric_sequentially by fastforce thus "\no. \n\no. xn n = x" using h(3) e(2) u(3) by(fastforce intro!: exI[where x=N]) qed (auto simp: derived_set_of_sequentially isolated_points_of_def, blast) lemma perfect_set_mball_infinite: assumes "perfect_set mtopology A" "a \ A" "e > 0" shows "infinite (mball a e)" proof safe assume h: "finite (mball a e)" have "a \ M" using assms perfect_setD(2)[OF assms(1)] by auto have "\e > 0. mball a e = {a}" proof - consider "mball a e = {a}" | "{a} \ mball a e" using \a \ M\ assms(3) by blast thus ?thesis proof cases case 1 with assms show ?thesis by auto next case 2 then have nen:"{d a b |b. b \ mball a e \ a \ b} \ {}" by auto have fin: "finite {d a b |b. b \ mball a e \ a \ b}" using h by (auto simp del: in_mball) define e' where "e' \ Min {d a b |b. b \ mball a e \ a \ b}" have "e' > 0" using mdist_pos_eq[OF \a \ M\] by(simp add: e'_def Min_gr_iff[OF fin nen] del: in_mball) auto have bd:"\b. b \ mball a e \ a \ b \ e' \ d a b" by(auto simp: e'_def Min_le_iff[OF fin nen] simp del: in_mball) have "e' \ e" unfolding e'_def Min_le_iff[OF fin nen] using nen by auto show ?thesis proof(safe intro!: exI[where x=e']) fix x assume x:"x \ mball a e'" then show "x = a" using \e' \ e\ bd by fastforce qed (use \a \ M\ \e' > 0\ in auto) qed qed then obtain e' where e':"e' > 0" "mball a e' = {a}" by auto show False using perfect_setD(3)[OF assms(1,2) centre_in_mball_iff[of a e',THEN iffD2]] \a \ M\ \e' > 0\ e'(2) by blast qed lemma MCauchy_dist_Cauchy: assumes "MCauchy xn" "MCauchy yn" shows "Cauchy (\n. d (xn n) (yn n))" unfolding metric_space_class.Cauchy_altdef2 dist_real_def proof safe have h:"\n. xn n \ M" "\n. yn n \ M" using assms by(auto simp: MCauchy_def) fix e :: real assume e:"0 < e" with assms obtain N1 N2 where N: "\n m. n \ N1 \ m \ N1 \ d (xn n) (xn m) < e / 2" "\n m. n \ N2 \ m \ N2 \ d (yn n) (yn m) < e / 2" by (metis MCauchy_def zero_less_divide_iff zero_less_numeral) define N where "N \ max N1 N2" then have N': "N \ N1" "N \ N2" by auto show "\N. \n\N. \d (xn n) (yn n) - d (xn N) (yn N)\ < e" proof(safe intro!: exI[where x=N]) fix n assume n:"N \ n" have "d (xn n) (yn n) \ d (xn n) (xn N) + d (xn N) (yn N) + d (yn N) (yn n)" "d (xn N) (yn N) \ d (xn N) (xn n) + d (xn n) (yn n) + d (yn n) (yn N)" using triangle[OF h(1)[of n] h(1)[of N] h(2)[of n]] triangle[OF h(1)[of N] h(2)[of N] h(2)[of n]] triangle[OF h(1)[of N] h(2)[of n] h(2)[of N]] triangle[OF h(1)[of N] h(1)[of n] h(2)[of n]] by auto thus "\d (xn n) (yn n) - d (xn N) (yn N)\ < e" using N(1)[OF N'(1) order.trans[OF N'(1) n]] N(2)[OF N'(2) order.trans[OF N'(2) n]] N(1)[OF order.trans[OF N'(1) n] N'(1)] N(2)[OF order.trans[OF N'(2) n] N'(2)] by auto qed qed subsubsection \ Dense in Metric Spaces\ abbreviation "mdense \ dense_in mtopology" text \\<^url>\https://people.bath.ac.uk/mw2319/ma30252/sec-dense.html\\ lemma mdense_def: "mdense U \ U \ M \ (\x\M. \\>0. mball x \ \ U \ {})" proof safe assume h:" U \ M" "(\x\M. \\>0. mball x \ \ U \ {})" show "dense_in mtopology U" proof(rule dense_inI) fix V assume h':"openin mtopology V" "V \ {}" then obtain x where 1:"x \ V" by auto then obtain \ where 2:"\>0" "mball x \ \ V" by (meson h'(1) openin_mtopology) have "mball x \ \ U \ {}" using h 1 2 openin_subset[OF h'(1)] by (auto simp del: in_mball) thus "U \ V \ {}" using 2 by auto qed(use h in auto) next fix x \ assume h:"x \ M" "(0 :: real) < \" "mball x \ \ U = {}" "mdense U" then have "mball x \ \ {}" "openin mtopology (mball x \)" by auto with h(4) have "mball x \ \ U \ {}" by(auto simp: dense_in_def) with h(3) show False by simp qed(auto simp: dense_in_def) corollary mdense_balls_cover: assumes "mdense U" and "e > 0" shows "(\u\U. mball u e) = M" using assms[simplified mdense_def] commute by fastforce lemma mdense_empty_iff: "mdense {} \ M = {}" by(auto simp: mdense_def) (use zero_less_one in blast) lemma mdense_M: "mdense M" by(auto simp: mdense_def) lemma mdense_def2: "mdense U \ U \ M \ (\x\M. \\>0.\y\U. d x y < \)" proof safe fix x e assume h: "mdense U" and hxe: "x \ M" "(0 :: real) < e" then have "x \ (\u\U. mball u e)" by(simp add: mdense_balls_cover) thus "\y\U. d x y < e" by (fastforce simp: commute) qed(fastforce simp: mdense_def)+ lemma mdense_def3: "mdense U \ U \ M \ (\x\M. \xn. range xn \ U \ limitin mtopology xn x sequentially)" unfolding mdense_def proof safe fix x assume h: "U \ M" "\x\M. \\>0. mball x \ \ U \ {}" "x \ M" then have "\n. mball x (1 / (real n + 1)) \ U \ {}" by auto hence "\n. \k. k \ mball x (1 / (real n + 1)) \ U" by (auto simp del: in_mball) hence "\a. \n. a n \ mball x (1 / (real n + 1)) \ U" by(rule choice) then obtain xn where xn: "\n. xn n \ mball x (1 / (real n + 1)) \ U" by auto show "\xn. range xn \ U \ limitin mtopology xn x sequentially" unfolding limitin_metric eventually_sequentially proof(safe intro!: exI[where x=xn]) fix \ :: real assume he:"0 < \" then obtain N where hn: "1 / \ < real N" using reals_Archimedean2 by blast have hn': "0 < real N" by(rule ccontr) (use hn he in fastforce) hence "1 / real N < \" using he hn by (metis divide_less_eq mult.commute) hence hn'':"1 / (real N + 1) < \" using hn' by(auto intro!: order.strict_trans[OF linordered_field_class.divide_strict_left_mono[of "real N" "real N + 1" 1]]) show "\N. \n\N. xn n \ M \ d (xn n) x < \" proof(safe intro!: exI[where x="N"]) fix n assume "N \ n" then have 1:"1 / (real n + 1) \ 1 / (real N + 1)" using hn' by(auto intro!: linordered_field_class.divide_left_mono) show "d (xn n) x < \" using xn[of n] order.strict_trans1[OF 1 hn''] by (auto simp: commute) qed(use xn in auto) qed(use xn h in auto) next fix x and e :: real assume h: "U \ M" " \x\M. \xn. range xn \ U \ limitin mtopology xn x sequentially" "x \ M" "0 < e" "mball x e \ U = {}" then obtain xn where xn:"range xn \ U" "limitin mtopology xn x sequentially" by auto with h(4) obtain N where N: "\n. n \ N \ d (xn n) x < e" by (meson limit_metric_sequentially) have "xn N \ mball x e \ U" using N[of N] xn(1) h(1,3) by (auto simp: commute) with h(5) show False by simp qed text \ Diameter\ definition mdiameter :: "'a set \ ennreal" where "mdiameter A \ \ {ennreal (d x y) | x y. x \ A \ M \ y \ A \ M}" lemma mdiameter_empty[simp]: "mdiameter {} = 0" by(simp add: mdiameter_def bot_ennreal) lemma mdiameter_def2: assumes "A \ M" shows "mdiameter A = \ {ennreal (d x y) | x y. x \ A \ y \ A}" using assms by(auto simp: mdiameter_def) (meson subset_eq) lemma mdiameter_subset: assumes "A \ B" shows "mdiameter A \ mdiameter B" unfolding mdiameter_def using assms by(auto intro!: Sup_subset_mono) lemma mdiameter_cball_leq: "mdiameter (mcball a \) \ ennreal (2 * \)" unfolding Sup_le_iff mdiameter_def proof safe fix x y assume h:"x \ mcball a \" "y \ mcball a \" "x \ M" "y \ M" have "d x y \ 2 * \" using h(1) h(2) triangle'' by fastforce thus "ennreal (d x y) \ ennreal (2 * \)" using ennreal_leI by blast qed lemma mdiameter_ball_leq: "mdiameter (mball a \) \ ennreal (2 * \)" using mdiameter_subset[OF mball_subset_mcball[of a \]] mdiameter_cball_leq[of a \] by auto lemma mdiameter_is_sup: assumes "x \ A \ M" "y \ A \ M" shows "d x y \ mdiameter A" using assms by(auto simp: mdiameter_def intro!: Sup_upper) lemma mdiameter_is_sup': assumes "x \ A \ M" "y \ A \ M" "mdiameter A \ ennreal r" "r \ 0" shows "d x y \ r" using order.trans[OF mdiameter_is_sup[OF assms(1,2)] assms(3)] assms(4) by simp lemma mdiameter_le: assumes "\x y. x \ A \ y \ A \ d x y \ r" shows "mdiameter A \ r" using assms by(auto simp: mdiameter_def Sup_le_iff ennreal_leI) lemma mdiameter_eq_closure: "mdiameter (mtopology closure_of A) = mdiameter A" proof(rule antisym) show "mdiameter A \ mdiameter (mtopology closure_of A)" by(fastforce intro!: Sup_subset_mono simp: mdiameter_def metric_closure_of) next have "{ennreal (d x y) |x y. x \ A \ M \ y \ A \ M} = ennreal ` {d x y |x y. x \ A \ M \ y \ A \ M}" by auto also have "mdiameter (mtopology closure_of A) \ \ ..." unfolding le_Sup_iff_less proof safe fix r assume "r < mdiameter (mtopology closure_of A)" then obtain x y where xy:"x \ mtopology closure_of A" "x \ M" "y \ mtopology closure_of A" "y \ M" "r < ennreal (d x y)" by(auto simp: mdiameter_def less_Sup_iff) hence "r < \" using dual_order.strict_trans ennreal_less_top by blast define e where "e \ (d x y - enn2real r)/2" have "e > 0" using xy(5) \r < \\ by(simp add: e_def) then obtain x' y' where xy':"x' \ mball x e" "x'\ A" "y' \ mball y e" "y'\ A" using xy by(fastforce simp: metric_closure_of) show "\i\{d x y |x y. x \ A \ M \ y \ A \ M}. r \ ennreal i" proof(safe intro!: bexI[where x="d x' y'"]) have "d x y \ d x x' + d x' y' + d y y'" using triangle[OF xy(2) _ xy(4),of x'] xy' triangle[of x' y' y] by(fastforce simp add: commute) also have "... < d x y - enn2real r + d x' y'" using xy'(1) xy'(3) by(simp add: e_def) finally have "enn2real r < d x' y'" by simp thus "r \ ennreal (d x' y')" by (simp add: \r < \\) qed(use xy'(1) xy'(3) xy'(2,4) in auto) qed finally show "mdiameter (mtopology closure_of A) \ mdiameter A" by(simp add: mdiameter_def) qed lemma mbounded_finite_mdiameter: "mbounded A \ A \ M \ mdiameter A < \" proof safe assume "mbounded A" then obtain x B where "A \ mcball x B" by(auto simp: mbounded_def) then have "mdiameter A \ mdiameter (mcball x B)" by(rule mdiameter_subset) also have "... \ ennreal (2 * B)" by(rule mdiameter_cball_leq) also have "... < \" by auto finally show "mdiameter A < \" . next assume h:"mdiameter A < \" "A \ M" consider "A = {}" | "A \ {}" by auto then show "mbounded A" proof cases case h2:2 then have 1:"{d x y |x y. x \ A \ y \ A} \ {}" by auto have eq:"{ennreal (d x y) | x y. x \ A \ y \ A} = ennreal ` {d x y | x y. x \ A \ y \ A}" by auto hence 2:"mdiameter A = \ (ennreal ` {d x y | x y. x \ A \ y \ A})" using h by(auto simp add: mdiameter_def2) obtain x y where hxy: "x \ A" "y \ A" "mdiameter A < ennreal (d x y + 1)" using SUP_approx_ennreal[OF _ 1 2,of 1] h by(fastforce simp: diameter_def) show ?thesis unfolding mbounded_alt proof(safe intro!: exI[where x="d x y + 1"]) fix w z assume "w \ A" "z \ A " with SUP_lessD[OF hxy(3)[simplified 2]] have "ennreal (d w z) < ennreal (d x y + 1)" by blast thus "d w z \ d x y + 1" by (metis canonically_ordered_monoid_add_class.lessE ennreal_le_iff2 ennreal_neg le_iff_add not_less_zero) qed (use h in auto) qed(auto simp: mbounded_def) qed(auto simp: mbounded_def) text \ Distance between a point and a set. \ definition d_set :: "'a set \ 'a \ real" where "d_set A \ (\x. if A \ {} \ A \ M \ x \ M then Inf {d x y |y. y \ A} else 0)" lemma d_set_nonneg[simp]: "d_set A x \ 0" proof - have "{d x y |y. y \ A} = d x ` A" by auto thus ?thesis by(auto simp: d_set_def intro!: cINF_greatest[of _ _ "d x"]) qed lemma d_set_bdd_below[simp]: "bdd_below {d x y |y. y \ A}" by(auto simp: bdd_below_def intro!: exI[where x=0]) lemma d_set_singleton[simp]: "x \ M \ y \ M \ d_set {y} x = d x y" by(auto simp: d_set_def) lemma d_set_empty[simp]: "d_set {} x = 0" by(simp add: d_set_def) lemma d_set_notin: "x \ M \ d_set A x = 0" by(auto simp: d_set_def) lemma d_set_inA: assumes "x \ A" shows "d_set A x = 0" proof - { assume "x \ M" "A \ M" then have "0 \ {d x y |y. y \ A}" using assms by(auto intro: exI[where x=x]) moreover have "A \ {}" using assms by auto ultimately have "Inf {d x y |y. y \ A} = 0" using cInf_lower[OF \0 \ {d x y |y. y \ A}\] d_set_nonneg[of A x] \A \ M\ \x \ M\ by(auto simp: d_set_def) } thus ?thesis using assms by(auto simp: d_set_def) qed lemma d_set_nzeroD: assumes "d_set A x \ 0" shows "A \ M" "x \ A" "A \ {}" by(rule ccontr, use assms d_set_inA d_set_def in auto) lemma d_set_antimono: assumes "A \ B" "A \ {}" "B \ M" shows "d_set B x \ d_set A x" proof(cases "B = {}") case h:False with assms have "{d x y |y. y \ B} \ {}" "{d x y |y. y \ A} \ {d x y |y. y \ B}" "A \ M" by auto with assms(3) show ?thesis by(simp add: d_set_def cInf_superset_mono assms(2)) qed(use assms in simp) lemma d_set_bounded: assumes "\y. y \ A \ d x y < K" "K > 0" shows "d_set A x < K" proof - consider "A = {}" | "\ A \ M" | "x \ M" | "A \ {}" "A \ M" "x \ M" by blast then show ?thesis proof cases case 4 then have 2:"{d x y |y. y \ A} \ {}" by auto show ?thesis using assms by (auto simp add: d_set_def cInf_lessD[OF 2] cInf_less_iff[OF 2]) qed(auto simp: d_set_def assms) qed lemma d_set_tr: assumes "x \ M" "y \ M" shows "d_set A x \ d x y + d_set A y" proof - consider "A = {}" | "\ A \ M" | "A \ {}" "A \ M" by blast then show ?thesis proof cases case 3 have "d_set A x \ Inf {d x y + d y a |a. a\A}" proof - have "\ {d x y |y. y \ A} \ \ {d x y + d y a |a. a \ A}" by(rule cInf_mono) (use 3 assms triangle in fastforce)+ thus ?thesis by(simp add: d_set_def assms 3) qed also have "... = (\a\A. d x y + d y a)" by (simp add: Setcompr_eq_image) also have "... = d x y + \ (d y ` A)" using 3 by(auto intro!: Inf_add_eq bdd_belowI[where m=0]) also have "... = d x y + d_set A y" using assms 3 by(auto simp: d_set_def Setcompr_eq_image) finally show ?thesis . qed(auto simp: d_set_def) qed lemma d_set_abs_le: assumes "x \ M" "y \ M" shows "\d_set A x - d_set A y\ \ d x y" using d_set_tr[OF assms,of A] d_set_tr[OF assms(2,1),of A,simplified commute[of y x]] by auto lemma d_set_inA_le: assumes "y \ A" shows "d_set A x \ d x y" proof - consider "A \ {}" "A \ M" "x \ M" | "\ A \ M" | "x \ M" using assms by blast then show ?thesis proof cases case 1 with assms have "y \ M" by auto from d_set_tr[OF 1(3) this,of A] show ?thesis by(simp add: d_set_inA[OF assms]) qed(auto simp: d_set_def) qed lemma d_set_ball_empty: assumes "A \ {}" "A \ M" "e > 0" "x \ M" "mball x e \ A = {}" shows "d_set A x \ e" using assms by(fastforce simp: d_set_def assms(1) le_cInf_iff) lemma d_set_closed_pos: assumes "closedin mtopology A" "A \ {}" "x \ M" "x \ A" shows "d_set A x > 0" proof - have a:"A \ M" "openin mtopology (M - A)" using closedin_subset[OF assms(1)] assms(1) by auto with assms(3,4) obtain e where e: "e > 0" "mball x e \ M - A" using openin_mtopology by auto thus ?thesis by(auto intro!: order.strict_trans2[OF e(1) d_set_ball_empty[OF assms(2) a(1) e(1) assms(3)]]) qed lemma gdelta_in_closed: assumes "closedin mtopology M" shows "gdelta_in mtopology M" by(auto intro!: closed_imp_gdelta_in metrizable_space_mtopology) text \ Oscillation \ definition osc_on :: "['b set, 'b topology, 'b \ 'a, 'b] \ ennreal" where "osc_on A X f \ (\y. \ {mdiameter (f ` (A \ U)) |U. y \ U \ openin X U})" abbreviation "osc X \ osc_on (topspace X) X" lemma osc_def: "osc X f = (\y. \ {mdiameter (f ` U) |U. y \ U \ openin X U})" by(standard,auto simp: osc_on_def) (metis (no_types, opaque_lifting) inf.absorb2 openin_subset) lemma osc_on_less_iff: "osc_on A X f x < t \ (\v. x \ v \ openin X v \ mdiameter (f ` (A \ v)) < t)" by(auto simp add: osc_on_def Inf_less_iff) lemma osc_less_iff: "osc X f x < t \ (\v. x \ v \ openin X v \ mdiameter (f ` v) < t)" by(auto simp add: osc_def Inf_less_iff) end definition mdist_set :: "'a metric \ 'a set \ 'a \ real" where "mdist_set m \ Metric_space.d_set (mspace m) (mdist m)" lemma(in Metric_space) mdist_set_Self: "mdist_set Self = d_set" by(simp add: mdist_set_def) lemma mdist_set_nonneg[simp]: "mdist_set m A x \ 0" by(auto simp: mdist_set_def Metric_space.d_set_nonneg) lemma mdist_set_singleton[simp]: "x \ mspace m \ y \ mspace m \ mdist_set m {y} x = mdist m x y" by(auto simp: mdist_set_def Metric_space.d_set_singleton) lemma mdist_set_empty[simp]: "mdist_set m {} x = 0" by(auto simp: mdist_set_def Metric_space.d_set_empty) lemma mdist_set_inA: assumes "x \ A" shows "mdist_set m A x = 0" by(auto simp: assms mdist_set_def Metric_space.d_set_inA) lemma mdist_set_nzeroD: assumes "mdist_set m A x \ 0" shows "A \ mspace m" "x \ A" "A \ {}" using assms Metric_space.d_set_nzeroD[of "mspace m" "mdist m"] by(auto simp: mdist_set_def) lemma mdist_set_antimono: assumes "A \ B" "A \ {}" "B \ mspace m" shows "mdist_set m B x \ mdist_set m A x" by(auto simp: assms mdist_set_def Metric_space.d_set_antimono) lemma mdist_set_bounded: assumes "\y. y \ A \ mdist m x y < K" "K > 0" shows "mdist_set m A x < K" by(auto simp: assms mdist_set_def Metric_space.d_set_bounded) lemma mdist_set_tr: assumes "x \ mspace m" "y \ mspace m" shows "mdist_set m A x \ mdist m x y + mdist_set m A y" by(auto simp: assms mdist_set_def Metric_space.d_set_tr) lemma mdist_set_abs_le: assumes "x \ mspace m" "y \ mspace m" shows "\mdist_set m A x - mdist_set m A y\ \ mdist m x y" by(auto simp: assms mdist_set_def Metric_space.d_set_abs_le) lemma mdist_set_inA_le: assumes "y \ A" shows "mdist_set m A x \ mdist m x y" by(auto simp: assms mdist_set_def Metric_space.d_set_inA_le) lemma mdist_set_ball_empty: assumes "A \ {}" "A \ mspace m" "e > 0" "x \ mspace m" "mball_of m x e \ A = {}" shows "mdist_set m A x \ e" by (metis Metric_space.d_set_ball_empty Metric_space_mspace_mdist assms mball_of_def mdist_set_def) lemma mdist_set_closed_pos: assumes "closedin (mtopology_of m) A" "A \ {}" "x \ mspace m" "x \ A" shows "mdist_set m A x > 0" by (metis Metric_space.d_set_closed_pos Metric_space_mspace_mdist assms mdist_set_def mtopology_of_def) -lemma continuous_map_metric_limitin: - "continuous_map (mtopology_of X) (mtopology_of Y) f \ f \ mspace X \ mspace Y \ (\xn x. limitin (mtopology_of X) xn x sequentially \ limitin (mtopology_of Y) (\n. f (xn n)) (f x) sequentially)" -proof - - interpret m1: Metric_space "mspace X" "mdist X" - by simp - interpret m2: Metric_space "mspace Y" "mdist Y" - by simp - have "continuous_map m1.mtopology m2.mtopology f \ f \ mspace X \ mspace Y \ (\xn x. limitin (mtopology_of X) xn x sequentially \ limitin (mtopology_of Y) (\n. f (xn n)) (f x) sequentially)" - unfolding mtopology_of_def - proof safe - fix xn x - assume "continuous_map m1.mtopology m2.mtopology f" - then have h: "f ` mspace X \ mspace Y" "\a\mspace X. \\>0. \\>0. \x. x \ mspace X \ mdist X a x < \ \ mdist Y (f a) (f x) < \" - by(auto simp: Metric_space.metric_continuous_map[OF m1.Metric_space_axioms m2.Metric_space_axioms]) - assume h': "limitin m1.mtopology xn x sequentially" - then have x:"x \ mspace X" - by (simp add: m1.limitin_mspace) - show "limitin m2.mtopology (\n. f (xn n)) (f x) sequentially" - unfolding m2.limit_metric_sequentially - proof safe - fix e :: real - assume e:"0 < e" - with h(2) x obtain \ where \: "\ > 0" "\y. y \ mspace X \ mdist X x y < \ \ mdist Y (f x) (f y) < e" - by fastforce - then obtain N where N: "\n. n \ N \ xn n \ mspace X" "\n. n \ N \ mdist X (xn n) x < \" - using h' by(fastforce simp add: m1.limit_metric_sequentially) - show "\N. \n\N. f (xn n) \ mspace Y \ mdist Y (f (xn n)) (f x) < e" - using N h(1) by(auto intro!: exI[where x=N] simp: mdist_commute \(2)) - qed(use x h in auto) - next - fix a and e :: real - assume h:"f \ mspace X \ mspace Y" " \xn x. limitin m1.mtopology xn x sequentially \ limitin m2.mtopology (\n. f (xn n)) (f x) sequentially" - show "continuous_map m1.mtopology m2.mtopology f" - unfolding continuous_map_def - proof safe - fix U - assume U:"openin m2.mtopology U" - show "openin m1.mtopology {x \ topspace m1.mtopology. f x \ U}" - unfolding m1.openin_mtopology2 m1.topspace_mtopology - proof safe - fix xn x - assume xnx:"limitin m1.mtopology xn x sequentially" "x \ mspace X" "f x \ U" - then have "limitin m2.mtopology (\n. f (xn n)) (f x) sequentially" - using h by auto - with U xnx obtain N where N:"\n. n\N \ f (xn n) \ U" - by(fastforce simp: m2.openin_mtopology2) - from xnx obtain N' where N': "\n. n \ N' \ xn n \ mspace X" - by (meson lessI m1.limit_metric_sequentially of_nat_0_less_iff) - show "\N. \n\N. xn n \ {x \ mspace X. f x \ U}" - by(auto intro!: exI[where x="max N N'"] N N') - qed - qed(use h in auto) - qed(auto simp: continuous_map_def) - thus ?thesis - by (simp add: mtopology_of_def) -qed - lemma mdist_set_uniformly_continuous: "uniformly_continuous_map m euclidean_metric (mdist_set m A)" unfolding uniformly_continuous_map_def proof safe fix e :: real assume "0 < e" then show "\\>0. \x\mspace m. \y\mspace m. mdist m y x < \ \ mdist euclidean_metric (mdist_set m A y) (mdist_set m A x) < e" using order.strict_trans1[OF mdist_set_abs_le] by(auto intro!: exI[where x=e] simp: dist_real_def) qed simp lemma uniformly_continuous_map_add: fixes f :: "'a \ 'b::real_normed_vector" assumes "uniformly_continuous_map m euclidean_metric f" "uniformly_continuous_map m euclidean_metric g" shows "uniformly_continuous_map m euclidean_metric (\x. f x + g x)" unfolding uniformly_continuous_map_def proof safe fix e :: real assume "e > 0" from half_gt_zero[OF this] assms obtain d1 d2 where d: "d1 > 0" "d2 > 0" "\x y. x \ mspace m \ y \ mspace m \ mdist m x y < d1 \ dist (f x) (f y) < e / 2" "\x y. x \ mspace m \ y \ mspace m \ mdist m x y < d2 \ dist (g x) (g y) < e / 2" by(simp only: uniformly_continuous_map_def mdist_euclidean_metric) metis show "\\>0. \y\mspace m. \x\mspace m. mdist m x y < \ \ mdist euclidean_metric (f x + g x) (f y + g y) < e" using d by(fastforce intro!: exI[where x="min d1 d2"] order.strict_trans1[OF dist_triangle_add]) qed simp lemma uniformly_continuous_map_real_divide: fixes f :: "'a \ real" assumes "uniformly_continuous_map m euclidean_metric f" "uniformly_continuous_map m euclidean_metric g" and "\x. x \ mspace m \ g x \ 0" "\x. x \ mspace m \ \g x\ \ a" "a > 0" "\x. x \ mspace m \ \g x\ < Kg" and "\x. x \ mspace m \ \f x\ < Kf" shows "uniformly_continuous_map m euclidean_metric (\x. f x / g x)" unfolding uniformly_continuous_map_def proof safe fix e :: real assume e[arith]:"e > 0" consider "mspace m \ {}" | "mspace m = {}" by auto then show "\\>0. \x\mspace m. \y\mspace m. mdist m y x < \ \ mdist euclidean_metric (f y / g y) (f x / g x) < e" proof cases case m:1 then have Kfg_pos[arith]: "Kg > 0" "Kf \ 0" using assms(4-7) by auto fastforce+ define e' where "e' \ a^2 / (Kf + Kg) * e" have e':"e' > 0" using assms(5) by(auto simp: e'_def) with assms obtain d1 d2 where d: "d1 > 0" "d2 > 0" "\x y. x \ mspace m \ y \ mspace m \ mdist m x y < d1 \ \f x - f y\ < e'" "\x y. x \ mspace m \ y \ mspace m \ mdist m x y < d2 \ \g x - g y\ < e'" by(simp only: uniformly_continuous_map_def mdist_euclidean_metric dist_real_def) metis show ?thesis unfolding mdist_euclidean_metric dist_real_def proof(safe intro!: exI[where x="min d1 d2"]) fix x y assume x:"x \ mspace m" and y:"y \ mspace m" and "mdist m x y < min d1 d2" then have dist[arith]: "mdist m x y < d1" "mdist m x y < d2" by auto note [arith] = assms(3,4,6,7)[OF x] assms(3,4,6,7)[OF y] have "\f x / g x - f y / g y\ = \(f x * g y - f y * g x) / (g x * g y)\" by(simp add: diff_frac_eq) also have "... = \(f x * g y - f x * g x + (f x * g x - f y * g x)) / (g x * g y)\" by simp also have "... = \(f x - f y) * g x - f x * (g x - g y)\ / (\g x\ * \g y\)" by(simp add: left_diff_distrib right_diff_distrib abs_mult) also have "... \ (\f x - f y\ * \g x\ + \f x\ * \g x - g y\) / (\g x\ * \g y\)" by(rule divide_right_mono) (use abs_triangle_ineq4 abs_mult in metis,auto) also have "... < (e' * Kg + Kf * e') / (\g x\ * \g y\)" by(rule divide_strict_right_mono[OF add_less_le_mono]) (auto intro!: mult_mono' mult_strict_mono, use d(3,4)[OF x y] in auto) also have "... \ (e' * Kg + Kf * e') / a^2" by(auto intro!: divide_left_mono simp: power2_eq_square) (insert assms(5) e', auto simp: \a \ \g x\\ mult_mono') also have "... = (Kf + Kg) / a^2 * e'" by (simp add: distrib_left mult.commute) also have "... = e" using assms(5) by(auto simp: e'_def) finally show " \f x / g x - f y / g y\ < e" . qed(use d in auto) qed (auto intro!: exI[where x=1]) qed simp lemma assumes "e > 0" shows uniformly_continuous_map_from_capped_metric:"uniformly_continuous_map (capped_metric e m1) m2 f \ uniformly_continuous_map m1 m2 f" (is ?g1) and uniformly_continuous_map_to_capped_metric:"uniformly_continuous_map m1 (capped_metric e m2) f \ uniformly_continuous_map m1 m2 f" (is ?g2) proof - have [simp]:"(\n. min e (X n)) \ 0 \ X \ 0" for X proof assume h:"(\n. min e (X n)) \ 0" show "X \ 0" unfolding LIMSEQ_def dist_real_def proof safe fix r :: real assume "0 < r" then have "min (e / 2) r > 0" using assms by auto from LIMSEQ_D[OF h this] obtain N where N:"\n. n \ N \ \min e (X n)\ < min (e / 2) r" by auto have "min e (X n) = X n" if h:"n \ N " for n proof(rule ccontr) assume "min e (X n) \ X n" then have "min e (X n) = e" by auto with N[OF h] show False by auto qed with N show "\no. \n\no. \X n - 0\ < r" by(auto intro!: exI[where x=N]) qed qed(auto intro!: tendsto_eq_intros(4)[of "\x. e" e sequentially X _ ] simp: assms) show ?g1 ?g2 using assms by(auto simp: uniformly_continuous_map_sequentially capped_metric_mdist) qed lemma Urysohn_lemma_uniform: assumes "closedin (mtopology_of m) T" "closedin (mtopology_of m) U" "T \ U = {}" "\x y. x \ T \ y \ U \ mdist m x y \ e" "e > 0" obtains f :: "'a \ real" where "uniformly_continuous_map m euclidean_metric f" "\x. f x \ 0" "\x. f x \ 1" "\x. x \ T \ f x = 1" "\x. x \ U \ f x = 0" proof - consider "T = {}" | "U = {}" | "T \ {}" "U \ {}" by auto then show ?thesis proof cases case 1 define f where "f \ (\x::'a. 0::real)" with 1 have "uniformly_continuous_map m euclidean_metric f" "\x. f x \ 0" "\x. f x \ 1" "\x. x \ T \ f x = 1" "\x. x \ U \ f x = 0" by auto then show ?thesis using that by auto next case 2 define f where "f \ (\x::'a. 1::real)" with 2 have "uniformly_continuous_map m euclidean_metric f" "\x. f x \ 0" "\x. f x \ 1" "\x. x \ T \ f x = 1" "\x. x \ U \ f x = 0" by auto then show ?thesis using that by auto next case TU:3 then have STU:"mspace m \ {}" "T \ mspace m" "U \ mspace m" using assms(1,2) closedin_topspace_empty closedin_subset by fastforce+ interpret m: Metric_space "mspace m" "mdist (capped_metric e m)" by (metis Metric_space_mspace_mdist capped_metric_mspace) have e:"x \ T \ y \ U \ mdist (capped_metric e m) x y \ e" for x y by (simp add: assms(4) capped_metric_mdist) define f where "f \ (\x. mdist_set (capped_metric e m) U x / (mdist_set (capped_metric e m) U x + mdist_set (capped_metric e m) T x))" have "uniformly_continuous_map m euclidean_metric f" unfolding uniformly_continuous_map_from_capped_metric[OF assms(5),of m,symmetric] f_def proof(rule uniformly_continuous_map_real_divide[where Kf="e + 1" and Kg="2 * e + 1" and a="e / 2"]) show "uniformly_continuous_map (capped_metric e m) euclidean_metric (mdist_set (capped_metric e m) U)" "uniformly_continuous_map (capped_metric e m) euclidean_metric (\x. mdist_set (capped_metric e m) U x + mdist_set (capped_metric e m) T x)" by(auto intro!: mdist_set_uniformly_continuous uniformly_continuous_map_add) next fix x assume x:"x \ mspace (capped_metric e m)" then consider "x \ T" | "x \ U" using TU assms by auto thus "mdist_set (capped_metric e m) U x + mdist_set (capped_metric e m) T x \ 0" proof cases case 1 with TU x assms have "mdist_set (capped_metric e m) T x > 0" by(auto intro!: mdist_set_closed_pos simp: mtopology_capped_metric) thus?thesis by (metis add_less_same_cancel2 linorder_not_less mdist_set_nonneg) next case 2 with TU x assms have "mdist_set (capped_metric e m) U x > 0" by(auto intro!: mdist_set_closed_pos simp: mtopology_capped_metric) thus?thesis by (metis add_less_same_cancel1 linorder_not_less mdist_set_nonneg) qed next fix x assume x:"x \ mspace (capped_metric e m)" consider "x \ (\a\U. m.mball a (e / 2))" | "x \ (\a\U. m.mball a (e / 2))" by auto then show "e / 2 \ \mdist_set (capped_metric e m) U x + mdist_set (capped_metric e m) T x\" proof cases case 1 have "m.mball x (e / 2) \ T = {}" proof(rule ccontr) assume "m.mball x (e / 2) \ T \ {}" then obtain y where y: "y \ m.mball x (e / 2)" "y \ T" by blast then obtain u where u:"u \ U" "x \ m.mball u (e / 2)" using 1 by auto have "mdist (capped_metric e m) y u \ mdist (capped_metric e m) y x + mdist (capped_metric e m) x u" using STU u y x by(auto intro!: m.triangle) also have "... < e / 2 + e / 2" using y u by(auto simp: m.commute) also have "... = e" using assms(5) by linarith finally show False using e[OF y(2) u(1)] by simp qed from m.d_set_ball_empty[OF _ _ _ _ this] TU STU x assms(1,5) have "e / 2 \ m.d_set T x" using STU(2) x by auto also have "... \ \mdist_set (capped_metric e m) U x + mdist_set (capped_metric e m) T x\" by (simp add: mdist_set_def) finally show ?thesis . next case 2 then have "m.mball x (e / 2) \ U = {}" using x by (auto simp: m.commute) hence "e / 2 \ m.d_set U x" by (metis STU(3) TU(2) assms(5) capped_metric_mspace half_gt_zero m.d_set_ball_empty x) also have "... \ \mdist_set (capped_metric e m) U x + mdist_set (capped_metric e m) T x\" by (simp add: mdist_set_def) finally show ?thesis . qed next fix x assume "x \ mspace (capped_metric e m)" have "\mdist_set (capped_metric e m) U x + mdist_set (capped_metric e m) T x\ = mdist_set (capped_metric e m) U x + mdist_set (capped_metric e m) T x" using mdist_set_nonneg by auto also have "... < (e + 1 / 2) + (e + 1 / 2)" apply(intro add_strict_mono mdist_set_bounded) using assms(5) add_strict_increasing[of "1 / 2",OF _ mdist_capped[OF assms(5),of m x]] by (auto simp add: add.commute) also have "... = 2 * e + 1" by auto finally show "\mdist_set (capped_metric e m) U x + mdist_set (capped_metric e m) T x\ < 2 * e + 1" . show " \mdist_set (capped_metric e m) U x\ < e + 1" using assms(5) add_strict_increasing[of 1,OF _ mdist_capped[OF assms(5),of m x]] by (auto simp add: add.commute intro!: mdist_set_bounded) qed(use assms in auto) moreover have "\x. f x \{0..1}" proof - fix x have "f x \ mdist_set (capped_metric e m) U x / mdist_set (capped_metric e m) U x" proof - consider "mdist_set (capped_metric e m) U x = 0" | "mdist_set (capped_metric e m) U x > 0" using antisym_conv1 by fastforce thus ?thesis proof cases case 2 show ?thesis unfolding f_def by(rule divide_left_mono[OF _ _ mult_pos_pos]) (auto simp: 2 add_strict_increasing) qed (simp add: f_def) qed also have "... \ 1" by simp finally show "f x \ {0..1}" by (auto simp: f_def) qed moreover have "f x = 1" if x:"x \ T" for x proof - { assume h:"mdist_set (capped_metric e m) U x = 0" then have "x \ U" using assms STU x by blast hence False by (metis STU(2) TU(2) assms(2) capped_metric_mspace h mdist_set_closed_pos mtopology_capped_metric order_less_irrefl subset_eq x) } thus ?thesis by (metis add.right_neutral divide_self_if f_def mdist_set_nzeroD(2) x) qed moreover have "\x. x \ U \ f x = 0" by (simp add: f_def m.d_set_inA mdist_set_def) ultimately show ?thesis using that by auto qed qed -lemma upper_semicontinuous_map_limsup_iff: - fixes f :: "'a \ ('b :: {complete_linorder,linorder_topology})" - shows "upper_semicontinuous_map (mtopology_of m) f \ (\xn x. limitin (mtopology_of m) xn x sequentially \ limsup (\n. f (xn n)) \ f x)" - unfolding upper_semicontinuous_map_def -proof safe - fix xn x - assume h:"\a. openin (mtopology_of m) {x \ topspace (mtopology_of m). f x < a}" "limitin (mtopology_of m) xn x sequentially" - show "limsup (\n. f (xn n)) \ f x" - unfolding Limsup_le_iff eventually_sequentially - proof safe - fix y - assume "f x < y" - show "\N. \n\N. f (xn n) < y" - proof(rule ccontr) - assume "\N. \n\N. f (xn n) < y" - then have hc:"\N. \n\N. f (xn n) \ y" - using linorder_not_less by blast - define a :: "nat \ nat" where "a \ rec_nat (SOME n. f (xn n) \ y) (\n an. SOME m. m > an \ f (xn m) \ y)" - have "strict_mono a" - proof(rule strict_monoI_Suc) - fix n - have [simp]:"a (Suc n) = (SOME m. m > a n \ f (xn m) \ y)" - by(auto simp: a_def) - show "a n < a (Suc n)" - by simp (metis (mono_tags, lifting) Suc_le_lessD hc someI) - qed - have *:"f (xn (a n)) \ y" for n - proof(cases n) - case 0 - then show ?thesis - using hc[of 0] by(auto simp: a_def intro!: someI_ex) - next - case (Suc n') - then show ?thesis - by(simp add: a_def) (metis (mono_tags, lifting) Suc_le_lessD hc someI_ex) - qed - interpret Metric_space "mspace m" "mdist m" - by simp - have [simp]: "mtopology = mtopology_of m" - by (simp add: mtopology_of_def) - show False - by (metis (mono_tags, lifting) limitin_subsequence[OF \strict_mono a\ h(2)] openin_mtopology2[THEN iffD1,simplified,OF h(1)[rule_format,of y]] \f x < y\ h(2) hc limitin_def linorder_not_less mem_Collect_eq) - qed - qed -next - fix a - assume h:" \xn x. limitin (mtopology_of m) xn x sequentially \ limsup (\n. f (xn n)) \ f x" - interpret Metric_space "mspace m" "mdist m" - by simp - have [simp]: "mtopology = mtopology_of m" - by (simp add: mtopology_of_def) - have "openin mtopology {x \ topspace (mtopology_of m). f x < a}" - unfolding openin_mtopology2 topspace_mtopology_of - proof safe - fix x xn - assume h':"limitin mtopology xn x sequentially" "x \ mspace m" "f x < a" - with h have "limsup (\n. f (xn n)) \ f x" - by auto - with h'(3) obtain N where N:"\n. n\N \ f (xn n) < a" - by(auto simp: Limsup_le_iff eventually_sequentially) - obtain N' where N': "\n. n\N' \ xn n \ mspace m" - by (meson h' limitin_sequentially open_in_mspace) - thus "\N. \n\N. xn n \ {x \ mspace m. f x < a}" - using h'(3) N by(auto intro!: exI[where x="max N N'"]) - qed - thus "openin (mtopology_of m) {x \ topspace (mtopology_of m). f x < a}" - by simp -qed - -lemma upper_semicontinuous_map_limsup_real: - fixes f :: "'a \ real" - shows "upper_semicontinuous_map (mtopology_of m) f \ (\xn x. limitin (mtopology_of m) xn x sequentially \ limsup (\n. f (xn n)) \ f x)" - unfolding upper_semicontinuous_map_real_iff upper_semicontinuous_map_limsup_iff by simp - -lemma lower_semicontinuous_map_liminf_iff: - fixes f :: "'a \ ('b :: {complete_linorder,linorder_topology})" - shows "lower_semicontinuous_map (mtopology_of m) f \ (\xn x. limitin (mtopology_of m) xn x sequentially \ f x \ liminf (\n. f (xn n)))" - unfolding lower_semicontinuous_map_def -proof safe - fix xn x - assume h:"\a. openin (mtopology_of m) {x \ topspace (mtopology_of m). a < f x}" "limitin (mtopology_of m) xn x sequentially" - show "f x \ liminf (\n. f (xn n))" - unfolding le_Liminf_iff eventually_sequentially - proof safe - fix y - assume "y < f x" - show "\N. \n\N. y < f (xn n)" - proof(rule ccontr) - assume "\N. \n\N. y < f (xn n)" - then have hc:"\N. \n\N. y \ f (xn n)" - by (meson verit_comp_simplify1(3)) - define a :: "nat \ nat" where "a \ rec_nat (SOME n. f (xn n) \ y) (\n an. SOME m. m > an \ f (xn m) \ y)" - have "strict_mono a" - proof(rule strict_monoI_Suc) - fix n - have [simp]:"a (Suc n) = (SOME m. m > a n \ f (xn m) \ y)" - by(auto simp: a_def) - show "a n < a (Suc n)" - by simp (metis (no_types, lifting) Suc_le_lessD \\N. \n\N. y < f (xn n)\ not_le someI_ex) - qed - have *:"f (xn (a n)) \ y" for n - proof(cases n) - case 0 - then show ?thesis - using hc[of 0] by(auto simp: a_def intro!: someI_ex) - next - case (Suc n') - then show ?thesis - by(simp add: a_def) (metis (mono_tags, lifting) Suc_le_lessD hc someI_ex) - qed - interpret Metric_space "mspace m" "mdist m" - by simp - have [simp]: "mtopology = mtopology_of m" - by (simp add: mtopology_of_def) - show False - by (metis (mono_tags, lifting) limitin_subsequence[OF \strict_mono a\ h(2)] openin_mtopology2[THEN iffD1,simplified,OF h(1)[rule_format,of y]] \f x > y\ h(2) hc limitin_def linorder_not_less mem_Collect_eq) - qed - qed -next - fix a - assume h:" \xn x. limitin (mtopology_of m) xn x sequentially \ f x \ liminf (\n. f (xn n))" - interpret Metric_space "mspace m" "mdist m" - by simp - have [simp]: "mtopology = mtopology_of m" - by (simp add: mtopology_of_def) - have "openin mtopology {x \ topspace (mtopology_of m). a < f x}" - unfolding openin_mtopology2 topspace_mtopology_of - proof safe - fix x xn - assume h':"limitin mtopology xn x sequentially" "x \ mspace m" "f x > a" - with h have "f x \ liminf (\n. f (xn n))" - by auto - with h'(3) obtain N where N:"\n. n\N \ f (xn n) > a" - by(auto simp: le_Liminf_iff eventually_sequentially) - obtain N' where N': "\n. n\N' \ xn n \ mspace m" - by (meson h' limitin_sequentially open_in_mspace) - thus "\N. \n\N. xn n \ {x \ mspace m. f x > a}" - using h'(3) N by(auto intro!: exI[where x="max N N'"]) - qed - thus "openin (mtopology_of m) {x \ topspace (mtopology_of m). f x > a}" - by simp -qed - -lemma lower_semicontinuous_map_liminf_real: - fixes f :: "'a \ real" - shows "lower_semicontinuous_map (mtopology_of m) f \ (\xn x. limitin (mtopology_of m) xn x sequentially \ f x \ liminf (\n. f (xn n)))" - unfolding lower_semicontinuous_map_real_iff lower_semicontinuous_map_liminf_iff by simp - text \ Open maps \ lemma Metric_space_open_map_from_dist: assumes "f \ mspace m1 \ mspace m2" and "\x \. x \ mspace m1 \ \ > 0 \ \\>0. \y\mspace m1. mdist m2 (f x) (f y) < \ \ mdist m1 x y < \" shows "open_map (mtopology_of m1) (subtopology (mtopology_of m2) (f ` mspace m1)) f" proof - interpret m1: Metric_space "mspace m1" "mdist m1" by simp interpret m2: Metric_space "mspace m2" "mdist m2" by simp interpret m2': Submetric "mspace m2" "mdist m2" "f ` mspace m1" using assms(1) by(auto simp: Submetric_def Submetric_axioms_def) have "open_map m1.mtopology m2'.sub.mtopology f" proof(safe intro!: open_map_with_base[OF m1.mtopology_base_in_balls]) fix a and e :: real assume h:"a \ mspace m1" "0 < e" show "openin m2'.sub.mtopology (f ` m1.mball a e)" unfolding m2'.sub.openin_mtopology proof safe fix x assume x:"x \ m1.mball a e" then have xs:"x \ mspace m1" by auto have "0 < e - mdist m1 a x" using x by auto from assms(2)[OF xs this] obtain d where d:"d > 0" "\y. y \ mspace m1 \ mdist m2 (f x) (f y) < d \ mdist m1 x y < e - mdist m1 a x" by auto show "\r>0. m2'.sub.mball (f x) r \ f ` m1.mball a e" proof(safe intro!: exI[where x=d]) fix z assume z:"z \ m2'.sub.mball (f x) d" then obtain y where y:"z = f y" "y \ mspace m1" by auto hence "mdist m2 (f x) (f y) < d" using z by simp hence "mdist m1 x y < e - mdist m1 a x" using d y by auto hence "mdist m1 a y < e" using h(1) x y m1.triangle[of a x y] by auto with h(1) show "z \ f ` m1.mball a e" by(auto simp: y) qed fact qed auto qed thus "open_map (mtopology_of m1) (subtopology (mtopology_of m2) (f ` mspace m1)) f" by (simp add: mtopology_of_def m2'.mtopology_submetric) qed subsubsection \ Separability in Metric Spaces \ context Metric_space begin text \ For a metric space $M$, $M$ is separable iff $M$ is second countable.\ lemma generated_by_countable_balls: assumes "countable U" and "mdense U" shows "mtopology = topology_generated_by {mball y (1 / real n) | y n. y \ U}" proof - have hu: "U \ M" "\x \. x \ M \ \ > 0 \ mball x \ \ U \ {}" using assms by(auto simp: mdense_def) show ?thesis unfolding base_is_subbase[OF mtopology_base_in_balls,simplified subbase_in_def] proof(rule topology_generated_by_eq) fix K assume "K \ {mball y (1 / real n) |y n. y \ U}" then show "openin (topology_generated_by {mball a \ |a \. a \ M \ 0 < \}) K" by(auto simp: base_is_subbase[OF mtopology_base_in_balls,simplified subbase_in_def,symmetric]) next have h0:"\x \. x \ M \ \ > 0 \ \y\U. \n. x \ mball y (1 / real n) \ mball y (1 / real n) \ mball x \" proof - fix x \ assume h: "x \ M" "(0 :: real) < \" then obtain N where hn: "1 / \ < real N" using reals_Archimedean2 by blast have hn0: "0 < real N" by(rule ccontr) (use hn h in fastforce) hence hn':"1 / real N < \" using h hn by (metis divide_less_eq mult.commute) have "mball x (1 / (2 * real N)) \ U \ {}" using mdense_def[of U] assms(2) h(1) hn0 by fastforce then obtain y where hy: "y\U" "y \ M" "y \ mball x (1 / (real (2 * N)))" using hu by auto show "\y\U. \n. x \ mball y (1 / real n) \ mball y (1 / real n) \ mball x \" proof(intro bexI[where x=y] exI[where x="2 * N"] conjI) show "x \ mball y (1 / real (2 * N))" using hy(3) by (auto simp: commute) next show "mball y (1 / real (2 * N)) \ mball x \" proof fix y' assume hy':"y' \ mball y (1 / real (2 * N))" have "d x y' < \" (is "?lhs < ?rhs") proof - have "?lhs \ d x y + d y y'" using hy(2) hy' h(1) triangle by auto also have "... < 1 / real (2 * N) + 1 / real (2 * N)" by(rule strict_ordered_ab_semigroup_add_class.add_strict_mono) (use hy(3) hy(2) hy' h(1) hy' in auto) finally show ?thesis using hn' by auto qed thus "y' \ mball x \" using hy' h(1) by auto qed qed fact qed fix K assume hk: "K \ {mball a \ |a \. a \ M \ 0 < \}" then obtain x \x where hxe: "x \ M" "0 < \x" "K = mball x \x" by auto have gh:"K = (\{mball y (1 / real n) | y n. y \ U \ mball y (1 / real n) \ K})" proof show "K \ (\ {mball y (1 / real n) |y n. y \ U \ mball y (1 / real n) \ K})" proof fix k assume hkink:"k \ K" then have hkinS:"k \ M" by(simp add: hxe) obtain \k where hek: "\k > 0" "mball k \k \ K" by (metis Metric_space.openin_mtopology Metric_space_axioms hxe(3) hkink openin_mball) obtain y n where hyey: "y \ U" "k \ mball y (1 / real n)" "mball y (1 / real n) \ mball k \k" using h0[OF hkinS hek(1)] by auto show "k \ \ {mball y (1 / real n) |y n. y \ U \ mball y (1 / real n) \ K}" using hek(2) hyey by blast qed qed auto show "openin (topology_generated_by {mball y (1 / real n) |y n. y \ U}) K" unfolding openin_topology_generated_by_iff apply(rule generate_topology_on.UN[of "{mball y (1 / real n) |y n. y \ U \ mball y (1 / real n) \ K}", simplified gh[symmetric]]) apply(rule generate_topology_on.Basis) by auto qed qed lemma separable_space_imp_second_countable: assumes "separable_space mtopology" shows "second_countable mtopology" proof - obtain U where hu: "countable U" "mdense U" using assms separable_space_def2 by blast show ?thesis proof(rule countable_base_from_countable_subbase[where \="{mball y (1 / real n) | y n. y \ U}"]) have "{mball y (1 / real n) |y n. y \ U} = (\(y,n). mball y (1 / real n)) ` (U \ UNIV)" by auto also have "countable ..." using hu by auto finally show "countable {mball y (1 / real n) |y n. y \ U}" . qed(use subbase_in_def generated_by_countable_balls[of U] hu in auto) qed corollary separable_space_iff_second_countable: "separable_space mtopology \ second_countable mtopology" using second_countable_imp_separable_space separable_space_imp_second_countable by auto lemma Lindelof_mdiameter: assumes "separable_space mtopology" "0 < e" shows "\U. countable U \ \ U = M \ (\u\U. mdiameter u < ennreal e)" proof - have "(\u. u \ {mball x (e / 3) |x. x \ M} \ openin mtopology u)" by auto moreover have "\ {mball x (e / 3) |x. x \ M} = M" using assms by auto ultimately have "\U'. countable U' \ U' \ {mball x (e / 3) |x. x \ M} \ \ U' = M" using second_countable_imp_Lindelof_space[OF assms(1)[simplified separable_space_iff_second_countable]] by(auto simp: Lindelof_space_def) then obtain U' where U': "countable U'" "U' \ {mball x (e / 3) |x. x \ M}" "\ U' = M" by auto show ?thesis proof(safe intro!: exI[where x=U']) fix u assume "u \ U'" then obtain x where u:"u = mball x (e / 3)" using U' by auto have "mdiameter u \ ennreal (2 * (e / 3))" by(simp only: u mdiameter_ball_leq) also have "... < ennreal e" by(auto intro!: ennreal_lessI assms) finally show "mdiameter u < ennreal e" . qed(use U' in auto) qed end lemma metrizable_space_separable_iff_second_countable: assumes "metrizable_space X" shows "separable_space X \ second_countable X" proof - obtain d where "Metric_space (topspace X) d" "Metric_space.mtopology (topspace X) d = X" by (metis assms(1) Metric_space.topspace_mtopology metrizable_space_def) thus ?thesis using Metric_space.separable_space_iff_second_countable by fastforce qed abbreviation "mdense_of m U \ dense_in (mtopology_of m) U" lemma mdense_of_def: "mdense_of m U \ (U \ mspace m \ (\x\mspace m. \\>0. mball_of m x \ \ U \ {}))" using Metric_space.mdense_def[of "mspace m" "mdist m"] by (simp add: mball_of_def mtopology_of_def) lemma mdense_of_def2: "mdense_of m U \ (U \ mspace m \ (\x\mspace m. \\>0. \y\U. mdist m x y < \))" using Metric_space.mdense_def2[of "mspace m" "mdist m"] by (simp add: mtopology_of_def) lemma mdense_of_def3: "mdense_of m U \ (U \ mspace m \ (\x\mspace m. \xn. range xn \ U \ limitin (mtopology_of m) xn x sequentially))" using Metric_space.mdense_def3[of "mspace m" "mdist m"] by (simp add: mtopology_of_def) subsubsection \ Compact Metric Spaces\ context Metric_space begin lemma mtotally_bounded_eq_compact_closedin: assumes "mcomplete" "closedin mtopology S" shows "mtotally_bounded S \ S \ M \ compactin mtopology S" by (metis assms closure_of_eq mtotally_bounded_eq_compact_closure_of) lemma mtotally_bounded_def2: "mtotally_bounded S \ (\\>0. \K. finite K \ K \ M \ S \ (\x\K. mball x \))" unfolding mtotally_bounded_def proof safe fix e :: real assume h:"e > 0" "\e>0. \K. finite K \ K \ S \ S \ (\x\K. mball x e)" then show "\K. finite K \ K \ M \ S \ (\x\K. mball x e)" by (metis Metric_space.mbounded_subset Metric_space.mtotally_bounded_imp_mbounded Metric_space_axioms mbounded_subset_mspace mtotally_bounded_def) next fix e :: real assume "e > 0" "\\>0. \K. finite K \ K \ M \ S \ (\x\K. mball x \)" then obtain K where K: "finite K" "K \ M" "S \ (\x\K. mball x (e / 2))" by (meson half_gt_zero) define K' where "K' \ {x\K. mball x (e / 2) \ S \ {}}" have K'1:"finite K'" "K' \ M" using K by(auto simp: K'_def) have K'2: "S \ (\x\K'. mball x (e / 2))" proof fix x assume x:"x \ S" then obtain k where k:"k \ K" "x \ mball k (e / 2)" using K by auto with x have "k \ K'" by(auto simp: K'_def) with k show "x \ (\x\K'. mball x (e / 2))" by auto qed have "\k\K'. \y. y \ mball k (e / 2) \ S" by(auto simp: K'_def) then obtain xk where xk: "\k. k \ K' \ xk k \ mball k (e / 2)" "\k. k \ K' \ xk k \ S" by (metis IntD2 inf_commute) hence "\k. k \ K' \ mball k (e / 2) \ mball (xk k) e" using triangle commute by fastforce hence "(\x\K'. mball x (e / 2)) \ (\x\xk ` K'. mball x e)" by blast with K'2 have "S \ (\x\xk ` K'. mball x e)" by blast thus "\K. finite K \ K \ S \ S \ (\x\K. mball x e)" by(intro exI[where x="xk ` K'"]) (use xk(2) K'1(1) in blast) qed lemma compact_space_imp_separable: assumes "compact_space mtopology" shows "separable_space mtopology" proof - have "\A. finite A \ A \ M \ M \ \ ((\a. mball a (1 / real (Suc n))) ` A)" for n using assms by(auto simp: compact_space_eq_mcomplete_mtotally_bounded mtotally_bounded_def) then obtain A where A: "\n. finite (A n)" "\n. A n \ M" "\n. M \ \ ((\a. mball a (1 / real (Suc n))) ` (A n))" by metis define K where "K \ \ (range A)" have 1: "countable K" using A(1) by(auto intro!: countable_UN[of _ id,simplified] simp: K_def countable_finite) show "separable_space mtopology" unfolding mdense_def2 separable_space_def2 proof(safe intro!: exI[where x=K] 1) fix x and \ :: real assume h: "x \ M" "0 < \" then obtain n where n:"1 / real (Suc n) \ \" by (meson nat_approx_posE order.strict_iff_not) then obtain y where y: "y \ A n" "x \ mball y (1 / real (Suc n))" using h(1) A(3)[of n] by auto thus "\y\K. d x y < \" using n by(auto intro!: bexI[where x=y] simp: commute K_def) qed(use K_def A(2) in auto) qed lemma separable_space_cfunspace: assumes "separable_space mtopology" mcomplete and "metrizable_space X" "compact_space X" shows "separable_space (mtopology_of (cfunspace X Self))" proof - consider "topspace X = {}" | "topspace X \ {}" "M = {}" | "topspace X \ {}" "M \ {}" by auto thus ?thesis proof cases case 1 then show ?thesis by(auto intro!: countable_space_separable_space) next case 2 then have [simp]:"mtopology = trivial_topology" using topspace_mtopology by auto have 1:"topspace (mtopology_of (cfunspace X Self)) = {}" apply simp using 2(1) by(auto simp: mtopology_of_def) show ?thesis by(rule countable_space_separable_space, simp only: 1) simp next case XS_nem:3 have cm: "mcomplete_of (cfunspace X Self)" by (simp add: assms(2) mcomplete_cfunspace) show ?thesis proof - obtain dx where dx: "Metric_space (topspace X) dx" "Metric_space.mtopology (topspace X) dx = X" by (metis Metric_space.topspace_mtopology assms(3) metrizable_space_def) interpret dx: Metric_space "topspace X" dx by fact have dx_c: dx.mcomplete using assms by(auto intro!: dx.compact_space_imp_mcomplete simp: dx(2)) have "\B. finite B \ B \ topspace X \ topspace X \ (\a\B. dx.mball a (1 / Suc m))" for m using dx.compactin_imp_mtotally_bounded[of "topspace X"] assms(4) by(fastforce simp: dx(2) compact_space_def dx.mtotally_bounded_def2) then obtain Xm where Xm: "\m. finite (Xm m)" "\m. (Xm m) \ topspace X" "\m. topspace X \ (\a\Xm m. dx.mball a (1 / Suc m))" by metis hence Xm_eq: "\m. topspace X = (\a\Xm m. dx.mball a (1 / Suc m))" by fastforce have Xm_nem:"\m. (Xm m) \ {}" using XS_nem Xm_eq by blast define xmk where "xmk \ (\m. from_nat_into (Xm m))" define km where "km \ (\m. card (Xm m))" have km_pos:"km m > 0" for m by(simp add: km_def card_gt_0_iff Xm Xm_nem) have xmk_bij: "bij_betw (xmk m) {.. Xm m" for m i by (simp add: Xm_nem from_nat_into xmk_def) have "\U. countable U \ \ U = M \ (\u\U. mdiameter u < 1 / (Suc l))" for l by(rule Lindelof_mdiameter[OF assms(1)]) auto then obtain U where U: "\l. countable (U l)" "\l. (\ (U l)) = M" "\l u. u \ U l \ mdiameter u < 1 / Suc l" by metis have Ul_nem: "U l \ {}" for l using XS_nem U(2) by auto define uli where "uli \ (\l. from_nat_into (U l))" have uli_into:"uli l i \ U l" for l i by (simp add: Ul_nem from_nat_into uli_def) hence uli_diam: "mdiameter (uli l i) < 1 / Suc l" for l i using U(3) by auto have uli_un:"M = (\i. uli l i)" for l by(auto simp: range_from_nat_into[OF Ul_nem[of l] U(1)] uli_def U) define Cmn where "Cmn \ (\m n. {f \ mspace (cfunspace X Self). \x\topspace X. \y\topspace X. dx x y < 1 / (Suc m) \ d (f x) (f y) < 1 / Suc n})" define fmnls where "fmnls \ (\m n l s. SOME f. f \ Cmn m n \ (\j uli l (s j)))" define Dmnl where "Dmnl \ (\m n l. {fmnls m n l s |s. s \ {..\<^sub>E UNIV \ (\f \ Cmn m n. (\j uli l (s j))) })" have in_Dmnl: "fmnls m n l s \ Dmnl m n l" if "s\{..\<^sub>E UNIV" "f\ Cmn m n" "\j uli l (s j)"for m n l s f using Dmnl_def that by blast define Dmn where "Dmn \ (\m n. \l. Dmnl m n l)" have Dmn_subset: "Dmn m n \ Cmn m n" for m n proof - have "Dmnl m n l \ Cmn m n" for l by(auto simp: Dmnl_def fmnls_def someI[of "\f. f \ Cmn m n \ (\j uli l (_ j))"]) thus ?thesis by(auto simp: Dmn_def) qed have c_Dmn: "countable (Dmn m n)" for m n proof - have "countable (Dmnl m n l)" for l proof - have 1:"Dmnl m n l \ (\s. fmnls m n l s) ` ({..\<^sub>E UNIV)" by(auto simp: Dmnl_def) have "countable ..." by(auto intro!: countable_PiE) with 1 show ?thesis using countable_subset by blast qed thus ?thesis by(auto simp: Dmn_def) qed have claim: "\g\Dmn m n. \y\Xm m. d (f y) (g y) < e" if f:"f \ Cmn m n" and e:"e > 0" for f m n e proof - obtain l where l:"1 / Suc l < e" using e nat_approx_posE by blast define s where "s \ (\i\{.. uli l j)" have s1:"s\{..\<^sub>E UNIV" by(simp add: s_def) have s2:"\i uli l (s i)" proof - fix i have "f (xmk m i) \ uli l (SOME j. f (xmk m i) \ uli l j)" for i proof(rule someI_ex) have "xmk m i \ topspace X" using Xm(2) xmk_into by auto hence "f (xmk m i) \ M" using f by(auto simp: Cmn_def continuous_map_def) thus "\x. f (xmk m i) \ uli l x" using uli_un by auto qed thus ?thesis by (auto simp: s_def) qed have fmnls:"fmnls m n l s \ Cmn m n \ (\j uli l (s j))" by(simp add: fmnls_def,rule someI[where x=f],auto simp: s2 f) show "\g\Dmn m n. \y\Xm m. d (f y) (g y) < e" proof(safe intro!: bexI[where x="fmnls m n l s"]) fix y assume y:"y \ Xm m" then obtain i where i:"i < km m" "xmk m i = y" by (meson xmk_bij[of m] bij_betw_iff_bijections lessThan_iff) have "f y \ uli l (s i)" "fmnls m n l s y \ uli l (s i)" using i(1) s2 fmnls by(auto simp: i(2)[symmetric]) moreover have "f y \ M" "fmnls m n l s y \ M" using f fmnls y Xm(2)[of m] by(auto simp: Cmn_def continuous_map_def) ultimately have "ennreal (d (f y) (fmnls m n l s y)) \ mdiameter (uli l (s i))" by(auto intro!: mdiameter_is_sup) also have "... < ennreal (1 / Suc l)" by(rule uli_diam) also have "... < ennreal e" using l e by(auto intro!: ennreal_lessI) finally show "d (f y) (fmnls m n l s y) < e" by(simp add: ennreal_less_iff[OF nonneg]) qed(use in_Dmnl[OF s1 f s2] Dmn_def in auto) qed show "separable_space (mtopology_of (cfunspace X Self))" unfolding separable_space_def2 mdense_of_def2 proof(safe intro!: exI[where x="\n. (\m. Dmn m n)"]) fix f and e :: real assume h:"f \ mspace (cfunspace X Self)" "0 < e" then obtain n where n:"1 / Suc n < e / 4" by (metis zero_less_divide_iff zero_less_numeral nat_approx_posE) have "\m. \l\ m. f \ Cmn l n" proof - have "uniformly_continuous_map dx.Self Self f" using continuous_eq_uniformly_continuous_map[of dx.Self Self f] h assms(4) by(auto simp: mtopology_of_def dx) then obtain \ where \:"\ > 0" "\x y. x\topspace X \ y\topspace X \ dx x y < \ \ d (f x) (f y) < 1 / (Suc n)" by(simp only: uniformly_continuous_map_def) (metis dx.mdist_Self dx.mspace_Self mdist_Self of_nat_0_less_iff zero_less_Suc zero_less_divide_1_iff) then obtain m where m:"1 / Suc m < \" using nat_approx_posE by blast have l: "l \ m \ 1 / Suc l \ 1 / Suc m" for l by (simp add: frac_le) show ?thesis using \(2)[OF _ _ order.strict_trans[OF _ order.strict_trans1[OF l m]]] h by(auto simp: Cmn_def) qed then obtain m where m:"f \ Cmn m n" by auto obtain g where g:"g\Dmn m n" "\y. y\Xm m \ d (f y) (g y) < e / 4" by (metis claim[OF m] h(2) zero_less_divide_iff zero_less_numeral) have "\n m. \g\Dmn m n. mdist (cfunspace X Self) f g < e" proof(rule exI[where x=n]) show "\m. \g\Dmn m n. mdist (cfunspace X Self) f g < e" proof(intro exI[where x=m] bexI[OF _ g(1)]) have g_cm:"g \ mspace (cfunspace X Self)" using g(1) Dmn_subset[of m n] by(auto simp: Cmn_def) have "mdist (cfunspace X Self) f g \ 3 * e / 4" proof(rule mdist_cfunspace_le) fix x assume x:"x \ topspace X" obtain y where y:"y \ Xm m" "x \ dx.mball y (1 / real (Suc m))" using Xm(3) x by fastforce hence ytop:"y \ topspace X" by auto have "mdist Self (f x) (g x) \ d (f x) (f y) + d (f y) (g x)" using x g_cm h(1) ytop by(auto intro!: triangle simp: continuous_map_def) also have "... \ d (f x) (f y) + d (f y) (g y) + d (g y) (g x)" using x g_cm h(1) ytop by(auto intro!: triangle simp: continuous_map_def) also have "... \ e / 4 + e / 4 + e / 4" proof - have dxy: "dx x y < 1 / Suc m" "dx y x < 1 / Suc m" using y(2) by(auto simp: dx.commute) hence "d (f x) (f y) < 1 / (Suc n)" "d (g y) (g x) < 1 / (Suc n)" using m x ytop g Dmn_subset[of m n] by(auto simp: Cmn_def) hence "d (f x) (f y) < e / 4" "d (g y) (g x) < e / 4" using n by auto thus ?thesis using g(2)[OF y(1)] by auto qed finally show "mdist Self (f x) (g x) \ 3 * e / 4" by simp qed(use h in auto) also have "... < e" using h by auto finally show "mdist (cfunspace X Self) f g < e" . qed qed thus "\y\\n m. Dmn m n. mdist (cfunspace X Self) f y < e" by auto qed(insert Dmn_subset c_Dmn,unfold Cmn_def, blast)+ qed qed qed end context Submetric begin lemma separable_sub: assumes "separable_space mtopology" shows "separable_space sub.mtopology" using assms unfolding separable_space_iff_second_countable sub.separable_space_iff_second_countable by(auto simp: second_countable_subtopology mtopology_submetric) end subsubsection \Discrete Distance\ lemma(in discrete_metric) separable_space_iff: "separable_space disc.mtopology \ countable M" by(simp add: mtopology_discrete_metric separable_space_discrete_topology) subsubsection \Binary Product Metric Spaces\ text \ We define the $L^{1}$-distance. $L^{1}$-distance and $L^{2}$ distance (Euclid distance) generate the same topological space.\ definition "prod_dist_L1 \ \d1 d2 (x,y) (x',y'). d1 x x' + d2 y y'" context Metric_space12 begin lemma prod_L1_metric: "Metric_space (M1 \ M2) (prod_dist_L1 d1 d2)" proof fix x y z assume "x \ M1 \ M2" "y \ M1 \ M2" "z \ M1 \ M2" then show "prod_dist_L1 d1 d2 x z \ prod_dist_L1 d1 d2 x y + prod_dist_L1 d1 d2 y z" by(auto simp: prod_dist_L1_def) (metis M1.commute M1.triangle'' M2.commute M2.triangle'' ab_semigroup_add_class.add_ac(1) add.left_commute add_mono) qed(auto simp: prod_dist_L1_def add_nonneg_eq_0_iff M1.commute M2.commute) sublocale Prod_metric_L1: Metric_space "M1 \ M2" "prod_dist_L1 d1 d2" by(simp add: prod_L1_metric) lemma prod_dist_L1_geq: shows "d1 x y \ prod_dist_L1 d1 d2 (x,x') (y,y')" "d2 x' y' \ prod_dist_L1 d1 d2 (x,x') (y,y')" by(auto simp: prod_dist_L1_def) lemma prod_dist_L1_ball: assumes "(x,x') \ Prod_metric_L1.mball (a,a') \" shows "x \ M1.mball a \" and "x' \ M2.mball a' \" using assms prod_dist_L1_geq order.strict_trans1 by simp_all blast+ lemma prod_dist_L1_ball': assumes "z \ Prod_metric_L1.mball a \" shows "fst z \ M1.mball (fst a) \" and "snd z \ M2.mball (snd a) \" using prod_dist_L1_ball[of "fst z" "snd z" "fst a" "snd a" \] assms by auto lemma prod_dist_L1_ball1': "Prod_metric_L1.mball (a1,a2) (min e1 e2) \ M1.mball a1 e1 \ M2.mball a2 e2" using prod_dist_L1_geq order.strict_trans1 by auto blast+ lemma prod_dist_L1_ball1: assumes "b1 \ M1.mball a1 e1" "b2 \ M2.mball a2 e2" shows "\e12>0. Prod_metric_L1.mball (b1,b2) e12 \ M1.mball a1 e1 \ M2.mball a2 e2" proof - obtain ea1 ea2 where ea: "ea1 > 0" "ea2 > 0" "M1.mball b1 ea1 \ M1.mball a1 e1" "M2.mball b2 ea2 \ M2.mball a2 e2" using assms by (meson M1.openin_mball M1.openin_mtopology M2.openin_mball M2.openin_mtopology) thus ?thesis using prod_dist_L1_ball1'[of b1 b2 ea1 ea2] by(auto intro!: exI[where x="min ea1 ea2"]) qed lemma prod_dist_L1_ball2': "M1.mball a1 e1 \ M2.mball a2 e2 \ Prod_metric_L1.mball (a1,a2) (e1 + e2)" by auto (auto simp: prod_dist_L1_def) lemma prod_dist_L1_ball2: assumes "(b1,b2) \ Prod_metric_L1.mball (a1,a2) e12" shows "\e1>0. \e2>0. M1.mball b1 e1 \ M2.mball b2 e2 \ Prod_metric_L1.mball (a1,a2) e12" proof - obtain e12' where "e12' > 0" "Prod_metric_L1.mball (b1,b2) e12' \ Prod_metric_L1.mball (a1,a2) e12" by (metis assms Prod_metric_L1.openin_mball Prod_metric_L1.openin_mtopology) thus ?thesis using prod_dist_L1_ball2'[of b1 "e12' / 2" b2 "e12' / 2"] by(auto intro!: exI[where x="e12' / 2"]) qed lemma prod_dist_L1_mtopology: "Prod_metric_L1.mtopology = prod_topology M1.mtopology M2.mtopology" proof - have "Prod_metric_L1.mtopology = topology_generated_by { M1.mball a1 e1 \ M2.mball a2 e2 | a1 a2 e1 e2. a1 \ M1 \ a2 \ M2 \ e1 > 0 \ e2 > 0}" unfolding base_is_subbase[OF Prod_metric_L1.mtopology_base_in_balls,simplified subbase_in_def] proof(rule topology_generated_by_eq) fix U assume "U \ {M1.mball a1 e1 \ M2.mball a2 e2 | a1 a2 e1 e2. a1 \ M1 \ a2 \ M2 \ 0 < e1 \ 0 < e2}" then obtain a1 e1 a2 e2 where hae: "U = M1.mball a1 e1 \ M2.mball a2 e2" "a1 \ M1" "a2 \ M2" "0 < e1" "0 < e2" by auto show "openin (topology_generated_by {Prod_metric_L1.mball a \ |a \. a \ M1 \ M2 \ 0 < \}) U" unfolding base_is_subbase[OF Prod_metric_L1.mtopology_base_in_balls,simplified subbase_in_def,symmetric] Prod_metric_L1.openin_mtopology hae(1) using prod_dist_L1_ball1[of _ a1 e1 _ a2 e2] by fastforce next fix U assume "U \ {Prod_metric_L1.mball a \ |a \. a \ M1 \ M2 \ 0 < \}" then obtain a1 a2 \ where hae: "U = Prod_metric_L1.mball (a1,a2) \" "a1 \ M1" "a2 \ M2" "0 < \" by auto show "openin (topology_generated_by {M1.mball a1 e1 \ M2.mball a2 e2 | a1 a2 e1 e2. a1 \ M1 \ a2 \ M2 \ 0 < e1 \ 0 < e2}) U" unfolding openin_subopen[of _ "Prod_metric_L1.mball (a1,a2) \"] hae(1) proof safe fix b1 b2 assume h:"(b1, b2) \ Prod_metric_L1.mball (a1, a2) \" from prod_dist_L1_ball2[OF this] obtain e1 e2 where "e1 > 0" "e2 > 0" "M1.mball b1 e1 \ M2.mball b2 e2 \ Prod_metric_L1.mball (a1, a2) \" by metis with h show "\T. openin (topology_generated_by {M1.mball a1 e1 \ M2.mball a2 e2 | a1 a2 e1 e2. a1 \ M1 \ a2 \ M2 \ 0 < e1 \ 0 < e2}) T \ (b1, b2) \ T \ T \ Prod_metric_L1.mball (a1, a2) \" unfolding openin_topology_generated_by_iff by(auto intro!: generate_topology_on.Basis exI[where x="M1.mball b1 e1 \ M2.mball b2 e2"]) qed qed also have "... = prod_topology M1.mtopology M2.mtopology" proof - have "{M1.mball a \ \ M2.mball a' \' |a a' \ \'. a \ M1 \ a' \ M2 \ 0 < \ \ 0 < \'} = {U \ V |U V. U \ {M1.mball a \ |a \. a \ M1 \ 0 < \} \ V \ {M2.mball a \ |a \. a \ M2 \ 0 < \}}" by blast thus ?thesis unfolding base_is_subbase[OF M1.mtopology_base_in_balls,simplified subbase_in_def] base_is_subbase[OF M2.mtopology_base_in_balls,simplified subbase_in_def] by(simp only: prod_topology_generated_by[symmetric]) qed finally show ?thesis . qed lemma prod_dist_L1_limitin_iff: "limitin Prod_metric_L1.mtopology zn z sequentially \ limitin M1.mtopology (\n. fst (zn n)) (fst z) sequentially \ limitin M2.mtopology (\n. snd (zn n)) (snd z) sequentially" proof safe assume h:"limitin Prod_metric_L1.mtopology zn z sequentially" show "limitin M1.mtopology (\n. fst (zn n)) (fst z) sequentially" "limitin M2.mtopology (\n. snd (zn n)) (snd z) sequentially" unfolding M1.limit_metric_sequentially M2.limit_metric_sequentially proof safe fix e :: real assume e: "0 < e" with h obtain N where N:"\n. n \ N \ zn n \ M1 \ M2" "\n. n \ N \ prod_dist_L1 d1 d2 (zn n) z < e" by(simp only: Prod_metric_L1.limit_metric_sequentially) metis show "\N. \n\N. fst (zn n) \ M1 \ d1 (fst (zn n)) (fst z) < e" "\N. \n\N. snd (zn n) \ M2 \ d2 (snd (zn n)) (snd z) < e" proof(safe intro!: exI[where x=N]) fix n assume "N \ n" from N[OF this] show "d1 (fst (zn n)) (fst z) < e" " d2 (snd (zn n)) (snd z) < e" using order.strict_trans1[OF prod_dist_L1_geq(1)[of "fst (zn n)" "fst z" "snd (zn n)" "snd z"]] order.strict_trans1[OF prod_dist_L1_geq(2)[of "snd (zn n)" "snd z" "fst (zn n)" "fst z"]] by auto qed(use N(1)[simplified mem_Times_iff] in auto) qed(use h Prod_metric_L1.limit_metric_sequentially in auto) next assume h:"limitin M1.mtopology (\n. fst (zn n)) (fst z) sequentially" "limitin M2.mtopology (\n. snd (zn n)) (snd z) sequentially" show "limitin Prod_metric_L1.mtopology zn z sequentially" unfolding Prod_metric_L1.limit_metric_sequentially proof safe fix e :: real assume e: "0 < e" with h obtain N1 N2 where N: "\n. n \ N1 \ fst (zn n) \ M1" "\n. n \ N1 \ d1 (fst (zn n)) (fst z) < e / 2" "\n. n \ N2 \ snd (zn n) \ M2" "\n. n \ N2 \ d2 (snd (zn n)) (snd z) < e / 2" unfolding M1.limit_metric_sequentially M2.limit_metric_sequentially using half_gt_zero by metis thus "\N. \n\N. zn n \ M1 \ M2 \ prod_dist_L1 d1 d2 (zn n) z < e" by(fastforce intro!: exI[where x="max N1 N2"] simp: prod_dist_L1_def split_beta' mem_Times_iff) qed(auto simp: mem_Times_iff h[simplified M1.limit_metric_sequentially M2.limit_metric_sequentially]) qed lemma prod_dist_L1_MCauchy_iff: "Prod_metric_L1.MCauchy zn \ M1.MCauchy (\n. fst (zn n)) \ M2.MCauchy (\n. snd (zn n))" proof safe assume h:"Prod_metric_L1.MCauchy zn" show "M1.MCauchy (\n. fst (zn n))" "M2.MCauchy (\n. snd (zn n))" unfolding M1.MCauchy_def M2.MCauchy_def proof safe fix e :: real assume "0 < e" with h obtain N where N:"\n m. N \ n \ N \ m \ prod_dist_L1 d1 d2 (zn n) (zn m) < e" by(fastforce simp: Prod_metric_L1.MCauchy_def) show "\N. \n n'. N \ n \ N \ n' \ d1 (fst (zn n)) (fst (zn n')) < e" " \N. \n n'. N \ n \ N \ n' \ d2 (snd (zn n)) (snd (zn n')) < e" proof(safe intro!: exI[where x=N]) fix n m assume "n \ N" "m \ N" from N[OF this] show "d1 (fst (zn n)) (fst (zn m)) < e" "d2 (snd (zn n)) (snd (zn m)) < e" using order.strict_trans1[OF prod_dist_L1_geq(1)[of "fst (zn n)" "fst (zn m)" "snd (zn n)" "snd (zn m)"]] order.strict_trans1[OF prod_dist_L1_geq(2)[of "snd (zn n)" "snd (zn m)" "fst (zn n)" "fst (zn m)"]] by auto qed next have "\n. zn n \ M1 \ M2" using h by(auto simp: Prod_metric_L1.MCauchy_def) thus "fst (zn n) \ M1" "snd (zn n) \ M2" for n by (auto simp: mem_Times_iff) qed next assume h:"M1.MCauchy (\n. fst (zn n))" "M2.MCauchy (\n. snd (zn n))" show "Prod_metric_L1.MCauchy zn" unfolding Prod_metric_L1.MCauchy_def proof safe fix e :: real assume "0 < e" with h obtain N1 N2 where "\n m. n \ N1 \ m \ N1 \ d1 (fst (zn n)) (fst (zn m)) < e / 2" "\n m. n \ N2 \ m \ N2 \ d2 (snd (zn n)) (snd (zn m)) < e / 2" unfolding M1.MCauchy_def M2.MCauchy_def using half_gt_zero by metis thus "\N. \n n'. N \ n \ N \ n' \ prod_dist_L1 d1 d2 (zn n) (zn n') < e" by(fastforce intro!: exI[where x="max N1 N2"] simp: prod_dist_L1_def split_beta') next fix x y n assume 1:"(x,y) = zn n" have "fst (zn n) \ M1" "snd (zn n) \ M2" using h[simplified M1.MCauchy_def M2.MCauchy_def] by auto thus "x \ M1" "y \ M2" by(simp_all add: 1[symmetric]) qed qed end subsubsection \Sum Metric Spaces\ locale Sum_metric = fixes I :: "'i set" and Mi :: "'i \ 'a set" and di :: "'i \ 'a \ 'a \ real" assumes Mi_disj: "disjoint_family_on Mi I" and d_nonneg: "\i x y. 0 \ di i x y" and d_bounded: "\i x y. di i x y < 1" and Md_metric: "\i. i \ I \ Metric_space (Mi i) (di i)" begin abbreviation "M \ \i\I. Mi i" lemma Mi_inj_on: assumes "i \ I" "j \ I" "a \ Mi i" "a \ Mi j" shows "i = j" using Mi_disj assms by(auto simp: disjoint_family_on_def) definition sum_dist :: "['a, 'a] \ real" where "sum_dist x y \ (if x \ M \ y \ M then (if \i\I. x \ Mi i \ y \ Mi i then di (THE i. i \ I \ x \ Mi i \ y \ Mi i) x y else 1) else 0)" lemma sum_dist_simps: shows "\i. \i \ I; x \ Mi i; y \ Mi i\ \ sum_dist x y = di i x y" and "\i j. \i \ I; j \ I; i \ j; x \ Mi i; y \ Mi j\ \ sum_dist x y = 1" and "\i. \i \ I; y \ M; x \ Mi i; y \ Mi i\ \ sum_dist x y = 1" and "\i. \i \ I; x \ M; y \ Mi i; x \ Mi i\ \ sum_dist x y = 1" and "x \ M \ sum_dist x y = 0" "y \ M \ sum_dist x y = 0" proof - { fix i assume h:"i \ I" "x \ Mi i" "y \ Mi i" then have "sum_dist x y = di (THE i. i \ I \ x \ Mi i \ y \ Mi i) x y" by(auto simp: sum_dist_def) also have "... = di i x y" proof - have "(THE i. i \ I \ x \ Mi i \ y \ Mi i) = i" using Mi_disj h by(auto intro!: the1_equality simp: disjoint_family_on_def) thus ?thesis by simp qed finally show "sum_dist x y = di i x y" . } show "\i j. \i \ I; j \ I; i \ j; x \ Mi i; y \ Mi j\ \ sum_dist x y = 1" "\i. \i \ I; y \ M; x \ Mi i; y \ Mi i\ \ sum_dist x y = 1" "\i. \i \ I; x \ M; y \ Mi i; x \ Mi i\ \ sum_dist x y = 1" "x \ M \ sum_dist x y = 0" "y \ M \ sum_dist x y = 0" using Mi_disj by(auto simp: sum_dist_def disjoint_family_on_def dest: Mi_inj_on) qed lemma sum_dist_if_less1: assumes "i \ I" "x \ Mi i" "y \ M" "sum_dist x y < 1" shows "y \ Mi i" using assms sum_dist_simps(3) by fastforce lemma inM_cases: assumes "x \ M" "y \ M" and "\i. \i \ I; x \ Mi i; y \ Mi i\ \ P x y" and "\i j. \i \ I; j \ I; i \ j; x \ Mi i; y \ Mi j; x \ y\ \ P x y" shows "P x y" using assms by auto sublocale Sum_metric: Metric_space M sum_dist proof fix x y assume "x \ M" "y \ M" then show "sum_dist x y = 0 \ x = y" by(rule inM_cases, insert Md_metric) (auto simp: sum_dist_simps Metric_space_def) next { fix i x y assume h: "i \ I" "x \ Mi i" "y \ Mi i" then have "sum_dist x y = di i x y" "sum_dist y x = di i x y" using Md_metric by(auto simp: sum_dist_simps Metric_space_def) } thus "\x y. sum_dist x y = sum_dist y x" by (metis (no_types, lifting) sum_dist_def) next show 1:"\x y. 0 \ sum_dist x y" using d_nonneg by(simp add: sum_dist_def) fix x y z assume h: "x \ M" "y \ M" "z \ M" show "sum_dist x z \ sum_dist x y + sum_dist y z" (is "?lhs \ ?rhs") proof(rule inM_cases[OF h(1,3)]) fix i assume h':"i \ I" "x \ Mi i" "z \ Mi i" consider "y \ Mi i" | "y \ Mi i" by auto thus "?lhs \ ?rhs" proof cases case 1 with h' Md_metric [OF h'(1)] show ?thesis by(simp add: sum_dist_simps Metric_space_def) next case 2 with h' h(2) d_bounded[of i x z] 1[of y z] show ?thesis by(auto simp: sum_dist_simps) qed next fix i j assume h': "i \ I" "j \ I" "i \ j" "x \ Mi i" "z \ Mi j" consider "y \ Mi i" | "y \ Mi j" using h' h(2) Mi_disj by(auto simp: disjoint_family_on_def) thus "?lhs \ ?rhs" by (cases, insert 1[of x y] 1[of y z] h' h(2)) (auto simp: sum_dist_simps) qed qed lemma sum_dist_le1: "sum_dist x y \ 1" using d_bounded[of _ x y] by(auto simp: sum_dist_def less_eq_real_def) lemma sum_dist_ball_eq_ball: assumes "i \ I" "e \ 1" "x \ Mi i" shows "Metric_space.mball (Mi i) (di i) x e = Sum_metric.mball x e" proof - interpret m: Metric_space "Mi i" "di i" by(simp add: assms Md_metric) show ?thesis using assms sum_dist_simps(1)[OF assms(1) assms(3)] sum_dist_if_less1[OF assms(1,3)] by(fastforce simp: Sum_metric.mball_def) qed lemma ball_le_sum_dist_ball: assumes "i \ I" shows "Metric_space.mball (Mi i) (di i) x e \ Sum_metric.mball x e" proof - interpret m: Metric_space "Mi i" "di i" by(simp add: assms Md_metric) show ?thesis using assms by(auto simp: sum_dist_simps) qed lemma openin_mtopology_iff: "openin Sum_metric.mtopology U \ U \ M \ (\i\I. openin (Metric_space.mtopology (Mi i) (di i)) (U \ Mi i))" proof safe fix i assume h:"openin Sum_metric.mtopology U" "i \ I" then interpret m: Metric_space "Mi i" "di i" using Md_metric by simp show "openin m.mtopology (U \ Mi i)" unfolding m.openin_mtopology proof safe fix x assume x:"x \ U" "x \ Mi i" with h obtain e where e: "e > 0" "Sum_metric.mball x e \ U" by(auto simp: Sum_metric.openin_mtopology) show "\\>0. m.mball x \ \ U \ Mi i" proof(safe intro!: exI[where x=e]) fix y assume "y \ m.mball x e" with h(2) have "y \ Sum_metric.mball x e" by(auto simp:sum_dist_simps) with e show "y \ U" by blast qed(use e in auto) qed next show "\x. openin Sum_metric.mtopology U \ x \ U \ x \ M" by(auto simp: Sum_metric.openin_mtopology) next assume h: "U \ M" "\i\I. openin (Metric_space.mtopology (Mi i) (di i)) (U \ Mi i)" show "openin Sum_metric.mtopology U" unfolding Sum_metric.openin_mtopology proof safe fix x assume x: "x \ U" then obtain i where i: "i \ I" "x \ Mi i" using h(1) by auto then interpret m: Metric_space "Mi i" "di i" using Md_metric by simp obtain e where e: "e > 0" "m.mball x e \ U \ Mi i" using i h(2) by (meson IntI m.openin_mtopology x) show "\\>0. Sum_metric.mball x \ \ U" proof(safe intro!: exI[where x="min e 1"]) fix y assume y:"y \ Sum_metric.mball x (min e 1)" then show "y \ U" using sum_dist_ball_eq_ball[OF i(1) _ i(2),of "min e 1"] e by fastforce qed(simp add: e) qed(use h(1) in auto) qed corollary openin_mtopology_Mi: assumes "i \ I" shows "openin Sum_metric.mtopology (Mi i)" unfolding openin_mtopology_iff proof safe fix j assume j:"j \ I" then interpret m: Metric_space "Mi j" "di j" by(simp add: Md_metric) show "openin m.mtopology (Mi i \ Mi j)" by (cases "i = j", insert assms Mi_disj j) (auto simp: disjoint_family_on_def) qed(use assms in auto) corollary subtopology_mtopology_Mi: assumes "i \ I" shows "subtopology Sum_metric.mtopology (Mi i) = Metric_space.mtopology (Mi i) (di i)" proof - interpret Mi: Metric_space "Mi i" "di i" by (simp add: Md_metric assms) show ?thesis unfolding topology_eq openin_subtopology proof safe fix T assume "openin Sum_metric.mtopology T" thus "openin Mi.mtopology (T \ Mi i)" using assms by(auto simp: openin_mtopology_iff) next fix S assume h:"openin Mi.mtopology S" show "\T. openin Sum_metric.mtopology T \ S = T \ Mi i" proof(safe intro!: exI[where x=S]) show "openin Sum_metric.mtopology S" unfolding openin_mtopology_iff proof safe fix j assume j:"j \ I" then interpret Mj: Metric_space "Mi j" "di j" using Md_metric by auto have "i \ j \ S \ Mi j = {}" using openin_subset[OF h] Mi_disj j assms by(auto simp: disjoint_family_on_def) thus "openin Mj.mtopology (S \ Mi j)" by(cases "i = j") (use openin_subset[OF h] h in auto) qed(use openin_subset[OF h] assms in auto) qed(use openin_subset[OF h] assms in auto) qed qed lemma limitin_Mi_limitin_M: assumes "i \ I" "limitin (Metric_space.mtopology (Mi i) (di i)) xn x sequentially" shows "limitin Sum_metric.mtopology xn x sequentially" proof - interpret m: Metric_space "Mi i" "di i" by(simp add: assms Md_metric) { fix e :: real assume "e > 0" then obtain N where "\n. n \ N \ xn n \ m.mball x e" using assms(2) m.commute m.limit_metric_sequentially by fastforce hence "\N. \n\N. xn n \ Sum_metric.mball x e" using ball_le_sum_dist_ball[OF assms(1),of x e] by(fastforce intro!: exI[where x=N]) } thus ?thesis by (metis Sum_metric.commute Sum_metric.in_mball Sum_metric.limit_metric_sequentially UN_I m.limitin_mspace assms) qed lemma limitin_M_limitin_Mi: assumes "limitin Sum_metric.mtopology xn x sequentially" shows "\i\I. limitin (Metric_space.mtopology (Mi i) (di i)) xn x sequentially" proof - obtain i where i: "i \ I" "x \ Mi i" using assms by (meson Sum_metric.limitin_mspace UN_E) then interpret m: Metric_space "Mi i" "di i" by(simp add: Md_metric) obtain N where N: "\n. n \ N \ sum_dist (xn n) x < 1" "\n. n \ N \ (xn n) \ M" using assms by (metis d_bounded i(2) m.mdist_zero Sum_metric.limit_metric_sequentially) hence xni: "n \ N \ xn n \ Mi i" for n using assms by(auto intro!: sum_dist_if_less1[OF i,of "xn n"] simp: Sum_metric.commute) show ?thesis proof(safe intro!: bexI[where x=i] i) show "limitin m.mtopology xn x sequentially" unfolding m.limit_metric_sequentially proof safe fix e :: real assume e: "0 < e" then obtain N' where N': "\n. n \ N' \ sum_dist (xn n) x < e" using assms by (meson Sum_metric.limit_metric_sequentially) hence "n \ max N N' \ di i (xn n) x < e" for n using sum_dist_simps(1)[OF i(1) xni[of n] i(2),symmetric] by auto thus "\N. \n\N. xn n \ Mi i \ di i (xn n) x < e" using xni by(auto intro!: exI[where x="max N N'"]) qed fact qed qed lemma MCauchy_Mi_MCauchy_M: assumes "i \ I" "Metric_space.MCauchy (Mi i) (di i) xn" shows "Sum_metric.MCauchy xn" proof - interpret m: Metric_space "Mi i" "di i" by(simp add: assms Md_metric) have [simp]:"sum_dist (xn n) (xn m) = di i (xn n) (xn m)" for n m using assms sum_dist_simps(1)[OF assms(1),of "xn n" "xn m"] by(auto simp: m.MCauchy_def) show ?thesis using assms by(auto simp: m.MCauchy_def Sum_metric.MCauchy_def) qed lemma MCauchy_M_MCauchy_Mi: assumes "Sum_metric.MCauchy xn" shows "\m. \i\I. Metric_space.MCauchy (Mi i) (di i) (\n. xn (n + m))" proof - obtain N where N: "\n m. n \ N \ m \ N \ sum_dist (xn n) (xn m) < 1" using assms by(fastforce simp: Sum_metric.MCauchy_def) obtain i where i: "i \ I" "xn N \ Mi i" by (metis assms Sum_metric.MCauchy_def UNIV_I UN_E image_eqI subsetD) then have xn:"\n. n \ N \ xn n \ Mi i" by (metis N Sum_metric.MCauchy_def assms dual_order.refl range_subsetD sum_dist_if_less1) interpret m: Metric_space "Mi i" "di i" using i Md_metric by auto show ?thesis proof(safe intro!: exI[where x=N] bexI[where x=i]) show "m.MCauchy (\n. xn (n + N))" unfolding m.MCauchy_def proof safe show 1: "\n. xn (n + N) \ Mi i" by(auto intro!: xn) fix e :: real assume "0 < e" then obtain N' where N': "\n m. n \ N' \ m \ N' \ sum_dist (xn n) (xn m) < e" using Sum_metric.MCauchy_def assms by blast thus "\N'. \n n'. N' \ n \ N' \ n' \ di i (xn (n + N)) (xn (n' + N)) < e" by(auto intro!: exI[where x="N'"] simp: sum_dist_simps(1)[OF i(1) xn xn,symmetric]) qed qed fact qed lemma separable_Mi_separable_M: assumes "countable I" "\i. i \ I \ separable_space (Metric_space.mtopology (Mi i) (di i))" shows "separable_space Sum_metric.mtopology" proof - obtain Ui where Ui: "\i. i \ I \ countable (Ui i)" "\i. i \ I \ dense_in (Metric_space.mtopology (Mi i) (di i)) (Ui i)" using assms by(simp only: separable_space_def2) metis define U where "U \ \i\I. Ui i" show "separable_space Sum_metric.mtopology" unfolding separable_space_def2 proof(safe intro!: exI[where x=U]) show "countable U" using Ui(1) assms by(auto simp: U_def) next show "Sum_metric.mdense U" unfolding Sum_metric.mdense_def U_def proof safe fix i x assume "i \ I" "x \ Ui i" then show "x \ M" using Ui(2) by(auto intro!: bexI[where x=i] simp: Md_metric Metric_space.mdense_def2) next fix i x e assume h:"i \ I" "x \ Mi i" "(0 :: real) < e" "Sum_metric.mball x e \ \ (Ui ` I) = {}" then interpret m: Metric_space "Mi i" "di i" by(simp add: Md_metric) have "m.mball x e \ Ui i \ {}" using Ui(2)[OF h(1)] h by(auto simp: m.mdense_def) hence "m.mball x e \ \ (Ui ` I) \ {}" using h(1) by blast thus False using ball_le_sum_dist_ball[OF \i \ I\,of x e] h(4) by blast qed qed qed lemma separable_M_separable_Mi: assumes "separable_space Sum_metric.mtopology" "\i. i \ I" shows "separable_space (Metric_space.mtopology (Mi i) (di i))" using subtopology_mtopology_Mi[OF assms(2)] separable_space_open_subset[OF assms(1) openin_mtopology_Mi[OF assms(2)]] by simp lemma mcomplete_Mi_mcomplete_M: assumes "\i. i \ I \ Metric_space.mcomplete (Mi i) (di i)" shows Sum_metric.mcomplete unfolding Sum_metric.mcomplete_def proof safe fix xn assume "Sum_metric.MCauchy xn" from MCauchy_M_MCauchy_Mi[OF this] obtain m i where mi: "i \ I" "Metric_space.MCauchy (Mi i) (di i) (\n. xn (n + m))" by metis then interpret m: Metric_space "Mi i" "di i" by(simp add: Md_metric) from assms[OF mi(1)] mi(2) obtain x where x: "limitin m.mtopology (\n. xn (n + m)) x sequentially" by(auto simp: m.mcomplete_def) from limitin_Mi_limitin_M[OF mi(1) limitin_sequentially_offset_rev[OF this]] show "\x. limitin Sum_metric.mtopology xn x sequentially" by auto qed lemma mcomplete_M_mcomplete_Mi: assumes Sum_metric.mcomplete "i \ I" shows "Metric_space.mcomplete (Mi i) (di i)" proof - interpret Mi: Metric_space "Mi i" "di i" using assms(2) Md_metric by fastforce show ?thesis unfolding Mi.mcomplete_def proof safe fix xn assume xn:"Mi.MCauchy xn" from MCauchy_Mi_MCauchy_M[OF assms(2) this] assms(1) obtain x where "limitin Sum_metric.mtopology xn x sequentially" by(auto simp: Sum_metric.mcomplete_def) from limitin_M_limitin_Mi[OF this] obtain j where j:"j \ I" "limitin (Metric_space.mtopology (Mi j) (di j)) xn x sequentially" by auto have "j = i" proof - obtain N where "\n. n \ N \ xn n \ Mi j" by (metis Md_metric Metric_space.limitin_metric_dist_null eventually_sequentially j) hence "xn N \ Mi i \ Mi j" using xn by(auto simp: Mi.MCauchy_def) with Mi_disj j(1) assms(2) show ?thesis by(fastforce simp: disjoint_family_on_def) qed thus "\x. limitin Mi.mtopology xn x sequentially" using j(2) by(auto intro!: exI[where x=x]) qed qed end lemma sum_metricI: fixes Si assumes "disjoint_family_on Si I" and "\i x y. i \ I \ 0 \ di i x y" and "\i x y. di i x y < 1" and "\i. i \ I \ Metric_space (Si i) (di i)" shows "Sum_metric I Si di" using assms by (metis Metric_space.nonneg Sum_metric_def) (* TDODO?: Sum metric space for Abstract type *) 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,292 +1,342 @@ \<^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 begin paragraph \Summary\ text \Basic functions on binary relations.\ -definition "rel_comp R S x y \ \z. R x z \ S z y" - -bundle rel_comp_syntax begin notation rel_comp (infixl "\\" 55) end -bundle no_rel_comp_syntax begin no_notation rel_comp (infixl "\\" 55) end -unbundle rel_comp_syntax - -lemma rel_compI [intro]: - assumes "R x y" - and "S y z" - shows "(R \\ S) x z" - using assms unfolding rel_comp_def by blast - -lemma rel_compE [elim]: - assumes "(R \\ S) x y" - obtains z where "R x z" "S z y" - using assms unfolding rel_comp_def by blast - -lemma rel_comp_assoc: "R \\ (S \\ T) = (R \\ S) \\ T" - by (intro ext) blast - -definition rel_inv :: "('a \ 'b \ bool) \ 'b \ 'a \ bool" - where "rel_inv R x y \ R y x" - -bundle rel_inv_syntax begin notation rel_inv ("(_\)" [1000]) end -bundle no_rel_inv_syntax begin no_notation rel_inv ("(_\)" [1000]) end -unbundle rel_inv_syntax - -lemma rel_invI [intro]: - assumes "R x y" - shows "R\ y x" - using assms unfolding rel_inv_def . - -lemma rel_invD [dest]: - assumes "R\ x y" - shows "R y x" - using assms unfolding rel_inv_def . - -lemma rel_inv_iff_rel [simp]: "R\ x y \ R y x" - by blast - -lemma rel_inv_comp_eq [simp]: "(R \\ S)\ = S\ \\ R\" - by (intro ext) blast - -lemma rel_inv_inv_eq_self [simp]: "R\\ = R" - by blast - -lemma rel_inv_eq_iff_eq [iff]: "R\ = S\ \ R = S" - by (blast dest: fun_cong) +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 -lemma in_dom_if_in_dom_rel_comp: - assumes "in_dom (R \\ S) x" - shows "in_dom R x" - using assms by blast - definition "in_codom R y \ \x. R x y" lemma in_codomI [intro]: assumes "R x y" shows "in_codom R y" using assms unfolding in_codom_def by blast lemma in_codomE [elim]: assumes "in_codom R y" obtains x where "R x y" using assms unfolding in_codom_def by blast -lemma in_codom_if_in_codom_rel_comp: - assumes "in_codom (R \\ S) y" - shows "in_codom S y" - using assms by blast - -lemma in_codom_rel_inv_eq_in_dom [simp]: "in_codom (R\) = in_dom R" - by (intro ext) blast - -lemma in_dom_rel_inv_eq_in_codom [simp]: "in_dom (R\) = in_codom R" - by (intro ext) blast - definition "in_field R x \ in_dom R x \ in_codom R x" lemma in_field_if_in_dom: assumes "in_dom R x" shows "in_field R x" unfolding in_field_def using assms by blast lemma in_field_if_in_codom: assumes "in_codom R x" shows "in_field R x" unfolding in_field_def using assms by blast lemma in_fieldE [elim]: assumes "in_field R x" obtains (in_dom) x' where "R x x'" | (in_codom) x' where "R x' x" using assms unfolding in_field_def by blast lemma in_fieldE': assumes "in_field R x" obtains (in_dom) "in_dom R x" | (in_codom) "in_codom R x" using assms by blast lemma in_fieldI [intro]: assumes "R x y" shows "in_field R x" "in_field R y" using assms by (auto intro: in_field_if_in_dom in_field_if_in_codom) lemma in_field_iff_in_dom_or_in_codom: - "in_field L x \ in_dom L x \ in_codom L x" + "in_field R x \ in_dom R x \ in_codom R x" by blast -lemma in_field_rel_inv_eq [simp]: "in_field R\ = in_field R" - by (intro ext) auto +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') + + +subsubsection \Composition\ + +definition "rel_comp R S x y \ \z. R x z \ S z y" + +bundle rel_comp_syntax begin notation rel_comp (infixl "\\" 55) end +bundle no_rel_comp_syntax begin no_notation rel_comp (infixl "\\" 55) end +unbundle rel_comp_syntax + +lemma rel_compI [intro]: + assumes "R x y" + and "S y z" + shows "(R \\ S) x z" + using assms unfolding rel_comp_def by blast + +lemma rel_compE [elim]: + assumes "(R \\ S) x y" + obtains z where "R x z" "S z y" + using assms unfolding rel_comp_def by blast + +lemma rel_comp_assoc: "R \\ (S \\ T) = (R \\ S) \\ T" + by (intro ext) blast + +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 -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') + +subsubsection \Inverse\ + +definition rel_inv :: "('a \ 'b \ bool) \ 'b \ 'a \ bool" + where "rel_inv R x y \ R y x" + +bundle rel_inv_syntax begin notation rel_inv ("(_\)" [1000]) end +bundle no_rel_inv_syntax begin no_notation rel_inv ("(_\)" [1000]) end +unbundle rel_inv_syntax + +lemma rel_invI [intro]: + assumes "R x y" + shows "R\ y x" + using assms unfolding rel_inv_def . + +lemma rel_invD [dest]: + assumes "R\ x y" + shows "R y x" + using assms unfolding rel_inv_def . + +lemma rel_inv_iff_rel [simp]: "R\ x y \ R y x" + by blast + +lemma 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" + by (intro ext) blast + +lemma in_field_rel_inv_eq [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) + +subsubsection \Restrictions\ definition "rel_if B R x y \ B \ R x y" bundle rel_if_syntax begin notation (output) rel_if (infixl "\" 50) end bundle no_rel_if_syntax begin no_notation (output) rel_if (infixl "\" 50) end unbundle rel_if_syntax lemma rel_if_eq_rel_if_pred [simp]: assumes "B" shows "(rel_if B R) = R" unfolding rel_if_def using assms by blast lemma rel_if_eq_top_if_not_pred [simp]: assumes "\B" shows "(rel_if B R) = (\_ _. True)" unfolding rel_if_def using assms by blast lemma rel_if_if_impI [intro]: assumes "B \ R x y" shows "(rel_if B R) x y" unfolding rel_if_def using assms by blast lemma rel_ifE [elim]: assumes "(rel_if B R) x y" obtains "\B" | "B" "R x y" using assms unfolding rel_if_def by blast lemma rel_ifD: assumes "(rel_if B R) x y" and "B" shows "R x y" using assms by blast -consts restrict_left :: "'a \ 'b \ 'a" -consts restrict_right :: "'a \ 'b \ 'a" +consts rel_restrict_left :: "'a \ 'b \ 'a" +consts rel_restrict_right :: "'a \ 'b \ 'a" overloading - bin_rel_restrict_left_pred \ - "restrict_left :: ('a \ 'b \ bool) \ ('a \ bool) \ 'a \ 'b \ bool" - bin_rel_restrict_right_pred \ - "restrict_right :: ('a \ 'b \ bool) \ ('b \ bool) \ 'a \ 'b \ bool" + 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 "bin_rel_restrict_left_pred R P x y \ P x \ R x y" - definition "bin_rel_restrict_right_pred R P x y \ P y \ R x y" + 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 restrict_syntax +bundle rel_restrict_syntax begin -notation restrict_left ("(_)\(\<^bsub>_\<^esub>)" [1000]) -notation restrict_right ("(_)\(\<^bsub>_\<^esub>)" [1000]) +notation rel_restrict_left ("(_)\(\<^bsub>_\<^esub>)" [1000]) +notation rel_restrict_right ("(_)\(\<^bsub>_\<^esub>)" [1000]) end -bundle no_restrict_syntax +bundle no_rel_restrict_syntax begin -no_notation restrict_left ("(_)\(\<^bsub>_\<^esub>)" [1000]) -no_notation restrict_right ("(_)\(\<^bsub>_\<^esub>)" [1000]) +no_notation rel_restrict_left ("(_)\(\<^bsub>_\<^esub>)" [1000]) +no_notation rel_restrict_right ("(_)\(\<^bsub>_\<^esub>)" [1000]) end -unbundle restrict_syntax +unbundle rel_restrict_syntax -lemma bin_rel_restrict_leftI [intro]: +lemma rel_restrict_leftI [intro]: assumes "R x y" and "P x" shows "R\\<^bsub>P\<^esub> x y" - using assms unfolding bin_rel_restrict_left_pred_def by blast + using assms unfolding rel_restrict_left_pred_def by blast -lemma bin_rel_restrict_leftE [elim]: +lemma rel_restrict_leftE [elim]: assumes "R\\<^bsub>P\<^esub> x y" obtains "P x" "R x y" - using assms unfolding bin_rel_restrict_left_pred_def by blast + using assms unfolding rel_restrict_left_pred_def by blast -lemma bin_rel_restrict_rightI [intro]: +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" + by (intro ext rel_restrict_leftI) auto + +lemma rel_restrict_left_top_eq_uhint [uhint]: + assumes "P \ (\ ::'a \ bool)" + shows "(R :: 'a \ 'b \ bool)\\<^bsub>P\<^esub> \ R" + using assms by simp + +lemma rel_restrict_left_le_self: "(R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'a \ bool)\<^esub> \ R" + by (auto intro: predicate2I dest: predicate2D) + +lemma rel_restrict_rightI [intro]: assumes "R x y" and "P y" shows "R\\<^bsub>P\<^esub> x y" - using assms unfolding bin_rel_restrict_right_pred_def by blast + using assms unfolding rel_restrict_right_pred_def by blast -lemma bin_rel_restrict_rightE [elim]: +lemma rel_restrict_rightE [elim]: assumes "R\\<^bsub>P\<^esub> x y" obtains "P y" "R x y" - using assms unfolding bin_rel_restrict_right_pred_def by blast + 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_le_self: "(R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'b \ bool)\<^esub> \ R" + by (auto intro: predicate2I dest: predicate2D) context fixes R :: "'a \ 'b \ bool" begin -lemma bin_rel_restrict_right_eq: "R\\<^bsub>P :: 'b \ bool\<^esub> = ((R\)\\<^bsub>P\<^esub>)\" +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 bin_rel_restrict_right_iff_restrict_left: "R\\<^bsub>P :: 'b \ bool\<^esub> x y \ (R\)\\<^bsub>P\<^esub> y x" - unfolding bin_rel_restrict_right_eq by simp +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_bin_rel_restrict_left_inv_bin_rel_restrict_left_eq: +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 bin_rel_restrict_leftI rel_invI) auto + by (intro ext iffI rel_restrict_leftI rel_invI) auto -lemma bin_rel_restrict_left_right_eq_restrict_right_left: +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 bin_rel_restrict_right_eq - by (fact rel_inv_bin_rel_restrict_left_inv_bin_rel_restrict_left_eq) + unfolding rel_restrict_right_eq + by (fact rel_inv_rel_restrict_left_inv_rel_restrict_left_eq) -lemma in_dom_bin_rel_restrict_leftI [intro]: +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_bin_rel_restrict_left_if_in_dom: +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_bin_rel_restrict_leftE [elim]: +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_bin_rel_restrict_leftI [intro]: +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_bin_rel_restrict_leftE [elim]: +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_Lattice.thy b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Lattice.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Lattice.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Binary_Relations_Lattice.thy @@ -1,92 +1,92 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Lattice\ theory Binary_Relations_Lattice imports Binary_Relations_Order_Base HOL.Boolean_Algebras begin paragraph \Summary\ text \Basic results about the lattice structure on binary relations.\ lemma rel_infI [intro]: assumes "R x y" and "S x y" shows "(R \ S) x y" using assms by (rule inf2I) lemma rel_infE [elim]: assumes "(R \ S) x y" obtains "R x y" "S x y" using assms by (rule inf2E) lemma rel_infD: assumes "(R \ S) x y" shows "R x y" and "S x y" using assms by auto lemma in_dom_rel_infI [intro]: assumes "R x y" and "S x y" shows "in_dom (R \ S) x" using assms by blast lemma in_dom_rel_infE [elim]: assumes "in_dom (R \ S) x" obtains y where "R x y" "S x y" using assms by blast lemma in_codom_rel_infI [intro]: assumes "R x y" and "S x y" shows "in_codom (R \ S) y" using assms by blast lemma in_codom_rel_infE [elim]: assumes "in_codom (R \ S) y" obtains x where "R x y" "S x y" using assms by blast lemma in_field_eq_in_dom_sup_in_codom: "in_field L = (in_dom L \ in_codom L)" by (intro ext) (simp add: in_field_iff_in_dom_or_in_codom) -lemma in_dom_bin_rel_restrict_left_eq [simp]: "in_dom R\\<^bsub>P\<^esub> = (in_dom R \ P)" +lemma in_dom_rel_restrict_left_eq [simp]: "in_dom R\\<^bsub>P\<^esub> = (in_dom R \ P)" by (intro ext) auto -lemma in_codom_bin_rel_restrict_left_eq [simp]: "in_codom R\\<^bsub>P\<^esub> = (in_codom R \ P)" +lemma in_codom_rel_restrict_left_eq [simp]: "in_codom R\\<^bsub>P\<^esub> = (in_codom R \ P)" by (intro ext) auto -lemma bin_rel_restrict_left_restrict_left_eq [simp]: +lemma rel_restrict_left_restrict_left_eq [simp]: fixes R :: "'a \ 'b \ bool" and P Q :: "'a \ bool" shows "R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub>" - by (intro ext iffI bin_rel_restrict_leftI) auto + by (intro ext iffI rel_restrict_leftI) auto -lemma bin_rel_restrict_left_restrict_right_eq [simp]: +lemma rel_restrict_left_restrict_right_eq [simp]: fixes R :: "'a \ 'b \ bool" and P :: "'a \ bool" and Q :: "'b \ bool" shows "R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub>" - by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI) auto + by (intro ext iffI rel_restrict_leftI rel_restrict_rightI) auto -lemma bin_rel_restrict_right_restrict_left_eq [simp]: +lemma rel_restrict_right_restrict_left_eq [simp]: fixes R :: "'a \ 'b \ bool" and P :: "'b \ bool" and Q :: "'a \ bool" shows "R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub>" - by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI) auto + by (intro ext iffI rel_restrict_leftI rel_restrict_rightI) auto -lemma bin_rel_restrict_right_restrict_right_eq [simp]: +lemma rel_restrict_right_restrict_right_eq [simp]: fixes R :: "'a \ 'b \ bool" and P Q :: "'b \ bool" shows "R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub>" by (intro ext iffI) auto -lemma bin_rel_restrict_left_sup_eq [simp]: +lemma rel_restrict_left_sup_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>((P :: 'a \ bool) \ Q)\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub> " - by (intro antisym le_relI) (auto elim!: bin_rel_restrict_leftE) + by (intro antisym le_relI) (auto elim!: rel_restrict_leftE) -lemma bin_rel_restrict_left_inf_eq [simp]: +lemma rel_restrict_left_inf_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>((P :: 'a \ bool) \ Q)\<^esub> = R\\<^bsub>P\<^esub> \ R\\<^bsub>Q\<^esub> " - by (intro antisym le_relI) (auto elim!: bin_rel_restrict_leftE) + by (intro antisym le_relI) (auto elim!: rel_restrict_leftE) lemma inf_rel_bimap_and_eq_restrict_left_restrict_right: "R \ (rel_bimap P Q (\)) = R\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub>" by (intro ext) auto end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order.thy b/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order.thy @@ -1,81 +1,81 @@ \<^marker>\creator "Kevin Kappelmann"\ theory Binary_Relations_Order imports Binary_Relations_Order_Base Binary_Relations_Reflexive Binary_Relations_Symmetric Binary_Relations_Transitive begin paragraph \Summary\ text \Basic results about the order on binary relations.\ lemma in_dom_if_rel_if_rel_comp_le: assumes "(R \\ S) \ (S \\ R)" and "R x y" "S y z" shows "in_dom S x" using assms by (blast intro: in_dom_if_in_dom_rel_comp) lemma in_codom_if_rel_if_rel_comp_le: assumes "(R \\ S) \ (S \\ R)" and "R x y" "S y z" shows "in_codom R z" using assms by (blast intro: in_codom_if_in_codom_rel_comp) lemma rel_comp_le_rel_inv_if_rel_comp_le_if_symmetric: assumes symms: "symmetric R1" "symmetric R2" and le: "(R1 \\ R2) \ R3" shows "(R2 \\ R1) \ R3\" proof - from le have "(R1 \\ R2)\ \ R3\" by blast with symms show ?thesis by simp qed lemma rel_inv_le_rel_comp_if_le_rel_comp_if_symmetric: assumes symms: "symmetric R1" "symmetric R2" and le: "R3 \ (R1 \\ R2)" shows "R3\ \ (R2 \\ R1)" proof - from le have "R3\ \ (R1 \\ R2)\" by blast with symms show ?thesis by simp qed corollary rel_comp_le_rel_comp_if_rel_comp_le_rel_comp_if_symmetric: - assumes "symmetric R1" "symmetric R2" "symmetric R3" "symmetric R4" + assumes "symmetric (R1 :: 'a \ 'a \ bool)" "symmetric R2" "symmetric R3" "symmetric R4" and "(R1 \\ R2) \ (R3 \\ R4)" shows "(R2 \\ R1) \ (R4 \\ R3)" proof - from assms have "(R2 \\ R1) \ (R3 \\ R4)\" by (intro rel_comp_le_rel_inv_if_rel_comp_le_if_symmetric) with assms show ?thesis by simp qed corollary rel_comp_le_rel_comp_iff_if_symmetric: - assumes "symmetric R1" "symmetric R2" "symmetric R3" "symmetric R4" + assumes "symmetric (R1 :: 'a \ 'a \ bool)" "symmetric R2" "symmetric R3" "symmetric R4" shows "(R1 \\ R2) \ (R3 \\ R4) \ (R2 \\ R1) \ (R4 \\ R3)" using assms by (blast intro: rel_comp_le_rel_comp_if_rel_comp_le_rel_comp_if_symmetric) corollary eq_if_le_rel_if_symmetric: assumes "symmetric R" "symmetric S" and "(R \\ S) \ (S \\ R)" shows "(R \\ S) = (S \\ R)" using assms rel_comp_le_rel_comp_iff_if_symmetric[of R S] by (intro antisym) auto lemma rel_comp_le_rel_comp_if_le_rel_if_reflexive_on_in_codom_if_transitive: assumes trans: "transitive S" and refl_on: "reflexive_on (in_codom S) R" and le_rel: "R \ S" shows "R \\ S \ S \\ R" proof (rule le_relI) fix x1 x2 assume"(R \\ S) x1 x2" then obtain x3 where "R x1 x3" "S x3 x2" by blast then have "S x1 x3" using le_rel by blast with \S x3 x2\ have "S x1 x2" using trans by blast with refl_on have "R x2 x2" by blast then show "(S \\ R) x1 x2" using \S x1 x2\ by blast qed end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order_Base.thy b/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order_Base.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order_Base.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Order/Binary_Relations_Order_Base.thy @@ -1,42 +1,26 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Order\ theory Binary_Relations_Order_Base imports Binary_Relation_Functions HOL.Orderings begin lemma le_relI [intro]: assumes "\x y. R x y \ S x y" shows "R \ S" using assms by (rule predicate2I) lemma le_relD [dest]: assumes "R \ S" and "R x y" shows "S x y" using assms by (rule predicate2D) lemma le_relE: assumes "R \ S" and "R x y" obtains "S x y" using assms by blast -lemma rel_inv_le_rel_inv_iff [iff]: "R\ \ S\ \ R \ S" - by blast - -lemma bin_rel_restrict_left_le_self: "(R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'a \ bool)\<^esub> \ R" - by blast - -lemma bin_rel_restrict_right_le_self: "(R :: 'a \ 'b \ bool)\\<^bsub>(P :: 'b \ bool)\<^esub> \ R" - by blast - -lemma bin_rel_restrict_left_top_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>(\ :: 'a \ bool)\<^esub> = R" - by (intro ext) auto - -lemma bin_rel_restrict_right_top_eq [simp]: "(R :: 'a \ 'b \ bool)\\<^bsub>(\ :: 'b \ bool)\<^esub> = R" - by (intro ext) auto - - end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_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,17 +1,19 @@ \<^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_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,65 +1,75 @@ \<^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 \ 'b \ bool) \ bool" +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" end lemma antisymmetric_onI [intro]: assumes "\x y. P x \ P y \ R x y \ R y x \ x = y" shows "antisymmetric_on P R" unfolding antisymmetric_on_pred_def using assms by blast lemma antisymmetric_onD: assumes "antisymmetric_on P R" and "P x" "P y" and "R x y" "R y x" shows "x = y" using assms unfolding antisymmetric_on_pred_def by blast -definition "(antisymmetric :: ('a \ _) \ _) \ antisymmetric_on (\ :: 'a \ bool)" +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 \ _) \ _) = antisymmetric_on (\ :: 'a \ bool)" + "(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" - unfolding antisymmetric_eq_antisymmetric_on using assms - by (intro antisymmetric_onI) + using assms by (urule antisymmetric_onI) lemma antisymmetricD: assumes "antisymmetric R" and "R x y" "R y x" shows "x = y" - using assms unfolding antisymmetric_eq_antisymmetric_on - by (auto dest: antisymmetric_onD) + using assms by (urule (d) antisymmetric_onD where chained = insert) simp_all lemma antisymmetric_on_if_antisymmetric: - fixes P :: "'a \ bool" and R :: "'a \ _" + 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 [simp]: +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_Bi_Total.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Bi_Total.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Bi_Total.thy @@ -0,0 +1,45 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsubsection \Bi-Total\ +theory Binary_Relations_Bi_Total + imports + Binary_Relations_Left_Total + Binary_Relations_Surjective +begin + +definition "bi_total_on P Q \ left_total_on P \ rel_surjective_at Q" + +lemma bi_total_onI [intro]: + assumes "left_total_on P R" + and "rel_surjective_at Q R" + shows "bi_total_on P Q R" + unfolding bi_total_on_def using assms by auto + +lemma bi_total_onE [elim]: + assumes "bi_total_on P Q R" + obtains "left_total_on P R" "rel_surjective_at Q R" + using assms unfolding bi_total_on_def by auto + +definition "bi_total \ bi_total_on (\ :: 'a \ bool) (\ :: 'b \ bool) :: ('a \ 'b \ bool) \ bool" + +lemma bi_total_eq_bi_total_on: + "(bi_total :: ('a \ 'b \ bool) \ bool) = bi_total_on (\ :: 'a \ bool) (\ :: 'b \ bool)" + unfolding bi_total_def .. + +lemma bi_total_eq_bi_total_on_uhint [uhint]: + assumes "P \ (\ :: 'a \ bool)" + and "Q \ (\ :: 'b \ bool)" + shows "(bi_total :: ('a \ 'b \ bool) \ bool) \ bi_total_on P Q" + using assms by (simp add: bi_total_eq_bi_total_on) + +lemma bi_totalI [intro]: + assumes "left_total R" + and "rel_surjective R" + shows "bi_total R" + using assms by (urule bi_total_onI) + +lemma bi_totalE [elim]: + assumes "bi_total R" + obtains "left_total R" "rel_surjective R" + using assms by (urule (e) bi_total_onE) + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Bi_Unique.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Bi_Unique.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Bi_Unique.thy @@ -0,0 +1,53 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsubsection \Bi-Unique\ +theory Binary_Relations_Bi_Unique + imports + Binary_Relations_Injective + Binary_Relations_Right_Unique +begin + +definition "bi_unique_on \ right_unique_on \ rel_injective_on" + +lemma bi_unique_onI [intro]: + assumes "right_unique_on P R" + and "rel_injective_on P R" + shows "bi_unique_on P R" + unfolding bi_unique_on_def using assms by auto + +lemma bi_unique_onE [elim]: + assumes "bi_unique_on P R" + obtains "right_unique_on P R" "rel_injective_on P R" + using assms unfolding bi_unique_on_def by auto + +lemma bi_unique_on_rel_inv_if_Fun_Rel_iff_if_bi_unique_on: + assumes "bi_unique_on P R" + and "(R \ (\)) P Q" + shows "bi_unique_on Q R\" + using assms by (intro bi_unique_onI + rel_injective_on_if_Fun_Rel_imp_if_rel_injective_at + right_unique_on_if_Fun_Rel_imp_if_right_unique_at) + (auto 0 3) + +definition "bi_unique \ bi_unique_on (\ :: 'a \ bool) :: ('a \ 'b \ bool) \ bool" + +lemma bi_unique_eq_bi_unique_on: + "bi_unique = (bi_unique_on (\ :: 'a \ bool) :: ('a \ 'b \ bool) \ bool)" + unfolding bi_unique_def .. + +lemma bi_unique_eq_bi_unique_on_uhint [uhint]: + assumes "P \ (\ :: 'a \ bool)" + shows "bi_unique \ (bi_unique_on P :: ('a \ 'b \ bool) \ bool)" + using assms by (simp add: bi_unique_eq_bi_unique_on) + +lemma bi_uniqueI [intro]: + assumes "right_unique R" + and "rel_injective R" + shows "bi_unique R" + using assms by (urule bi_unique_onI) + +lemma bi_uniqueE [elim]: + assumes "bi_unique R" + obtains "right_unique R" "rel_injective R" + using assms by (urule (e) bi_unique_onE) + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Injective.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Injective.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Injective.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Injective.thy @@ -1,106 +1,152 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Injective\ theory Binary_Relations_Injective imports - Binary_Relation_Functions - HOL_Syntax_Bundles_Lattices + Reverse_Implies + Functions_Monotone ML_Unification.ML_Unification_HOL_Setup + ML_Unification.Unify_Resolve_Tactics begin -consts rel_injective_on :: "'a \ ('b \ 'c \ bool) \ bool" +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'" end lemma rel_injective_onI [intro]: assumes "\x x' y. P x \ P x' \ R x y \ R x' y \ x = x'" shows "rel_injective_on P R" unfolding rel_injective_on_pred_def using assms by blast lemma rel_injective_onD: assumes "rel_injective_on P R" and "P x" "P x'" and "R x y" "R x' y" shows "x = x'" using assms unfolding rel_injective_on_pred_def by blast -consts rel_injective_at :: "'a \ ('b \ 'c \ bool) \ bool" +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) + + +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" + and "(R \ (\)) P Q" + shows "rel_injective_on P R" + using assms by (intro rel_injective_onI) (auto dest: rel_injective_atD) -definition "(rel_injective :: ('a \ _) \ _) \ rel_injective_on (\ :: 'a \ bool)" +lemma rel_injective_at_if_Fun_Rel_rev_imp_if_rel_injective_on: + assumes "rel_injective_on P R" + 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 \ _) \ _) = rel_injective_on (\ :: 'a \ bool)" + "(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" - unfolding rel_injective_eq_rel_injective_on using assms by blast + 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 unfolding rel_injective_eq_rel_injective_on - by (auto dest: rel_injective_onD) + using assms by (urule (d) rel_injective_onD where chained = insert) simp_all lemma rel_injective_eq_rel_injective_at: - "rel_injective (R :: 'a \ 'b \ bool) = rel_injective_at (\ :: 'b \ bool) R" - by (intro iffI rel_injectiveI) (auto dest: rel_injective_atD rel_injectiveD) + "(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 \ _" + 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_Irreflexive.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Irreflexive.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Irreflexive.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Irreflexive.thy @@ -1,51 +1,63 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Irreflexive\ theory Binary_Relations_Irreflexive imports - Binary_Relation_Functions HOL_Syntax_Bundles_Lattices + ML_Unification.ML_Unification_HOL_Setup + ML_Unification.Unify_Resolve_Tactics begin -consts irreflexive_on :: "'a \ ('b \ 'b \ bool) \ bool" +consts irreflexive_on :: "'a \ 'b \ bool" overloading irreflexive_on_pred \ "irreflexive_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool" begin definition "irreflexive_on_pred P R \ \x. P x \ \(R x x)" end lemma irreflexive_onI [intro]: assumes "\x. P x \ \(R x x)" shows "irreflexive_on P R" using assms unfolding irreflexive_on_pred_def by blast lemma irreflexive_onD [dest]: assumes "irreflexive_on P R" and "P x" shows "\(R x x)" using assms unfolding irreflexive_on_pred_def by blast -definition "(irreflexive :: ('a \ _) \ _) \ irreflexive_on (\ :: 'a \ bool)" +consts irreflexive :: "'a \ bool" + +overloading + irreflexive \ "irreflexive :: ('a \ 'a \ bool) \ bool" +begin + definition "(irreflexive :: ('a \ 'a \ bool) \ bool) \ irreflexive_on (\ :: 'a \ bool)" +end lemma irreflexive_eq_irreflexive_on: - "(irreflexive :: ('a \ _) \ _) = irreflexive_on (\ :: 'a \ bool)" + "(irreflexive :: ('a \ 'a \ bool) \ _) = irreflexive_on (\ :: 'a \ bool)" unfolding irreflexive_def .. +lemma irreflexive_eq_irreflexive_on_uhint [uhint]: + assumes "P \ (\ :: 'a \ bool)" + shows "irreflexive :: ('a \ 'a \ bool) \ _ \ irreflexive_on P" + using assms by (simp add: irreflexive_eq_irreflexive_on) + lemma irreflexiveI [intro]: assumes "\x. \(R x x)" shows "irreflexive R" - unfolding irreflexive_eq_irreflexive_on using assms by (intro irreflexive_onI) + using assms by (urule irreflexive_onI) lemma irreflexiveD: assumes "irreflexive R" shows "\(R x x)" - using assms unfolding irreflexive_eq_irreflexive_on by auto + using assms by (urule (d) irreflexive_onD where chained = insert) simp lemma irreflexive_on_if_irreflexive: - fixes P :: "'a \ bool" and R :: "'a \ _" + fixes P :: "'a \ bool" and R :: "'a \ 'a \ bool" assumes "irreflexive R" shows "irreflexive_on P R" using assms by (intro irreflexive_onI) (blast dest: irreflexiveD) end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Left_Total.thy b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Left_Total.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Left_Total.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Properties/Binary_Relations_Left_Total.thy @@ -1,61 +1,98 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Left Total\ theory Binary_Relations_Left_Total imports - Binary_Relation_Functions - HOL_Syntax_Bundles_Lattices + Functions_Monotone + ML_Unification.ML_Unification_HOL_Setup + ML_Unification.Unify_Resolve_Tactics begin -consts left_total_on :: "'a \ ('b \ 'c \ bool) \ bool" +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" end lemma left_total_onI [intro]: assumes "\x. P x \ in_dom R x" shows "left_total_on P R" unfolding left_total_on_pred_def using assms by blast lemma left_total_onE [elim]: assumes "left_total_on P R" and "P x" obtains y where "R x y" using assms unfolding left_total_on_pred_def by blast lemma in_dom_if_left_total_on: assumes "left_total_on P R" - and "P x" - shows "in_dom R x" - using assms by blast + shows "P \ in_dom R" + using assms by force -definition "(left_total :: ('a \ _) \ _) \ left_total_on (\ :: 'a \ bool)" +lemma left_total_on_if_le_in_dom: + assumes "P \ in_dom R" + shows "left_total_on P R" + using assms by fastforce -lemma left_total_eq_left_total_on: "(left_total :: ('a \ _) \ _) = left_total_on (\ :: 'a \ bool)" +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 + +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 + +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 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 + +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" - unfolding left_total_eq_left_total_on using assms by (intro left_total_onI) + using assms by (urule left_total_onI) lemma left_totalE: assumes "left_total R" obtains y where "R x y" - using assms unfolding left_total_eq_left_total_on by (blast intro: top1I) + 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 \ _" + 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,104 +1,124 @@ \<^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 \ 'b \ bool) \ bool" +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 :: 'a \ bool) (R :: 'a \ _)" + "reflexive_on P R\ \ reflexive_on P R" by blast -lemma antimono_reflexive_on [iff]: - "antimono (\(P :: 'a \ bool). reflexive_on P (R :: 'a \ _))" - by (intro antimonoI) auto +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 lemma reflexive_on_if_le_pred_if_reflexive_on: - fixes P P' :: "'a \ bool" and R :: "'a \ _" + 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 \ _) \ _) ((P :: 'a \ bool) \ Q) + "(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) + (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 -definition "(reflexive :: ('a \ _) \ _) \ reflexive_on (\ :: 'a \ bool)" -lemma reflexive_eq_reflexive_on: "(reflexive :: ('a \ _) \ _) = reflexive_on (\ :: 'a \ bool)" +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" - unfolding reflexive_eq_reflexive_on using assms by (intro reflexive_onI) + using assms by (urule reflexive_onI) lemma reflexiveD: assumes "reflexive R" shows "R x x" - using assms unfolding reflexive_eq_reflexive_on by (blast intro: top1I) + using assms by (urule (d) reflexive_onD where chained = insert) simp lemma reflexive_on_if_reflexive: - fixes P :: "'a \ bool" and R :: "'a \ _" + 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\ \ reflexive R" + "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 simp + by auto paragraph \Instantiations\ lemma reflexive_eq: "reflexive (=)" by (rule reflexiveI) (rule refl) -lemma reflexive_top: "reflexive \" +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,136 +1,172 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Right Unique\ theory Binary_Relations_Right_Unique - imports - Binary_Relations_Injective - HOL_Syntax_Bundles_Lattices + imports Binary_Relations_Injective begin -consts right_unique_on :: "'a \ ('b \ 'c \ bool) \ bool" +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'" end lemma right_unique_onI [intro]: assumes "\x y y'. P x \ R x y \ R x y' \ y = y'" shows "right_unique_on P R" using assms unfolding right_unique_on_pred_def by blast lemma right_unique_onD: assumes "right_unique_on P R" and "P x" and "R x y" "R x y'" shows "y = y'" using assms unfolding right_unique_on_pred_def by blast -consts right_unique_at :: "'a \ ('b \ 'c \ bool) \ bool" +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) + +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) + +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" + and "(R \ (\)) P Q" + shows "right_unique_on P R" + using assms by (intro right_unique_onI) (auto dest: right_unique_atD) -definition "(right_unique :: ('a \ _) \ _) \ right_unique_on (\ :: 'a \ bool)" +lemma right_unique_at_if_Fun_Rel_rev_imp_if_right_unique_on: + assumes "right_unique_on P R" + 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 \ _) \ _) = right_unique_on (\ :: 'a \ bool)" + "(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" - unfolding right_unique_eq_right_unique_on using assms by blast + 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 unfolding right_unique_eq_right_unique_on - by (auto dest: right_unique_onD) + using assms by (urule (d) right_unique_onD where chained = insert) simp_all lemma right_unique_eq_right_unique_at: - "right_unique (R :: 'a \ 'b \ bool) = right_unique_at (\ :: 'b \ bool) R" - by (intro iffI right_uniqueI) (auto dest: right_unique_atD right_uniqueD) + "(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 \ _" + 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,79 +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 \ 'c \ bool) \ bool" +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" 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 -definition "(rel_surjective :: (_ \ 'a \ _) \ _) \ rel_surjective_at (\ :: 'a \ bool)" +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 + +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 :: (_ \ 'a \ _) \ _) = rel_surjective_at (\ :: 'a \ bool)" + "(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" - unfolding rel_surjective_eq_rel_surjective_at using assms by (intro rel_surjective_atI) + using assms by (urule rel_surjective_atI) lemma rel_surjectiveE: assumes "rel_surjective R" obtains x where "R x y" - using assms unfolding rel_surjective_eq_rel_surjective_at - by (blast intro: top1I) + 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" - unfolding rel_surjective_eq_rel_surjective_at left_total_eq_left_total_on - by simp + 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" - unfolding rel_surjective_eq_rel_surjective_at left_total_eq_left_total_on - by simp + by (urule left_total_on_rel_inv_iff_rel_surjective_at) lemma rel_surjective_at_if_surjective: - fixes P :: "'a \ bool" and R :: "_ \ 'a \ _" + 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,106 +1,117 @@ \<^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 \ 'b \ bool) \ bool" +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" 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 \ _)" + "symmetric_on P R\ \ symmetric_on (P :: 'a \ bool) (R :: 'a \ 'a \ bool)" by (blast dest: symmetric_onD) -lemma antimono_symmetric_on [iff]: - "antimono (\(P :: 'a \ bool). symmetric_on P (R :: 'a \ _))" +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 P' :: "'a \ bool" and R :: "'a \ _" + fixes P' :: "'a \ bool" and R :: "'a \ 'a \ bool" assumes "symmetric_on P R" and "P' \ P" shows "symmetric_on P' R" - using assms by (blast dest: symmetric_onD) + using assms antimono_symmetric_on by blast -definition "(symmetric :: ('a \ _) \ _) \ symmetric_on (\ :: 'a \ bool)" +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 \ _) \ _) = symmetric_on (\ :: 'a \ bool)" + "(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" - unfolding symmetric_eq_symmetric_on using assms by (intro symmetric_onI) + using assms by (urule symmetric_onI) lemma symmetricD: assumes "symmetric R" and "R x y" shows "R y x" - using assms unfolding symmetric_eq_symmetric_on by (auto dest: symmetric_onD) + using assms by (urule (d) symmetric_onD where chained = insert) simp_all lemma symmetric_on_if_symmetric: - fixes P :: "'a \ bool" and R :: "'a \ _" + 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" +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 [simp]: +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 \" +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,114 +1,115 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Transitive\ theory Binary_Relations_Transitive imports - Binary_Relation_Functions Functions_Monotone + ML_Unification.ML_Unification_HOL_Setup + ML_Unification.Unify_Resolve_Tactics begin -consts transitive_on :: "'a \ ('b \ 'b \ bool) \ bool" +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" 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 \ _)" + "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 [iff]: - "antimono (\(P :: 'a \ bool). transitive_on P (R :: 'a \ _))" +lemma antimono_transitive_on: "antimono (transitive_on :: ('a \ bool) \ ('a \ 'a \ bool) \ bool)" by (intro antimonoI) (auto dest: transitive_onD) -lemma transitive_on_if_le_pred_if_transitive_on: - fixes P P' :: "'a \ bool" and R :: "'a \ _" - assumes "transitive_on P R" - and "P' \ P" - shows "transitive_on P' R" - using assms by (auto dest: transitive_onD) +consts transitive :: "'a \ bool" -definition "(transitive :: ('a \ _) \ _) \ transitive_on (\ :: '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 \ _) \ _) = transitive_on (\ :: 'a \ bool)" + "(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" - unfolding transitive_eq_transitive_on using assms by (intro transitive_onI) + 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 unfolding transitive_eq_transitive_on - by (auto dest: transitive_onD) + using assms by (urule (d) transitive_onD where chained = insert) simp_all lemma transitive_on_if_transitive: - fixes P :: "'a \ bool" and R :: "'a \ _" + 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" - using assms unfolding transitive_eq_transitive_on - by (intro transitive_on_if_rel_comp_self_imp) blast + 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 [simp]: +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" +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 \" +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/Binary_Relations/Restricted_Equality.thy b/thys/Transport/HOL_Basics/Binary_Relations/Restricted_Equality.thy --- a/thys/Transport/HOL_Basics/Binary_Relations/Restricted_Equality.thy +++ b/thys/Transport/HOL_Basics/Binary_Relations/Restricted_Equality.thy @@ -1,68 +1,68 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Restricted Equality\ theory Restricted_Equality imports Binary_Relations_Order_Base Binary_Relation_Functions Equivalence_Relations Partial_Orders begin paragraph \Summary\ text \Introduces notations and theorems for restricted equalities. An equality @{term "(=)"} can be restricted to only apply to a subset of its elements. The restriction can be formulated, for example, by a predicate or a set.\ -bundle eq_restrict_syntax +bundle eq_rel_restrict_syntax begin syntax - "_eq_restrict_infix" :: "'a \ ('a \ bool) \ 'a \ bool" ("(_) =(\<^bsub>_\<^esub>) (_)" [51,51,51] 50) - "_eq_restrict" :: "('a \ bool) \ 'a \ bool" ("'(=(\<^bsub>_\<^esub>)')") + "_eq_restrict_infix" :: "'a \ 'b \ 'a \ bool" ("(_) =(\<^bsub>_\<^esub>) (_)" [51,51,51] 50) + "_eq_restrict" :: "'b \ 'a \ bool" ("'(=(\<^bsub>_\<^esub>)')") end -bundle no_eq_restrict_syntax +bundle no_eq_rel_restrict_syntax begin no_syntax - "_eq_restrict_infix" :: "'a \ ('a \ bool) \ 'a \ bool" ("(_) =(\<^bsub>_\<^esub>) (_)" [51,51,51] 50) - "_eq_restrict" :: "('a \ bool) \ 'a \ bool" ("'(=(\<^bsub>_\<^esub>)')") + "_eq_restrict_infix" :: "'a \ 'b \ 'a \ bool" ("(_) =(\<^bsub>_\<^esub>) (_)" [51,51,51] 50) + "_eq_restrict" :: "'b \ 'a \ bool" ("'(=(\<^bsub>_\<^esub>)')") end -unbundle eq_restrict_syntax +unbundle eq_rel_restrict_syntax translations - "(=\<^bsub>P\<^esub>)" \ "CONST restrict_left (=) P" - "x =\<^bsub>P\<^esub> y" \ "CONST restrict_left (=) P x y" + "(=\<^bsub>P\<^esub>)" \ "CONST rel_restrict_left (=) P" + "x =\<^bsub>P\<^esub> y" \ "CONST rel_restrict_left (=) P x y" lemma in_dom_eq_restrict_eq [simp]: "in_dom (=\<^bsub>P\<^esub>) = P" by auto lemma in_codom_eq_restrict_eq [simp]: "in_codom (=\<^bsub>P\<^esub>) = P" by auto lemma in_field_eq_restrict_eq [simp]: "in_field (=\<^bsub>P\<^esub>) = P" by auto paragraph \Order Properties\ context fixes P :: "'a \ bool" begin context begin lemma reflexive_on_eq_restrict: "reflexive_on P ((=\<^bsub>P\<^esub>) :: 'a \ _)" by auto lemma transitive_eq_restrict: "transitive ((=\<^bsub>P\<^esub>) :: 'a \ _)" by auto lemma symmetric_eq_restrict: "symmetric ((=\<^bsub>P\<^esub>) :: 'a \ _)" by auto lemma antisymmetric_eq_restrict: "antisymmetric ((=\<^bsub>P\<^esub>) :: 'a \ _)" by auto end context begin lemma preorder_on_eq_restrict: "preorder_on P ((=\<^bsub>P\<^esub>) :: 'a \ _)" using reflexive_on_eq_restrict transitive_eq_restrict by auto lemma partial_equivalence_rel_eq_restrict: "partial_equivalence_rel ((=\<^bsub>P\<^esub>) :: 'a \ _)" using symmetric_eq_restrict transitive_eq_restrict by auto end lemma partial_order_on_eq_restrict: "partial_order_on P ((=\<^bsub>P\<^esub>) :: 'a \ _)" using preorder_on_eq_restrict antisymmetric_eq_restrict by auto lemma equivalence_rel_on_eq_restrict: "equivalence_rel_on P ((=\<^bsub>P\<^esub>) :: 'a \ _)" using partial_equivalence_rel_eq_restrict reflexive_on_eq_restrict by blast end end diff --git a/thys/Transport/HOL_Basics/Binary_Relations/Reverse_Implies.thy b/thys/Transport/HOL_Basics/Binary_Relations/Reverse_Implies.thy new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Binary_Relations/Reverse_Implies.thy @@ -0,0 +1,27 @@ +\<^marker>\creator "Kevin Kappelmann"\ +theory Reverse_Implies + imports Binary_Relation_Functions +begin + +definition "rev_implies \ (\)\" + +bundle rev_implies_syntax begin notation rev_implies (infixr "\" 25) end +bundle no_rev_implies_syntax begin no_notation rev_implies (infixr "\" 25) end +unbundle rev_implies_syntax + +lemma rev_imp_eq_imp_inv [simp]: "(\) = (\)\" + unfolding rev_implies_def by simp + +lemma rev_impI [intro!]: + assumes "Q \ P" + shows "P \ Q" + using assms by auto + +lemma rev_impD [dest!]: + assumes "P \ Q" + shows "Q \ P" + using assms by auto + +lemma rev_imp_iff_imp: "(P \ Q) \ (Q \ P)" by 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,163 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Relators\ theory Function_Relators imports Binary_Relation_Functions Functions_Base Predicates_Lattice begin paragraph \Summary\ text \Introduces the concept of function relators. The slogan of function relators is "related functions map related inputs to related outputs".\ definition "Dep_Fun_Rel_rel R S f g \ \x y. R x y \ S x y (f x) (g y)" abbreviation "Fun_Rel_rel R S \ Dep_Fun_Rel_rel R (\_ _. S)" definition "Dep_Fun_Rel_pred P R f g \ \x. P x \ R x (f x) (g x)" abbreviation "Fun_Rel_pred P R \ Dep_Fun_Rel_pred P (\_. R)" bundle Dep_Fun_Rel_syntax begin syntax "_Fun_Rel_rel" :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("(_) \ (_)" [41, 40] 40) "_Dep_Fun_Rel_rel" :: "idt \ idt \ ('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _] \ (_)" [41, 41, 41, 40] 40) "_Dep_Fun_Rel_rel_if" :: "idt \ idt \ ('a \ 'b \ bool) \ bool \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _/ |/ _] \ (_)" [41, 41, 41, 41, 40] 40) "_Fun_Rel_pred" :: "('a \ bool) \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_] \ (_)" [41, 40] 40) "_Dep_Fun_Rel_pred" :: "idt \ ('a \ bool) \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_/ \/ _] \ (_)" [41, 41, 40] 40) "_Dep_Fun_Rel_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_/ \/ _/ |/ _] \ (_)" [41, 41, 41, 40] 40) end bundle no_Dep_Fun_Rel_syntax begin no_syntax "_Fun_Rel_rel" :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("(_) \ (_)" [41, 40] 40) "_Dep_Fun_Rel_rel" :: "idt \ idt \ ('a \ 'b \ bool) \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _] \ (_)" [41, 41, 41, 40] 40) "_Dep_Fun_Rel_rel_if" :: "idt \ idt \ ('a \ 'b \ bool) \ bool \ ('c \ 'd \ bool) \ ('a \ 'c) \ ('b \ 'd) \ bool" ("[_/ _/ \/ _/ |/ _] \ (_)" [41, 41, 41, 41, 40] 40) "_Fun_Rel_pred" :: "('a \ bool) \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_] \ (_)" [41, 40] 40) "_Dep_Fun_Rel_pred" :: "idt \ ('a \ bool) \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_/ \/ _] \ (_)" [41, 41, 40] 40) "_Dep_Fun_Rel_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'c \ bool) \ ('a \ 'b) \ ('a \ 'c) \ bool" ("[_/ \/ _/ |/ _] \ (_)" [41, 41, 41, 40] 40) end unbundle Dep_Fun_Rel_syntax translations "R \ S" \ "CONST Fun_Rel_rel R S" "[x y \ R] \ S" \ "CONST Dep_Fun_Rel_rel R (\x y. S)" "[x y \ R | B] \ S" \ "CONST Dep_Fun_Rel_rel R (\x y. CONST rel_if B S)" "[P] \ R" \ "CONST Fun_Rel_pred P R" "[x \ P] \ R" \ "CONST Dep_Fun_Rel_pred P (\x. R)" "[x \ P | B] \ R" \ "CONST Dep_Fun_Rel_pred P (\x. CONST rel_if B R)" lemma Dep_Fun_Rel_relI [intro]: assumes "\x y. R x y \ S x y (f x) (g y)" shows "([x y \ R] \ S x y) f g" unfolding Dep_Fun_Rel_rel_def using assms by blast lemma Dep_Fun_Rel_relD: assumes "([x y \ R] \ S x y) f g" and "R x y" shows "S x y (f x) (g y)" using assms unfolding Dep_Fun_Rel_rel_def by blast lemma Dep_Fun_Rel_relE [elim]: assumes "([x y \ R] \ S x y) f g" and "R x y" obtains "S x y (f x) (g y)" using assms unfolding Dep_Fun_Rel_rel_def by blast lemma Dep_Fun_Rel_predI [intro]: assumes "\x. P x \ R x (f x) (g x)" shows "([x \ P] \ R x) f g" unfolding Dep_Fun_Rel_pred_def using assms by blast lemma Dep_Fun_Rel_predD: assumes "([x \ P] \ R x) f g" and "P x" shows "R x (f x) (g x)" using assms unfolding Dep_Fun_Rel_pred_def by blast lemma Dep_Fun_Rel_predE [elim]: assumes "([x \ P] \ R x) f g" and "P x" obtains "R x (f x) (g x)" using assms unfolding Dep_Fun_Rel_pred_def by blast lemma rel_inv_Dep_Fun_Rel_rel_eq [simp]: "([x y \ R] \ S x y)\ = ([y x \ R\] \ (S x y)\)" by (intro ext) auto lemma rel_inv_Dep_Fun_Rel_pred_eq [simp]: "([x \ P] \ R x)\ = ([x \ P] \ (R x)\)" by (intro ext) auto lemma Dep_Fun_Rel_pred_eq_Dep_Fun_Rel_rel: "([x \ P] \ R x) = ([x _ \ (((\) P) \ (=))] \ R x)" by (intro ext) (auto intro!: Dep_Fun_Rel_predI Dep_Fun_Rel_relI) lemma Fun_Rel_eq_eq_eq [simp]: "((=) \ (=)) = (=)" by (intro ext) auto paragraph \Composition\ lemma Dep_Fun_Rel_rel_compI: assumes Dep_Fun_Rel1: "([x y \ R] \ S x y) f g" and Dep_Fun_Rel2: "\x y. R x y \ ([x' y' \ T x y] \ U x y x' y') f' g'" and le: "\x y. R x y \ S x y (f x) (g y) \ T x y (f x) (g y)" shows "([x y \ R] \ U x y (f x) (g y)) (f' \ f) (g' \ g)" using assms by (intro Dep_Fun_Rel_relI) (auto 6 0) corollary Dep_Fun_Rel_rel_compI': assumes "([x y \ R] \ S x y) f g" and "\x y. R x y \ ([x' y' \ S x y] \ T x y x' y') f' g'" shows "([x y \ R] \ T x y (f x) (g y)) (f' \ f) (g' \ g)" using assms by (intro Dep_Fun_Rel_rel_compI) lemma Dep_Fun_Rel_pred_comp_Dep_Fun_Rel_rel_compI: assumes Dep_Fun_Rel1: "([x \ P] \ R x) f g" and Dep_Fun_Rel2: "\x. P x \ ([x' y' \ S x] \ T x x' y') f' g'" and le: "\x. P x \ R x (f x) (g x) \ S x (f x) (g x)" shows "([x \ P] \ T x (f x) (g x)) (f' \ f) (g' \ g)" using assms by (intro Dep_Fun_Rel_predI) (auto 6 0) corollary Dep_Fun_Rel_pred_comp_Dep_Fun_Rel_rel_compI': assumes "([x \ P] \ R x) f g" and "\x. P x \ ([x' y' \ R x] \ S x x' y') f' g'" shows "([x \ P] \ S x (f x) (g x)) (f' \ f) (g' \ g)" using assms by (intro Dep_Fun_Rel_pred_comp_Dep_Fun_Rel_rel_compI) paragraph \Restrictions\ lemma restrict_left_Dep_Fun_Rel_rel_restrict_left_eq: fixes R :: "'a1 \ 'a2 \ bool" and S :: "'a1 \ 'a2 \ 'b1 \ 'b2 \ bool" and P :: "'a1 \ 'a2 \ 'b1 \ bool" assumes "\f x y. Q f \ R x y \ P x y (f x)" shows "([x y \ R] \ (S x y)\\<^bsub>P x y\<^esub>)\\<^bsub>Q\<^esub> = ([x y \ R] \ S x y)\\<^bsub>Q\<^esub>" - using assms by (intro ext iffI bin_rel_restrict_leftI Dep_Fun_Rel_relI) + 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>)" - unfolding bin_rel_restrict_right_eq + 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)\"] by simp 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 new file mode 100644 --- /dev/null +++ b/thys/Transport/HOL_Basics/Functions/Functions_Restrict.thy @@ -0,0 +1,40 @@ +\<^marker>\creator "Kevin Kappelmann"\ +theory Functions_Restrict + imports HOL.HOL +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 + +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 + +end + + +end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Bijection.thy b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Bijection.thy --- a/thys/Transport/HOL_Basics/Functions/Properties/Functions_Bijection.thy +++ b/thys/Transport/HOL_Basics/Functions/Properties/Functions_Bijection.thy @@ -1,149 +1,182 @@ \<^marker>\creator "Kevin Kappelmann"\ -subsubsection \Bijection\ +subsubsection \Bijections\ theory Functions_Bijection imports Functions_Inverse Functions_Monotone begin -consts bijection_on :: "'a \ 'b \ ('c \ 'd) \ ('d \ 'c) \ bool" +consts bijection_on :: "'a \ 'b \ 'c \ 'd \ bool" overloading bijection_on_pred \ "bijection_on :: ('a \ bool) \ ('b \ bool) \ ('a \ 'b) \ ('b \ 'a) \ bool" begin definition "bijection_on_pred P P' f g \ ([P] \\<^sub>m P') f \ ([P'] \\<^sub>m P) g \ inverse_on P f g \ inverse_on P' g f" end lemma bijection_onI [intro]: assumes "([P] \\<^sub>m P') f" and "([P'] \\<^sub>m P) g" and "inverse_on P f g" and "inverse_on P' g f" shows "bijection_on P P' f g" using assms unfolding bijection_on_pred_def by blast -lemma bijection_onE: +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" using assms unfolding bijection_on_pred_def by blast context - fixes P :: "'a \ bool" - and P' :: "'b \ bool" - and f :: "'a \ 'b" + 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" using assms by (elim bijection_onE) lemma mono_wrt_pred_if_bijection_on_right: assumes "bijection_on P P' f g" shows "([P'] \\<^sub>m P) g" using assms by (elim bijection_onE) lemma bijection_on_pred_right: assumes "bijection_on P P' f g" and "P x" shows "P' (f x)" - using assms by (blast elim: bijection_onE) + using assms by blast lemma bijection_on_pred_left: assumes "bijection_on P P' f g" and "P' y" shows "P (g y)" - using assms by (blast elim: bijection_onE) + using assms by blast lemma inverse_on_if_bijection_on_left_right: assumes "bijection_on P P' f g" shows "inverse_on P f g" using assms by (elim bijection_onE) lemma inverse_on_if_bijection_on_right_left: assumes "bijection_on P P' f g" shows "inverse_on P' g f" using assms by (elim bijection_onE) lemma bijection_on_left_right_eq_self: assumes "bijection_on P P' f g" and "P x" shows "g (f x) = x" using assms inverse_on_if_bijection_on_left_right by (intro inverse_onD) lemma bijection_on_right_left_eq_self': assumes "bijection_on P P' f g" and "P' y" shows "f (g y) = y" using assms inverse_on_if_bijection_on_right_left by (intro inverse_onD) lemma bijection_on_right_left_if_bijection_on_left_right: assumes "bijection_on P P' f g" shows "bijection_on P' P g f" - using assms by (auto elim: bijection_onE) + using assms by auto + +end + +context + fixes P :: "'a \ bool" and P' :: "'b \ bool" and f :: "'a \ 'b" and g :: "'b \ 'a" +begin lemma injective_on_if_bijection_on_left: assumes "bijection_on P P' f g" shows "injective_on P f" using assms by (intro injective_on_if_inverse_on inverse_on_if_bijection_on_left_right) lemma injective_on_if_bijection_on_right: assumes "bijection_on P P' f g" shows "injective_on P' g" by (intro injective_on_if_inverse_on) (fact inverse_on_if_bijection_on_right_left[OF assms]) end +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" + assumes "bijection_on P P' f g" + and "bijection_on P' P'' f' g'" + shows "bijection_on P P'' (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) -definition "(bijection :: ('a \ 'b) \ _) \ bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool)" + +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 lemma bijection_eq_bijection_on: - "(bijection :: ('a \ 'b) \ _) = bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool)" + "(bijection :: ('a \ 'b) \ ('b \ 'a) \ bool) = bijection_on (\ :: 'a \ bool) (\ :: 'b \ bool)" unfolding bijection_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" +begin + lemma bijectionI [intro]: assumes "inverse f g" and "inverse g f" shows "bijection f g" - unfolding bijection_eq_bijection_on using assms - by (intro bijection_onI inverse_on_if_inverse dep_mono_wrt_predI) simp_all + by (urule bijection_onI) (simp | urule assms)+ lemma bijectionE [elim]: assumes "bijection f g" obtains "inverse f g" "inverse g f" - using assms unfolding bijection_eq_bijection_on inverse_eq_inverse_on - by (blast elim: bijection_onE) + 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: assumes "bijection f g" shows "bijection g f" using assms by auto - paragraph \Instantiations\ -lemma bijection_on_self_id: - fixes P :: "'a \ bool" - shows "bijection_on P P (id :: 'a \ _) id" +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 +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,57 +1,71 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Injective\ theory Functions_Injective imports Functions_Base HOL_Syntax_Bundles_Lattices + ML_Unification.ML_Unification_HOL_Setup + ML_Unification.Unify_Resolve_Tactics begin -consts injective_on :: "'a \ ('b \ 'c) \ bool" +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 x \ P x' \ f x = f x' \ x = x'" end lemma injective_onI [intro]: assumes "\x x'. P x \ P x' \ f x = f x' \ x = x'" shows "injective_on P f" unfolding injective_on_pred_def using assms by blast lemma injective_onD: assumes "injective_on P f" and "P x" "P x'" and "f x = f x'" shows "x = x'" using assms unfolding injective_on_pred_def by blast -definition "(injective :: ('a \ _) \ _) \ injective_on (\ :: 'a \ bool)" +consts injective :: "'a \ bool" -lemma injective_eq_injective_on: "(injective :: ('a \ _) \ _) = injective_on (\ :: '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" - unfolding injective_eq_injective_on using assms by (intro injective_onI) + using assms by (urule injective_onI) lemma injectiveD: assumes "injective f" and "f x = f x'" shows "x = x'" - using assms unfolding injective_eq_injective_on by (auto dest: injective_onD) + 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,61 +1,93 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Inverse\ theory Functions_Inverse imports Functions_Injective + Functions_Monotone begin -consts inverse_on :: "'a \ ('b \ 'c) \ ('c \ 'b) \ bool" +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" end lemma inverse_onI [intro]: assumes "\x. P x \ g (f x) = x" shows "inverse_on P f g" unfolding inverse_on_pred_def using assms by blast lemma inverse_onD: assumes "inverse_on P f g" and "P x" shows "g (f x) = x" using assms unfolding inverse_on_pred_def by blast lemma injective_on_if_inverse_on: - assumes inv: "inverse_on (P :: 'a \ bool) (f :: 'a \ _) g" + 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 -definition "(inverse :: ('a \ _) \ _) \ inverse_on (\ :: 'a \ bool)" +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" + shows "inverse_on P (f' \ f) (g \ g')" + using assms by (intro inverse_onI) (auto dest!: inverse_onD) -lemma inverse_eq_inverse_on: "(inverse :: ('a \ _) \ _) = inverse_on (\ :: 'a \ bool)" +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" - unfolding inverse_eq_inverse_on using assms by (intro inverse_onI) + using assms by (urule inverse_onI) lemma inverseD: assumes "inverse f g" shows "g (f x) = x" - using assms unfolding inverse_eq_inverse_on by (auto dest: inverse_onD) + 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" + 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: + assumes "injective f" + shows "inverse f (\z. THE y. z = f y)" + using assms injectiveD by fastforce + +lemma injective_inverseE: + assumes "injective (f :: 'a \ 'b)" + obtains g :: "'b \ 'a" where "inverse f g" + using assms inverse_THE_eq_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,340 +1,350 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Monotonicity\ theory Functions_Monotone imports Binary_Relations_Order_Base Function_Relators Predicates begin paragraph \Summary\ text \Introduces the concept of monotone functions. A function is monotone if it is related to itself - see \<^term>\Dep_Fun_Rel_rel\.\ declare le_funI[intro] declare le_funE[elim] definition "dep_mono_wrt_rel R S f \ ([x y \ R] \ S x y) f f" abbreviation "mono_wrt_rel R S \ dep_mono_wrt_rel R (\_ _. S)" definition "dep_mono_wrt_pred P Q f \ ([x \ P] \ (\_. Q x)) f f" abbreviation "mono_wrt_pred P Q \ dep_mono_wrt_pred P (\_. Q)" bundle dep_mono_wrt_syntax begin syntax "_mono_wrt_rel" :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("(_) \\<^sub>m (_)" [41, 40] 40) "_dep_mono_wrt_rel" :: "idt \ idt \ ('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ _/ \/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) "_dep_mono_wrt_rel_if" :: "idt \ idt \ ('a \ 'a \ bool) \ bool \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ _/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 41, 40] 40) "_mono_wrt_pred" :: "('a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_] \\<^sub>m (_)" [41, 40] 40) "_dep_mono_wrt_pred" :: "idt \ ('a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ \/ _] \\<^sub>m (_)" [41, 41, 40] 40) "_dep_mono_wrt_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) end bundle no_dep_mono_wrt_syntax begin no_syntax "_mono_wrt_rel" :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("(_) \\<^sub>m (_)" [41, 40] 40) "_dep_mono_wrt_rel" :: "idt \ idt \ ('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ _/ \/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) "_dep_mono_wrt_rel_if" :: "idt \ idt \ ('a \ 'a \ bool) \ bool \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ _/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 41, 40] 40) "_mono_wrt_pred" :: "('a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_] \\<^sub>m (_)" [41, 40] 40) "_dep_mono_wrt_pred" :: "idt \ ('a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ \/ _] \\<^sub>m (_)" [41, 41, 40] 40) "_dep_mono_wrt_pred_if" :: "idt \ ('a \ bool) \ bool \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" ("[_/ \/ _/ |/ _] \\<^sub>m (_)" [41, 41, 41, 40] 40) 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)" lemma dep_mono_wrt_relI [intro]: assumes "\x y. R x y \ S x y (f x) (f y)" shows "([x y \ R] \\<^sub>m S x y) f" using assms unfolding dep_mono_wrt_rel_def by blast lemma dep_mono_wrt_relE [elim]: assumes "([x y \ R] \\<^sub>m S x y) f" obtains "\x y. R x y \ S x y (f x) (f y)" using assms unfolding dep_mono_wrt_rel_def by blast lemma dep_mono_wrt_relD: assumes "([x y \ R] \\<^sub>m S x y) f" and "R x y" shows "S x y (f x) (f y)" using assms by blast lemma dep_mono_wrt_predI [intro]: assumes "\x. P x \ Q x (f x)" shows "([x \ P] \\<^sub>m Q x) f" using assms unfolding dep_mono_wrt_pred_def by blast lemma dep_mono_wrt_predE [elim]: assumes "([x \ P] \\<^sub>m Q x) f" obtains "\x. P x \ Q x (f x)" using assms unfolding dep_mono_wrt_pred_def by blast lemma dep_mono_wrt_predD: assumes "([x \ P] \\<^sub>m Q x) f" and "P x" shows "Q x (f x)" using assms by blast lemma dep_mono_wrt_rel_if_Dep_Fun_Rel_rel_self: assumes "([x y \ R] \ S x y) f f" shows "([x y \ R] \\<^sub>m S x y) f" using assms by blast lemma dep_mono_wrt_pred_if_Dep_Fun_Rel_pred_self: assumes "([x \ P] \ (\_. Q x)) f f" shows "([x \ P] \\<^sub>m Q x) f" using assms by blast lemma Dep_Fun_Rel_rel_self_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" shows "([x y \ R] \ S x y) f f" using assms by blast lemma Dep_Fun_Rel_pred_self_if_dep_mono_wrt_pred: assumes "([x \ P] \\<^sub>m Q x) f" shows "([x \ P] \ (\_. Q x)) f f" using assms by blast corollary Dep_Fun_Rel_rel_self_iff_dep_mono_wrt_rel: "([x y \ R] \ S x y) f f \ ([x y \ R] \\<^sub>m S x y) f" using dep_mono_wrt_rel_if_Dep_Fun_Rel_rel_self Dep_Fun_Rel_rel_self_if_dep_mono_wrt_rel by blast corollary Dep_Fun_Rel_pred_self_iff_dep_mono_wrt_pred: "([x \ P] \ (\_. Q x)) f f \ ([x \ P] \\<^sub>m Q x) f" using dep_mono_wrt_pred_if_Dep_Fun_Rel_pred_self Dep_Fun_Rel_pred_self_if_dep_mono_wrt_pred by blast lemma dep_mono_wrt_rel_inv_eq [simp]: "([y x \ R\] \\<^sub>m (S x y)\) = ([x y \ R] \\<^sub>m S x y)" by (intro ext) force lemma in_dom_if_rel_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" and "R x y" shows "in_dom (S x y) (f x)" using assms by (intro in_domI) blast corollary in_dom_if_in_dom_if_mono_wrt_rel: assumes "(R \\<^sub>m S) f" shows "([in_dom R] \\<^sub>m in_dom S) f" using assms in_dom_if_rel_if_dep_mono_wrt_rel by fast lemma in_codom_if_rel_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" and "R x y" shows "in_codom (S x y) (f y)" using assms by (intro in_codomI) blast corollary in_codom_if_in_codom_if_mono_wrt_rel: assumes "(R \\<^sub>m S) f" shows "([in_codom R] \\<^sub>m in_codom S) f" using assms in_dom_if_rel_if_dep_mono_wrt_rel by fast corollary in_field_if_in_field_if_mono_wrt_rel: assumes "(R \\<^sub>m S) f" shows "([in_field R] \\<^sub>m in_field S) f" using assms by (intro dep_mono_wrt_predI) blast lemma le_rel_map_if_mono_wrt_rel: assumes "(R \\<^sub>m S) f" shows "R \ rel_map f S" using assms by (intro le_relI) auto lemma le_pred_map_if_mono_wrt_pred: assumes "([P] \\<^sub>m Q) f" shows "P \ pred_map f Q" using assms by (intro le_predI) auto lemma mono_wrt_rel_if_le_rel_map: assumes "R \ rel_map f S" shows "(R \\<^sub>m S) f" using assms by (intro dep_mono_wrt_relI) auto lemma mono_wrt_pred_if_le_pred_map: assumes "P \ pred_map f Q" shows "([P] \\<^sub>m Q) f" using assms by (intro dep_mono_wrt_predI) auto corollary mono_wrt_rel_iff_le_rel_map: "(R \\<^sub>m S) f \ R \ rel_map f S" using mono_wrt_rel_if_le_rel_map le_rel_map_if_mono_wrt_rel by auto corollary mono_wrt_pred_iff_le_pred_map: "([P] \\<^sub>m Q) f \ P \ pred_map f Q" using mono_wrt_pred_if_le_pred_map le_pred_map_if_mono_wrt_pred by auto definition "mono \ ((\) \\<^sub>m (\))" definition "antimono \ ((\) \\<^sub>m (\))" lemma monoI [intro]: assumes "\x y. x \ y \ f x \ f y" shows "mono f" unfolding mono_def using assms by blast lemma monoE [elim]: assumes "mono f" obtains "\x y. x \ y \ f x \ f y" using assms unfolding mono_def by blast lemma monoD: assumes "mono f" and "x \ y" shows "f x \ f y" using assms by blast lemma antimonoI [intro]: assumes "\x y. x \ y \ f y \ f x" shows "antimono f" unfolding antimono_def using assms by blast lemma antimonoE [elim]: assumes "antimono f" obtains "\x y. x \ y \ f y \ f x" using assms unfolding antimono_def by blast lemma antimonoD: assumes "antimono f" and "x \ y" shows "f y \ f x" using assms by blast lemma antimono_Dep_Fun_Rel_rel_left: "antimono (\R. [x y \ R] \ S x y)" by (intro antimonoI) auto lemma antimono_Dep_Fun_Rel_pred_left: "antimono (\P. [x \ P] \ Q x)" by (intro antimonoI) auto lemma antimono_dep_mono_wrt_rel_left: "antimono (\R. [x y \ R] \\<^sub>m S x y)" by (intro antimonoI) blast lemma antimono_dep_mono_wrt_pred_left: "antimono (\P. [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" using assms by blast lemma Dep_Fun_Rel_pred_if_le_left_if_Dep_Fun_Rel_pred: assumes "([x \ P] \ Q x) f g" and "T \ P" shows "([x \ T] \ Q x) f g" using assms by blast lemma dep_mono_wrt_rel_if_le_left_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" and "T \ R" shows "([x y \ T] \\<^sub>m S x y) f" using assms by blast lemma dep_mono_wrt_pred_if_le_left_if_dep_mono_wrt_pred: assumes "([x \ P] \\<^sub>m Q x) f" and "T \ P" shows "([x \ T] \\<^sub>m Q x) f" using assms by blast lemma mono_Dep_Fun_Rel_rel_right: "mono (\S. [x y \ R] \ S x y)" by (intro monoI) blast lemma mono_Dep_Fun_Rel_pred_right: "mono (\Q. [x \ P] \ Q x)" by (intro monoI) blast lemma mono_dep_mono_wrt_rel_right: "mono (\S. [x y \ R] \\<^sub>m S x y)" by (intro monoI) blast lemma mono_dep_mono_wrt_pred_right: "mono (\Q. [x \ P] \\<^sub>m Q x)" by (intro monoI) blast lemma Dep_Fun_Rel_rel_if_le_right_if_Dep_Fun_Rel_rel: assumes "([x y \ R] \ S x y) f g" and "\x y. R x y \ S x y (f x) (g y) \ T x y (f x) (g y)" shows "([x y \ R] \ T x y) f g" using assms by (intro Dep_Fun_Rel_relI) blast lemma Dep_Fun_Rel_pred_if_le_right_if_Dep_Fun_Rel_pred: assumes "([x \ P] \ Q x) f g" and "\x. P x \ Q x (f x) (g x) \ T x (f x) (g x)" shows "([x \ P] \ T x) f g" using assms by blast lemma dep_mono_wrt_rel_if_le_right_if_dep_mono_wrt_rel: assumes "([x y \ R] \\<^sub>m S x y) f" and "\x y. R x y \ S x y (f x) (f y) \ T x y (f x) (f y)" shows "([x y \ R] \\<^sub>m T x y) f" using assms by (intro dep_mono_wrt_relI) blast lemma dep_mono_wrt_pred_if_le_right_if_dep_mono_wrt_pred: assumes "([x \ P] \\<^sub>m Q x) f" and "\x. P x \ Q x (f x) \ T x (f x)" shows "([x \ P] \\<^sub>m T x) f" using assms by blast paragraph \Composition\ lemma dep_mono_wrt_rel_compI: assumes "([x y \ R] \\<^sub>m S x y) f" and "\x y. R x y \ ([x' y' \ T x y] \\<^sub>m U x y x' y') f'" and "\x y. R x y \ S x y (f x) (f y) \ T x y (f x) (f y)" shows "([x y \ R] \\<^sub>m U x y (f x) (f y)) (f' \ f)" using assms by (intro dep_mono_wrt_relI) force corollary dep_mono_wrt_rel_compI': assumes "([x y \ R] \\<^sub>m S x y) f" and "\x y. R x y \ ([x' y' \ S x y] \\<^sub>m T x y x' y') f'" shows "([x y \ R] \\<^sub>m T x y (f x) (f y)) (f' \ f)" using assms by (intro dep_mono_wrt_rel_compI) lemma dep_mono_wrt_pred_comp_dep_mono_wrt_rel_compI: assumes "([x \ P] \\<^sub>m Q x) f" and "\x. P x \ ([x' y' \ R x] \\<^sub>m S x x' y') f'" and "\x. P x \ Q x (f x) \ R x (f x) (f x)" shows "([x \ P] \\<^sub>m (\y. S x (f x) (f x) y y)) (f' \ f)" using assms by (intro dep_mono_wrt_predI) force lemma dep_mono_wrt_pred_comp_dep_mono_wrt_pred_compI: assumes "([x \ P] \\<^sub>m Q x) f" and "\x. P x \ ([x' \ R x] \\<^sub>m S x x') f'" and "\x. P x \ Q x (f x) \ R x (f x)" shows "([x \ P] \\<^sub>m S x (f x)) (f' \ f)" using assms by (intro dep_mono_wrt_predI) force corollary dep_mono_wrt_pred_comp_dep_mono_wrt_pred_compI': assumes "([x \ P] \\<^sub>m Q x) f" and "\x. P x \ ([x' \ Q x] \\<^sub>m S x x') f'" shows "([x \ P] \\<^sub>m S x (f x)) (f' \ f)" using assms by (intro dep_mono_wrt_pred_comp_dep_mono_wrt_pred_compI) +lemma mono_wrt_rel_top [iff]: "(R \\<^sub>m \) f" by auto +lemma mono_wrt_pred_top [iff]: "([P] \\<^sub>m \) 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_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_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 + +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 + 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,49 +1,63 @@ \<^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 begin -consts surjective_at :: "'a \ ('b \ 'c) \ bool" +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)" end lemma surjective_atI [intro]: assumes "\y. P y \ \x. y = f x" shows "surjective_at P f" unfolding surjective_at_pred_def using assms by blast lemma surjective_atE [elim]: assumes "surjective_at P f" and "P y" obtains x where "y = f x" using assms unfolding surjective_at_pred_def by blast -definition "(surjective :: (_ \ 'a) \ _) \ surjective_at (\ :: 'a \ bool)" +consts surjective :: "'a \ bool" -lemma surjective_eq_surjective_at: "(surjective :: (_ \ 'a) \ _) = surjective_at (\ :: '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" shows "surjective f" - unfolding surjective_eq_surjective_at using assms by (intro surjective_atI) + using assms by (urule surjective_atI) lemma surjectiveE: assumes "surjective f" obtains x where "y = f x" - using assms unfolding surjective_eq_surjective_at by (blast intro: top1I) + using assms by (urule (e) surjective_atE where chained = insert) simp lemma surjective_at_if_surjective: - fixes P :: "'a \ bool" and f :: "_ \ 'a" + 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_Relator_Base.thy b/thys/Transport/HOL_Basics/Galois/Galois_Relator_Base.thy --- a/thys/Transport/HOL_Basics/Galois/Galois_Relator_Base.thy +++ b/thys/Transport/HOL_Basics/Galois/Galois_Relator_Base.thy @@ -1,62 +1,62 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Basics For Relator For Galois Connections\ theory Galois_Relator_Base imports Galois_Base begin locale galois_rel = orders L R for L :: "'a \ 'b \ bool" and R :: "'c \ 'd \ bool" and r :: "'d \ 'b" begin text \Morally speaking, the Galois relator characterises when two terms \<^term>\x :: 'a\ and \<^term>\y :: 'b\ are "similar".\ definition "Galois x y \ in_codom (\\<^bsub>R\<^esub>) y \ x \\<^bsub>L\<^esub> r y" abbreviation "left_Galois \ Galois" notation left_Galois (infix "\<^bsub>L\<^esub>\" 50) abbreviation (input) "ge_Galois_left \ (\<^bsub>L\<^esub>\)\" notation ge_Galois_left (infix "\\<^bsub>L\<^esub>" 50) text \Here we only introduced the (left) Galois relator @{term "(\<^bsub>L\<^esub>\)"}. All other variants can be introduced by considering suitable flipped and inversed interpretations (see @{file "Half_Galois_Property.thy"}).\ lemma left_GaloisI [intro]: assumes "in_codom (\\<^bsub>R\<^esub>) y" and "x \\<^bsub>L\<^esub> r y" shows "x \<^bsub>L\<^esub>\ y" unfolding Galois_def using assms by blast lemma left_GaloisE [elim]: assumes "x \<^bsub>L\<^esub>\ y" obtains "in_codom (\\<^bsub>R\<^esub>) y" "x \\<^bsub>L\<^esub> r y" using assms unfolding Galois_def by blast corollary in_dom_left_if_left_Galois: assumes "x \<^bsub>L\<^esub>\ y" shows "in_dom (\\<^bsub>L\<^esub>) x" using assms by blast corollary left_Galois_iff_in_codom_and_left_rel_right: "x \<^bsub>L\<^esub>\ y \ in_codom (\\<^bsub>R\<^esub>) y \ x \\<^bsub>L\<^esub> r y" by blast lemma left_Galois_restrict_left_eq_left_Galois_left_restrict_left: "(\<^bsub>L\<^esub>\)\\<^bsub>P :: 'a \ bool\<^esub> = galois_rel.Galois (\\<^bsub>L\<^esub>)\\<^bsub>P\<^esub> (\\<^bsub>R\<^esub>) r" - by (intro ext iffI galois_rel.left_GaloisI bin_rel_restrict_leftI) + by (intro ext iffI galois_rel.left_GaloisI rel_restrict_leftI) (auto elim: galois_rel.left_GaloisE) lemma left_Galois_restrict_right_eq_left_Galois_right_restrict_right: "(\<^bsub>L\<^esub>\)\\<^bsub>P :: 'd \ bool\<^esub> = galois_rel.Galois (\\<^bsub>L\<^esub>) (\\<^bsub>R\<^esub>)\\<^bsub>P\<^esub> r" - by (intro ext iffI galois_rel.left_GaloisI bin_rel_restrict_rightI) - (auto elim!: galois_rel.left_GaloisE bin_rel_restrict_rightE) + by (intro ext iffI galois_rel.left_GaloisI rel_restrict_rightI) + (auto elim!: galois_rel.left_GaloisE rel_restrict_rightE) end end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Galois/Half_Galois_Property.thy b/thys/Transport/HOL_Basics/Galois/Half_Galois_Property.thy --- a/thys/Transport/HOL_Basics/Galois/Half_Galois_Property.thy +++ b/thys/Transport/HOL_Basics/Galois/Half_Galois_Property.thy @@ -1,283 +1,285 @@ \<^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 R. inflationary_on P R\ \ deflationary_on P R" - and "\P R. deflationary_on P R\ \ inflationary_on P 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. transitive R\ \ transitive R" + 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) 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,81 +1,79 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Alignment With Definitions from HOL-Algebra\ theory HOL_Algebra_Alignment_Galois imports "HOL-Algebra.Galois_Connection" HOL_Algebra_Alignment_Orders Galois begin named_theorems HOL_Algebra_galois_alignment context galois_connection begin context fixes L R l r defines "L \ (\\<^bsub>\\<^esub>)\\<^bsub>carrier \\<^esub>\\<^bsub>carrier \\<^esub>" and "R \ (\\<^bsub>\\<^esub>)\\<^bsub>carrier \\<^esub>\\<^bsub>carrier \\<^esub>" and "l \ \\<^sup>*" and "r \ \\<^sub>*" - notes defs[simp] = L_def R_def l_def r_def and bin_rel_restrict_right_eq[simp] - and bin_rel_restrict_leftI[intro!] bin_rel_restrict_leftE[elim!] + 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 bin_rel_restrict_right_eq[simp] - and bin_rel_restrict_leftI[intro!] bin_rel_restrict_leftE[elim!] - in_codom_bin_rel_restrict_leftE[elim!] + 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" +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" +lemma half_galois_prop_right_right_left [HOL_Algebra_galois_alignment]: "(R \\<^sub>h L) r l" using gal_bij_conn.left lower_closed lower_inv_eq upper_inv_eq by (intro half_galois_prop_rightI; elim Galois_rightE) (auto; metis) lemma prop_right_right_left [HOL_Algebra_galois_alignment]: "(R \ L) r l" using half_galois_prop_left_right_left half_galois_prop_right_right_left by blast lemma galois_equivalence [HOL_Algebra_galois_alignment]: "(L \\<^sub>G R) l r" using gal_bij_conn.galois_connection prop_right_right_left by (intro galois.galois_equivalenceI) auto end end end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Binary_Relations.thy b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Binary_Relations.thy --- a/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Binary_Relations.thy +++ b/thys/Transport/HOL_Basics/HOL_Alignments/HOL_Alignment_Binary_Relations.thy @@ -1,296 +1,386 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Alignment With Definitions from HOL.Main\ theory HOL_Alignment_Binary_Relations imports Main HOL_Mem_Of HOL_Syntax_Bundles_Relations LBinary_Relations begin unbundle no_HOL_relation_syntax named_theorems HOL_bin_rel_alignment paragraph \Properties\ subparagraph \Antisymmetric\ overloading antisymmetric_on_set \ "antisymmetric_on :: 'a set \ ('a \ 'a \ bool) \ bool" begin - definition "antisymmetric_on_set (S :: 'a set) :: ('a \ _) \ _ \ + 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 \ _) \ bool) = - antisymmetric_on (mem_of S)" + "(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 \ _) \ antisymmetric_on (mem_of S) R" + "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 \ _) \ _ \ + 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 \ _) \ _ \ + 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 \ _) \ bool) = + "(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 \ _) \ rel_injective_on (mem_of S) R" + "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 \ _) \ _ \ + 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 \ _) \ bool) = + "(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 \ _) \ - irreflexive_on (mem_of S) R" + "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" +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 \ _) \ _ \ + 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 \ _) \ bool) = + "(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 \ _) \ left_total_on (mem_of S) R" + "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 \ _) \ _ \ + 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)" + "(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" + "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 \ _) \ _ \ + 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 \ _) \ _ \ + 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 \ _) \ bool) = - right_unique_on (mem_of S)" + "(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 \ _) \ right_unique_on (mem_of S) R" + "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 \ _) \ _ \ + 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) = - rel_surjective_at (mem_of S)" + "(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 \ _) \ rel_surjective_at (mem_of S) R" + "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 \ _) \ _ \ + 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)" + "(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" + "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 \ _) \ _ \ + 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)" + "(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" + "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 eq_onp_eq_eq_restrict [HOL_bin_rel_alignment]: "eq_onp = restrict_left (=)" +lemma eq_onp_eq_eq_restrict [HOL_bin_rel_alignment]: "eq_onp = rel_restrict_left (=)" unfolding eq_onp_def by (intro ext) auto overloading - bin_rel_restrict_left_set \ "restrict_left :: ('a \ 'b \ bool) \ 'a set \ 'a \ 'b \ bool" - bin_rel_restrict_right_set \ "restrict_right :: ('a \ 'b \ bool) \ 'b set \ 'a \ 'b \ bool" + 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 "bin_rel_restrict_left_set (R :: 'a \ 'b \ bool) (S :: 'a set) \ R\\<^bsub>mem_of S\<^esub>" - definition "bin_rel_restrict_right_set (R :: 'a \ 'b \ bool) (S :: 'b set) \ R\\<^bsub>mem_of S\<^esub>" + 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 -lemma bin_rel_restrict_left_set_eq_restrict_left_pred [simp]: +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>" - unfolding bin_rel_restrict_left_set_def by simp + unfolding rel_restrict_left_set_def by simp -lemma bin_rel_restrict_right_set_eq_restrict_right_pred [simp]: - "(R\\<^bsub>S :: 'b set\<^esub> :: 'a \ 'b \ bool) = R\\<^bsub>mem_of S\<^esub>" - unfolding bin_rel_restrict_right_set_def by simp +lemma 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>" + 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" 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>" + 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>" + 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" 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,160 +1,180 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Alignment With Definitions from HOL.Main\ theory HOL_Alignment_Functions imports HOL_Alignment_Binary_Relations HOL_Syntax_Bundles_Functions LFunctions begin unbundle no_HOL_function_syntax named_theorems HOL_fun_alignment paragraph \Functions\ subparagraph \Bijection\ overloading bijection_on_set \ "bijection_on :: 'a set \ 'b set \ ('a \ 'b) \ ('b \ 'a) \ bool" begin definition "bijection_on_set (S :: 'a set) (S' :: 'b set) :: ('a \ 'b) \ ('b \ 'a) \ bool \ bijection_on (mem_of S) (mem_of S')" end lemma bijection_on_set_eq_bijection_on_pred [simp]: - "(bijection_on (S :: 'a set) (S' :: 'b set) :: ('a \ 'b) \ _) = + "(bijection_on (S :: 'a set) (S' :: 'b set) :: ('a \ 'b) \ ('b \ 'a) \ _) = 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" + 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 \ + "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 S S'" - obtains g where "bijection_on S S' f g" + 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 g" + 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 S S' \ (\g. bijection_on S S' f g)" - by (intro iffI) - (auto elim!: bij_betw_bijection_onE intro: bij_betw_if_bijection_on) + "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) \ _ \ + 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) \ _) = inverse_on (mem_of S)" + "(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 \ inverse_on (mem_of S) f g" + "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" 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,78 +1,143 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Alignment With 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)" + 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" + 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 \ 'a \ bool) \ ('a \ 'a) \ bool" + inflationary_on_set \ "inflationary_on :: 'a set \ ('a \ 'b \ bool) \ ('a \ 'b) \ bool" begin - definition "inflationary_on_set (S :: 'a set) :: ('a \ _) \ _ \ + 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 \ _) \ _) = inflationary_on (mem_of S)" + "(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 \ _) f \ inflationary_on (mem_of S) R f" + "inflationary_on (S :: 'a set) (R :: 'a \ 'b \ bool) (f :: 'a \ 'b) \ + inflationary_on (mem_of S) R f" by simp -text \Terms like @{term deflationary_on}, @{term rel_equivalence_on}, -and @{term idempotent_on} are automatically overloaded. One can get similar -correspondence lemmas by unfolding the corresponding definitional theorems, -e.g. @{thm deflationary_on_eq_inflationary_on_rel_inv}.\ + +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) + 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_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,11 +1,18 @@ \<^marker>\creator "Kevin Kappelmann"\ theory HOL_Mem_Of imports HOL.Set begin definition "mem_of A x \ x \ A" -lemma mem_of_eq [simp]: "mem_of \ \A x. x \ A" unfolding mem_of_def by simp -lemma mem_of_iff [iff]: "mem_of A x \ x \ A" by simp +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_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/Equivalence_Relations.thy b/thys/Transport/HOL_Basics/Orders/Equivalence_Relations.thy --- a/thys/Transport/HOL_Basics/Orders/Equivalence_Relations.thy +++ b/thys/Transport/HOL_Basics/Orders/Equivalence_Relations.thy @@ -1,71 +1,81 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Equivalences\ theory Equivalence_Relations imports Partial_Equivalence_Relations begin definition "equivalence_rel_on \ partial_equivalence_rel_on \ reflexive_on" lemma equivalence_rel_onI [intro]: assumes "partial_equivalence_rel_on P R" and "reflexive_on P R" shows "equivalence_rel_on P R" unfolding equivalence_rel_on_def using assms by auto lemma equivalence_rel_onE [elim]: assumes "equivalence_rel_on P R" obtains "partial_equivalence_rel_on P R" "reflexive_on P R" using assms unfolding equivalence_rel_on_def by auto lemma equivalence_rel_on_in_field_if_partial_equivalence_rel: assumes "partial_equivalence_rel R" shows "equivalence_rel_on (in_field R) R" using assms by (intro equivalence_rel_onI reflexive_on_in_field_if_partial_equivalence_rel) auto corollary partial_equivalence_rel_iff_equivalence_rel_on_in_field: "partial_equivalence_rel R \ equivalence_rel_on (in_field R) R" using equivalence_rel_on_in_field_if_partial_equivalence_rel by auto +consts equivalence_rel :: "'a \ bool" -definition "(equivalence_rel :: ('a \ _) \ bool) \ equivalence_rel_on (\ :: 'a \ bool)" +overloading + equivalence_rel \ "equivalence_rel :: ('a \ 'a \ bool) \ bool" +begin + definition "(equivalence_rel :: ('a \ 'a \ bool) \ bool) \ equivalence_rel_on (\ :: 'a \ bool)" +end lemma equivalence_rel_eq_equivalence_rel_on: - "(equivalence_rel :: ('a \ _) \ bool) = equivalence_rel_on (\ :: 'a \ bool)" + "(equivalence_rel :: ('a \ 'a \ bool) \ bool) = equivalence_rel_on (\ :: 'a \ bool)" unfolding equivalence_rel_def .. +lemma equivalence_rel_eq_equivalence_rel_on_uhint [uhint]: + assumes "P \ \ :: 'a \ bool" + shows "equivalence_rel :: ('a \ 'a \ bool) \ bool \ equivalence_rel_on P" + using assms by (simp add: equivalence_rel_eq_equivalence_rel_on) + +context + fixes R :: "'a \ 'a \ bool" +begin + lemma equivalence_relI [intro]: assumes "partial_equivalence_rel R" and "reflexive R" shows "equivalence_rel R" - unfolding equivalence_rel_eq_equivalence_rel_on using assms - by (intro equivalence_rel_onI partial_equivalence_rel_on_if_partial_equivalence_rel - reflexive_on_if_reflexive) + using assms by (urule equivalence_rel_onI) lemma equivalence_relE [elim]: assumes "equivalence_rel R" obtains "partial_equivalence_rel R" "reflexive R" - using assms unfolding equivalence_rel_eq_equivalence_rel_on - by (elim equivalence_rel_onE) - (simp only: partial_equivalence_rel_eq_partial_equivalence_rel_on - reflexive_eq_reflexive_on) + using assms by (urule (e) equivalence_rel_onE) lemma equivalence_rel_on_if_equivalence: - fixes P :: "'a \ bool" and R :: "'a \ _" + fixes P :: "'a \ bool" assumes "equivalence_rel R" shows "equivalence_rel_on P R" using assms by (elim equivalence_relE) (intro equivalence_rel_onI partial_equivalence_rel_on_if_partial_equivalence_rel reflexive_on_if_reflexive) +end paragraph \Instantiations\ lemma equivalence_eq: "equivalence_rel (=)" using partial_equivalence_rel_eq reflexive_eq by (rule equivalence_relI) -lemma equivalence_top: "equivalence_rel \" +lemma equivalence_top: "equivalence_rel (\ :: 'a \ 'a \ bool)" using partial_equivalence_rel_top reflexive_top by (rule equivalence_relI) end diff --git a/thys/Transport/HOL_Basics/Orders/Functions/Closure_Operators.thy b/thys/Transport/HOL_Basics/Orders/Functions/Closure_Operators.thy --- a/thys/Transport/HOL_Basics/Orders/Functions/Closure_Operators.thy +++ b/thys/Transport/HOL_Basics/Orders/Functions/Closure_Operators.thy @@ -1,85 +1,125 @@ \<^marker>\creator "Kevin Kappelmann"\ subsubsection \Closure Operators\ theory Closure_Operators imports Order_Functions_Base begin -definition "idempotent_on P R f \ rel_equivalence_on P (rel_map f R) f" +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_def using assms by fastforce + unfolding idempotent_on_pred_def using assms by fastforce lemma idempotent_onE [elim]: assumes "idempotent_on P R f" and "P x" - obtains "R (f (f x)) (f x)" "R (f x) (f (f x))" - using assms unfolding idempotent_on_def by fastforce + 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_def by simp + unfolding idempotent_on_pred_def by simp -lemma bi_related_if_idempotent_onD: +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 -definition "(idempotent :: ('a \ _) \ _) \ idempotent_on (\ :: 'a \ bool)" +end -lemma idempotent_eq_idempotent_on: "(idempotent :: ('a \ _) \ _) = idempotent_on (\ :: 'a \ bool)" +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. R (f (f x)) (f x)" - and "\x. R (f x) (f (f x))" + assumes "\x. f x \\<^bsub>R\<^esub> f (f x)" shows "idempotent R f" - unfolding idempotent_eq_idempotent_on using assms by blast + using assms by (urule idempotent_onI) -lemma idempotentE [elim]: +lemma idempotentD [dest]: assumes "idempotent R f" - obtains "R (f (f x)) (f x)" "R (f x) (f (f x))" - using assms unfolding idempotent_eq_idempotent_on by (blast intro: top1I) + 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: - fixes P :: "'a \ bool" and R :: "'a \ _" assumes "idempotent R f" shows "idempotent_on P R f" using assms by (intro idempotent_onI) auto -definition "closure_operator R \ +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" 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,447 +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\ -definition "bi_related R x y \ R x y \ R y x" +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"*) bundle bi_related_syntax begin syntax - "_bi_related" :: "'a \ ('a \ 'a \ bool) \ 'a \ bool" ("(_) \\<^bsub>(_)\<^esub> (_)" [51,51,51] 50) + "_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 \ ('a \ 'a \ bool) \ 'a \ bool" ("(_) \\<^bsub>(_)\<^esub> (_)" [51,51,51] 50) + "_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 -lemma symmetric_bi_related [iff]: "symmetric (\\<^bsub>R\<^esub>)" +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>)" + 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>)" + shows "transitive ((\\<^bsub>R\<^esub>) :: 'a \ 'a \ bool)" using assms by (intro transitiveI bi_relatedI) auto -lemma mono_bi_related [iff]: "mono bi_related" +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_le_eq_if_antisymmetric_on_in_field: - assumes "antisymmetric_on (in_field R) R" - shows "(\\<^bsub>R\<^esub>) \ (=)" - using assms - by (intro le_relI eq_if_bi_related_if_in_field_le_if_antisymmetric_on) blast+ - lemma bi_related_if_all_rel_iff_if_reflexive_on: assumes "reflexive_on P R" and "\z. P z \ R x z \ R y z" and "P x" "P y" shows "x \\<^bsub>R\<^esub> y" using assms by blast lemma bi_related_if_all_rel_iff_if_reflexive_on': assumes "reflexive_on P R" and "\z. P z \ R z x \ R z y" and "P x" "P y" shows "x \\<^bsub>R\<^esub> y" using assms by blast corollary eq_if_all_rel_iff_if_antisymmetric_on_if_reflexive_on: assumes "reflexive_on P R" and "antisymmetric_on P R" and "\z. P z \ R x z \ R y z" and "P x" "P y" shows "x = y" using assms by (blast intro: eq_if_bi_related_if_antisymmetric_on bi_related_if_all_rel_iff_if_reflexive_on) corollary eq_if_all_rel_iff_if_antisymmetric_on_if_reflexive_on': assumes "reflexive_on P R" and "antisymmetric_on P R" and "\z. P z \ R z x \ R z y" and "P x" "P y" shows "x = y" using assms by (blast intro: eq_if_bi_related_if_antisymmetric_on bi_related_if_all_rel_iff_if_reflexive_on') +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 \ 'b \ bool) \ ('b \ 'b) \ bool" +consts inflationary_on :: "'a \ 'b \ 'c \ bool" overloading - inflationary_on_pred \ "inflationary_on :: - ('a \ bool) \ ('a \ 'a \ bool) \ ('a \ 'a) \ bool" + 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 \ 'a \ _) f \ \x. P x \ R x (f x)" + definition "inflationary_on_pred P (R :: 'a \ 'b \ bool) (f :: 'a \ 'b) \ \x. P x \ R x (f x)" end lemma inflationary_onI [intro]: assumes "\x. P x \ R x (f x)" shows "inflationary_on P R f" unfolding inflationary_on_pred_def using assms by blast lemma inflationary_onD [dest]: assumes "inflationary_on P R f" and "P x" shows "R x (f x)" using assms unfolding inflationary_on_pred_def by blast lemma inflationary_on_eq_dep_mono_wrt_pred: "inflationary_on = dep_mono_wrt_pred" by blast -lemma antimono_inflationary_on_pred [iff]: - "antimono (\(P :: 'a \ bool). inflationary_on P (R :: 'a \ _))" - by (intro antimonoI) auto - -lemma inflationary_on_if_le_pred_if_inflationary_on: - fixes P P' :: "'a \ bool" and R :: "'a \ _" - assumes "inflationary_on P R f" - and "P' \ P" - shows "inflationary_on P' R f" - using assms by blast - -lemma mono_inflationary_on_rel [iff]: - "mono (\(R :: 'a \ _). inflationary_on (P :: 'a \ bool) R)" - by (intro monoI) auto - lemma inflationary_on_if_le_rel_if_inflationary_on: assumes "inflationary_on P R f" and "\x. P x \ R x (f x) \ R' x (f x)" shows "inflationary_on P R' f" using assms by blast +lemma 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 + +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 \ _) \ _) ((P :: 'a \ bool) \ Q) + "(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" -definition "(inflationary :: ('a \ _) \ _) \ inflationary_on (\ :: 'a \ 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 \ _) \ _) = inflationary_on (\ :: 'a \ bool)" + "(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" - unfolding inflationary_eq_inflationary_on using assms - by (intro inflationary_onI) + using assms by (urule inflationary_onI) lemma inflationaryD: assumes "inflationary R f" shows "R x (f x)" - using assms unfolding inflationary_eq_inflationary_on by auto + using assms by (urule (d) inflationary_onD where chained = insert) simp lemma inflationary_on_if_inflationary: - fixes P :: "'a \ bool" and R :: "'a \ _" + 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\ -definition "deflationary_on P R \ inflationary_on P R\" +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 = inflationary_on P R\" - unfolding deflationary_on_def .. + "(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] -corollary deflationary_on_rel_inv_eq_inflationary_on [simp]: - "deflationary_on P R\ = inflationary_on P R" - unfolding deflationary_on_eq_inflationary_on_rel_inv by simp - lemma deflationary_onI [intro]: assumes "\x. P x \ R (f x) x" shows "deflationary_on P R f" unfolding deflationary_on_eq_inflationary_on_rel_inv using assms by (intro inflationary_onI rel_invI) lemma deflationary_onD [dest]: assumes "deflationary_on P R f" and "P x" shows "R (f x) x" using assms unfolding deflationary_on_eq_inflationary_on_rel_inv by blast -lemma deflationary_on_eq_dep_mono_wrt_pred_rel_inv: - "deflationary_on P R = ([x \ P] \\<^sub>m R\ x)" - by blast - -lemma antimono_deflationary_on_pred [iff]: - "antimono (\(P :: 'a \ bool). deflationary_on P (R :: 'a \ _))" - by (intro antimonoI) auto +end -lemma deflationary_on_if_le_pred_if_deflationary_on: - fixes P P' :: "'a \ bool" and R :: "'a \ _" - assumes "deflationary_on P R f" - and "P' \ P" - shows "deflationary_on P' R f" - using assms by blast +context + fixes P P' :: "'a \ bool" and R :: "'b \ 'a \ bool" and f :: "'a \ 'b" +begin -lemma mono_deflationary_on_rel [iff]: - "mono (\(R :: 'a \ _). deflationary_on (P :: 'a \ bool) R)" - by (intro monoI) auto +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)" + 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) \ ('a \ _) \ _) ((P :: 'a \ bool) \ Q) + "(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 -definition "(deflationary :: ('a \ _) \ _) \ deflationary_on (\ :: 'a \ bool)" +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 :: ('a \ _) \ _) = deflationary_on (\ :: 'a \ bool)" + "(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" - unfolding deflationary_eq_deflationary_on using assms by (intro deflationary_onI) + using assms by (urule deflationary_onI) lemma deflationaryD: assumes "deflationary R f" shows "R (f x) x" - using assms unfolding deflationary_eq_deflationary_on by auto + using assms by (urule (d) deflationary_onD where chained = insert) simp lemma deflationary_on_if_deflationary: - fixes P :: "'a \ bool" and f :: "'a \ _" + 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 "R x (f x)" "R (f x) 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\ = rel_equivalence_on P R" + "(rel_equivalence_on P R\ :: ('a \ 'a) \ bool) = rel_equivalence_on P R" by (intro ext) fastforce -lemma antimono_rel_equivalence_on_pred [iff]: - "antimono (\(P :: 'a \ bool). rel_equivalence_on P (R :: 'a \ _))" - by (intro antimonoI) blast +lemma 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: - fixes P P' :: "'a \ bool" and R :: "'a \ _" assumes "rel_equivalence_on P R f" and "P' \ P" shows "rel_equivalence_on P' R f" using assms by blast lemma rel_equivalence_on_sup_eq [simp]: - "(rel_equivalence_on :: ('a \ bool) \ ('a \ _) \ _) ((P :: 'a \ bool) \ Q) + "(rel_equivalence_on :: ('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 "transitive_on P R" shows "reflexive_on P R" using assms by (blast dest: transitive_onD) lemma inflationary_on_eq_rel_equivalence_on_if_symmetric: assumes "symmetric R" - shows "inflationary_on P R = rel_equivalence_on P R" + 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 = rel_equivalence_on P 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 -definition "(rel_equivalence :: ('a \ _) \ _ ) \ rel_equivalence_on (\ :: 'a \ bool)" +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 \ _) \ _ ) = rel_equivalence_on (\ :: 'a \ bool)" + "(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" - unfolding rel_equivalence_eq_rel_equivalence_on using assms - by (intro rel_equivalence_onI) - (auto dest: inflationary_on_if_inflationary deflationary_on_if_deflationary) + 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 unfolding rel_equivalence_eq_rel_equivalence_on - by (elim rel_equivalence_onE) - (simp only: inflationary_eq_inflationary_on deflationary_eq_deflationary_on) + 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 (elim rel_equivalenceE) + using assms by (rule rel_equivalenceE) lemma deflationary_if_rel_equivalence: assumes "rel_equivalence R f" shows "deflationary R f" - using assms by (elim rel_equivalenceE) + using assms by (rule rel_equivalenceE) lemma rel_equivalence_on_if_rel_equivalence: - fixes P :: "'a \ bool" and R :: "'a \ _" assumes "rel_equivalence R f" shows "rel_equivalence_on P R f" using assms by (intro rel_equivalence_onI) (auto dest: inflationary_on_if_inflationary deflationary_on_if_deflationary) lemma bi_related_if_rel_equivalence: assumes "rel_equivalence R f" shows "x \\<^bsub>R\<^esub> f x" using assms by (intro bi_relatedI) (auto dest: inflationaryD deflationaryD) lemma rel_equivalence_if_all_bi_related: assumes "\x. x \\<^bsub>R\<^esub> f x" shows "rel_equivalence R f" using assms by auto lemma rel_equivalenceD: assumes "rel_equivalence R f" shows "R x (f x)" "R (f x) x" using assms by (auto dest: bi_related_if_rel_equivalence) lemma reflexive_on_in_field_if_transitive_if_rel_equivalence_on: assumes "rel_equivalence_on (in_field R) R f" and "transitive R" shows "reflexive_on (in_field R) R" using assms by (intro reflexive_onI) blast corollary preorder_on_in_field_if_transitive_if_rel_equivalence_on: assumes "rel_equivalence_on (in_field R) R f" and "transitive R" shows "preorder_on (in_field R) R" using assms reflexive_on_in_field_if_transitive_if_rel_equivalence_on by blast +end end diff --git a/thys/Transport/HOL_Basics/Orders/Functors/Order_Functors_Base.thy b/thys/Transport/HOL_Basics/Orders/Functors/Order_Functors_Base.thy --- a/thys/Transport/HOL_Basics/Orders/Functors/Order_Functors_Base.thy +++ b/thys/Transport/HOL_Basics/Orders/Functors/Order_Functors_Base.thy @@ -1,196 +1,196 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Order Functors\ subsubsection \Basic Setup and Results\ theory Order_Functors_Base imports Functions_Inverse Order_Functions_Base begin text \In the following, we do not add any assumptions to our locales but rather add them as needed to the theorem statements. This allows consumers to state preciser results; particularly, the development of Transport depends on this setup.\ locale orders = fixes L :: "'a \ 'b \ bool" and R :: "'c \ 'd \ bool" begin notation L (infix "\\<^bsub>L\<^esub>" 50) notation R (infix "\\<^bsub>R\<^esub>" 50) text\We call @{term "(\\<^bsub>L\<^esub>)"} the \<^emph>\left relation\ and @{term "(\\<^bsub>R\<^esub>)"} the \<^emph>\right relation\.\ abbreviation (input) "ge_left \ (\\<^bsub>L\<^esub>)\" notation ge_left (infix "\\<^bsub>L\<^esub>" 50) abbreviation (input) "ge_right \ (\\<^bsub>R\<^esub>)\" notation ge_right (infix "\\<^bsub>R\<^esub>" 50) end text \Homogeneous orders\ locale hom_orders = orders L R for L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" locale order_functor = hom_orders L R for L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" begin lemma left_right_rel_left_self_if_reflexive_on_left_if_mono_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on P (\\<^bsub>L\<^esub>)" and "P x" shows "l x \\<^bsub>R\<^esub> l x" using assms by blast lemma left_right_rel_left_self_if_reflexive_on_in_dom_right_if_mono_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on (in_dom (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) x" shows "l x \\<^bsub>R\<^esub> l x" using assms by blast lemma left_right_rel_left_self_if_reflexive_on_in_codom_right_if_mono_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on (in_codom (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" and "in_codom (\\<^bsub>L\<^esub>) x" shows "l x \\<^bsub>R\<^esub> l x" using assms by blast lemma left_right_rel_left_self_if_reflexive_on_in_field_right_if_mono_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" and "in_field (\\<^bsub>L\<^esub>) x" shows "l x \\<^bsub>R\<^esub> l x" using assms by blast lemma mono_wrt_rel_left_if_reflexive_on_if_le_eq_if_mono_wrt_in_field: assumes "([in_field (\\<^bsub>L\<^esub>)] \\<^sub>m P) l" and "(\\<^bsub>L\<^esub>) \ (=)" and "reflexive_on P (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" using assms by (intro dep_mono_wrt_relI) auto end locale order_functors = order_functor L R l + flip_of : order_functor R L r for L R l r begin text \We call the composition \<^term>\r \ l\ the \<^emph>\unit\ and the term \<^term>\l \ r\ the \<^emph>\counit\ of the order functors pair. This is terminology is borrowed from category theory - the functors are an \<^emph>\adjoint\.\ definition "unit \ r \ l" notation unit ("\") lemma unit_eq_comp: "\ = r \ l" unfolding unit_def by simp lemma unit_eq [simp]: "\ x = r (l x)" by (simp add: unit_eq_comp) context begin text \Note that by flipping the roles of the left and rights functors, we obtain a flipped interpretation of @{locale order_functors}. In many cases, this allows us to obtain symmetric definitions and theorems for free. As such, in many cases, we do we do not explicitly state those free results but users can obtain them as needed by creating said flipped interpretation.\ interpretation flip : order_functors R L r l . definition "counit \ flip.unit" notation counit ("\") lemma counit_eq_comp: "\ = l \ r" unfolding counit_def flip.unit_def by simp lemma counit_eq [simp]: "\ x = l (r x)" by (simp add: counit_eq_comp) end context begin interpretation flip : order_functors R L r l . lemma flip_counit_eq_unit: "flip.counit = \" by (intro ext) simp lemma flip_unit_eq_counit: "flip.unit = \" by (intro ext) simp lemma inflationary_on_unit_if_left_rel_right_if_left_right_relI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "reflexive_on P (\\<^bsub>L\<^esub>)" and "\x y. P x \ l x \\<^bsub>R\<^esub> y \ x \\<^bsub>L\<^esub> r y" shows "inflationary_on P (\\<^bsub>L\<^esub>) \" using assms by (intro inflationary_onI) 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) - (auto dest!: inverse_onD) + (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/Partial_Equivalence_Relations.thy b/thys/Transport/HOL_Basics/Orders/Partial_Equivalence_Relations.thy --- a/thys/Transport/HOL_Basics/Orders/Partial_Equivalence_Relations.thy +++ b/thys/Transport/HOL_Basics/Orders/Partial_Equivalence_Relations.thy @@ -1,116 +1,129 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Partial Equivalence Relations\ theory Partial_Equivalence_Relations imports Binary_Relations_Symmetric Preorders begin definition "partial_equivalence_rel_on \ transitive_on \ symmetric_on" lemma partial_equivalence_rel_onI [intro]: assumes "transitive_on P R" and "symmetric_on P R" shows "partial_equivalence_rel_on P R" unfolding partial_equivalence_rel_on_def using assms by auto lemma partial_equivalence_rel_onE [elim]: assumes "partial_equivalence_rel_on P R" obtains "transitive_on P R" "symmetric_on P R" using assms unfolding partial_equivalence_rel_on_def by auto lemma partial_equivalence_rel_on_rel_self_if_rel_dom: assumes "partial_equivalence_rel_on (P :: 'a \ bool) (R :: 'a \ 'a \ bool)" and "P x" "P y" and "R x y" shows "R x x" using assms by (blast dest: symmetric_onD transitive_onD) lemma partial_equivalence_rel_on_rel_self_if_rel_codom: assumes "partial_equivalence_rel_on (P :: 'a \ bool) (R :: 'a \ 'a \ bool)" and "P x" "P y" and "R x y" shows "R y y" using assms by (blast dest: symmetric_onD transitive_onD) lemma partial_equivalence_rel_on_rel_inv_iff_partial_equivalence_rel_on [iff]: - "partial_equivalence_rel_on P R\ \ partial_equivalence_rel_on (P :: 'a \ bool) (R :: 'a \ _)" + "partial_equivalence_rel_on P R\ \ partial_equivalence_rel_on (P :: 'a \ bool) (R :: 'a \ 'a \ bool)" by blast -definition "(partial_equivalence_rel :: ('a \ _) \ bool) \ - partial_equivalence_rel_on (\ :: 'a \ bool)" +consts partial_equivalence_rel :: "'a \ bool" + +overloading + partial_equivalence_rel \ "partial_equivalence_rel :: ('a \ 'a \ bool) \ bool" +begin + definition "(partial_equivalence_rel :: ('a \ 'a \ bool) \ bool) \ partial_equivalence_rel_on (\ :: 'a \ bool)" +end lemma partial_equivalence_rel_eq_partial_equivalence_rel_on: - "(partial_equivalence_rel :: ('a \ _) \ bool) = partial_equivalence_rel_on (\ :: 'a \ bool)" + "(partial_equivalence_rel :: ('a \ 'a \ bool) \ bool) = partial_equivalence_rel_on (\ :: 'a \ bool)" unfolding partial_equivalence_rel_def .. +lemma partial_equivalence_rel_eq_partial_equivalence_rel_on_uhint [uhint]: + assumes "P \ \ :: 'a \ bool" + shows "(partial_equivalence_rel :: ('a \ 'a \ bool) \ bool) \ partial_equivalence_rel_on P" + using assms by (simp add: partial_equivalence_rel_eq_partial_equivalence_rel_on) + +context + fixes R :: "'a \ 'a \ bool" +begin + lemma partial_equivalence_relI [intro]: assumes "transitive R" and "symmetric R" shows "partial_equivalence_rel R" - unfolding partial_equivalence_rel_eq_partial_equivalence_rel_on using assms - by (intro partial_equivalence_rel_onI transitive_on_if_transitive symmetric_on_if_symmetric) + using assms by (urule partial_equivalence_rel_onI) lemma reflexive_on_in_field_if_partial_equivalence_rel: assumes "partial_equivalence_rel R" shows "reflexive_on (in_field R) R" using assms unfolding partial_equivalence_rel_eq_partial_equivalence_rel_on by (intro reflexive_onI) (blast intro: top1I partial_equivalence_rel_on_rel_self_if_rel_dom partial_equivalence_rel_on_rel_self_if_rel_codom) lemma partial_equivalence_relE [elim]: assumes "partial_equivalence_rel R" obtains "preorder_on (in_field R) R" "symmetric R" - using assms unfolding partial_equivalence_rel_eq_partial_equivalence_rel_on - by (elim partial_equivalence_rel_onE) + using assms by (urule (e) partial_equivalence_rel_onE where chained = insert) (auto intro: reflexive_on_in_field_if_partial_equivalence_rel simp flip: transitive_eq_transitive_on symmetric_eq_symmetric_on) lemma partial_equivalence_rel_on_if_partial_equivalence_rel: - fixes P :: "'a \ bool" and R :: "'a \ _" + fixes P :: "'a \ bool" assumes "partial_equivalence_rel R" shows "partial_equivalence_rel_on P R" using assms by (elim partial_equivalence_relE preorder_on_in_fieldE) (intro partial_equivalence_rel_onI transitive_on_if_transitive symmetric_on_if_symmetric) lemma partial_equivalence_rel_rel_inv_iff_partial_equivalence_rel [iff]: "partial_equivalence_rel R\ \ partial_equivalence_rel R" unfolding partial_equivalence_rel_eq_partial_equivalence_rel_on by blast corollary in_codom_eq_in_dom_if_partial_equivalence_rel: assumes "partial_equivalence_rel R" shows "in_codom R = in_dom R" using assms reflexive_on_in_field_if_partial_equivalence_rel in_codom_eq_in_dom_if_reflexive_on_in_field by auto lemma partial_equivalence_rel_rel_comp_self_eq_self: assumes "partial_equivalence_rel R" shows "(R \\ R) = R" using assms by (intro ext) (blast dest: symmetricD) lemma partial_equivalence_rel_if_partial_equivalence_rel_on_in_field: assumes "partial_equivalence_rel_on (in_field R) R" shows "partial_equivalence_rel R" using assms by (intro partial_equivalence_relI) (auto intro: transitive_if_transitive_on_in_field symmetric_if_symmetric_on_in_field) corollary partial_equivalence_rel_on_in_field_iff_partial_equivalence_rel [iff]: "partial_equivalence_rel_on (in_field R) R \ partial_equivalence_rel R" using partial_equivalence_rel_if_partial_equivalence_rel_on_in_field partial_equivalence_rel_on_if_partial_equivalence_rel by blast +end paragraph \Instantiations\ lemma partial_equivalence_rel_eq: "partial_equivalence_rel (=)" using transitive_eq symmetric_eq by (rule partial_equivalence_relI) -lemma partial_equivalence_rel_top: "partial_equivalence_rel \" +lemma partial_equivalence_rel_top: "partial_equivalence_rel (\ :: 'a \ 'a \ bool)" using transitive_top symmetric_top by (rule partial_equivalence_relI) end diff --git a/thys/Transport/HOL_Basics/Orders/Partial_Orders.thy b/thys/Transport/HOL_Basics/Orders/Partial_Orders.thy --- a/thys/Transport/HOL_Basics/Orders/Partial_Orders.thy +++ b/thys/Transport/HOL_Basics/Orders/Partial_Orders.thy @@ -1,61 +1,74 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Partial Orders\ theory Partial_Orders imports Binary_Relations_Antisymmetric Preorders begin definition "partial_order_on \ preorder_on \ antisymmetric_on" lemma partial_order_onI [intro]: assumes "preorder_on P R" and "antisymmetric_on P R" shows "partial_order_on P R" unfolding partial_order_on_def using assms by auto lemma partial_order_onE [elim]: assumes "partial_order_on P R" obtains "preorder_on P R" "antisymmetric_on P R" using assms unfolding partial_order_on_def by auto lemma transitive_if_partial_order_on_in_field: assumes "partial_order_on (in_field R) R" shows "transitive R" using assms by (elim partial_order_onE) (rule transitive_if_preorder_on_in_field) lemma antisymmetric_if_partial_order_on_in_field: assumes "partial_order_on (in_field R) R" shows "antisymmetric R" using assms by (elim partial_order_onE) (rule antisymmetric_if_antisymmetric_on_in_field) -definition "(partial_order :: ('a \ _) \ bool) \ partial_order_on (\ :: 'a \ bool)" +consts partial_order :: "'a \ bool" + +overloading + partial_order \ "partial_order :: ('a \ 'a \ bool) \ bool" +begin + definition "(partial_order :: ('a \ 'a \ bool) \ bool) \ partial_order_on (\ :: 'a \ bool)" +end lemma partial_order_eq_partial_order_on: - "(partial_order :: ('a \ _) \ bool) = partial_order_on (\ :: 'a \ bool)" + "(partial_order :: ('a \ 'a \ bool) \ bool) = partial_order_on (\ :: 'a \ bool)" unfolding partial_order_def .. +lemma partial_order_eq_partial_order_on_uhint [uhint]: + assumes "P \ \ :: 'a \ bool" + shows "(partial_order :: ('a \ 'a \ bool) \ bool) \ partial_order_on P" + using assms by (simp add: partial_order_eq_partial_order_on) + +context + fixes R :: "'a \ 'a \ bool" +begin + lemma partial_orderI [intro]: assumes "preorder R" and "antisymmetric R" shows "partial_order R" - unfolding partial_order_eq_partial_order_on using assms - by (intro partial_order_onI preorder_on_if_preorder antisymmetric_on_if_antisymmetric) + using assms by (urule partial_order_onI) lemma partial_orderE [elim]: assumes "partial_order R" obtains "preorder R" "antisymmetric R" - using assms unfolding partial_order_eq_partial_order_on - by (elim partial_order_onE) - (simp only: preorder_eq_preorder_on antisymmetric_eq_antisymmetric_on) + using assms by (urule (e) partial_order_onE) lemma partial_order_on_if_partial_order: - fixes P :: "'a \ bool" and R :: "'a \ _" + fixes P :: "'a \ bool" assumes "partial_order R" shows "partial_order_on P R" using assms by (elim partial_orderE) (intro partial_order_onI preorder_on_if_preorder antisymmetric_on_if_antisymmetric) +end end \ No newline at end of file diff --git a/thys/Transport/HOL_Basics/Orders/Preorders.thy b/thys/Transport/HOL_Basics/Orders/Preorders.thy --- a/thys/Transport/HOL_Basics/Orders/Preorders.thy +++ b/thys/Transport/HOL_Basics/Orders/Preorders.thy @@ -1,94 +1,108 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Preorders\ theory Preorders imports Binary_Relations_Reflexive Binary_Relations_Transitive begin definition "preorder_on \ reflexive_on \ transitive_on" lemma preorder_onI [intro]: assumes "reflexive_on P R" and "transitive_on P R" shows "preorder_on P R" unfolding preorder_on_def using assms by auto lemma preorder_onE [elim]: assumes "preorder_on P R" obtains "reflexive_on P R" "transitive_on P R" using assms unfolding preorder_on_def by auto lemma reflexive_on_if_preorder_on: assumes "preorder_on P R" shows "reflexive_on P R" using assms by (elim preorder_onE) lemma transitive_on_if_preorder_on: assumes "preorder_on P R" shows "transitive_on P R" using assms by (elim preorder_onE) lemma transitive_if_preorder_on_in_field: assumes "preorder_on (in_field R) R" shows "transitive R" using assms by (elim preorder_onE) (rule transitive_if_transitive_on_in_field) corollary preorder_on_in_fieldE [elim]: assumes "preorder_on (in_field R) R" obtains "reflexive_on (in_field R) R" "transitive R" using assms by (blast dest: reflexive_on_if_preorder_on transitive_if_preorder_on_in_field) lemma preorder_on_rel_inv_if_preorder_on [iff]: - "preorder_on P R\ \ preorder_on (P :: 'a \ bool) (R :: 'a \ _)" + "preorder_on P R\ \ preorder_on (P :: 'a \ bool) (R :: 'a \ 'a \ bool)" by auto lemma rel_if_all_rel_if_rel_if_reflexive_on: assumes "reflexive_on P R" and "\z. P z \ R x z \ R y z" and "P x" shows "R y x" using assms by blast lemma rel_if_all_rel_if_rel_if_reflexive_on': assumes "reflexive_on P R" and "\z. P z \ R z x \ R z y" and "P x" shows "R x y" using assms by blast -definition "(preorder :: ('a \ _) \ bool) \ preorder_on (\ :: 'a \ bool)" +consts preorder :: "'a \ bool" + +overloading + preorder \ "preorder :: ('a \ 'a \ bool) \ bool" +begin + definition "(preorder :: ('a \ 'a \ bool) \ bool) \ preorder_on (\ :: 'a \ bool)" +end lemma preorder_eq_preorder_on: - "(preorder :: ('a \ _) \ bool) = preorder_on (\ :: 'a \ bool)" + "(preorder :: ('a \ 'a \ bool) \ bool) = preorder_on (\ :: 'a \ bool)" unfolding preorder_def .. +lemma preorder_eq_preorder_on_uhint [uhint]: + assumes "P \ \ :: 'a \ bool" + shows "(preorder :: ('a \ 'a \ bool) \ bool) \ preorder_on P" + using assms by (simp add: preorder_eq_preorder_on) + +context + fixes R :: "'a \ 'a \ bool" +begin + lemma preorderI [intro]: assumes "reflexive R" and "transitive R" shows "preorder R" - unfolding preorder_eq_preorder_on using assms - by (intro preorder_onI reflexive_on_if_reflexive transitive_on_if_transitive) + using assms by (urule preorder_onI) lemma preorderE [elim]: assumes "preorder R" obtains "reflexive R" "transitive R" - using assms unfolding preorder_eq_preorder_on by (elim preorder_onE) - (simp only: reflexive_eq_reflexive_on transitive_eq_transitive_on) + using assms by (urule (e) preorder_onE) lemma preorder_on_if_preorder: - fixes P :: "'a \ bool" and R :: "'a \ _" + fixes P :: "'a \ bool" assumes "preorder R" shows "preorder_on P R" using assms by (elim preorderE) (intro preorder_onI reflexive_on_if_reflexive transitive_on_if_transitive) +end paragraph \Instantiations\ lemma preorder_eq: "preorder (=)" using reflexive_eq transitive_eq by (rule preorderI) end diff --git a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Order_Equivalence.thy b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Order_Equivalence.thy --- a/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Order_Equivalence.thy +++ b/thys/Transport/Transport/Compositions/Agree/Transport_Compositions_Agree_Order_Equivalence.thy @@ -1,131 +1,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" 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 "\R. inflationary_on P R\ \ deflationary_on P R" - and "\R. transitive R\ \ transitive R" + 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 lemma deflationary_on_in_field_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "deflationary_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "deflationary_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ r2 (l x) \\<^bsub>L2\<^esub> l1 x \ r2 (l x) \\<^bsub>R1\<^esub> l1 x" shows "deflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms by (intro inv.inflationary_on_in_field_unitI[simplified rel_inv_iff_rel]) auto end text \Relational Equivalence\ corollary rel_equivalence_on_in_field_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "rel_equivalence_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ l1 x \\<^bsub>L2\<^esub> r2 (l x) \ l1 x \\<^bsub>R1\<^esub> r2 (l x)" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ r2 (l x) \\<^bsub>L2\<^esub> l1 x \ r2 (l x) \\<^bsub>R1\<^esub> l1 x" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms by (intro rel_equivalence_onI inflationary_on_in_field_unitI deflationary_on_in_field_unitI) auto subsubsection \Counit\ text \Corresponding lemmas for the counit can be obtained by flipping the interpretation of the locale.\ subsubsection \Order Equivalence\ interpretation flip : transport_comp_agree R2 L2 r2 l2 R1 L1 r1 l1 rewrites "flip.g1.unit \ \\<^sub>2" and "flip.g2.unit \ \\<^sub>1" and "flip.unit \ \" by (simp_all only: g1.flip_unit_eq_counit g2.flip_unit_eq_counit flip_unit_eq_counit) lemma order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>o (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>R2\<^esub>)" and "\x y. x \\<^bsub>L1\<^esub> y \ l1 x \\<^bsub>R1\<^esub> l1 y \ l1 x \\<^bsub>L2\<^esub> l1 y" and "\x y. x \\<^bsub>R2\<^esub> y \ r2 x \\<^bsub>L2\<^esub> r2 y \ r2 x \\<^bsub>R1\<^esub> r2 y" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ l1 x \\<^bsub>L2\<^esub> r2 (l x) \ l1 x \\<^bsub>R1\<^esub> r2 (l x)" and "\x. in_field (\\<^bsub>L1\<^esub>) x \ r2 (l x) \\<^bsub>L2\<^esub> l1 x \ r2 (l x) \\<^bsub>R1\<^esub> l1 x" and "\x. in_field (\\<^bsub>R2\<^esub>) x \ r2 x \\<^bsub>R1\<^esub> l1 (r x) \ r2 x \\<^bsub>L2\<^esub> l1 (r x)" and "\x. in_field (\\<^bsub>R2\<^esub>) x \ l1 (r x) \\<^bsub>R1\<^esub> r2 x \ l1 (r x) \\<^bsub>L2\<^esub> r2 x" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalenceI rel_equivalence_on_in_field_unitI flip.rel_equivalence_on_in_field_unitI mono_wrt_rel_leftI flip.mono_wrt_rel_leftI dep_mono_wrt_relI) (auto elim!: g1.order_equivalenceE g2.order_equivalenceE) end context transport_comp_same begin lemma order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R1\<^esub>) \\<^sub>o (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>R2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (rule order_equivalenceI) auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Connection.thy b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Connection.thy --- a/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Connection.thy +++ b/thys/Transport/Transport/Compositions/Generic/Transport_Compositions_Generic_Galois_Connection.thy @@ -1,88 +1,87 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Connection\ theory Transport_Compositions_Generic_Galois_Connection imports Transport_Compositions_Generic_Galois_Property Transport_Compositions_Generic_Monotone 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 galois_connection_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_field (\\<^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_connectionI galois_prop_left_rightI mono_wrt_rel_leftI flip.mono_wrt_rel_leftI) (auto intro: inflationary_on_if_le_pred_if_inflationary_on in_field_if_in_dom in_field_if_in_codom) lemma galois_connection_left_rightI': assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>R1\<^esub>) \\<^sub>h (\\<^bsub>L1\<^esub>)) r1 l1" and "inflationary_on (in_dom (\\<^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>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>L2\<^esub>) \\<^sub>h (\\<^bsub>R2\<^esub>)) l2 r2" 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" using assms by (intro galois_connectionI galois_prop_left_rightI' mono_wrt_rel_leftI' flip.mono_wrt_rel_leftI') (auto elim!: preorder_on_in_fieldE intro!: reflexive_on_in_field_if_transitive_if_rel_equivalence_on intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) corollary galois_connection_left_right_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>L2\<^esub>) \\<^sub>G (\\<^bsub>R2\<^esub>)) l2 r2" 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_connection_left_rightI) (auto elim!: galois.galois_connectionE intro!: flip.t2.rel_equivalence_on_unit_if_reflexive_on_if_galois_equivalence t2.inflationary_on_unit_if_reflexive_on_if_galois_equivalence - intro: rel_equivalence_on_if_le_pred_if_rel_equivalence_on - in_field_if_in_codom) + intro: in_field_if_in_codom) corollary galois_connection_left_right_if_order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and "transitive (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>o (\\<^bsub>R2\<^esub>)) l2 r2" and "transitive (\\<^bsub>L2\<^esub>)" and "middle_compatible_codom" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_connection_left_rightI') (auto elim!: rel_equivalence_onE intro!: 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 t2.half_galois_prop_right_left_right_if_transitive_if_inflationary_on_if_mono_wrt_rel flip.t2.half_galois_prop_right_left_right_if_transitive_if_inflationary_on_if_mono_wrt_rel preorder_on_in_field_if_transitive_if_rel_equivalence_on rel_comp_comp_le_assms_if_in_codom_rel_comp_comp_leI intro: inflationary_on_if_le_pred_if_inflationary_on deflationary_on_if_le_pred_if_deflationary_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/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,195 @@ \<^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" 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" 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 "((\\<^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) 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 "((\\<^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 auto + 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/Examples/Prototype/Transport_Prototype.thy b/thys/Transport/Transport/Examples/Prototype/Transport_Prototype.thy --- a/thys/Transport/Transport/Examples/Prototype/Transport_Prototype.thy +++ b/thys/Transport/Transport/Examples/Prototype/Transport_Prototype.thy @@ -1,168 +1,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] 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, + 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) |> K) (Transport_Unification_Combine.default_metadata Transport_Unification_Hints.binding)\]] paragraph \Prototype\ ML_file\transport.ML\ declare transport_Dep_Fun_Rel.transport_defs[trp_def] transport_Fun_Rel.transport_defs[trp_def] declare (*dependent case currently disabled by default since they easily make the unifier enumerate many undesired instantiations*) (* transport_Dep_Fun_Rel.partial_equivalence_rel_equivalenceI[per_intro] *) (* transport.rel_if_partial_equivalence_rel_equivalence_if_iff_if_partial_equivalence_rel_equivalenceI[rotated, per_intro] transport_Dep_Fun_Rel_no_dep_fun.partial_equivalence_rel_equivalenceI [ML_Krattr \Conversion_Util.move_prems_to_front_conv [1] |> Conversion_Util.thm_conv\, ML_Krattr \Conversion_Util.move_prems_to_front_conv [2,3] |> Conversion_Util.thm_conv\, per_intro] *) transport_Fun_Rel.partial_equivalence_rel_equivalenceI[rotated, per_intro] transport_eq_id.partial_equivalence_rel_equivalenceI[per_intro] transport_eq_restrict_id.partial_equivalence_rel_equivalence[per_intro] declare transport_id.left_Galois_eq_left[trp_relator_rewrite] transport_Fun_Rel.left_Galois_eq_Fun_Rel_left_Galois[trp_relator_rewrite] end diff --git a/thys/Transport/Transport/Examples/Prototype/Transport_Rel_If.thy b/thys/Transport/Transport/Examples/Prototype/Transport_Rel_If.thy --- a/thys/Transport/Transport/Examples/Prototype/Transport_Rel_If.thy +++ b/thys/Transport/Transport/Examples/Prototype/Transport_Rel_If.thy @@ -1,218 +1,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" and "B \ B'" shows "([x y \ (rel_if B R)] \\<^sub>m (rel_if B' (S x y))) f" using assms by (intro dep_mono_wrt_relI) auto -end - corollary reflexive_rel_if_if_reflexiveI [intro]: assumes "B \ reflexive R" shows "reflexive (rel_if B R)" using assms unfolding reflexive_eq_reflexive_on by blast corollary transitive_rel_if_if_transitiveI [intro]: assumes "B \ transitive R" shows "transitive (rel_if B R)" - using assms unfolding transitive_eq_transitive_on by blast + 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" shows "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" proof - have per2I: "((\\<^bsub>L2 x1 (r1 x2')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2 r2" if hyps: "x1 \\<^bsub>L1\<^esub> x2" "x2 \<^bsub>L1\<^esub>\ x1'" "x1' \\<^bsub>R1\<^esub> x2'" for x1 x2 x1' x2' proof - from hyps have "x1 \<^bsub>L1\<^esub>\ x2'" using per_equiv1 t1.left_Galois_if_left_Galois_if_left_relI t1.left_Galois_if_right_rel_if_left_GaloisI by fast with per_equiv2 show ?thesis by blast qed have "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) (\_ _. l2)" by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI Dep_Fun_Rel_predI rel_if_if_impI) (auto 10 0 dest!: per2I) moreover have "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) (\_ _. r2)" by (intro dep_mono_wrt_relI Dep_Fun_Rel_relI Dep_Fun_Rel_predI rel_if_if_impI) (auto 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,408 +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 (*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) 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_Unification.e_match Unification_Util.match_types (Higher_Order_Pattern_Decomp_Unification.e_match unif Unification_Combinator.fail_match) - Unification_Combinator.fail_match binders + 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,85 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Example Transports for Dependent Function Relator\ theory Transport_Dep_Fun_Rel_Examples imports Transport_Prototype Transport_Syntax "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] 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) (-) (-)" 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] \ (=)" (*fastforce discharges the remaining side-conditions*) 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) (!) (!)" by (fastforce simp: list_all2_lengthD list_all2_nthD2) context - fixes R :: "'a \ _" assumes [per_intro]: "partial_equivalence_rel R" + 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) (*fastforce discharges the remaining side-conditions*) (fastforce simp: list_all2_lengthD elim: iarray.rel_cases)+ end end end diff --git a/thys/Transport/Transport/Functions/Reflexive_Relator.thy b/thys/Transport/Transport/Functions/Reflexive_Relator.thy --- a/thys/Transport/Transport/Functions/Reflexive_Relator.thy +++ b/thys/Transport/Transport/Functions/Reflexive_Relator.thy @@ -1,275 +1,276 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Reflexive Relator\ theory Reflexive_Relator imports Galois_Equivalences Galois_Relator begin definition "Refl_Rel R x y \ R x x \ R y y \ R x y" bundle Refl_Rel_syntax begin notation Refl_Rel ("(_\<^sup>\)" [1000]) end bundle no_Refl_Rel_syntax begin no_notation Refl_Rel ("(_\<^sup>\)" [1000]) end unbundle Refl_Rel_syntax lemma Refl_RelI [intro]: assumes "R x x" and "R y y" and "R x y" shows "R\<^sup>\ x y" using assms unfolding Refl_Rel_def by blast lemma Refl_Rel_selfI [intro]: assumes "R x x" shows "R\<^sup>\ x x" using assms by blast lemma Refl_RelE [elim]: assumes "R\<^sup>\ x y" obtains "R x x" "R y y" "R x y" using assms unfolding Refl_Rel_def by blast lemma Refl_Rel_reflexive_on_in_field [iff]: "reflexive_on (in_field R\<^sup>\) R\<^sup>\" by (rule reflexive_onI) auto lemma Refl_Rel_le_self [iff]: "R\<^sup>\ \ R" by blast lemma Refl_Rel_eq_self_if_reflexive_on [simp]: assumes "reflexive_on (in_field R) R" shows "R\<^sup>\ = R" using assms by blast lemma reflexive_on_in_field_if_Refl_Rel_eq_self: assumes "R\<^sup>\ = R" shows "reflexive_on (in_field R) R" by (fact Refl_Rel_reflexive_on_in_field[of R, simplified assms]) corollary Refl_Rel_eq_self_iff_reflexive_on: "R\<^sup>\ = R \ reflexive_on (in_field R) R" using Refl_Rel_eq_self_if_reflexive_on reflexive_on_in_field_if_Refl_Rel_eq_self by blast lemma Refl_Rel_Refl_Rel_eq [simp]: "(R\<^sup>\)\<^sup>\ = R\<^sup>\" by (intro ext) auto lemma rel_inv_Refl_Rel_eq [simp]: "(R\<^sup>\)\ = (R\)\<^sup>\" by (intro ext iffI Refl_RelI rel_invI) auto lemma Refl_Rel_transitive_onI [intro]: assumes "transitive_on (P :: 'a \ bool) (R :: 'a \ _)" shows "transitive_on P R\<^sup>\" using assms by (intro transitive_onI) (blast dest: transitive_onD) corollary Refl_Rel_transitiveI [intro]: assumes "transitive R" shows "transitive R\<^sup>\" using assms by blast corollary Refl_Rel_preorder_onI: assumes "transitive_on P R" and "P \ in_field R\<^sup>\" shows "preorder_on P R\<^sup>\" using assms by (intro preorder_onI reflexive_on_if_le_pred_if_reflexive_on[where ?P="in_field R\<^sup>\" and ?P'=P]) auto corollary Refl_Rel_preorder_on_in_fieldI [intro]: assumes "transitive R" shows "preorder_on (in_field R\<^sup>\) R\<^sup>\" using assms by (intro Refl_Rel_preorder_onI) auto lemma Refl_Rel_symmetric_onI [intro]: assumes "symmetric_on (P :: 'a \ bool) (R :: 'a \ _)" shows "symmetric_on P R\<^sup>\" using assms by (intro symmetric_onI) (auto dest: symmetric_onD) lemma Refl_Rel_symmetricI [intro]: assumes "symmetric R" shows "symmetric R\<^sup>\" - using assms by (fold symmetric_on_in_field_iff_symmetric) - (blast intro: symmetric_on_if_le_pred_if_symmetric_on) + 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 context galois begin interpretation gR : galois "(\\<^bsub>L\<^esub>)\<^sup>\" "(\\<^bsub>R\<^esub>)\<^sup>\" l r . lemma Galois_Refl_RelI: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "in_field (\\<^bsub>L\<^esub>)\<^sup>\ x" and "in_field (\\<^bsub>R\<^esub>)\<^sup>\ y" and "in_codom (\\<^bsub>R\<^esub>) y \ x \<^bsub>L\<^esub>\ y" shows "(galois_rel.Galois ((\\<^bsub>L\<^esub>)\<^sup>\) ((\\<^bsub>R\<^esub>)\<^sup>\) r) x y" using assms by (intro gR.left_GaloisI in_codomI Refl_Rel_app_rightI[where ?f=r]) auto lemma half_galois_prop_left_Refl_Rel_left_rightI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \<^sub>h\ (\\<^bsub>R\<^esub>)\<^sup>\) l r" using assms by (intro gR.half_galois_prop_leftI Refl_RelI) (auto elim!: in_codomE gR.left_GaloisE Refl_RelE) interpretation flip_inv : galois "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" r l rewrites "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) \ ((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>))" and "\R. (R\)\<^sup>\ \ (R\<^sup>\)\" and "\R S f g. (R\ \<^sub>h\ S\) f g \ (S \\<^sub>h R) g f" by (simp_all add: galois_prop.half_galois_prop_left_rel_inv_iff_half_galois_prop_right) lemma half_galois_prop_right_Refl_Rel_right_leftI: assumes "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \\<^sub>h (\\<^bsub>R\<^esub>)\<^sup>\) l r" using assms by (fact flip_inv.half_galois_prop_left_Refl_Rel_left_rightI) corollary galois_prop_Refl_Rel_left_rightI: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \ (\\<^bsub>R\<^esub>)\<^sup>\) l r" using assms by (intro gR.galois_propI half_galois_prop_left_Refl_Rel_left_rightI half_galois_prop_right_Refl_Rel_right_leftI) auto lemma galois_connection_Refl_Rel_left_rightI: assumes "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \ (\\<^bsub>R\<^esub>)\<^sup>\) l r" using assms by (intro gR.galois_connectionI galois_prop_Refl_Rel_left_rightI) auto lemma galois_equivalence_Refl_RelI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \\<^sub>G (\\<^bsub>R\<^esub>)\<^sup>\) l r" proof - interpret flip : galois R L r l . show ?thesis using assms by (intro gR.galois_equivalenceI galois_connection_Refl_Rel_left_rightI flip.galois_prop_Refl_Rel_left_rightI) auto qed end context order_functors begin lemma inflationary_on_in_field_Refl_Rel_left: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "inflationary_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "inflationary_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms by (intro inflationary_onI Refl_RelI) (auto 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 R. inflationary_on P R\ \ deflationary_on P R" + and "\(P :: 'c \ bool) (R :: 'd \ 'c \ bool). + (inflationary_on P R\ :: ('c \ 'd) \ bool) \ deflationary_on P R" by simp_all lemma deflationary_on_in_field_Refl_Rel_leftI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "deflationary_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "deflationary_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms by (fact inv.inflationary_on_in_field_Refl_Rel_left') lemma deflationary_on_in_field_Refl_RelI_left': assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "deflationary_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "deflationary_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms by (fact inv.inflationary_on_in_field_Refl_Rel_left) lemma rel_equivalence_on_in_field_Refl_Rel_leftI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "rel_equivalence_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms by (intro rel_equivalence_onI inflationary_on_in_field_Refl_Rel_left deflationary_on_in_field_Refl_Rel_leftI) auto lemma rel_equivalence_on_in_field_Refl_Rel_leftI': assumes "((\\<^bsub>L\<^esub>) \\<^sub>m (\\<^bsub>R\<^esub>)) l" and "((\\<^bsub>R\<^esub>) \\<^sub>m (\\<^bsub>L\<^esub>)) r" and "rel_equivalence_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)\<^sup>\) (\\<^bsub>L\<^esub>)\<^sup>\ \" using assms by (intro rel_equivalence_onI inflationary_on_in_field_Refl_Rel_left' deflationary_on_in_field_Refl_RelI_left') auto interpretation oR : order_functors "(\\<^bsub>L\<^esub>)\<^sup>\" "(\\<^bsub>R\<^esub>)\<^sup>\" l r . lemma order_equivalence_Refl_RelI: assumes "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" shows "((\\<^bsub>L\<^esub>)\<^sup>\ \\<^sub>o (\\<^bsub>R\<^esub>)\<^sup>\) l r" proof - interpret flip : galois R L r l rewrites "flip.unit \ \" by (simp only: flip_unit_eq_counit) show ?thesis using assms by (intro oR.order_equivalenceI mono_wrt_rel_Refl_Rel_Refl_Rel_if_mono_wrt_rel rel_equivalence_on_in_field_Refl_Rel_leftI flip.rel_equivalence_on_in_field_Refl_Rel_leftI) (auto intro: rel_equivalence_on_if_le_pred_if_rel_equivalence_on in_field_if_in_dom) qed end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions_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,474 @@ \<^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>)) (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) 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 \ _) (P :: 'z \ bool). reflexive_on P R\ \ reflexive_on P R" - and "\R x1 x2. transitive (flip2 R x1 x2) \ 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 "\(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>))\" 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) 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 "\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>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" 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" 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" 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 "\x. x \\<^bsub>L1\<^esub> x \ ([in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" 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>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" 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 "\x. x \\<^bsub>L1\<^esub> x \ ([in_dom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)] \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ([in_codom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)] \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using assms by (intro galois_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) 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) 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,865 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Relator\ theory Transport_Functions_Galois_Relator imports Transport_Functions_Relation_Simplifications begin paragraph \Dependent Function Relator\ context transport_Dep_Fun_Rel begin interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.counit \ \\<^sub>1" by (simp only: t1.flip_counit_eq_unit) lemma Dep_Fun_Rel_left_Galois_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and mono_r2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and L2_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and ge_L2_r2_le2: "\x x' y'. x \<^bsub>L1\<^esub>\ x' \ in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>) y' \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x (l1 x)\<^esub> y') \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x x'\<^esub> y')" and trans_L2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" and "f \<^bsub>L\<^esub>\ g" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" proof (intro Dep_Fun_Rel_relI) fix x x' assume "x \<^bsub>L1\<^esub>\ x'" show "f x \<^bsub>L2 x x'\<^esub>\ g x'" proof (intro t2.left_GaloisI) from \x \<^bsub>L1\<^esub>\ x'\ \((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1\ have "x \\<^bsub>L1\<^esub> r1 x'" "l1 x \\<^bsub>R1\<^esub> x'" by auto with \g \\<^bsub>R\<^esub> g\ have "g (l1 x) \\<^bsub>R2 (l1 x) x'\<^esub> g x'" by blast then show "in_codom (\\<^bsub>R2 (l1 x) x'\<^esub>) (g x')" by blast from \f \<^bsub>L\<^esub>\ g\ have "f \\<^bsub>L\<^esub> r g" by blast moreover from refl_L1 \x \<^bsub>L1\<^esub>\ x'\ have "x \\<^bsub>L1\<^esub> x" by blast ultimately have "f x \\<^bsub>L2 x x\<^esub> (r g) x" by blast with L2_le2 \x \\<^bsub>L1\<^esub> r1 x'\ have "f x \\<^bsub>L2 x (r1 x')\<^esub> (r g) x" by blast then have "f x \\<^bsub>L2 x (r1 x')\<^esub> r2\<^bsub>x (l1 x)\<^esub> (g (l1 x))" by simp with ge_L2_r2_le2 have "f x \\<^bsub>L2 x (r1 x')\<^esub> r2\<^bsub>x x'\<^esub> (g (l1 x))" using \x \<^bsub>L1\<^esub>\ x'\ \g (l1 x) \\<^bsub>R2 (l1 x) x'\<^esub> _\ by blast moreover have "... \\<^bsub>L2 x (r1 x')\<^esub> r2\<^bsub>x x'\<^esub> (g x')" using mono_r2 \x \<^bsub>L1\<^esub>\ x'\ \g (l1 x) \\<^bsub>R2 (l1 x) x'\<^esub> g x'\ by blast ultimately show "f x \\<^bsub>L2 x (r1 x')\<^esub> r2\<^bsub>x x'\<^esub> (g x')" using trans_L2 \x \<^bsub>L1\<^esub>\ x'\ by blast qed qed lemma left_rel_right_if_Dep_Fun_Rel_left_GaloisI: assumes mono_l1: "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and half_galois_prop_right1: "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and L2_unit_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and ge_L2_r2_le1: "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and rel_f_g: "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "f \\<^bsub>L\<^esub> r g" proof (intro left_relI) fix x1 x2 assume "x1 \\<^bsub>L1\<^esub> x2" with mono_l1 half_galois_prop_right1 have "x1 \<^bsub>L1\<^esub>\ l1 x2" by (intro t1.left_Galois_left_if_left_relI) auto with rel_f_g have "f x1 \<^bsub>L2 x1 (l1 x2)\<^esub>\ g (l1 x2)" by blast then have "in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) (g (l1 x2))" "f x1 \\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub> r2\<^bsub>x1 (l1 x2)\<^esub> (g (l1 x2))" by auto with L2_unit_le2 \x1 \\<^bsub>L1\<^esub> x2\ have "f x1 \\<^bsub>L2 x1 x2\<^esub> r2\<^bsub>x1 (l1 x2)\<^esub> (g (l1 x2))" by blast with ge_L2_r2_le1 \x1 \\<^bsub>L1\<^esub> x2\ \in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) (g (l1 x2))\ have "f x1 \\<^bsub>L2 x1 x2\<^esub> r2\<^bsub>x2 (l1 x2)\<^esub> (g (l1 x2))" by blast then show "f x1 \\<^bsub>L2 x1 x2\<^esub> r g x2" by simp qed lemma left_Galois_if_Dep_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and "in_codom (\\<^bsub>R\<^esub>) g" and "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro left_GaloisI left_rel_right_if_Dep_Fun_Rel_left_GaloisI) auto lemma left_right_rel_if_Dep_Fun_Rel_left_GaloisI: assumes mono_r1: "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and half_galois_prop_left2: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and R2_le1: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and R2_l2_le1: "\x1' x2' y. x1' \\<^bsub>R1\<^esub> x2' \ in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) y \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub> y) \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x1' (r1 x1')\<^esub> y)" and rel_f_g: "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "l f \\<^bsub>R\<^esub> g" proof (rule flip.left_relI) fix x1' x2' assume "x1' \\<^bsub>R1\<^esub> x2'" with mono_r1 have "r1 x1' \<^bsub>L1\<^esub>\ x2'" by blast with rel_f_g have "f (r1 x1') \<^bsub>L2 (r1 x1') x2'\<^esub>\ g x2'" by blast with half_galois_prop_left2[OF \x1' \\<^bsub>R1\<^esub> x2'\] have "l2\<^bsub>x2' (r1 x1')\<^esub> (f (r1 x1')) \\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub> g x2'" by auto with R2_le1 \x1' \\<^bsub>R1\<^esub> x2'\ have "l2\<^bsub>x2' (r1 x1')\<^esub> (f (r1 x1')) \\<^bsub>R2 x1' x2'\<^esub> g x2'" by blast with R2_l2_le1 \x1' \\<^bsub>R1\<^esub> x2'\ \f (r1 x1') \<^bsub>L2 (r1 x1') x2'\<^esub>\ _\ have "l2\<^bsub>x1' (r1 x1')\<^esub> (f (r1 x1')) \\<^bsub>R2 x1' x2'\<^esub> g x2'" by blast then show "l f x1' \\<^bsub>R2 x1' x2'\<^esub> g x2'" by simp qed lemma left_Galois_if_Dep_Fun_Rel_left_GaloisI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and "\x1' x2' y. x1' \\<^bsub>R1\<^esub> x2' \ in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) y \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub> y) \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x1' (r1 x1')\<^esub> y)" and "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro left_Galois_if_Dep_Fun_Rel_left_GaloisI in_codomI[where ?x="l f"]) (auto intro!: left_right_rel_if_Dep_Fun_Rel_left_GaloisI) lemma left_Galois_iff_Dep_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and "\x x' y'. x \<^bsub>L1\<^esub>\ x' \ in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>) y' \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x (l1 x)\<^esub> y') \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x x'\<^esub> y')" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro iffI) (auto intro!: Dep_Fun_Rel_left_Galois_if_left_GaloisI left_Galois_if_Dep_Fun_Rel_left_GaloisI) lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI: assumes "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" using assms by blast lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI': assumes "\x x'. x \<^bsub>L1\<^esub>\ x' \ ([in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>)] \ (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "\x x' y'. x \<^bsub>L1\<^esub>\ x' \ in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>) y' \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x (l1 x)\<^esub> y') \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x x'\<^esub> y')" using assms by blast lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_codom_rightI: assumes mono_l1: "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and half_galois_prop_right1: "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and L2_le_unit2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "x1 \\<^bsub>L1\<^esub> x2" shows "([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" proof (intro Dep_Fun_Rel_predI) from mono_l1 half_galois_prop_right1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "l1 x2 \\<^bsub>R1\<^esub> l1 x2" "x2 \<^bsub>L1\<^esub>\ l1 x2" by (blast intro: t1.left_Galois_left_if_left_relI)+ fix y' assume "in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y'" with Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x1 \\<^bsub>L1\<^esub> x2\] \l1 x2 \\<^bsub>R1\<^esub> l1 x2\] have "r2\<^bsub>x1 (l1 x2)\<^esub> y' \\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub> r2\<^bsub>x2 (l1 x2)\<^esub> y'" using \x2 \<^bsub>L1\<^esub>\ l1 x2\ by (auto dest: in_field_if_in_codom) with L2_le_unit2 \x1 \\<^bsub>L1\<^esub> x2\ show "r2\<^bsub>x1 (l1 x2)\<^esub> y' \\<^bsub>L2 x1 x2\<^esub> r2\<^bsub>x2 (l1 x2)\<^esub> y'" by blast qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_dom_rightI: assumes mono_l1: "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and half_galois_prop_right1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "x \<^bsub>L1\<^esub>\ x'" shows "([in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>)] \ (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>) (r2\<^bsub>x x'\<^esub>)" proof - from mono_l1 half_galois_prop_right1 refl_L1 \x \<^bsub>L1\<^esub>\ x'\ have "x \\<^bsub>L1\<^esub> x" "l1 x \\<^bsub>R1\<^esub> x'" "x \<^bsub>L1\<^esub>\ l1 x" by (auto intro!: t1.half_galois_prop_leftD t1.left_Galois_left_if_left_relI) with Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x \\<^bsub>L1\<^esub> x\] \l1 x \\<^bsub>R1\<^esub> x'\] show ?thesis by blast qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_GaloisI left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI' left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_dom_rightI left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_codom_rightI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI: assumes refl_L1: "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and mono_L2: "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "x1 \\<^bsub>L1\<^esub> x2" shows "(\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" proof - from refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "x1 \\<^bsub>L1\<^esub> x1" by blast with dep_mono_wrt_relD[OF dep_mono_wrt_predD[OF mono_L2] \x1 \\<^bsub>L1\<^esub> x2\] show "(\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" by auto qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assmI: assumes mono_l1: "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and half_galois_prop_right1: "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and antimono_L2: "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "x1 \\<^bsub>L1\<^esub> x2" shows "(\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" proof - from mono_l1 half_galois_prop_right1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast intro: t1.rel_unit_if_reflexive_on_if_half_galois_prop_right_if_mono_wrt_rel) with refl_L1 have "\\<^sub>1 x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by blast with dep_mono_wrt_relD[OF dep_mono_wrt_predD[OF antimono_L2] \x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2\] show "(\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" using \x1 \\<^bsub>L1\<^esub> x2\ by auto qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assmI left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom in_field_if_in_dom) corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI') auto interpretation flip_inv : galois "(\\<^bsub>R1\<^esub>)" "(\\<^bsub>L1\<^esub>)" r1 l1 . lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assm_if_galois_equivI: assumes galois_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and mono_L2: "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "x1 \\<^bsub>L1\<^esub> x2" shows "(\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" proof - from refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "x1 \\<^bsub>L1\<^esub> x1" by blast from galois_equiv1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "\\<^sub>1 x2 \\<^bsub>L1\<^esub> x2" by (intro flip.t1.counit_rel_if_reflexive_on_if_half_galois_prop_left_if_mono_wrt_rel) blast+ have "x1 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (rule t1.rel_unit_if_left_rel_if_mono_wrt_relI) (insert galois_equiv1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\, auto) with dep_mono_wrt_relD[OF dep_mono_wrt_predD[OF mono_L2] \\\<^sub>1 x2 \\<^bsub>L1\<^esub> x2\] show "(\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" by auto qed lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_monoI left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assm_if_galois_equivI) auto corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI reflexive_on_in_field_mono_assm_left2I) auto corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI') auto corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI transitive_left2_if_preorder_equivalenceI) (auto 5 0) subparagraph \Simplification of Restricted Function Relator\ lemma Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2' y. x1' \\<^bsub>R1\<^esub> x2' \ in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) y \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub> y) \ (\\<^bsub>R2 x1' x2'\<^esub>) (l2\<^bsub>x1' (r1 x1')\<^esub> y)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))" - using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI + 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. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))" using assms by (intro Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI reflexive_on_in_field_mono_assm_left2I left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI mono_wrt_rel_left_in_dom_mono_left_assm galois_connection_left_right_if_galois_connection_mono_assms_leftI galois_connection_left_right_if_galois_connection_mono_assms_rightI left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI) auto text \Simplification of Restricted Function Relator for Nested Transports\ lemma Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq: fixes S :: "'a1 \ 'a2 \ 'b1 \ 'b2 \ bool" assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (S x x')\\<^bsub>in_dom (\\<^bsub>L2 x (r1 x')\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R2 (l1 x) x'\<^esub>)\<^esub>) \\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ S x x')\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" (is "?lhs = ?rhs") proof - have "?lhs = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (S x x')\\<^bsub>in_codom (\\<^bsub>R2 (l1 x) x'\<^esub>)\<^esub>) \\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - by (subst bin_rel_restrict_left_right_eq_restrict_right_left, + 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 bin_rel_restrict_left_right_eq_restrict_right_left, + 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: bin_rel_restrict_left_right_eq_restrict_right_left) + 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 corollary left_Galois_if_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "in_codom (\\<^bsub>R\<^esub>) g" and "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro tdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI) simp_all lemma left_Galois_if_Fun_Rel_left_GaloisI': assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \<^sub>h\ (\\<^bsub>R2\<^esub>)) l2 r2" and "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro tdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI') simp_all corollary left_Galois_iff_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "transitive (\\<^bsub>L2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "f \<^bsub>L\<^esub>\ g \ ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" using assms by (intro tdfr.left_Galois_iff_Dep_Fun_Rel_left_GaloisI) simp_all subparagraph \Simplification of Restricted Function Relator\ lemma Fun_Rel_left_Galois_restrict_left_right_eq_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \<^sub>h\ (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))" using assms by (intro tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI) simp_all text \Simplification of Restricted Function Relator for Nested Transports\ lemma Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq: fixes S :: "'b1 \ 'b2 \ bool" assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" shows "((\<^bsub>L1\<^esub>\) \ S\\<^bsub>in_dom (\\<^bsub>L2\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R2\<^esub>)\<^esub>)\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ((\<^bsub>L1\<^esub>\) \ S)\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" using assms by (intro tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq) simp_all end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin lemma Dep_Fun_Rel_left_Galois_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x x' y'. x \<^bsub>L1\<^esub>\ x' \ in_dom (\\<^bsub>R2 (l1 x) x'\<^esub>) y' \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x (l1 x)\<^esub> y') \ (\\<^bsub>L2 x (r1 x')\<^esub>) (r2\<^bsub>x x'\<^esub> y')" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "f \<^bsub>L\<^esub>\ g" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel by (intro tdfr.Dep_Fun_Rel_left_Galois_if_left_GaloisI tdfr.left_GaloisI) (auto elim!: galois_rel.left_GaloisE in_codomE) lemma left_Galois_if_Dep_Fun_Rel_left_GaloisI: assumes "(tdfr.R \\<^sub>m tdfr.L) r" and "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2 y'. x1 \\<^bsub>L1\<^esub> x2 \ in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y' \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub> y')" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" and "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel by (intro tdfr.Galois_Refl_RelI tdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI) (auto simp: in_codom_eq_in_dom_if_reflexive_on_in_field) lemma left_Galois_iff_Dep_Fun_Rel_left_GaloisI: assumes "(tdfr.R \\<^sub>m tdfr.L) r" and "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro iffI Dep_Fun_Rel_left_Galois_if_left_GaloisI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI' tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_dom_rightI) (auto intro!: left_Galois_if_Dep_Fun_Rel_left_GaloisI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_codom_rightI intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) lemma left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI: assumes galois_conn1: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and L2_le_unit2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and trans_L2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" (is "?lhs \ ?rhs") proof - have "(\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub> y') \ (\\<^bsub>L2 x1 x2\<^esub>) (r2\<^bsub>x1 (l1 x1)\<^esub> y')" if hyps: "x1 \\<^bsub>L1\<^esub> x2" "in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y'" for x1 x2 y' proof - have "([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" proof (intro Dep_Fun_Rel_predI) from galois_conn1 refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have "x1 \\<^bsub>L1\<^esub> x1" "l1 x1 \\<^bsub>R1\<^esub> l1 x2" "x1 \<^bsub>L1\<^esub>\ l1 x1" by (blast intro: t1.left_Galois_left_if_left_relI)+ fix y' assume "in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>) y'" with Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x1 \\<^bsub>L1\<^esub> x1\] \l1 x1 \\<^bsub>R1\<^esub> l1 x2\] have "r2\<^bsub>x1 (l1 x1)\<^esub> y' \\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub> r2\<^bsub>x1 (l1 x2)\<^esub> y'" using \x1 \<^bsub>L1\<^esub>\ l1 x1\ by (auto dest: in_field_if_in_dom) with L2_le_unit2 \x1 \\<^bsub>L1\<^esub> x2\ show "r2\<^bsub>x1 (l1 x1)\<^esub> y' \\<^bsub>L2 x1 x2\<^esub> r2\<^bsub>x1 (l1 x2)\<^esub> y'" by blast qed with hyps show ?thesis using trans_L2 by blast qed then show ?thesis using assms using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_GaloisI tdfr.mono_wrt_rel_rightI tdfr.mono_wrt_rel_right2_if_mono_wrt_rel_right2_if_left_GaloisI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_ge_left_rel2_assmI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_mono_assm_in_codom_rightI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom) qed corollary left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI': assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" (is "?lhs \ ?rhs") using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assmI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI) auto corollary left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_mono_if_galois_connectionI: assumes "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 \ \] \\<^sub>m [x2 _ \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x2] \\<^sub>m (\)) L2" and "([x1 \ \] \\<^sub>m [x2 x3 \ (\\<^bsub>L1\<^esub>) | (x1 \\<^bsub>L1\<^esub> x2 \ x3 \\<^bsub>L1\<^esub> \\<^sub>1 x2)] \\<^sub>m (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI + 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. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_mono_if_galois_connectionI tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_le_assmI tdfr.reflexive_on_in_field_mono_assm_left2I tdfr.left_Galois_iff_Dep_Fun_Rel_left_Galois_left_rel2_unit1_le_assm_if_galois_equivI) auto theorem left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_galois_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>G (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI + 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. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_galois_equivalenceI) auto corollary left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI + 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 "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\)) f g" using assms by (intro left_Galois_iff_Dep_Fun_Rel_left_Galois_if_preorder_equivalenceI tdfr.transitive_left2_if_preorder_equivalenceI) (auto 5 0) corollary left_Galois_eq_Dep_Fun_Rel_left_Galois_restrict_if_preorder_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" shows "(\<^bsub>L\<^esub>\) = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI + 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. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))" using assms by (auto simp only: left_rel_eq_tdfr_left_rel_if_reflexive_on right_rel_eq_tdfr_right_rel_if_reflexive_on intro!: tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI') interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.unit \ \\<^sub>1" by (simp only: t1.flip_unit_eq_counit) lemma Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ((\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>) \<^sub>h\ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (r2\<^bsub>(r1 x1') x2'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and PERS: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>L2 x1 x2\<^esub>)" "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ partial_equivalence_rel (\\<^bsub>R2 x1' x2'\<^esub>)" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ (\<^bsub>L2 x x'\<^esub>\))" using assms by (intro Dep_Fun_Rel_left_Galois_restrict_left_right_eq_Dep_Fun_Rel_left_Galois_if_reflexive_onI tdfr.reflexive_on_in_field_left_if_equivalencesI flip.reflexive_on_in_field_left_if_equivalencesI tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI flip.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI) (auto dest!: PERS) text \Simplification of Restricted Function Relator for Nested Transports\ lemma Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq: fixes S :: "'a1 \ 'a2 \ 'b1 \ 'b2 \ bool" assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" shows "([x x' \ (\<^bsub>L1\<^esub>\)] \ (S x x')\\<^bsub>in_dom (\\<^bsub>L2 x (r1 x')\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R2 (l1 x) x'\<^esub>)\<^esub>) \\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub> = ([x x' \ (\<^bsub>L1\<^esub>\)] \ S x x')\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" (is "?lhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> = ?rhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub>") proof (intro ext) fix f g have "?lhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> f g \ ?lhs f g \ ?DL f \ ?CR g" by blast also have "... \ ?lhs\\<^bsub>in_dom tdfr.L\<^esub>\\<^bsub>in_codom tdfr.R\<^esub> f g \ ?DL f \ ?CR g" unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel by blast also with assms have "... \ ?rhs\\<^bsub>in_dom tdfr.L\<^esub>\\<^bsub>in_codom tdfr.R\<^esub> f g \ ?DL f \ ?CR g" by (simp only: tdfr.Dep_Fun_Rel_left_Galois_restrict_left_right_restrict_left_right_eq) also have "... \ ?rhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> f g" unfolding left_rel_eq_tdfr_left_Refl_Rel right_rel_eq_tdfr_right_Refl_Rel by blast finally show "?lhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> f g \ ?rhs\\<^bsub>?DL\<^esub>\\<^bsub>?CR\<^esub> f g" . qed end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin corollary Fun_Rel_left_Galois_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \<^sub>h\ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) (r2)" and "transitive (\\<^bsub>L2\<^esub>)" and "f \<^bsub>L\<^esub>\ g" shows "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" using assms by (intro tpdfr.Dep_Fun_Rel_left_Galois_if_left_GaloisI) simp_all interpretation flip : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma left_Galois_if_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" and "((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" shows "f \<^bsub>L\<^esub>\ g" using assms by (intro tpdfr.left_Galois_if_Dep_Fun_Rel_left_GaloisI flip.tfr.mono_wrt_rel_leftI) simp_all corollary left_Galois_iff_Fun_Rel_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) (r2)" and "transitive (\\<^bsub>L2\<^esub>)" and "in_dom (\\<^bsub>L\<^esub>) f" and "in_codom (\\<^bsub>R\<^esub>) g" shows "f \<^bsub>L\<^esub>\ g \ ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\)) f g" using assms by (intro iffI Fun_Rel_left_Galois_if_left_GaloisI) (auto intro!: left_Galois_if_Fun_Rel_left_GaloisI) theorem left_Galois_eq_Fun_Rel_left_Galois_restrictI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" and "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "transitive (\\<^bsub>L2\<^esub>)" shows "(\<^bsub>L\<^esub>\) = ((\<^bsub>L1\<^esub>\) \ (\<^bsub>L2\<^esub>\))\\<^bsub>in_dom (\\<^bsub>L\<^esub>)\<^esub>\\<^bsub>in_codom (\\<^bsub>R\<^esub>)\<^esub>" - using assms by (intro ext iffI bin_rel_restrict_leftI bin_rel_restrict_rightI + 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 end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions_Order_Base.thy b/thys/Transport/Transport/Functions/Transport_Functions_Order_Base.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Order_Base.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Order_Base.thy @@ -1,428 +1,428 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Basic Order Properties\ theory Transport_Functions_Order_Base imports Transport_Functions_Base begin paragraph \Dependent Function Relator\ context hom_Dep_Fun_Rel_orders begin lemma reflexive_on_in_domI: assumes refl_L: "reflexive_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x2 x2\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and pequiv_R: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "reflexive_on (in_dom DFR) DFR" proof (intro reflexive_onI Dep_Fun_Rel_relI) fix f x1 x2 assume "in_dom DFR f" then obtain g where "DFR f g" by auto moreover assume "x1 \\<^bsub>L\<^esub> x2" moreover with refl_L have "x2 \\<^bsub>L\<^esub> x2" by blast ultimately have "f x1 \\<^bsub>R x1 x2\<^esub> g x2" "f x2 \\<^bsub>R x1 x2\<^esub> g x2" using R_le_R_if_L by auto moreover with pequiv_R \x1 \\<^bsub>L\<^esub> x2\ have "g x2 \\<^bsub>R x1 x2\<^esub> f x2" by (blast dest: symmetricD) ultimately show "f x1 \\<^bsub>R x1 x2\<^esub> f x2" using pequiv_R \x1 \\<^bsub>L\<^esub> x2\ by blast qed lemma reflexive_on_in_codomI: assumes refl_L: "reflexive_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x1\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and pequiv_R: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "reflexive_on (in_codom DFR) DFR" proof (intro reflexive_onI Dep_Fun_Rel_relI) fix f x1 x2 assume "in_codom DFR f" then obtain g where "DFR g f" by auto moreover assume "x1 \\<^bsub>L\<^esub> x2" moreover with refl_L have "x1 \\<^bsub>L\<^esub> x1" by blast ultimately have "g x1 \\<^bsub>R x1 x2\<^esub> f x2" "g x1 \\<^bsub>R x1 x2\<^esub> f x1" using R_le_R_if_L by auto moreover with pequiv_R \x1 \\<^bsub>L\<^esub> x2\ have "f x1 \\<^bsub>R x1 x2\<^esub> g x1" by (blast dest: symmetricD) ultimately show "f x1 \\<^bsub>R x1 x2\<^esub> f x2" using pequiv_R \x1 \\<^bsub>L\<^esub> x2\ by blast qed corollary reflexive_on_in_fieldI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x2 x2\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x1\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "reflexive_on (in_field DFR) DFR" proof - from assms have "reflexive_on (in_dom DFR) DFR" by (intro reflexive_on_in_domI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_codom) moreover from assms have "reflexive_on (in_codom DFR) DFR" by (intro reflexive_on_in_codomI) (auto intro: reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom) ultimately show ?thesis by (auto iff: in_field_iff_in_dom_or_in_codom) qed lemma transitiveI: assumes refl_L: "reflexive_on (in_dom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x1\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and trans: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ transitive (\\<^bsub>R x1 x2\<^esub>)" shows "transitive DFR" proof (intro transitiveI Dep_Fun_Rel_relI) fix f1 f2 f3 x1 x2 assume "x1 \\<^bsub>L\<^esub> x2" with refl_L have "x1 \\<^bsub>L\<^esub> x1" by blast moreover assume "DFR f1 f2" ultimately have "f1 x1 \\<^bsub>R x1 x1\<^esub> f2 x1" by blast with R_le_R_if_L have "f1 x1 \\<^bsub>R x1 x2\<^esub> f2 x1" using \x1 \\<^bsub>L\<^esub> x2\ by blast assume "DFR f2 f3" with \x1 \\<^bsub>L\<^esub> x2\ have "f2 x1 \\<^bsub>R x1 x2\<^esub> f3 x2" by blast with \f1 x1 \\<^bsub>R x1 x2\<^esub> f2 x1\ show "f1 x1 \\<^bsub>R x1 x2\<^esub> f3 x2" using trans \x1 \\<^bsub>L\<^esub> x2\ by blast qed lemma transitiveI': assumes refl_L: "reflexive_on (in_codom (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x2 x2\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and trans: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ transitive (\\<^bsub>R x1 x2\<^esub>)" shows "transitive DFR" proof (intro Binary_Relations_Transitive.transitiveI Dep_Fun_Rel_relI) fix f1 f2 f3 x1 x2 assume "DFR f1 f2" "x1 \\<^bsub>L\<^esub> x2" then have "f1 x1 \\<^bsub>R x1 x2\<^esub> f2 x2" by blast from \x1 \\<^bsub>L\<^esub> x2\ refl_L have "x2 \\<^bsub>L\<^esub> x2" by blast moreover assume "DFR f2 f3" ultimately have "f2 x2 \\<^bsub>R x2 x2\<^esub> f3 x2" by blast with R_le_R_if_L have "f2 x2 \\<^bsub>R x1 x2\<^esub> f3 x2" using \x1 \\<^bsub>L\<^esub> x2\ by blast with \f1 x1 \\<^bsub>R x1 x2\<^esub> f2 x2\ show "f1 x1 \\<^bsub>R x1 x2\<^esub> f3 x2" using trans \x1 \\<^bsub>L\<^esub> x2\ by blast qed lemma preorder_on_in_fieldI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x2 x2\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x1\<^esub>) \ (\\<^bsub>R x1 x2\<^esub>)" and pequiv_R: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "preorder_on (in_field DFR) DFR" using assms by (intro preorder_onI reflexive_on_in_fieldI) (auto intro!: transitiveI dest: pequiv_R elim!: partial_equivalence_relE) lemma symmetricI: assumes sym_L: "symmetric (\\<^bsub>L\<^esub>)" and R_le_R_if_L: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ (\\<^bsub>R x1 x2\<^esub>) \ (\\<^bsub>R x2 x1\<^esub>)" and sym_R: "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ symmetric (\\<^bsub>R x1 x2\<^esub>)" shows "symmetric DFR" proof (intro symmetricI Dep_Fun_Rel_relI) fix f g x y assume "x \\<^bsub>L\<^esub> y" with sym_L have "y \\<^bsub>L\<^esub> x" by (rule symmetricD) moreover assume "DFR f g" ultimately have "f y \\<^bsub>R y x\<^esub> g x" by blast with sym_R \y \\<^bsub>L\<^esub> x\ have "g x \\<^bsub>R y x\<^esub> f y" by (blast dest: symmetricD) with R_le_R_if_L \y \\<^bsub>L\<^esub> x\ show "g x \\<^bsub>R x y\<^esub> f y" by blast qed corollary partial_equivalence_relI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and sym_L: "symmetric (\\<^bsub>L\<^esub>)" and mono_R: "([x1 x2 \ (\\<^bsub>L\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L\<^esub>) | x1 \\<^bsub>L\<^esub> x3] \ (\)) R" and "\x1 x2. x1 \\<^bsub>L\<^esub> x2 \ partial_equivalence_rel (\\<^bsub>R x1 x2\<^esub>)" shows "partial_equivalence_rel DFR" proof - have "(\\<^bsub>R x1 x2\<^esub>) \ (\\<^bsub>R x2 x1\<^esub>)" if "x1 \\<^bsub>L\<^esub> x2" for x1 x2 proof - from sym_L \x1 \\<^bsub>L\<^esub> x2\ have "x2 \\<^bsub>L\<^esub> x1" by (rule symmetricD) with mono_R \x1 \\<^bsub>L\<^esub> x2\ show ?thesis by blast qed with assms show ?thesis by (intro partial_equivalence_relI transitiveI symmetricI) (blast elim: partial_equivalence_relE[OF assms(4)])+ qed end context transport_Dep_Fun_Rel begin lemmas reflexive_on_in_field_leftI = dfro1.reflexive_on_in_fieldI [folded left_rel_eq_Dep_Fun_Rel] lemmas transitive_leftI = dfro1.transitiveI[folded left_rel_eq_Dep_Fun_Rel] lemmas transitive_leftI' = dfro1.transitiveI'[folded left_rel_eq_Dep_Fun_Rel] lemmas preorder_on_in_field_leftI = dfro1.preorder_on_in_fieldI [folded left_rel_eq_Dep_Fun_Rel] lemmas symmetric_leftI = dfro1.symmetricI[folded left_rel_eq_Dep_Fun_Rel] lemmas partial_equivalence_rel_leftI = dfro1.partial_equivalence_relI [folded left_rel_eq_Dep_Fun_Rel] subparagraph \Introduction Rules for Assumptions\ lemma transitive_left2_if_transitive_left2_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and L2_eq: "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ transitive (\\<^bsub>L2 x (r1 x')\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "transitive (\\<^bsub>L2 x1 x2\<^esub>)" by (subst L2_eq) (auto intro!: assms t1.left_Galois_left_if_left_relI) lemma symmetric_left2_if_symmetric_left2_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and L2_eq: "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ symmetric (\\<^bsub>L2 x (r1 x')\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "symmetric (\\<^bsub>L2 x1 x2\<^esub>)" by (subst L2_eq) (auto intro!: assms t1.left_Galois_left_if_left_relI) lemma partial_equivalence_rel_left2_if_partial_equivalence_rel_left2_if_left_GaloisI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>L1\<^esub>) \\<^sub>h (\\<^bsub>R1\<^esub>)) l1 r1" and L2_eq: "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ partial_equivalence_rel (\\<^bsub>L2 x (r1 x')\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "partial_equivalence_rel (\\<^bsub>L2 x1 x2\<^esub>)" by (subst L2_eq) (auto intro!: assms t1.left_Galois_left_if_left_relI) context assumes galois_prop: "((\\<^bsub>L1\<^esub>) \ (\\<^bsub>R1\<^esub>)) l1 r1" begin interpretation flip_inv : transport_Dep_Fun_Rel "(\\<^bsub>R1\<^esub>)" "(\\<^bsub>L1\<^esub>)" r1 l1 "flip2 R2" "flip2 L2" r2 l2 rewrites "flip_inv.t1.unit \ \\<^sub>1" and "\R x y. (flip2 R x y) \ (R y x)\" and "\R S. R\ = S\ \ R = S" and "\R S. (R\ \\<^sub>m S\) \ (R \\<^sub>m S)" and "\x x'. x' \<^bsub>R1\<^esub>\ x \ x \<^bsub>L1\<^esub>\ x'" and "((\\<^bsub>R1\<^esub>) \\<^sub>h (\\<^bsub>L1\<^esub>)) r1 l1 \ True" - and "\R. transitive R\ \ transitive R" - and "\R. symmetric R\ \ symmetric R" - and "\R. partial_equivalence_rel R\ \ partial_equivalence_rel R" + and "\(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 galois_prop.half_galois_prop_right_rel_inv_iff_half_galois_prop_left simp del: rel_inv_iff_rel) lemma transitive_right2_if_transitive_right2_if_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "(\\<^bsub>R2 x1 x2\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1) x2\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ transitive (\\<^bsub>R2 (l1 x) x'\<^esub>)" and "x1 \\<^bsub>R1\<^esub> x2" shows "transitive (\\<^bsub>R2 x1 x2\<^esub>)" using galois_prop assms by (intro flip_inv.transitive_left2_if_transitive_left2_if_left_GaloisI [simplified rel_inv_iff_rel, of x1]) auto lemma symmetric_right2_if_symmetric_right2_if_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "(\\<^bsub>R2 x1 x2\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1) x2\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ symmetric (\\<^bsub>R2 (l1 x) x'\<^esub>)" and "x1 \\<^bsub>R1\<^esub> x2" shows "symmetric (\\<^bsub>R2 x1 x2\<^esub>)" using galois_prop assms by (intro flip_inv.symmetric_left2_if_symmetric_left2_if_left_GaloisI [simplified rel_inv_iff_rel, of x1]) auto lemma partial_equivalence_rel_right2_if_partial_equivalence_rel_right2_if_left_GaloisI: assumes "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "(\\<^bsub>R2 x1 x2\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1) x2\<^esub>)" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ partial_equivalence_rel (\\<^bsub>R2 (l1 x) x'\<^esub>)" and "x1 \\<^bsub>R1\<^esub> x2" shows "partial_equivalence_rel (\\<^bsub>R2 x1 x2\<^esub>)" using galois_prop assms by (intro flip_inv.partial_equivalence_rel_left2_if_partial_equivalence_rel_left2_if_left_GaloisI [simplified rel_inv_iff_rel, of x1]) auto end lemma transitive_left2_if_preorder_equivalenceI: assumes pre_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "transitive (\\<^bsub>L2 x1 x2\<^esub>)" proof - from \x1 \\<^bsub>L1\<^esub> x2\ pre_equiv1 have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" by (intro left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro transitive_left2_if_transitive_left2_if_left_GaloisI[of x1]) blast+ qed lemma symmetric_left2_if_partial_equivalence_rel_equivalenceI: assumes PER_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "symmetric (\\<^bsub>L2 x1 x2\<^esub>)" proof - from \x1 \\<^bsub>L1\<^esub> x2\ PER_equiv1 have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" by (intro left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro symmetric_left2_if_symmetric_left2_if_left_GaloisI[of x1]) blast+ qed lemma partial_equivalence_rel_left2_if_partial_equivalence_rel_equivalenceI: assumes PER_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1 \\<^bsub>L1\<^esub> x2" shows "partial_equivalence_rel (\\<^bsub>L2 x1 x2\<^esub>)" proof - from \x1 \\<^bsub>L1\<^esub> x2\ PER_equiv1 have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>L2 x1 x2\<^esub>) = (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>)" by (intro left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro partial_equivalence_rel_left2_if_partial_equivalence_rel_left2_if_left_GaloisI[of x1]) blast+ qed interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.t1.counit \ \\<^sub>1" and "flip.t1.unit \ \\<^sub>1" by (simp_all only: t1.flip_counit_eq_unit t1.flip_unit_eq_counit) lemma transitive_right2_if_preorder_equivalenceI: assumes pre_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1' \\<^bsub>R1\<^esub> x2'" shows "transitive (\\<^bsub>R2 x1' x2'\<^esub>)" proof - from \x1' \\<^bsub>R1\<^esub> x2'\ pre_equiv1 have "x1' \\<^bsub>R1\<^esub> \\<^sub>1 x1'" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>R2 x1' x2'\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)" by (intro flip.left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro transitive_right2_if_transitive_right2_if_left_GaloisI[of x1']) blast+ qed lemma symmetric_right2_if_partial_equivalence_rel_equivalenceI: assumes PER_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1' \\<^bsub>R1\<^esub> x2'" shows "symmetric (\\<^bsub>R2 x1' x2'\<^esub>)" proof - from \x1' \\<^bsub>R1\<^esub> x2'\ PER_equiv1 have "x1' \\<^bsub>R1\<^esub> \\<^sub>1 x1'" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>R2 x1' x2'\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)" by (intro flip.left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro symmetric_right2_if_symmetric_right2_if_left_GaloisI[of x1']) blast+ qed lemma partial_equivalence_rel_right2_if_partial_equivalence_rel_equivalenceI: assumes PER_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "x1' \\<^bsub>R1\<^esub> x2'" shows "partial_equivalence_rel (\\<^bsub>R2 x1' x2'\<^esub>)" proof - from \x1' \\<^bsub>R1\<^esub> x2'\ PER_equiv1 have "x1' \\<^bsub>R1\<^esub> \\<^sub>1 x1'" by (blast elim: t1.preorder_equivalence_order_equivalenceE intro: bi_related_if_rel_equivalence_on) with assms have "(\\<^bsub>R2 x1' x2'\<^esub>) = (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>)" by (intro flip.left2_eq_if_bi_related_if_monoI) blast+ with assms show ?thesis by (intro partial_equivalence_rel_right2_if_partial_equivalence_rel_right2_if_left_GaloisI[of x1']) blast+ qed end paragraph \Function Relator\ context transport_Fun_Rel begin lemma reflexive_on_in_field_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" shows "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.reflexive_on_in_field_leftI) simp_all lemma transitive_leftI: assumes "reflexive_on (in_dom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L2\<^esub>)" shows "transitive (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.transitive_leftI) simp_all lemma transitive_leftI': assumes "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L2\<^esub>)" shows "transitive (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.transitive_leftI') simp_all lemma preorder_on_in_field_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" shows "preorder_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.preorder_on_in_field_leftI) simp_all lemma symmetric_leftI: assumes "symmetric (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>L2\<^esub>)" shows "symmetric (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.symmetric_leftI) simp_all corollary partial_equivalence_rel_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>L1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" shows "partial_equivalence_rel (\\<^bsub>L\<^esub>)" using assms by (intro tdfr.partial_equivalence_rel_leftI) auto end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin lemmas reflexive_on_in_field_leftI = Refl_Rel_reflexive_on_in_field[of tdfr.L, folded left_rel_eq_tdfr_left_Refl_Rel] lemmas transitive_leftI = Refl_Rel_transitiveI [of tdfr.L, folded left_rel_eq_tdfr_left_Refl_Rel] lemmas preorder_on_in_field_leftI = Refl_Rel_preorder_on_in_fieldI[of tdfr.L, folded left_rel_eq_tdfr_left_Refl_Rel] lemmas symmetric_leftI = Refl_Rel_symmetricI[of tdfr.L, OF tdfr.symmetric_leftI, folded left_rel_eq_tdfr_left_Refl_Rel] lemmas partial_equivalence_rel_leftI = Refl_Rel_partial_equivalence_relI[of tdfr.L, OF tdfr.partial_equivalence_rel_leftI, folded left_rel_eq_tdfr_left_Refl_Rel] end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin lemma symmetric_leftI: assumes "symmetric (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>L2\<^esub>)" shows "symmetric (\\<^bsub>L\<^esub>)" using assms by (intro tpdfr.symmetric_leftI) auto lemma partial_equivalence_rel_leftI: assumes "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "symmetric (\\<^bsub>L1\<^esub>)" and "partial_equivalence_rel (\\<^bsub>L2\<^esub>)" shows "partial_equivalence_rel (\\<^bsub>L\<^esub>)" using assms by (intro tpdfr.partial_equivalence_rel_leftI) auto end end \ No newline at end of file diff --git a/thys/Transport/Transport/Functions/Transport_Functions_Order_Equivalence.thy b/thys/Transport/Transport/Functions/Transport_Functions_Order_Equivalence.thy --- a/thys/Transport/Transport/Functions/Transport_Functions_Order_Equivalence.thy +++ b/thys/Transport/Transport/Functions/Transport_Functions_Order_Equivalence.thy @@ -1,726 +1,727 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Order Equivalence\ theory Transport_Functions_Order_Equivalence imports Transport_Functions_Monotone Transport_Functions_Galois_Equivalence begin paragraph \Dependent Function Relator\ context transport_Dep_Fun_Rel begin subparagraph \Inflationary\ lemma rel_unit_self_if_rel_selfI: assumes inflationary_unit1: "inflationary_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and refl_L1: "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and trans_L1: "transitive (\\<^bsub>L1\<^esub>)" and mono_l2: "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" and mono_r2: "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and inflationary_unit2: "\x. x \\<^bsub>L1\<^esub> x \ inflationary_on (in_codom (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and L2_le1: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and L2_unit_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and ge_R2_l2_le2: "\x y. x \\<^bsub>L1\<^esub> x \ in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and trans_L2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "f \\<^bsub>L\<^esub> f" shows "f \\<^bsub>L\<^esub> \ f" proof (intro left_relI) fix x1 x2 assume [iff]: "x1 \\<^bsub>L1\<^esub> x2" moreover with inflationary_unit1 have "x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by blast ultimately have "x1 \\<^bsub>L1\<^esub> \\<^sub>1 x2" using trans_L1 by blast with \f \\<^bsub>L\<^esub> f\ have "f x1 \\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub> f (\\<^sub>1 x2)" by blast with L2_unit_le2 have "f x1 \\<^bsub>L2 x1 x2\<^esub> f (\\<^sub>1 x2)" by blast moreover have "... \\<^bsub>L2 x1 x2\<^esub> \ f x2" proof - from refl_L1 \x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2\ have "\\<^sub>1 x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2" by blast with \f \\<^bsub>L\<^esub> f\ have "f (\\<^sub>1 x2) \\<^bsub>L2 (\\<^sub>1 x2) (\\<^sub>1 x2)\<^esub> f (\\<^sub>1 x2)" by blast with L2_le1 have "f (\\<^sub>1 x2) \\<^bsub>L2 x2 (\\<^sub>1 x2)\<^esub> f (\\<^sub>1 x2)" using \x2 \\<^bsub>L1\<^esub> \\<^sub>1 x2\ by blast moreover from refl_L1 \x1 \\<^bsub>L1\<^esub> x2\ have [iff]: "x2 \\<^bsub>L1\<^esub> x2" by blast ultimately have "f (\\<^sub>1 x2) \\<^bsub>L2 x2 x2\<^esub> f (\\<^sub>1 x2)" using L2_unit_le2 by blast with inflationary_unit2 have "f (\\<^sub>1 x2) \\<^bsub>L2 x2 x2\<^esub> \\<^bsub>2 x2 (l1 x2)\<^esub> (f (\\<^sub>1 x2))" by blast moreover have "... \\<^bsub>L2 x2 x2\<^esub> \ f x2" proof - from \f (\\<^sub>1 x2) \\<^bsub>L2 x2 x2\<^esub> f (\\<^sub>1 x2)\ mono_l2 have "l2\<^bsub>(l1 x2) x2\<^esub> (f (\\<^sub>1 x2)) \\<^bsub>R2 (l1 x2) (l1 x2)\<^esub> l2\<^bsub>(l1 x2) x2\<^esub> (f (\\<^sub>1 x2))" by blast with ge_R2_l2_le2 have "l2\<^bsub>(l1 x2) x2\<^esub> (f (\\<^sub>1 x2)) \\<^bsub>R2 (l1 x2) (l1 x2)\<^esub> l2\<^bsub>(l1 x2) (\\<^sub>1 x2)\<^esub> (f (\\<^sub>1 x2))" using \f (\\<^sub>1 x2) \\<^bsub>L2 x2 (\\<^sub>1 x2)\<^esub> f (\\<^sub>1 x2)\ by blast with mono_r2 have "\\<^bsub>2 x2 (l1 x2)\<^esub> (f (\\<^sub>1 x2)) \\<^bsub>L2 x2 (\\<^sub>1 x2)\<^esub> \ f x2" by auto with L2_unit_le2 show ?thesis by blast qed ultimately have "f (\\<^sub>1 x2) \\<^bsub>L2 x2 x2\<^esub> \ f x2" using trans_L2 by blast with L2_le1 show ?thesis by blast qed ultimately show "f x1 \\<^bsub>L2 x1 x2\<^esub> \ f x2" using trans_L2 by blast qed subparagraph \Deflationary\ interpretation flip_inv : transport_Dep_Fun_Rel "(\\<^bsub>R1\<^esub>)" "(\\<^bsub>L1\<^esub>)" r1 l1 "flip2 R2" "flip2 L2" r2 l2 rewrites "flip_inv.L \ (\\<^bsub>R\<^esub>)" and "flip_inv.R \ (\\<^bsub>L\<^esub>)" and "flip_inv.unit \ \" and "flip_inv.t1.unit \ \\<^sub>1" and "\x y. flip_inv.t2_unit x y \ \\<^bsub>2 y x\<^esub>" and "\R x y. (flip2 R x y)\ \ R y x" and "\R. in_codom R\ \ in_dom R" and "\R x1 x2. in_codom (flip2 R x1 x2) \ in_dom (R x2 x1)" and "\x1 x2 x1' x2'. (flip2 R2 x1' x2' \\<^sub>m flip2 L2 x1 x2) \ ((\\<^bsub>R2 x2' x1'\<^esub>) \\<^sub>m (\\<^bsub>L2 x2 x1\<^esub>))" and "\x1 x2 x1' x2'. (flip2 L2 x1 x2 \\<^sub>m flip2 R2 x1' x2') \ ((\\<^bsub>L2 x2 x1\<^esub>) \\<^sub>m (\\<^bsub>R2 x2' x1'\<^esub>))" - and "\P. inflationary_on P (\\<^bsub>R1\<^esub>) \ deflationary_on P (\\<^bsub>R1\<^esub>)" - and "\P x. inflationary_on P (flip2 R2 x x) \ deflationary_on P (\\<^bsub>R2 x x\<^esub>)" + and "\(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 \ _) (P :: 'z \ bool). reflexive_on P R\ \ reflexive_on P R" - and "\R. transitive R\ \ transitive R" + 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) lemma counit_rel_self_if_rel_selfI: assumes "deflationary_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "transitive (\\<^bsub>R1\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) \\<^sub>m (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>)) (l2\<^bsub> x' (r1 x')\<^esub>)" and "\x' x'. x' \\<^bsub>R1\<^esub> x' \ ((\\<^bsub>R2 x' x'\<^esub>) \\<^sub>m (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>)) (r2\<^bsub>(r1 x') x'\<^esub>)" and "\x'. x' \\<^bsub>R1\<^esub> x' \ deflationary_on (in_dom (\\<^bsub>R2 x' x'\<^esub>)) (\\<^bsub>R2 x' x'\<^esub>) (\\<^bsub>2 (r1 x') x'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 x1' x1'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_dom (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y' \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') x'\<^esub> y') \ (\\<^bsub>L2 (r1 x') (r1 x')\<^esub>) (r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y')" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ transitive (\\<^bsub>R2 x1' x2'\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "\ g \\<^bsub>R\<^esub> g" using assms by (intro flip_inv.rel_unit_self_if_rel_selfI[simplified rel_inv_iff_rel]) subparagraph \Relational Equivalence\ lemma bi_related_unit_self_if_rel_self_aux: assumes rel_equiv_unit1: "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and mono_r2: "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and rel_equiv_unit2: "\x. x \\<^bsub>L1\<^esub> x \ rel_equivalence_on (in_field (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and L2_le1: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and L2_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and [iff]: "x \\<^bsub>L1\<^esub> x" shows "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "deflationary_on (in_dom (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) \\<^bsub>2 x (l1 x)\<^esub>" and "inflationary_on (in_codom (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) \\<^bsub>2 x (l1 x)\<^esub>" proof - from rel_equiv_unit1 have "x \\<^bsub>L1\<^esub> \\<^sub>1 x" by blast with mono_r2 show "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" using L2_le1 L2_le2 by blast+ qed (insert rel_equiv_unit2, blast+) interpretation flip : transport_Dep_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 rewrites "flip.counit \ \" and "flip.t1.counit \ \\<^sub>1" and "\x y. flip.t2_counit x y \ \\<^bsub>2 y x\<^esub>" by (simp_all add: order_functors.flip_counit_eq_unit) lemma bi_related_unit_self_if_rel_selfI: assumes rel_equiv_unit1: "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and trans_L1: "transitive (\\<^bsub>L1\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x x\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ rel_equivalence_on (in_field (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 (\\<^sub>1 x1) x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 x1\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_dom (\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "f \\<^bsub>L\<^esub> f" shows "f \\<^bsub>L\<^esub> \ f" proof - from rel_equiv_unit1 trans_L1 have "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" by (intro reflexive_on_in_field_if_transitive_if_rel_equivalence_on) with assms show ?thesis by (intro bi_relatedI rel_unit_self_if_rel_selfI flip.counit_rel_self_if_rel_selfI bi_related_unit_self_if_rel_self_aux) (auto intro: inflationary_on_if_le_pred_if_inflationary_on deflationary_on_if_le_pred_if_deflationary_on reflexive_on_if_le_pred_if_reflexive_on in_field_if_in_dom in_field_if_in_codom) qed subparagraph \Lemmas for Monotone Function Relator\ lemma order_equivalence_if_order_equivalence_mono_assms_leftI: assumes order_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and refl_R1: "reflexive_on (in_field (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and R2_counit_le1: "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ (\\<^bsub>R2 (\\<^sub>1 x1') x2'\<^esub>) \ (\\<^bsub>R2 x1' x2'\<^esub>)" and mono_l2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and [iff]: "x1' \\<^bsub>R1\<^esub> x2'" shows "([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" and "([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" proof - from refl_R1 have "x1' \\<^bsub>R1\<^esub> x1'" "x2' \\<^bsub>R1\<^esub> x2'" by auto moreover with order_equiv1 have "r1 x1' \\<^bsub>L1\<^esub> r1 x2'" "r1 x1' \\<^bsub>L1\<^esub> r1 x1'" "r1 x2' \\<^bsub>L1\<^esub> r1 x2'" by auto ultimately have "r1 x1' \<^bsub>L1\<^esub>\ x1'" "r1 x2' \<^bsub>L1\<^esub>\ x2'" by blast+ note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_l2 \x1' \\<^bsub>R1\<^esub> x2'\] \r1 x1' \\<^bsub>L1\<^esub> r1 x1'\] with \r1 x1' \<^bsub>L1\<^esub>\ x1'\ R2_counit_le1 show "([in_dom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x1' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x1')\<^esub>)" by (intro Dep_Fun_Rel_predI) (auto dest!: in_field_if_in_dom) note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_l2 \x2' \\<^bsub>R1\<^esub> x2'\] \r1 x1' \\<^bsub>L1\<^esub> r1 x2'\] with \r1 x2' \<^bsub>L1\<^esub>\ x2'\ R2_counit_le1 show "([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" by (intro Dep_Fun_Rel_predI) (auto dest!: in_field_if_in_codom) qed lemma order_equivalence_if_order_equivalence_mono_assms_rightI: assumes order_equiv1: "((\\<^bsub>L1\<^esub>) \\<^sub>o (\\<^bsub>R1\<^esub>)) l1 r1" and refl_L1: "reflexive_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and L2_unit_le2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and [iff]: "x1 \\<^bsub>L1\<^esub> x2" shows "([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" proof - from refl_L1 have "x1 \\<^bsub>L1\<^esub> x1" "x2 \\<^bsub>L1\<^esub> x2" by auto moreover with order_equiv1 have "l1 x1 \\<^bsub>R1\<^esub> l1 x2" "l1 x1 \\<^bsub>R1\<^esub> l1 x1" "l1 x2 \\<^bsub>R1\<^esub> l1 x2" by auto ultimately have "x1 \<^bsub>L1\<^esub>\ l1 x1" "x2 \<^bsub>L1\<^esub>\ l1 x2" using order_equiv1 by (auto intro!: t1.left_Galois_left_if_in_codom_if_inflationary_onI) note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x1 \\<^bsub>L1\<^esub> x2\] \l1 x2 \\<^bsub>R1\<^esub> l1 x2\] with \x2 \<^bsub>L1\<^esub>\ l1 x2\ L2_unit_le2 show "([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" by (intro Dep_Fun_Rel_predI) (auto dest!: in_field_if_in_codom) note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \x1 \\<^bsub>L1\<^esub> x1\] \l1 x1 \\<^bsub>R1\<^esub> l1 x2\] with \x1 \<^bsub>L1\<^esub>\ l1 x1\ L2_unit_le2 show "([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" by (intro Dep_Fun_Rel_predI) (auto dest!: in_field_if_in_dom) qed lemma l2_unit_bi_rel_selfI: assumes pre_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and mono_L2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and mono_R2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | (x2' \\<^bsub>R1\<^esub> x3' \ x4' \\<^bsub>R1\<^esub> \\<^sub>1 x3')] \ (\)) R2" and mono_l2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "x \\<^bsub>L1\<^esub> x" and "in_field (\\<^bsub>L2 x x\<^esub>) y" shows "l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" proof (rule bi_relatedI) note t1.preorder_equivalence_order_equivalenceE[elim!] from \x \\<^bsub>L1\<^esub> x\ pre_equiv1 have "l1 x \\<^bsub>R1\<^esub> l1 x" "x \\<^bsub>L1\<^esub> \\<^sub>1 x" "\\<^sub>1 x \\<^bsub>L1\<^esub> x" by blast+ with pre_equiv1 have "x \<^bsub>L1\<^esub>\ l1 x" "\\<^sub>1 x \<^bsub>L1\<^esub>\ l1 x" by (auto 4 3) from pre_equiv1 \x \\<^bsub>L1\<^esub> \\<^sub>1 x\ have "x \\<^bsub>L1\<^esub> \\<^sub>1 (\\<^sub>1 x)" by fastforce moreover note \in_field (\\<^bsub>L2 x x\<^esub>) y\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_L2 \\\<^sub>1 x \\<^bsub>L1\<^esub> x\] \\\<^sub>1 x \\<^bsub>L1\<^esub> x\] Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_L2 \x \\<^bsub>L1\<^esub> x\] \\\<^sub>1 x \\<^bsub>L1\<^esub> x\] ultimately have "in_field (\\<^bsub>L2 (\\<^sub>1 x) (\\<^sub>1 x)\<^esub>) y" "in_field (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y" using \x \\<^bsub>L1\<^esub> \\<^sub>1 x\ by blast+ moreover note \x \<^bsub>L1\<^esub>\ l1 x\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_l2 \l1 x \\<^bsub>R1\<^esub> l1 x\] \\\<^sub>1 x \\<^bsub>L1\<^esub> x\] ultimately have "l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (\\<^sub>1 (l1 x)) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" by auto moreover from pre_equiv1 \l1 x \\<^bsub>R1\<^esub> l1 x\ have "\\<^sub>1 (l1 x) \\<^bsub>R1\<^esub> l1 x" "l1 x \\<^bsub>R1\<^esub> \\<^sub>1 (l1 x)" by fastforce+ moreover note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD [OF mono_R2 \l1 x \\<^bsub>R1\<^esub> \\<^sub>1 (l1 x)\] \l1 x \\<^bsub>R1\<^esub> l1 x\] ultimately show "l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" by blast note \\\<^sub>1 x \<^bsub>L1\<^esub>\ l1 x\ \in_field (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_l2 \l1 x \\<^bsub>R1\<^esub> l1 x\] \x \\<^bsub>L1\<^esub> \\<^sub>1 x\] then show "l2\<^bsub>(l1 x) x\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y" by auto qed lemma r2_counit_bi_rel_selfI: assumes pre_equiv1: "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and mono_L2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and mono_R2: "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | (x2' \\<^bsub>R1\<^esub> x3' \ x4' \\<^bsub>R1\<^esub> \\<^sub>1 x3')] \ (\)) R2" and mono_r2: "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "x' \\<^bsub>R1\<^esub> x'" and "in_field (\\<^bsub>R2 x' x'\<^esub>) y'" shows "r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') x'\<^esub> y'" proof (rule bi_relatedI) note t1.preorder_equivalence_order_equivalenceE[elim!] from \x' \\<^bsub>R1\<^esub> x'\ pre_equiv1 have "r1 x' \\<^bsub>L1\<^esub> r1 x'" "x' \\<^bsub>R1\<^esub> \\<^sub>1 x'" "\\<^sub>1 x' \\<^bsub>R1\<^esub> x'" by blast+ with pre_equiv1 have "r1 x' \<^bsub>L1\<^esub>\ x'" "r1 x' \<^bsub>L1\<^esub>\ \\<^sub>1 x'" by force+ from pre_equiv1 \x' \\<^bsub>R1\<^esub> \\<^sub>1 x'\ have "x' \\<^bsub>R1\<^esub> \\<^sub>1 (\\<^sub>1 x')" by fastforce moreover note \in_field (\\<^bsub>R2 x' x'\<^esub>) y'\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_R2 \\\<^sub>1 x' \\<^bsub>R1\<^esub> x'\] \\\<^sub>1 x' \\<^bsub>R1\<^esub> x'\] Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_R2 \\\<^sub>1 x' \\<^bsub>R1\<^esub> x'\] \x' \\<^bsub>R1\<^esub> x'\] ultimately have "in_field (\\<^bsub>R2 (\\<^sub>1 x') (\\<^sub>1 x')\<^esub>) y'" "in_field (\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) y'" using \x' \\<^bsub>R1\<^esub> \\<^sub>1 x'\ \x' \\<^bsub>R1\<^esub> x'\ by blast+ moreover note \r1 x' \<^bsub>L1\<^esub>\ \\<^sub>1 x'\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \r1 x' \\<^bsub>L1\<^esub> r1 x'\] \\\<^sub>1 x' \\<^bsub>R1\<^esub> x'\] ultimately show "r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') x'\<^esub> y'" by auto note \r1 x' \<^bsub>L1\<^esub>\ x'\ \in_field (\\<^bsub>R2 (\\<^sub>1 x') (\\<^sub>1 x')\<^esub>) y'\ Dep_Fun_Rel_relD[OF dep_mono_wrt_relD[OF mono_r2 \r1 x' \\<^bsub>L1\<^esub> r1 x'\] \x' \\<^bsub>R1\<^esub> \\<^sub>1 x'\] then have "r2\<^bsub>(r1 x') x'\<^esub> y' \\<^bsub>L2 (r1 x') (\\<^sub>1 (r1 x'))\<^esub> r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y'" by auto moreover from pre_equiv1 \r1 x' \\<^bsub>L1\<^esub> r1 x'\ have "\\<^sub>1 (r1 x') \\<^bsub>L1\<^esub> r1 x'" "r1 x' \\<^bsub>L1\<^esub> \\<^sub>1 (r1 x')" by fastforce+ moreover note Dep_Fun_Rel_relD[OF dep_mono_wrt_relD [OF mono_L2 \r1 x' \\<^bsub>L1\<^esub> r1 x'\] \r1 x' \\<^bsub>L1\<^esub> \\<^sub>1 (r1 x')\] ultimately show "r2\<^bsub>(r1 x') x'\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y'" using pre_equiv1 by blast qed end paragraph \Function Relator\ context transport_Fun_Rel begin corollary rel_unit_self_if_rel_selfI: assumes "inflationary_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "inflationary_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>L2\<^esub>)" and "f \\<^bsub>L\<^esub> f" shows "f \\<^bsub>L\<^esub> \ f" using assms by (intro tdfr.rel_unit_self_if_rel_selfI) simp_all corollary counit_rel_self_if_rel_selfI: assumes "deflationary_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "transitive (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "deflationary_on (in_dom (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>R2\<^esub>)" and "g \\<^bsub>R\<^esub> g" shows "\ g \\<^bsub>R\<^esub> g" using assms by (intro tdfr.counit_rel_self_if_rel_selfI) simp_all lemma bi_related_unit_self_if_rel_selfI: assumes "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "rel_equivalence_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>L2\<^esub>)" and "f \\<^bsub>L\<^esub> f" shows "f \\<^bsub>L\<^esub> \ f" using assms by (intro tdfr.bi_related_unit_self_if_rel_selfI) simp_all end paragraph \Monotone Dependent Function Relator\ context transport_Mono_Dep_Fun_Rel begin subparagraph \Inflationary\ lemma inflationary_on_unitI: assumes "(tdfr.L \\<^sub>m tdfr.R) l" and "(tdfr.R \\<^sub>m tdfr.L) r" and "inflationary_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L1\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ ((\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) \\<^sub>m (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>)) (r2\<^bsub>x (l1 x)\<^esub>)" and "\x. x \\<^bsub>L1\<^esub> x \ inflationary_on (in_codom (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x2 x2\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ (\\<^bsub>L2 x1 (\\<^sub>1 x2)\<^esub>) \ (\\<^bsub>L2 x1 x2\<^esub>)" and "\x y. x \\<^bsub>L1\<^esub> x \ in_codom (\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) y \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) x\<^esub> y) \ (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>) (l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" shows "inflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" unfolding left_rel_eq_tdfr_left_Refl_Rel using assms by (intro inflationary_onI Refl_RelI) (auto 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>)" and "\x1' x2'. x1' \\<^bsub>R1\<^esub> x2' \ ([in_codom (\\<^bsub>L2 (r1 x1') (r1 x2')\<^esub>)] \ (\\<^bsub>R2 x1' x2'\<^esub>)) (l2\<^bsub>x2' (r1 x1')\<^esub>) (l2\<^bsub>x2' (r1 x2')\<^esub>)" and l2_bi_rel: "\x y. x \\<^bsub>L1\<^esub> x \ in_field (\\<^bsub>L2 x x\<^esub>) y \ l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_codom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x2)\<^esub>) (r2\<^bsub>x2 (l1 x2)\<^esub>)" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ ([in_dom (\\<^bsub>R2 (l1 x1) (l1 x2)\<^esub>)] \ (\\<^bsub>L2 x1 x2\<^esub>)) (r2\<^bsub>x1 (l1 x1)\<^esub>) (r2\<^bsub>x1 (l1 x2)\<^esub>)" and r2_bi_rel: "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_field (\\<^bsub>R2 x' x'\<^esub>) y' \ r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') x'\<^esub> y'" and trans_L2: "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and trans_R2: "\x1 x2. x1 \\<^bsub>R1\<^esub> x2 \ transitive (\\<^bsub>R2 x1 x2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" proof - from pre_equiv1 L2_les have L2_unit_eq1: "(\\<^bsub>L2 (\\<^sub>1 x) x\<^esub>) = (\\<^bsub>L2 x x\<^esub>)" and L2_unit_eq2: "(\\<^bsub>L2 x (\\<^sub>1 x)\<^esub>) = (\\<^bsub>L2 x x\<^esub>)" if "x \\<^bsub>L1\<^esub> x" for x using \x \\<^bsub>L1\<^esub> x\ by (auto elim!: t1.preorder_equivalence_order_equivalenceE intro!: tdfr.left_rel2_unit_eqs_left_rel2I bi_related_if_rel_equivalence_on simp del: t1.unit_eq) from pre_equiv1 R2_les have R2_counit_eq1: "(\\<^bsub>R2 (\\<^sub>1 x') x'\<^esub>) = (\\<^bsub>R2 x' x'\<^esub>)" and R2_counit_eq2: "(\\<^bsub>R2 x' (\\<^sub>1 x')\<^esub>) = (\\<^bsub>R2 x' x'\<^esub>)" (is ?goal2) if "x' \\<^bsub>R1\<^esub> x'" for x' using \x' \\<^bsub>R1\<^esub> x'\ by (auto elim!: t1.preorder_equivalence_order_equivalenceE intro!: flip.tdfr.left_rel2_unit_eqs_left_rel2I bi_related_if_rel_equivalence_on simp del: t1.counit_eq) from order_equiv2 have mono_l2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>)" and mono_r2: "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>R2 (l1 x) x'\<^esub>) \\<^sub>m (\\<^bsub>L2 x (r1 x')\<^esub>)) (r2\<^bsub>x x'\<^esub>)" by auto moreover have "rel_equivalence_on (in_field (\\<^bsub>L2 x x\<^esub>)) (\\<^bsub>L2 x x\<^esub>) (\\<^bsub>2 x (l1 x)\<^esub>)" (is ?goal1) and "((\\<^bsub>L2 x x\<^esub>) \\<^sub>m (\\<^bsub>R2 (l1 x) (l1 x)\<^esub>)) (l2\<^bsub>(l1 x) x\<^esub>)" (is ?goal2) if [iff]: "x \\<^bsub>L1\<^esub> x" for x proof - - from pre_equiv1 have "x \<^bsub>L1\<^esub>\ l1 x" - by (auto intro!: t1.left_GaloisI - elim!: t1.preorder_equivalence_order_equivalenceE t1.order_equivalenceE) + 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 + 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 "\x y. x \\<^bsub>L1\<^esub> x \ in_field (\\<^bsub>L2 x x\<^esub>) y \ l2\<^bsub>(l1 x) (\\<^sub>1 x)\<^esub> y \\<^bsub>R2 (l1 x) (l1 x)\<^esub> l2\<^bsub>(l1 x) x\<^esub> y" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x' y'. x' \\<^bsub>R1\<^esub> x' \ in_field (\\<^bsub>R2 x' x'\<^esub>) y' \ r2\<^bsub>(r1 x') (\\<^sub>1 x')\<^esub> y' \\<^bsub>L2 (r1 x') (r1 x')\<^esub> r2\<^bsub>(r1 x') x'\<^esub> y'" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>R1\<^esub> x2 \ transitive (\\<^bsub>R2 x1 x2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalence_if_preorder_equivalenceI tdfr.order_equivalence_if_order_equivalence_mono_assms_leftI tdfr.order_equivalence_if_order_equivalence_mono_assms_rightI reflexive_on_in_field_if_transitive_if_rel_equivalence_on) (auto elim!: t1.preorder_equivalence_order_equivalenceE) lemma order_equivalence_if_mono_if_preorder_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^sub>o (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>) | \\<^sub>1 x2 \\<^bsub>L1\<^esub> x1] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x2 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | (x2 \\<^bsub>L1\<^esub> x3 \ x4 \\<^bsub>L1\<^esub> \\<^sub>1 x3)] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>) | \\<^sub>1 x2' \\<^bsub>R1\<^esub> x1'] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x2' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | (x2' \\<^bsub>R1\<^esub> x3' \ x4' \\<^bsub>R1\<^esub> \\<^sub>1 x3')] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" and "\x1 x2. x1 \\<^bsub>L1\<^esub> x2 \ transitive (\\<^bsub>L2 x1 x2\<^esub>)" and "\x1 x2. x1 \\<^bsub>R1\<^esub> x2 \ transitive (\\<^bsub>R2 x1 x2\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalence_if_preorder_equivalenceI' tdfr.l2_unit_bi_rel_selfI tdfr.r2_counit_bi_rel_selfI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI flip.tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_leftI tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI flip.tdfr.left_rel_right_iff_left_right_rel_if_galois_prop_le_assms_rightI t1.galois_connection_left_right_if_transitive_if_order_equivalence flip.t1.galois_connection_left_right_if_transitive_if_order_equivalence reflexive_on_in_field_if_transitive_if_rel_equivalence_on) (auto elim!: t1.preorder_equivalence_order_equivalenceE) theorem order_equivalence_if_mono_if_preorder_equivalenceI': assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "\x x'. x \<^bsub>L1\<^esub>\ x' \ ((\\<^bsub>L2 x (r1 x')\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2 (l1 x) x'\<^esub>)) (l2\<^bsub>x' x\<^esub>) (r2\<^bsub>x x'\<^esub>)" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x3 x4 \ (\\<^bsub>L1\<^esub>) | x1 \\<^bsub>L1\<^esub> x3] \ (\)) L2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x3' x4' \ (\\<^bsub>R1\<^esub>) | x1' \\<^bsub>R1\<^esub> x3'] \ (\)) R2" and "([x1' x2' \ (\\<^bsub>R1\<^esub>)] \\<^sub>m [x1 x2 \ (\\<^bsub>L1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>L2 x1 (r1 x2')\<^esub>)] \ (\\<^bsub>R2 (l1 x1) x2'\<^esub>)) l2" and "([x1 x2 \ (\\<^bsub>L1\<^esub>)] \\<^sub>m [x1' x2' \ (\\<^bsub>R1\<^esub>) | x2 \<^bsub>L1\<^esub>\ x1'] \ [in_field (\\<^bsub>R2 (l1 x1) x2'\<^esub>)] \ (\\<^bsub>L2 x1 (r1 x2')\<^esub>)) r2" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro order_equivalence_if_mono_if_preorder_equivalenceI tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI flip.tdfr.galois_equivalence_if_mono_if_galois_equivalence_mono_assms_leftI tdfr.transitive_left2_if_preorder_equivalenceI tdfr.transitive_right2_if_preorder_equivalenceI t1.preorder_on_in_field_left_if_transitive_if_order_equivalence flip.t1.preorder_on_in_field_left_if_transitive_if_order_equivalence t1.galois_equivalence_left_right_if_transitive_if_order_equivalence flip.t1.galois_equivalence_left_right_if_transitive_if_order_equivalence) (auto elim!: t1.preorder_equivalence_order_equivalenceE t2.preorder_equivalence_order_equivalenceE) end paragraph \Monotone Function Relator\ context transport_Mono_Fun_Rel begin interpretation flip : transport_Mono_Fun_Rel R1 L1 r1 l1 R2 L2 r2 l2 . lemma inflationary_on_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "inflationary_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "reflexive_on (in_codom (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>)" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "inflationary_on (in_codom (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>L2\<^esub>)" shows "inflationary_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms by (intro tpdfr.inflationary_on_unitI tfr.mono_wrt_rel_leftI flip.tfr.mono_wrt_rel_leftI) simp_all lemma deflationary_on_counitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "deflationary_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>) \\<^sub>1" and "reflexive_on (in_dom (\\<^bsub>R1\<^esub>)) (\\<^bsub>R1\<^esub>)" and "transitive (\\<^bsub>R1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "deflationary_on (in_dom (\\<^bsub>R2\<^esub>)) (\\<^bsub>R2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>R2\<^esub>)" shows "deflationary_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>) \" using assms by (intro tpdfr.deflationary_on_counitI tfr.mono_wrt_rel_leftI flip.tfr.mono_wrt_rel_leftI) simp_all lemma rel_equivalence_on_unitI: assumes "((\\<^bsub>L1\<^esub>) \\<^sub>m (\\<^bsub>R1\<^esub>)) l1" and "((\\<^bsub>R1\<^esub>) \\<^sub>m (\\<^bsub>L1\<^esub>)) r1" and "rel_equivalence_on (in_field (\\<^bsub>L1\<^esub>)) (\\<^bsub>L1\<^esub>) \\<^sub>1" and "transitive (\\<^bsub>L1\<^esub>)" and "((\\<^bsub>L2\<^esub>) \\<^sub>m (\\<^bsub>R2\<^esub>)) l2" and "((\\<^bsub>R2\<^esub>) \\<^sub>m (\\<^bsub>L2\<^esub>)) r2" and "rel_equivalence_on (in_field (\\<^bsub>L2\<^esub>)) (\\<^bsub>L2\<^esub>) \\<^sub>2" and "transitive (\\<^bsub>L2\<^esub>)" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms by (intro tpdfr.rel_equivalence_on_unitI tfr.mono_wrt_rel_leftI flip.tfr.mono_wrt_rel_leftI) simp_all lemma order_equivalenceI: assumes "((\\<^bsub>L1\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R1\<^esub>)) l1 r1" and "((\\<^bsub>L2\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R2\<^esub>)) l2 r2" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms by (intro tpdfr.order_equivalenceI tfr.mono_wrt_rel_leftI flip.tfr.mono_wrt_rel_leftI) (auto elim!: tdfrs.t1.preorder_equivalence_order_equivalenceE tdfrs.t2.preorder_equivalence_order_equivalenceE) end end \ No newline at end of file diff --git a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Base.thy b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Base.thy --- a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Base.thy +++ b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Base.thy @@ -1,663 +1,663 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Transport for Natural Functors\ subsection \Basic Setup\ theory Transport_Natural_Functors_Base imports HOL.BNF_Def HOL_Alignment_Functions Transport_Base begin paragraph \Summary\ text \Basic setup for closure proofs and simple lemmas.\ text \In the following, we willingly use granular apply-style proofs since, in practice, these theorems have to be automatically generated whenever we declare a new natural functor. Note that "HOL-Library" provides a command \bnf_axiomatization\ which allows one to axiomatically declare a bounded natural functor. However, we only need a subset of these axioms - the boundedness of the functor is irrelevant for our purposes. For this reason - and the sake of completeness - we state all the required axioms explicitly below.\ lemma Grp_UNIV_eq_eq_comp: "BNF_Def.Grp UNIV f = (=) \ f" by (intro ext) (auto elim: GrpE intro: GrpI) lemma eq_comp_rel_comp_eq_comp: "(=) \ f \\ R = R \ f" by (intro ext) auto lemma Domain_Collect_case_prod_eq_Collect_in_dom: "Domain {(x, y). R x y} = {x. in_dom R x}" by blast lemma ball_in_dom_iff_ball_ex: "(\x \ S. in_dom R x) \ (\x \ S. \y. R x y)" by blast lemma pair_mem_Collect_case_prod_iff: "(x, y) \ {(x, y). R x y} \ R x y" by blast paragraph \Natural Functor Axiomatisation\ typedecl ('d, 'a, 'b, 'c) F consts Fmap :: "('a1 \ 'a2) \ ('b1 \ 'b2) \ ('c1 \ 'c2) \ ('d, 'a1, 'b1, 'c1) F \ ('d, 'a2, 'b2, 'c2) F" Fset1 :: "('d, 'a, 'b, 'c) F \ 'a set" Fset2 :: "('d, 'a, 'b, 'c) F \ 'b set" Fset3 :: "('d, 'a, 'b, 'c) F \ 'c set" axiomatization where Fmap_id: "Fmap id id id = id" and Fmap_comp: "\f1 f2 f3 g1 g2 g3. Fmap (g1 \ f1) (g2 \ f2) (g3 \ f3) = Fmap g1 g2 g3 \ Fmap f1 f2 f3" and Fmap_cong: "\f1 f2 f3 g1 g2 g3 x. (\x1. x1 \ Fset1 x \ f1 x1 = g1 x1) \ (\x2. x2 \ Fset2 x \ f2 x2 = g2 x2) \ (\x3. x3 \ Fset3 x \ f3 x3 = g3 x3) \ Fmap f1 f2 f3 x = Fmap g1 g2 g3 x" and Fset1_natural: "\f1 f2 f3. Fset1 \ Fmap f1 f2 f3 = image f1 \ Fset1" and Fset2_natural: "\f1 f2 f3. Fset2 \ Fmap f1 f2 f3 = image f2 \ Fset2" and Fset3_natural: "\f1 f2 f3. Fset3 \ Fmap f1 f2 f3 = image f3 \ Fset3" lemma Fmap_id_eq_self: "Fmap id id id x = x" by (subst Fmap_id, subst id_eq_self, rule refl) lemma Fmap_comp_eq_Fmap_Fmap: "Fmap (g1 \ f1) (g2 \ f2) (g3 \ f3) x = Fmap g1 g2 g3 (Fmap f1 f2 f3 x)" by (fact fun_cong[OF Fmap_comp, simplified comp_eq]) lemma Fset1_Fmap_eq_image_Fset1: "Fset1 (Fmap f1 f2 f3 x) = f1 ` Fset1 x" by (fact fun_cong[OF Fset1_natural, simplified comp_eq]) lemma Fset2_Fmap_eq_image_Fset2: "Fset2 (Fmap f1 f2 f3 x) = f2 ` Fset2 x" by (fact fun_cong[OF Fset2_natural, simplified comp_eq]) lemma Fset3_Fmap_eq_image_Fset3: "Fset3 (Fmap f1 f2 f3 x) = f3 ` Fset3 x" by (fact fun_cong[OF Fset3_natural, simplified comp_eq]) lemmas Fset_Fmap_eqs = Fset1_Fmap_eq_image_Fset1 Fset2_Fmap_eq_image_Fset2 Fset3_Fmap_eq_image_Fset3 paragraph \Relator\ definition Frel :: "('a1 \ 'a2 \ bool) \ ('b1 \ 'b2 \ bool) \ ('c1 \ 'c2 \ bool) \ ('d, 'a1, 'b1, 'c1) F \ ('d, 'a2, 'b2, 'c2) F \ bool" where "Frel R1 R2 R3 x y \ (\z. z \ {x. Fset1 x \ {(x, y). R1 x y} \ Fset2 x \ {(x, y). R2 x y} \ Fset3 x \ {(x, y). R3 x y}} \ Fmap fst fst fst z = x \ Fmap snd snd snd z = y)" lemma FrelI: assumes "Fset1 z \ {(x, y). R1 x y}" and "Fset2 z \ {(x, y). R2 x y}" and "Fset3 z \ {(x, y). R3 x y}" and "Fmap fst fst fst z = x" and "Fmap snd snd snd z = y" shows "Frel R1 R2 R3 x y" apply (subst Frel_def) apply (intro exI conjI CollectI) apply (fact assms)+ done lemma FrelE: assumes "Frel R1 R2 R3 x y" obtains z where "Fset1 z \ {(x, y). R1 x y}" "Fset2 z \ {(x, y). R2 x y}" "Fset3 z \ {(x, y). R3 x y}" "Fmap fst fst fst z = x" "Fmap snd snd snd z = y" apply (insert assms) apply (subst (asm) Frel_def) apply (elim exE CollectE conjE) apply assumption done lemma Grp_UNIV_Fmap_eq_Frel_Grp: "BNF_Def.Grp UNIV (Fmap f1 f2 f3) = Frel (BNF_Def.Grp UNIV f1) (BNF_Def.Grp UNIV f2) (BNF_Def.Grp UNIV f3)" apply (intro ext iffI) apply (rule FrelI[where ?z="Fmap (BNF_Def.convol id f1) (BNF_Def.convol id f2) (BNF_Def.convol id f3) _"]) apply (subst Fset_Fmap_eqs, rule image_subsetI, rule convol_mem_GrpI[simplified Fun_id_eq_id], rule UNIV_I)+ apply (unfold Fmap_comp_eq_Fmap_Fmap[symmetric] fst_convol[simplified Fun_comp_eq_comp] snd_convol[simplified Fun_comp_eq_comp] Fmap_id id_eq_self) apply (rule refl) apply (subst (asm) Grp_UNIV_eq_eq_comp) apply (subst (asm) comp_eq) apply assumption apply (erule FrelE) apply hypsubst apply (subst Grp_UNIV_eq_eq_comp) apply (subst comp_eq) apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (rule Fmap_cong; rule Collect_case_prod_Grp_eqD[simplified Fun_comp_eq_comp], drule rev_subsetD, assumption+) done lemma Frel_Grp_UNIV_Fmap: "Frel (BNF_Def.Grp UNIV f1) (BNF_Def.Grp UNIV f2) (BNF_Def.Grp UNIV f3) x (Fmap f1 f2 f3 x)" apply (subst Grp_UNIV_Fmap_eq_Frel_Grp[symmetric]) apply (subst Grp_UNIV_eq_eq_comp) apply (subst comp_eq) apply (rule refl) done lemma Frel_Grp_UNIV_iff_eq_Fmap: "Frel (BNF_Def.Grp UNIV f1) (BNF_Def.Grp UNIV f2) (BNF_Def.Grp UNIV f3) x y \ (y = Fmap f1 f2 f3 x)" by (subst eq_commute[of y]) (fact fun_cong[OF fun_cong[OF Grp_UNIV_Fmap_eq_Frel_Grp], simplified Grp_UNIV_eq_eq_comp comp_eq, folded Grp_UNIV_eq_eq_comp, symmetric]) lemma Frel_eq: "Frel (=) (=) (=) = (=)" apply (unfold BNF_Def.eq_alt[simplified Fun_id_eq_id]) apply (subst Grp_UNIV_Fmap_eq_Frel_Grp[symmetric]) apply (subst Fmap_id) apply (fold BNF_Def.eq_alt[simplified Fun_id_eq_id]) apply (rule refl) done corollary Frel_eq_self: "Frel (=) (=) (=) x x" by (fact iffD2[OF fun_cong[OF fun_cong[OF Frel_eq]] refl]) lemma Frel_mono_strong: assumes "Frel R1 R2 R3 x y" and "\x1 y1. x1 \ Fset1 x \ y1 \ Fset1 y \ R1 x1 y1 \ S1 x1 y1" and "\x2 y2. x2 \ Fset2 x \ y2 \ Fset2 y \ R2 x2 y2 \ S2 x2 y2" and "\x3 y3. x3 \ Fset3 x \ y3 \ Fset3 y \ R3 x3 y3 \ S3 x3 y3" shows "Frel S1 S2 S3 x y" apply (insert assms(1)) apply (erule FrelE) apply (rule FrelI) apply (rule subsetI, frule rev_subsetD, assumption, frule imageI[of _ "Fset1 _" fst] imageI[of _ "Fset2 _" fst] imageI[of _ "Fset3 _" fst], drule imageI[of _ "Fset1 _" snd] imageI[of _ "Fset2 _" snd] imageI[of _ "Fset3 _" snd], (subst (asm) Fset_Fmap_eqs[symmetric])+, intro CollectI case_prodI2, rule assms; hypsubst, unfold fst_conv snd_conv, (elim CollectE case_prodE Pair_inject, hypsubst)?, assumption)+ apply assumption+ done corollary Frel_mono: assumes "R1 \ S1" "R2 \ S2" "R3 \ S3" shows "Frel R1 R2 R3 \ Frel S1 S2 S3" apply (intro le_relI) apply (rule Frel_mono_strong) apply assumption apply (insert assms) apply (drule le_relD[OF assms(1)] le_relD[OF assms(2)] le_relD[OF assms(3)], assumption)+ done lemma Frel_refl_strong: assumes "\x1. x1 \ Fset1 x \ R1 x1 x1" and "\x2. x2 \ Fset2 x \ R2 x2 x2" and "\x3. x3 \ Fset3 x \ R3 x3 x3" shows "Frel R1 R2 R3 x x" by (rule Frel_mono_strong[OF Frel_eq_self[of x]]; drule assms, hypsubst, assumption) lemma Frel_cong: assumes "\x1 y1. x1 \ Fset1 x \ y1 \ Fset1 y \ R1 x1 y1 \ R1' x1 y1" and "\x2 y2. x2 \ Fset2 x \ y2 \ Fset2 y \ R2 x2 y2 \ R2' x2 y2" and "\x3 y3. x3 \ Fset3 x \ y3 \ Fset3 y \ R3 x3 y3 \ R3' x3 y3" shows "Frel R1 R2 R3 x y = Frel R1' R2' R3' x y" by (rule iffI; rule Frel_mono_strong, assumption; rule iffD1[OF assms(1)] iffD1[OF assms(2)] iffD1[OF assms(3)] iffD2[OF assms(1)] iffD2[OF assms(2)] iffD2[OF assms(3)]; assumption) lemma Frel_rel_inv_eq_rel_inv_Frel: "Frel R1\ R2\ R3\ = (Frel R1 R2 R3)\" by (intro ext iffI; unfold rel_inv_iff_rel, erule FrelE, hypsubst, rule FrelI[where ?z="Fmap prod.swap prod.swap prod.swap _"]; ((subst Fset_Fmap_eqs, rule image_subsetI, drule rev_subsetD, assumption, elim CollectE case_prodE, hypsubst, subst swap_simp, subst pair_mem_Collect_case_prod_iff, assumption) | (subst Fmap_comp_eq_Fmap_Fmap[symmetric], rule Fmap_cong; unfold comp_eq fst_swap snd_swap, rule refl))) text \Given the former axioms, the following axiom - subdistributivity of the relator - is equivalent to the (F, Fmap) functor preserving weak pullbacks.\ axiomatization where Frel_comp_le_Frel_rel_comp: "\R1 R2 R3 S1 S2 S3. Frel R1 R2 R3 \\ Frel S1 S2 S3 \ Frel (R1 \\ S1) (R2 \\ S2) (R3 \\ S3)" lemma fst_sndOp_eq_snd_fstOp: "fst \ BNF_Def.sndOp P Q = snd \ BNF_Def.fstOp P Q" unfolding fstOp_def sndOp_def by (intro ext) simp lemma Frel_rel_comp_le_Frel_comp: "Frel (R1 \\ S1) (R2 \\ S2) (R3 \\ S3) \ (Frel R1 R2 R3 \\ Frel S1 S2 S3)" apply (rule le_relI) apply (erule FrelE) apply (rule rel_compI[where ?y="Fmap (snd \ BNF_Def.fstOp R1 S1) (snd \ BNF_Def.fstOp R2 S2) (snd \ BNF_Def.fstOp R3 S3) _"]) apply (rule FrelI[where ?z="Fmap (BNF_Def.fstOp R1 S1) (BNF_Def.fstOp R2 S2) (BNF_Def.fstOp R3 S3) _"]) apply (subst Fset_Fmap_eqs, intro image_subsetI, rule fstOp_in[unfolded relcompp_eq_rel_comp], drule rev_subsetD, assumption+)+ apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (fold ext[of fst "fst \ _", OF fst_fstOp[unfolded Fun_comp_eq_comp]]) apply hypsubst apply (rule refl) apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (rule refl) apply (rule FrelI[where ?z="Fmap (BNF_Def.sndOp R1 S1) (BNF_Def.sndOp R2 S2) (BNF_Def.sndOp R3 S3) _"]) apply (subst Fset_Fmap_eqs, intro image_subsetI, rule sndOp_in[unfolded relcompp_eq_rel_comp], drule rev_subsetD, assumption+)+ apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (unfold fst_sndOp_eq_snd_fstOp) apply (rule refl) apply (subst Fmap_comp_eq_Fmap_Fmap[symmetric]) apply (fold ext[of snd "snd \ _", OF snd_sndOp[unfolded Fun_comp_eq_comp]]) apply hypsubst apply (rule refl) done corollary Frel_comp_eq_Frel_rel_comp: "Frel R1 R2 R3 \\ Frel S1 S2 S3 = Frel (R1 \\ S1) (R2 \\ S2) (R3 \\ S3)" by (rule antisym; rule Frel_comp_le_Frel_rel_comp Frel_rel_comp_le_Frel_comp) lemma Frel_Fmap_eq1: "Frel R1 R2 R3 (Fmap f1 f2 f3 x) y = Frel (\x. R1 (f1 x)) (\x. R2 (f2 x)) (\x. R3 (f3 x)) x y" apply (rule iffI) apply (fold comp_eq[of R1] comp_eq[of R2] comp_eq[of R3]) apply (drule rel_compI[where ?R="Frel _ _ _" and ?S="Frel _ _ _", OF Frel_Grp_UNIV_Fmap]) apply (unfold Grp_UNIV_eq_eq_comp) apply (drule le_relD[OF Frel_comp_le_Frel_rel_comp]) apply (unfold eq_comp_rel_comp_eq_comp) apply assumption apply (fold eq_comp_rel_comp_eq_comp[where ?R=R1] eq_comp_rel_comp_eq_comp[where ?R=R2] eq_comp_rel_comp_eq_comp[where ?R=R3] Grp_UNIV_eq_eq_comp) apply (drule le_relD[OF Frel_rel_comp_le_Frel_comp]) apply (erule rel_compE) apply (subst (asm) Frel_Grp_UNIV_iff_eq_Fmap) apply hypsubst apply assumption done lemma Frel_Fmap_eq2: "Frel R1 R2 R3 x (Fmap g1 g2 g3 y) = Frel (\x y. R1 x (g1 y)) (\x y. R2 x (g2 y)) (\x y. R3 x (g3 y)) x y" apply (subst rel_inv_iff_rel[of "Frel _ _ _", symmetric]) apply (subst Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (subst Frel_Fmap_eq1) apply (rule sym) apply (subst rel_inv_iff_rel[of "Frel _ _ _", symmetric]) apply (subst Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (unfold rel_inv_iff_rel) apply (rule refl) done lemmas Frel_Fmap_eqs = Frel_Fmap_eq1 Frel_Fmap_eq2 paragraph \Predicator\ definition Fpred :: "('a \ bool) \ ('b \ bool) \ ('c \ bool) \ ('d, 'a, 'b, 'c) F \ bool" where "Fpred P1 P2 P3 x \ Frel ((=)\\<^bsub>P1\<^esub>) ((=)\\<^bsub>P2\<^esub>) ((=)\\<^bsub>P3\<^esub>) x x" lemma Fpred_mono_strong: assumes "Fpred P1 P2 P3 x" and "\x1. x1 \ Fset1 x \ P1 x1 \ Q1 x1" and "\x2. x2 \ Fset2 x \ P2 x2 \ Q2 x2" and "\x3. x3 \ Fset3 x \ P3 x3 \ Q3 x3" shows "Fpred Q1 Q2 Q3 x" apply (insert assms(1)) apply (unfold Fpred_def) apply (rule Frel_mono_strong, assumption; - erule bin_rel_restrict_leftE, - rule bin_rel_restrict_leftI, + 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 bin_rel_restrict_left_top_eq, + 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_bin_rel_restrict_left_eq, + 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) 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 bin_rel_restrict_leftI, + rule rel_restrict_leftI, rule refl, drule rev_subsetD, assumption, - elim CollectE case_prodE bin_rel_restrict_leftE, + 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 bin_rel_restrict_leftE) + apply (erule rel_restrict_leftE) apply (elim FpredE) apply (rule Frel_mono_strong, assumption; - rule bin_rel_restrict_leftI, + rule rel_restrict_leftI, assumption+) - apply (rule bin_rel_restrict_leftI) + apply (rule rel_restrict_leftI) apply (rule Frel_mono_strong, assumption; - erule bin_rel_restrict_leftE, + 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_bin_rel_restrict_left_eq inf_apply inf_bool_def; + 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 bin_rel_restrict_right_eq) + 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 bin_rel_restrict_right_eq) + 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 (unfold Frel_Fmap_eqs) apply (fold rel_map_eq) apply (rule le_relD[OF Frel_mono]) apply (subst mono_wrt_rel_iff_le_rel_map[symmetric], rule assms)+ apply assumption done end end \ No newline at end of file diff --git a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Galois_Relator.thy b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Galois_Relator.thy --- a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Galois_Relator.thy +++ b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Galois_Relator.thy @@ -1,55 +1,55 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Galois Relator\ theory Transport_Natural_Functors_Galois_Relator imports Transport_Natural_Functors_Base begin context transport_natural_functor begin lemma left_Galois_Frel_left_Galois: "(\<^bsub>L\<^esub>\) \ Frel (\<^bsub>L1\<^esub>\) (\<^bsub>L2\<^esub>\) (\<^bsub>L3\<^esub>\)" apply (rule le_relI) apply (erule left_GaloisE) apply (unfold left_rel_eq_Frel right_rel_eq_Frel right_eq_Fmap) apply (subst (asm) in_codom_Frel_eq_Fpred_in_codom) apply (erule FpredE) apply (subst (asm) Frel_Fmap_eq2) apply (rule Frel_mono_strong, assumption; rule t1.left_GaloisI t2.left_GaloisI t3.left_GaloisI; assumption) done lemma Frel_left_Galois_le_left_Galois: "Frel (\<^bsub>L1\<^esub>\) (\<^bsub>L2\<^esub>\) (\<^bsub>L3\<^esub>\) \ (\<^bsub>L\<^esub>\)" apply (rule le_relI) apply (unfold t1.left_Galois_iff_in_codom_and_left_rel_right t2.left_Galois_iff_in_codom_and_left_rel_right t3.left_Galois_iff_in_codom_and_left_rel_right) apply (fold - bin_rel_restrict_right_eq[of "\x y. x \\<^bsub>L1\<^esub> r1 y" "in_codom (\\<^bsub>R1\<^esub>)", - unfolded bin_rel_restrict_left_pred_def rel_inv_iff_rel] - bin_rel_restrict_right_eq[of "\x y. x \\<^bsub>L2\<^esub> r2 y" "in_codom (\\<^bsub>R2\<^esub>)", - unfolded bin_rel_restrict_left_pred_def rel_inv_iff_rel] - bin_rel_restrict_right_eq[of "\x y. x \\<^bsub>L3\<^esub> r3 y" "in_codom (\\<^bsub>R3\<^esub>)", - unfolded bin_rel_restrict_left_pred_def rel_inv_iff_rel]) + rel_restrict_right_eq[of "\x y. x \\<^bsub>L1\<^esub> r1 y" "in_codom (\\<^bsub>R1\<^esub>)", + unfolded rel_restrict_left_pred_def rel_inv_iff_rel] + rel_restrict_right_eq[of "\x y. x \\<^bsub>L2\<^esub> r2 y" "in_codom (\\<^bsub>R2\<^esub>)", + unfolded rel_restrict_left_pred_def rel_inv_iff_rel] + rel_restrict_right_eq[of "\x y. x \\<^bsub>L3\<^esub> r3 y" "in_codom (\\<^bsub>R3\<^esub>)", + unfolded rel_restrict_left_pred_def rel_inv_iff_rel]) apply (subst (asm) Frel_restrict_right_Fpred_eq_Frel_restrict_right[symmetric]) - apply (erule bin_rel_restrict_rightE) + apply (erule rel_restrict_rightE) apply (subst (asm) in_codom_Frel_eq_Fpred_in_codom[symmetric]) apply (erule in_codomE) apply (rule left_GaloisI) apply (rule in_codomI) apply (subst right_rel_eq_Frel) apply assumption apply (unfold left_rel_eq_Frel right_eq_Fmap Frel_Fmap_eq2) apply assumption done corollary left_Galois_eq_Frel_left_Galois: "(\<^bsub>L\<^esub>\) = Frel (\<^bsub>L1\<^esub>\) (\<^bsub>L2\<^esub>\) (\<^bsub>L3\<^esub>\)" by (intro antisym left_Galois_Frel_left_Galois Frel_left_Galois_le_left_Galois) end end \ No newline at end of file diff --git a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Base.thy b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Base.thy --- a/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Base.thy +++ b/thys/Transport/Transport/Natural_Functors/Transport_Natural_Functors_Order_Base.thy @@ -1,99 +1,105 @@ \<^marker>\creator "Kevin Kappelmann"\ subsection \Basic Order Properties\ theory Transport_Natural_Functors_Order_Base imports Transport_Natural_Functors_Base begin +context + fixes R1 :: "'a \ 'a \ bool" and R2 :: "'b \ 'b \ bool" and R3 :: "'c \ 'c \ bool" +begin + lemma reflexive_on_in_field_FrelI: assumes "reflexive_on (in_field R1) R1" and "reflexive_on (in_field R2) R2" and "reflexive_on (in_field R3) R3" defines "R \ Frel R1 R2 R3" shows "reflexive_on (in_field R) R" apply (subst reflexive_on_iff_eq_restrict_le) apply (subst Frel_eq[symmetric]) apply (unfold R_def) apply (subst in_field_Frel_eq_Fpred_in_in_field) - apply (subst bin_rel_restrict_left_sup_eq) + apply (subst rel_restrict_left_sup_eq) apply (subst Frel_restrict_left_Fpred_eq_Frel_restrict_left)+ apply (rule le_supI; rule Frel_mono; subst reflexive_on_iff_eq_restrict_le[symmetric], rule reflexive_on_if_le_pred_if_reflexive_on, rule assms, rule le_predI[OF in_field_if_in_dom] le_predI[OF in_field_if_in_codom], assumption) done lemma transitive_FrelI: assumes "transitive R1" and "transitive R2" and "transitive R3" shows "transitive (Frel R1 R2 R3)" apply (subst transitive_iff_rel_comp_le_self) apply (subst Frel_comp_eq_Frel_rel_comp) apply (rule Frel_mono; subst transitive_iff_rel_comp_le_self[symmetric], rule assms) done lemma preorder_on_in_field_FrelI: assumes "preorder_on (in_field R1) R1" and "preorder_on (in_field R2) R2" and "preorder_on (in_field R3) R3" defines "R \ Frel R1 R2 R3" shows "preorder_on (in_field R) R" apply (unfold R_def) apply (insert assms) apply (elim preorder_on_in_fieldE) apply (rule preorder_onI) apply (rule reflexive_on_in_field_FrelI; assumption) apply (subst transitive_on_in_field_iff_transitive) apply (rule transitive_FrelI; assumption) done lemma symmetric_FrelI: assumes "symmetric R1" and "symmetric R2" and "symmetric R3" shows "symmetric (Frel R1 R2 R3)" apply (subst symmetric_iff_rel_inv_eq_self) apply (subst Frel_rel_inv_eq_rel_inv_Frel[symmetric]) apply (subst rel_inv_eq_self_if_symmetric, fact)+ apply (rule refl) done lemma partial_equivalence_rel_FrelI: assumes "partial_equivalence_rel R1" and "partial_equivalence_rel R2" and "partial_equivalence_rel R3" shows "partial_equivalence_rel (Frel R1 R2 R3)" apply (insert assms) apply (elim partial_equivalence_relE preorder_on_in_fieldE) apply (rule partial_equivalence_relI; rule transitive_FrelI symmetric_FrelI; assumption) done +end + context transport_natural_functor begin lemmas reflexive_on_in_field_leftI = reflexive_on_in_field_FrelI [of L1 L2 L3, folded transport_defs] lemmas transitive_leftI = transitive_FrelI[of L1 L2 L3, folded transport_defs] lemmas preorder_on_in_field_leftI = preorder_on_in_field_FrelI [of L1 L2 L3, folded transport_defs] lemmas symmetricI = symmetric_FrelI[of L1 L2 L3, folded transport_defs] lemmas partial_equivalence_rel_leftI = partial_equivalence_rel_FrelI [of L1 L2 L3, folded transport_defs] end end \ No newline at end of file diff --git a/thys/Transport/Transport/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,165 +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" - defines "R \ Frel R1 R2 R3" - shows "inflationary_on (in_dom R) R (Fmap f1 f2 f3)" - apply (unfold R_def) + 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" - defines "R \ Frel R1 R2 R3" - shows "inflationary_on (in_codom R) R (Fmap f1 f2 f3)" - apply (unfold R_def) + 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" - defines "R \ Frel R1 R2 R3" - shows "inflationary_on (in_field R) R (Fmap f1 f2 f3)" - apply (unfold R_def) + 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" - defines "R \ Frel R1 R2 R3" - shows "deflationary_on (in_dom R) R (Fmap f1 f2 f3)" - apply (unfold R_def) + 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" - defines "R \ Frel R1 R2 R3" - shows "deflationary_on (in_codom R) R (Fmap f1 f2 f3)" - apply (unfold R_def) + 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" - defines "R \ Frel R1 R2 R3" - shows "deflationary_on (in_field R) R (Fmap f1 f2 f3)" - apply (unfold R_def) + 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 (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, 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" - defines "R \ Frel R1 R2 R3" - shows "rel_equivalence_on (in_field R) R (Fmap f1 f2 f3)" - apply (unfold R_def) + 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; - rule inflationary_on_in_field_FrelI deflationary_on_in_field_FrelI; - assumption) + 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 diff --git a/thys/Transport/Transport/Transport_Bijections.thy b/thys/Transport/Transport/Transport_Bijections.thy --- a/thys/Transport/Transport/Transport_Bijections.thy +++ b/thys/Transport/Transport/Transport_Bijections.thy @@ -1,199 +1,199 @@ \<^marker>\creator "Kevin Kappelmann"\ section \Transport using Bijections\ theory Transport_Bijections imports Restricted_Equality Functions_Bijection Transport_Base begin paragraph \Summary\ text \Setup for Transport using bijective transport functions.\ locale transport_bijection = fixes L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes mono_wrt_rel_left: "(L \\<^sub>m R) l" and mono_wrt_rel_right: "(R \\<^sub>m L) r" and inverse_left_right: "inverse_on (in_field L) l r" and inverse_right_left: "inverse_on (in_field R) r l" begin interpretation transport L R l r . interpretation g_flip_inv : galois "(\\<^bsub>R\<^esub>)" "(\\<^bsub>L\<^esub>)" r l . lemma bijection_on_in_field: "bijection_on (in_field (\\<^bsub>L\<^esub>)) (in_field (\\<^bsub>R\<^esub>)) l r" using mono_wrt_rel_left mono_wrt_rel_right inverse_left_right inverse_right_left by (intro bijection_onI in_field_if_in_field_if_mono_wrt_rel) auto lemma half_galois_prop_left: "((\\<^bsub>L\<^esub>) \<^sub>h\ (\\<^bsub>R\<^esub>)) l r" using mono_wrt_rel_left inverse_right_left by (intro half_galois_prop_leftI) (fastforce dest: inverse_onD) lemma half_galois_prop_right: "((\\<^bsub>L\<^esub>) \\<^sub>h (\\<^bsub>R\<^esub>)) l r" using mono_wrt_rel_right inverse_left_right by (intro half_galois_prop_rightI) (force dest: in_field_if_in_dom inverse_onD) lemma galois_prop: "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using half_galois_prop_left half_galois_prop_right by (intro galois_propI) lemma galois_connection: "((\\<^bsub>L\<^esub>) \ (\\<^bsub>R\<^esub>)) l r" using mono_wrt_rel_left mono_wrt_rel_right galois_prop by (intro galois_connectionI) lemma rel_equivalence_on_unitI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" shows "rel_equivalence_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>) \" using assms inverse_left_right by (subst rel_equivalence_on_unit_iff_reflexive_on_if_inverse_on) interpretation flip : transport_bijection R L r l rewrites "order_functors.unit r l \ \" using mono_wrt_rel_left mono_wrt_rel_right inverse_left_right inverse_right_left by unfold_locales (simp_all only: flip_unit_eq_counit) lemma galois_equivalence: "((\\<^bsub>L\<^esub>) \\<^sub>G (\\<^bsub>R\<^esub>)) l r" using galois_connection flip.galois_prop by (intro galois_equivalenceI) lemmas rel_equivalence_on_counitI = flip.rel_equivalence_on_unitI lemma order_equivalenceI: assumes "reflexive_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "reflexive_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^sub>o (\\<^bsub>R\<^esub>)) l r" using assms mono_wrt_rel_left mono_wrt_rel_right rel_equivalence_on_unitI rel_equivalence_on_counitI by (intro order_equivalenceI) lemma preorder_equivalenceI: assumes "preorder_on (in_field (\\<^bsub>L\<^esub>)) (\\<^bsub>L\<^esub>)" and "preorder_on (in_field (\\<^bsub>R\<^esub>)) (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^bsub>pre\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro preorder_equivalence_if_galois_equivalenceI galois_equivalence) simp_all lemma partial_equivalence_rel_equivalenceI: assumes "partial_equivalence_rel (\\<^bsub>L\<^esub>)" and "partial_equivalence_rel (\\<^bsub>R\<^esub>)" shows "((\\<^bsub>L\<^esub>) \\<^bsub>PER\<^esub> (\\<^bsub>R\<^esub>)) l r" using assms by (intro partial_equivalence_rel_equivalence_if_galois_equivalenceI galois_equivalence) simp_all end locale transport_reflexive_on_in_field_bijection = fixes L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes reflexive_on_in_field_left: "reflexive_on (in_field L) L" and reflexive_on_in_field_right: "reflexive_on (in_field R) R" and transport_bijection: "transport_bijection L R l r" begin sublocale tbij? : transport_bijection L R l r rewrites "reflexive_on (in_field L) L \ True" and "reflexive_on (in_field R) R \ True" and "\P. (True \ P) \ Trueprop P" using transport_bijection reflexive_on_in_field_left reflexive_on_in_field_right by auto lemmas rel_equivalence_on_unit = rel_equivalence_on_unitI lemmas rel_equivalence_on_counit = rel_equivalence_on_counitI lemmas order_equivalence = order_equivalenceI end locale transport_preorder_on_in_field_bijection = fixes L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes preorder_on_in_field_left: "preorder_on (in_field L) L" and preorder_on_in_field_right: "preorder_on (in_field R) R" and transport_bijection: "transport_bijection L R l r" begin sublocale trefl_bij? : transport_reflexive_on_in_field_bijection L R l r rewrites "preorder_on (in_field L) L \ True" and "preorder_on (in_field R) R \ True" and "\P. (True \ P) \ Trueprop P" using transport_bijection by (intro transport_reflexive_on_in_field_bijection.intro) (insert preorder_on_in_field_left preorder_on_in_field_right, auto) lemmas preorder_equivalence = preorder_equivalenceI end locale transport_partial_equivalence_rel_bijection = fixes L :: "'a \ 'a \ bool" and R :: "'b \ 'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes partial_equivalence_rel_left: "partial_equivalence_rel L" and partial_equivalence_rel_right: "partial_equivalence_rel R" and transport_bijection: "transport_bijection L R l r" begin sublocale tpre_bij? : transport_preorder_on_in_field_bijection L R l r rewrites "partial_equivalence_rel L \ True" and "partial_equivalence_rel R \ True" and "\P. (True \ P) \ Trueprop P" using transport_bijection by (intro transport_preorder_on_in_field_bijection.intro) (insert partial_equivalence_rel_left partial_equivalence_rel_right, auto) lemmas partial_equivalence_rel_equivalence = partial_equivalence_rel_equivalenceI end locale transport_eq_restrict_bijection = fixes P :: "'a \ bool" and Q :: "'b \ bool" and l :: "'a \ 'b" and r :: "'b \ 'a" assumes bijection_on_in_field: "bijection_on (in_field ((=\<^bsub>P\<^esub>) :: 'a \ _)) (in_field ((=\<^bsub>Q\<^esub>) :: 'b \ _)) l r" begin interpretation transport "(=\<^bsub>P\<^esub>)" "(=\<^bsub>Q\<^esub>)" l r . sublocale tper_bij? : transport_partial_equivalence_rel_bijection "(=\<^bsub>P\<^esub>)" "(=\<^bsub>Q\<^esub>)" l r using bijection_on_in_field partial_equivalence_rel_eq_restrict by unfold_locales (auto elim: bijection_onE intro!: mono_wrt_rel_left_if_reflexive_on_if_le_eq_if_mono_wrt_in_field[of "in_field (=\<^bsub>Q\<^esub>)"] flip_of.mono_wrt_rel_left_if_reflexive_on_if_le_eq_if_mono_wrt_in_field[of "in_field (=\<^bsub>P\<^esub>)"]) lemma left_Galois_eq_Galois_eq_eq_restrict: "(\<^bsub>L\<^esub>\) = (galois_rel.Galois (=) (=) r)\\<^bsub>P\<^esub>\\<^bsub>Q\<^esub>" by (subst galois_rel.left_Galois_restrict_left_eq_left_Galois_left_restrict_left galois_rel.left_Galois_restrict_right_eq_left_Galois_right_restrict_right - bin_rel_restrict_right_eq rel_inv_eq_self_if_symmetric)+ + rel_restrict_right_eq rel_inv_eq_self_if_symmetric)+ auto end locale transport_eq_bijection = fixes l :: "'a \ 'b" and r :: "'b \ 'a" assumes bijection_on_in_field: "bijection_on (in_field ((=) :: 'a \ _)) (in_field ((=) :: 'b \ _)) l r" begin sublocale teq_restr_bij? : transport_eq_restrict_bijection \ \ l r rewrites "(=\<^bsub>\ :: 'a \ bool\<^esub>) = ((=) :: 'a \ _)" and "(=\<^bsub>\ :: 'b \ bool\<^esub>) = ((=) :: 'b \ _)" using bijection_on_in_field by unfold_locales simp_all end end \ No newline at end of file diff --git a/tools/afp_submit.scala b/tools/afp_submit.scala --- a/tools/afp_submit.scala +++ b/tools/afp_submit.scala @@ -1,1622 +1,1619 @@ /* Author: Fabian Huch, TU Muenchen AFP submission system backend. */ package afp import isabelle.* import isabelle.HTML.* import afp.Web_App.{ACTION, API, FILE, Params} import afp.Web_App.Params.{List_Key, Nest_Key, empty} import afp.Web_App.HTML.* import afp.Metadata.{Affiliation, Author, DOI, Email, Entry, Formatted, Homepage, License, Orcid, Reference, Release, Topic, Unaffiliated} -import java.net.URL import java.text.Normalizer import java.time.LocalDate import scala.collection.immutable.StringView import scala.util.matching.Regex object AFP_Submit { case class Val_Opt[A] private(opt: Option[A], err: Option[String]) { def is_empty: Boolean = opt.isEmpty } object Val_Opt { def ok[A](value: A): Val_Opt[A] = Val_Opt(Some(value), None) def user_err[A](msg: String): Val_Opt[A] = Val_Opt(None, Some(msg)) def error[A]: Val_Opt[A] = Val_Opt(None, None) } case class Val[A] private(v: A, err: Option[String]) { def with_err(err: String): Val[A] = Val.err(v, err) def perhaps_err(opt: Val_Opt[_]): Val[A] = opt.err.map(with_err).getOrElse(this) } object Val { def ok[A](value: A): Val[A] = Val(value, None) def err[A](value: A, msg: String): Val[A] = Val(value, Some(msg)) } /* data model */ object Model { object Related extends Enumeration { val DOI, Plaintext = Value def from_string(s: String): Option[Value] = values.find(_.toString == s) def get(r: Reference): Value = r match { case afp.Metadata.DOI(_) => DOI case afp.Metadata.Formatted(_) => Plaintext } } case class Create_Entry( name: Val[String] = Val.ok(""), title: Val[String] = Val.ok(""), affils: Val[List[Affiliation]] = Val.ok(Nil), notifies: Val[List[Email]] = Val.ok(Nil), author_input: Option[Author] = None, notify_input: Option[Author] = None, topics: Val[List[Topic]] = Val.ok(Nil), topic_input: Option[Topic] = None, license: License, `abstract`: Val[String] = Val.ok(""), related: List[Reference] = Nil, related_kind: Option[Related.Value] = None, related_input: Val[String] = Val.ok("") ) { val used_affils: Set[Affiliation] = (affils.v ++ notifies.v).toSet val used_authors: Set[Author.ID] = used_affils.map(_.author) } case class Create( entries: Val[List[Create_Entry]] = Val.ok(Nil), new_authors: Val[List[Author]] = Val.ok(Nil), new_author_input: String = "", new_author_orcid: String = "", new_affils: Val[List[Affiliation]] = Val.ok(Nil), new_affils_author: Option[Author] = None, new_affils_input: String = "", ) extends T { def update_entry(num: Int, entry: Create_Entry): Create = this.copy(entries = Val.ok(entries.v.updated(num, entry))) def updated_authors(old_authors: Map[Author.ID, Author]): Map[Author.ID, Author] = (old_authors ++ new_authors.v.map(author => author.id -> author)).map { case (id, author) => val emails = author.emails ++ new_affils.v.collect { case e: Email if e.author == id => e } val homepages = author.homepages ++ new_affils.v.collect { case h: Homepage if h.author == id => h } id -> author.copy(emails = emails.distinct, homepages = homepages.distinct) } val used_affils: Set[Affiliation] = entries.v.toSet.flatMap(_.used_affils) val used_authors: Set[Author.ID] = new_affils.v.map(_.author).toSet ++ entries.v.flatMap(_.used_authors) def create(authors: Map[Author.ID, Author]): (Map[Author.ID, Author], List[Entry]) = (updated_authors(authors), entries.v.map(entry => Entry( name = entry.name.v, title = entry.title.v, authors = entry.affils.v, date = LocalDate.now(), topics = entry.topics.v, `abstract` = entry.`abstract`.v.trim, notifies = entry.notifies.v, license = entry.license, note = "", related = entry.related))) } object Build extends Enumeration { val Pending, Running, Aborted, Failed, Success = Value } object Status extends Enumeration { val Submitted, Review, Added, Rejected = Value def from_string(s: String): Option[Value] = values.find(_.toString == s) } case class Overview(id: String, date: LocalDate, name: String, status: Status.Value) case class Metadata(authors: Map[Author.ID, Author], entries: List[Entry]) { def new_authors(old_authors: Map[Author.ID, Author]): Set[Author] = entries.flatMap(_.authors).map(_.author).filterNot(old_authors.contains).toSet.map(authors) def new_affils(old_authors: Map[Author.ID, Author]): Set[Affiliation] = entries.flatMap(entry => entry.authors ++ entry.notifies).toSet.filter { case _: Unaffiliated => false case e: Email => !old_authors.get(e.author).exists(_.emails.contains(e)) case h: Homepage => !old_authors.get(h.author).exists(_.homepages.contains(h)) } } sealed trait T case object Invalid extends T case class Upload(meta: Metadata, message: String, error: String) extends T case class Created(id: String) extends T case class Submission( id: Handler.ID, meta: Metadata, message: String, build: Build.Value, status: Option[Status.Value], log: String) extends T case class Submission_List(submissions: List[Overview]) extends T } /* Physical submission handling */ trait Handler { def create( date: Date, meta: Model.Metadata, message: String, archive: Bytes, ext: String ): Handler.ID def list(): Model.Submission_List def get(id: Handler.ID, topics: Map[Topic.ID, Topic], licenses: Map[License.ID, License] ): Option[Model.Submission] def submit(id: Handler.ID): Unit def set_status(id: Handler.ID, status: Model.Status.Value): Unit def abort_build(id: Handler.ID): Unit def get_patch(id: Handler.ID): Option[Path] def get_archive(id: Handler.ID): Option[Path] } object Handler { import Model._ type ID = String object ID { private val format = Date.Format.make( List( Date.Formatter.pattern("uuuu-MM-dd_HH-mm-ss_SSS"), Date.Formatter.pattern("uuuu-MM-dd_HH-mm-ss_SSS_VV")), _ + "_" + Date.timezone().getId) def apply(submission_time: Date): ID = format(submission_time) def unapply(id: ID): Option[Date] = format.unapply(id) def check(id: ID): Option[ID] = unapply(id).map(apply) } /* Handler for local edits */ class Edit(afp_structure: AFP_Structure) extends Handler { val authors = afp_structure.load_authors.map(author => author.id -> author).toMap val topics = afp_structure.load_topics.flatMap(_.all_topics) val all_topics = topics.map(topic => topic.id -> topic).toMap val licenses = afp_structure.load_licenses.map(license => license.id -> license).toMap val releases = afp_structure.load_releases.groupBy(_.entry) val dates = afp_structure.load().map(entry => entry.name -> entry.date).toMap override def create( date: Date, meta: Metadata, message: String, archive: Bytes, ext: String ): ID = { val entry = meta.entries match { case e :: Nil => e case _ => isabelle.error("Must be a single entry") } val old = afp_structure.load_entry(entry.name, authors, all_topics, licenses, releases) val updated = old.copy( title = entry.title, authors = entry.authors, topics = entry.topics, `abstract` = entry.`abstract`, notifies = entry.notifies, license = entry.license, related = entry.related) afp_structure.save_entry(updated) // TODO what happens to the authors entry.name } override def list(): Submission_List = Submission_List(afp_structure.entries.sortBy(dates.get).reverse.map { entry => Overview(entry, dates(entry), entry, Status.Added) }) override def get( id: ID, topics: Map[ID, Topic], licenses: Map[ID, License] ): Option[Submission] = if (!afp_structure.entries.contains(id)) None else { val entry = afp_structure.load_entry(id, authors, all_topics, licenses, releases) val meta = Metadata(authors, List(entry)) Some(Submission(id, meta, "", Model.Build.Success, Some(Status.Added), "")) } override def submit(id: ID): Unit = () override def set_status(id: ID, status: Model.Status.Value): Unit = () override def abort_build(id: ID): Unit = () override def get_patch(id: ID): Option[Path] = None override def get_archive(id: ID): Option[Path] = None } /* Adapter to existing submission system */ class Adapter(submission_dir: Path, afp_structure: AFP_Structure) extends Handler { private val up: Path = submission_dir + Path.basic("up") private def up(id: ID): Path = up + Path.basic(id) private def down(id: ID): Path = submission_dir + Path.basic("down") + Path.basic(id) private def signal(id: ID, s: String): Unit = File.write(up(id) + Path.basic(s), s.toUpperCase) private def is_signal(id: ID, s: String): Boolean = (up(id) + Path.basic(s)).file.exists() private def read_build(id: ID): Build.Value = { val build = down(id) + Path.basic("result") if (!build.file.exists) Build.Pending else File.read(build).trim match { case "" => Build.Running case "NOT_FINISHED" => Build.Running case "FAILED" => if (is_signal(id, "kill")) Build.Aborted else Build.Failed case "SUCCESS" => Build.Success case s => isabelle.error("Unkown build status: " + quote(s)) } } private def status_file(id: ID): Path = up(id) + Path.basic("AFP_STATUS") private def read_status(id: ID): Option[Status.Value] = { val status = status_file(id) if (!status.file.exists()) None else File.read(status).trim match { case "SUBMITTED" => Some(Status.Submitted) case "PROCESSING" => Some(Status.Review) case "REJECTED" => Some(Status.Rejected) case "ADDED" => Some(Status.Added) case s => isabelle.error("Unknown status: " + quote(s)) } } private def info_file(id: ID): Path = up(id) + Path.basic("info.json") private def patch_file(id: ID): Path = up(id) + Path.basic("patch") private val archive_name: String = "archive" def make_partial_patch(base_dir: Path, src: Path, dst: Path): String = { val patch = Isabelle_System.make_patch(base_dir, src, dst, "--unidirectional-new-file") split_lines(patch).filterNot(_.startsWith("Only in")).mkString("\n") } def create(date: Date, meta: Metadata, message: String, archive: Bytes, ext: String): ID = { val id = ID(date) val dir = up(id) dir.file.mkdirs() val structure = AFP_Structure(dir) structure.save_authors(meta.authors.values.toList.sortBy(_.id)) meta.entries.foreach(structure.save_entry) val archive_file = dir + Path.basic(archive_name + ext) Bytes.write(archive_file, archive) val metadata_rel = File.relative_path(afp_structure.base_dir, afp_structure.metadata_dir).getOrElse( afp_structure.metadata_dir) val metadata_patch = make_partial_patch(afp_structure.base_dir, metadata_rel, structure.metadata_dir) File.write(patch_file(id), metadata_patch) val info = JSON.Format(JSON.Object( "comment" -> message, "entries" -> meta.entries.map(_.name), "notify" -> meta.entries.flatMap(_.notifies).map(_.address).distinct)) File.write(info_file(id), info) signal(id, "done") id } def list(): Submission_List = Submission_List( File.read_dir(up).flatMap(ID.unapply).reverse.flatMap { date => val id = ID(date) val day = date.rep.toLocalDate read_status(id).map(Overview(id, day, AFP_Structure(up(id)).entries_unchecked.head, _)) }) def get( id: ID, topics: Map[Topic.ID, Topic], licenses: Map[License.ID, License] ): Option[Submission] = ID.check(id).filter(up(_).file.exists).map { id => val structure = AFP_Structure(up(id)) val authors = structure.load_authors.map(author => author.id -> author).toMap val entries = structure.entries_unchecked.map( structure.load_entry(_, authors, topics, licenses, Map.empty)) val log_file = down(id) + Path.basic("isabelle.log") val log = if (log_file.file.exists) File.read(log_file) else "" JSON.parse(File.read(info_file(id))) match { case JSON.Object(m) if m.contains("comment") => val meta = Metadata(authors, entries) Submission(id, meta, m("comment").toString, read_build(id), read_status(id), log) case _ => isabelle.error("Could not read info") } } private def check(id: ID): Option[ID] = ID.check(id).filter(up(_).file.exists) def submit(id: ID): Unit = check(id).foreach(signal(_, "mail")) def set_status(id: ID, status: Status.Value): Unit = check(id).foreach { id => val content = status match { case Status.Submitted => "SUBMITTED" case Status.Review => "PROCESSING" case Status.Added => "ADDED" case Status.Rejected => "REJECTED" } File.write(status_file(id), content) } def abort_build(id: ID): Unit = check(id).foreach(signal(_, "kill")) def get_patch(id: ID): Option[Path] = check(id).map(patch_file) def get_archive(id: ID): Option[Path] = check(id).flatMap { id => val dir = up(id) File.read_dir(dir).find(_.startsWith(archive_name + ".")).map(dir + Path.basic(_)) } } } /* server */ object Mode extends Enumeration { val EDIT, SUBMISSION = Value } class Server( api: API, afp_structure: AFP_Structure, mode: Mode.Value, handler: Handler, devel: Boolean, verbose: Boolean, progress: Progress, port: Int = 0 ) extends Web_App.Server[Model.T](api, port, verbose, progress) { val repo = Mercurial.the_repository(afp_structure.base_dir) var authors: Map[Author.ID, Metadata.Author] = Map.empty var topics: List[Topic] = Nil var all_topics: Map[Topic.ID, Topic] = Map.empty var licenses: Map[License.ID, License] = Map.empty var releases: Map[Entry.Name, List[Release]] = Map.empty var entries: Set[Entry.Name] = Set.empty private def load(): Unit = synchronized { authors = afp_structure.load_authors.map(author => author.id -> author).toMap topics = afp_structure.load_topics.flatMap(_.all_topics) all_topics = topics.map(topic => topic.id -> topic).toMap licenses = afp_structure.load_licenses.map(license => license.id -> license).toMap releases = afp_structure.load_releases.groupBy(_.entry) entries = afp_structure.entries.toSet } load() /* endpoints */ val SUBMIT = Path.explode("submit") val SUBMISSION = Path.explode("submission") val SUBMISSIONS = Path.explode("submissions") val API_SUBMISSION = Path.explode("api/submission") val API_SUBMISSION_UPLOAD = Path.explode("api/submission/upload") val API_SUBMISSION_CREATE = Path.explode("api/submission/create") val API_SUBMISSION_STATUS = Path.explode("api/submission/status") val API_RESUBMIT = Path.explode("api/resubmit") val API_BUILD_ABORT = Path.explode("api/build/abort") val API_SUBMIT = Path.explode("api/submit") val API_SUBMISSION_AUTHORS_ADD = Path.explode("api/submission/authors/add") val API_SUBMISSION_AUTHORS_REMOVE = Path.explode("api/submission/authors/remove") val API_SUBMISSION_AFFILIATIONS_ADD = Path.explode("api/submission/affiliations/add") val API_SUBMISSION_AFFILIATIONS_REMOVE = Path.explode("api/submission/affiliations/remove") val API_SUBMISSION_ENTRIES_ADD = Path.explode("api/submission/entries/add") val API_SUBMISSION_ENTRIES_REMOVE = Path.explode("api/submission/entries/remove") val API_SUBMISSION_ENTRY_TOPICS_ADD = Path.explode("api/submission/entry/topics/add") val API_SUBMISSION_ENTRY_TOPICS_REMOVE = Path.explode("api/submission/entry/topics/remove") val API_SUBMISSION_ENTRY_AUTHORS_ADD = Path.explode("api/submission/entry/authors/add") val API_SUBMISSION_ENTRY_AUTHORS_REMOVE = Path.explode("api/submission/entry/authors/remove") val API_SUBMISSION_ENTRY_NOTIFIES_ADD = Path.explode("api/submission/entry/notifies/add") val API_SUBMISSION_ENTRY_NOTIFIES_REMOVE = Path.explode("api/submission/entry/notifies/remove") val API_SUBMISSION_ENTRY_RELATED_ADD = Path.explode("api/submission/entry/related/add") val API_SUBMISSION_ENTRY_RELATED_REMOVE = Path.explode("api/submission/entry/related/remove") val API_SUBMISSION_DOWNLOAD = Path.explode("api/download/patch") val API_SUBMISSION_DOWNLOAD_ZIP = Path.explode("api/download/archive.zip") val API_SUBMISSION_DOWNLOAD_TGZ = Path.explode("api/download/archive.tar.gz") val API_CSS = Path.explode("api/main.css") /* fields */ val ABSTRACT = "abstract" val ADDRESS = "address" val AFFILIATION = "affiliation" val ARCHIVE = "archive" val AUTHOR = "author" val DATE = "date" val ENTRY = "entry" val ID = "id" val INPUT = "input" val KIND = "kind" val LICENSE = "license" val MESSAGE = "message" val NAME = "name" val NOTIFY = "notify" val ORCID = "orcid" val RELATED = "related" val STATUS = "status" val TITLE = "title" val TOPIC = "topic" /* utils */ def keyed(id: String, value: String): String = "[" + id + "] " + value def author_string(author: Author): String = { val orcid = author.orcid.map(orcid => " (ORCID id: " + orcid.identifier + ")").getOrElse("") keyed(author.id, author.name) + orcid } def author_option(author: Author): XML.Elem = option(author.id, author_string(author)) def affil_id(affil: Affiliation): String = affil match { case Unaffiliated(_) => "" case Email(_, id, _) => id case Homepage(_, id, _) => id } def affil_address(affil: Affiliation): String = affil match { case Unaffiliated(_) => "" case Email(_, _, address) => address case Homepage(_, _, url) => url.toString } def affil_string(affil: Affiliation): String = affil match { case Unaffiliated(_) => "No email or homepage" case Email(_, id, address) => keyed(id, address) case Homepage(_, id, url) => keyed(id, url.toString) } def related_string(related: Reference): String = related match { case Metadata.DOI(identifier) => identifier case Metadata.Formatted(rep) => rep } def indexed[A, B](l: List[A], key: Params.Key, field: String, f: (A, Params.Key) => B) = l.zipWithIndex map { case (a, i) => f(a, List_Key(key, field, i)) } def fold[A](it: List[Params.Data], a: A)(f: (Params.Data, A) => Option[A]): Option[A] = it.foldLeft(Option(a)) { case (None, _) => None case (Some(a), param) => f(param, a) } def download_link(href: String, body: XML.Body): XML.Elem = class_("download")(link(href, body)) + ("target" -> "_blank") def frontend_link(path: Path, params: Properties.T, body: XML.Body): XML.Elem = link(api.app_url(path, params), body) + ("target" -> "_parent") def render_if(cond: Boolean, body: XML.Body): XML.Body = if (cond) body else Nil def render_if(cond: Boolean, elem: XML.Elem): XML.Body = if (cond) List(elem) else Nil def render_error(for_elem: String, validated: Val[_]): XML.Body = validated.err.map(error => break ::: List(css("color: red")(label(for_elem, error)))).getOrElse(Nil) /* view */ def render_create(model: Model.Create): XML.Body = { val updated_authors = model.updated_authors(authors) val authors_list = updated_authors.values.toList.sortBy(_.id) val (entry_authors, other_authors) = updated_authors.values.toList.sortBy(_.id).partition(author => model.used_authors.contains(author.id)) val email_authors = authors_list.filter(_.emails.nonEmpty) def render_topic(topic: Topic, key: Params.Key): XML.Elem = par( hidden(Nest_Key(key, ID), topic.id) :: text(topic.id) ::: action_button(api.api_url(API_SUBMISSION_ENTRY_TOPICS_REMOVE), "x", key) :: Nil) def render_affil(affil: Affiliation, key: Params.Key): XML.Elem = { val author = updated_authors(affil.author) val affils = author.emails ::: author.homepages ::: Unaffiliated(author.id) :: Nil par( hidden(Nest_Key(key, ID), affil.author) :: text(author_string(updated_authors(affil.author))) ::: selection(Nest_Key(key, AFFILIATION), Some(affil_id(affil)), affils.map(affil => option(affil_id(affil), affil_string(affil)))) :: action_button(api.api_url(API_SUBMISSION_ENTRY_AUTHORS_REMOVE), "x", key) :: Nil) } def render_notify(email: Email, key: Params.Key): XML.Elem = { val author = updated_authors(email.author) par( hidden(Nest_Key(key, ID), email.author) :: text(author_string(updated_authors(email.author))) ::: selection( Nest_Key(key, AFFILIATION), Some(affil_id(email)), author.emails.map(affil => option(affil_id(affil), affil_string(affil)))) :: action_button(api.api_url(API_SUBMISSION_ENTRY_NOTIFIES_REMOVE), "x", key) :: Nil) } def render_related(related: Reference, key: Params.Key): XML.Elem = par( hidden(Nest_Key(key, KIND), Model.Related.get(related).toString) :: hidden(Nest_Key(key, INPUT), related_string(related)) :: text(related_string(related)) ::: action_button(api.api_url(API_SUBMISSION_ENTRY_RELATED_REMOVE), "x", key) :: Nil) def render_entry(entry: Model.Create_Entry, key: Params.Key): XML.Elem = fieldset(legend("Entry") :: par( fieldlabel(Nest_Key(key, TITLE), "Title of article") :: textfield(Nest_Key(key, TITLE), "Example Submission", entry.title.v) :: render_error(Nest_Key(key, TITLE), entry.title)) :: par( fieldlabel(Nest_Key(key, NAME), "Short name") :: textfield(Nest_Key(key, NAME), "Example_Submission", entry.name.v) :: explanation(Nest_Key(key, NAME), "Name of the folder where your ROOT and theory files reside.") :: render_error(Nest_Key(key, NAME), entry.name)) :: fieldset(legend("Topics") :: indexed(entry.topics.v, key, TOPIC, render_topic) ::: selection(Nest_Key(key, TOPIC), entry.topic_input.map(_.id), topics.map(topic => option(topic.id, topic.id))) :: action_button(api.api_url(API_SUBMISSION_ENTRY_TOPICS_ADD), "add", key) :: render_error("", entry.topics)) :: par(List( fieldlabel(Nest_Key(key, LICENSE), "License"), radio(Nest_Key(key, LICENSE), entry.license.id, licenses.values.toList.map(license => license.id -> license.name)), explanation(Nest_Key(key, LICENSE), "Note: For LGPL submissions, please include the header \"License: LGPL\" in each file"))) :: par( fieldlabel(Nest_Key(key, ABSTRACT), "Abstract") :: placeholder("HTML and MathJax, no LaTeX")( textarea(Nest_Key(key, ABSTRACT), entry.`abstract`.v) + ("rows" -> "8") + ("cols" -> "70")) :: explanation(Nest_Key(key, ABSTRACT), "Note: You can use HTML or MathJax (not LaTeX!) to format your abstract.") :: render_error(Nest_Key(key, ABSTRACT), entry.`abstract`)) :: fieldset(legend("Authors") :: indexed(entry.affils.v, key, AUTHOR, render_affil) ::: selection(Nest_Key(key, AUTHOR), entry.author_input.map(_.id), authors_list.map(author => option(author.id, author_string(author)))) :: action_button(api.api_url(API_SUBMISSION_ENTRY_AUTHORS_ADD), "add", key) :: explanation(Nest_Key(key, AUTHOR), "Add an author from the list. Register new authors first below.") :: render_error(Nest_Key(key, AUTHOR), entry.affils)) :: fieldset(legend("Contact") :: indexed(entry.notifies.v, key, NOTIFY, render_notify) ::: selection(Nest_Key(key, NOTIFY), entry.notify_input.map(_.id), optgroup("Suggested", email_authors.filter(author => entry.used_authors.contains(author.id)).map(author_option)) :: email_authors.filter(author => !entry.used_authors.contains(author.id)).map(author_option)) :: action_button(api.api_url(API_SUBMISSION_ENTRY_NOTIFIES_ADD), "add", key) :: explanation(Nest_Key(key, NOTIFY), "These addresses serve two purposes: " + "1. They are used to send you updates about the state of your submission. " + "2. They are the maintainers of the entry once it is accepted. " + "Typically this will be one or more of the authors.") :: render_error("", entry.notifies)) :: fieldset(legend("Related Publications") :: indexed(entry.related, key, RELATED, render_related) ::: selection(Nest_Key(Nest_Key(key, RELATED), KIND), entry.related_kind.map(_.toString), Model.Related.values.toList.map(v => option(v.toString, v.toString))) :: textfield(Nest_Key(Nest_Key(key, RELATED), INPUT), "10.1109/5.771073", entry.related_input.v) :: action_button(api.api_url(API_SUBMISSION_ENTRY_RELATED_ADD), "add", key) :: explanation(Nest_Key(Nest_Key(key, RELATED), INPUT), "Publications related to the entry, as DOIs (10.1109/5.771073) or plaintext (HTML)." + "Typically a publication by the authors describing the entry," + " background literature (articles, books) or web resources. ") :: render_error(Nest_Key(Nest_Key(key, RELATED), INPUT), entry.related_input)) :: render_if(mode == Mode.SUBMISSION, action_button(api.api_url(API_SUBMISSION_ENTRIES_REMOVE), "remove entry", key))) def render_new_author(author: Author, key: Params.Key): XML.Elem = par( hidden(Nest_Key(key, ID), author.id) :: hidden(Nest_Key(key, NAME), author.name) :: hidden(Nest_Key(key, ORCID), author.orcid.map(_.identifier).getOrElse("")) :: text(author_string(author)) ::: render_if(!model.used_authors.contains(author.id), action_button(api.api_url(API_SUBMISSION_AUTHORS_REMOVE), "x", key))) def render_new_affil(affil: Affiliation, key: Params.Key): XML.Elem = par( hidden(Nest_Key(key, AUTHOR), affil.author) :: hidden(Nest_Key(key, ID), affil_id(affil)) :: hidden(Nest_Key(key, AFFILIATION), affil_address(affil)) :: text(author_string(updated_authors(affil.author)) + ": " + affil_string(affil)) ::: render_if(!model.used_affils.contains(affil), action_button(api.api_url(API_SUBMISSION_AFFILIATIONS_REMOVE), "x", key))) val (upload, preview) = mode match { case Mode.EDIT => ("Save", "preview and save >") case Mode.SUBMISSION => ("Upload", "preview and upload >") } List(submit_form(api.api_url(API_SUBMISSION), indexed(model.entries.v, Params.empty, ENTRY, render_entry) ::: render_error("", model.entries) ::: render_if(mode == Mode.SUBMISSION, par(List( explanation("", "You can submit multiple entries at once. " + "Put the corresponding folders in the archive " + "and use the button below to add more input fields for metadata. "), api_button(api.api_url(API_SUBMISSION_ENTRIES_ADD), "additional entry")))) ::: break ::: fieldset(legend("New Authors") :: explanation("", "If you are new to the AFP, add yourself here.") :: indexed(model.new_authors.v, Params.empty, AUTHOR, render_new_author) ::: fieldlabel(Nest_Key(AUTHOR, NAME), "Name") :: textfield(Nest_Key(AUTHOR, NAME), "Gerwin Klein", model.new_author_input) :: fieldlabel(Nest_Key(AUTHOR, ORCID), "ORCID id (optional)") :: textfield(Nest_Key(AUTHOR, ORCID), "0000-0002-1825-0097", model.new_author_orcid) :: api_button(api.api_url(API_SUBMISSION_AUTHORS_ADD), "add") :: render_error("", model.new_authors)) :: fieldset(legend("New email or homepage") :: explanation("", "Add new email or homepages here. " + "If you would like to update an existing, " + "submit with the old one and write to the editors.") :: indexed(model.new_affils.v, Params.empty, AFFILIATION, render_new_affil) ::: fieldlabel(AFFILIATION, "Author") :: selection(AFFILIATION, model.new_affils_author.map(_.id), optgroup("Entry authors", entry_authors.map(author_option)) :: other_authors.map(author_option)) :: fieldlabel(Nest_Key(AFFILIATION, ADDRESS), "Email or Homepage") :: textfield(Nest_Key(AFFILIATION, ADDRESS), "https://proofcraft.org", model.new_affils_input) :: api_button(api.api_url(API_SUBMISSION_AFFILIATIONS_ADD), "add") :: render_error("", model.new_affils)) :: break ::: fieldset(List(legend(upload), api_button(api.api_url(API_SUBMISSION_UPLOAD), preview))) :: Nil)) } def render_metadata(meta: Model.Metadata): XML.Body = { def render_topic(topic: Topic, key: Params.Key): XML.Elem = item(hidden(Nest_Key(key, ID), topic.id) :: text(topic.id)) def render_affil(affil: Affiliation, key: Params.Key): XML.Elem = item( hidden(Nest_Key(key, ID), affil.author) :: hidden(Nest_Key(key, AFFILIATION), affil_id(affil)) :: text(author_string(meta.authors(affil.author)) + ", " + affil_string(affil))) def render_related(related: Reference, key: Params.Key): XML.Elem = item( hidden(Nest_Key(key, KIND), Model.Related.get(related).toString) :: hidden(Nest_Key(key, INPUT), related_string(related)) :: unescape(related_string(related))) def render_entry(entry: Entry, key: Params.Key): XML.Elem = fieldset(List( legend("Entry"), par( fieldlabel(Nest_Key(key, TITLE), "Title") :: hidden(Nest_Key(key, TITLE), entry.title) :: text(entry.title)), par( fieldlabel(Nest_Key(key, NAME), "Short Name") :: hidden(Nest_Key(key, NAME), entry.name) :: text(entry.name)), par( fieldlabel(Nest_Key(key, DATE), "Date") :: hidden(Nest_Key(key, DATE), entry.date.toString) :: text(entry.date.toString)), par(List( fieldlabel("", "Topics"), list(indexed(entry.topics, key, TOPIC, render_topic)))), par( fieldlabel(Nest_Key(key, LICENSE), "License") :: hidden(Nest_Key(key, LICENSE), entry.license.id) :: text(entry.license.name)), par(List( fieldlabel(Nest_Key(key, ABSTRACT), "Abstract"), hidden(Nest_Key(key, ABSTRACT), entry.`abstract`), class_("mathjax_process")(span(unescape(entry.`abstract`))))), par(List( fieldlabel("", "Authors"), list(indexed(entry.authors, key, AUTHOR, render_affil)))), par(List( fieldlabel("", "Contact"), list(indexed(entry.notifies, key, NOTIFY, render_affil)))), par(List( fieldlabel("", "Related Publications"), list(indexed(entry.related, key, RELATED, render_related)))))) def render_new_author(author: Author, key: Params.Key): XML.Elem = par(List( hidden(Nest_Key(key, ID), author.id), hidden(Nest_Key(key, NAME), author.name), hidden(Nest_Key(key, ORCID), author.orcid.map(_.identifier).getOrElse("")))) def render_new_affil(affil: Affiliation, key: Params.Key): XML.Elem = par(List( hidden(Nest_Key(key, AUTHOR), affil.author), hidden(Nest_Key(key, ID), affil_id(affil)), hidden(Nest_Key(key, AFFILIATION), affil_address(affil)))) indexed(meta.entries, Params.empty, ENTRY, render_entry) ::: indexed(meta.new_authors(authors).toList, Params.empty, AUTHOR, render_new_author) ::: indexed(meta.new_affils(authors).toList, Params.empty, AFFILIATION, render_new_affil) } def render_submission(submission: Model.Submission): XML.Body = { def status_text(status: Option[Model.Status.Value]): String = status.map { case Model.Status.Submitted => "AFP editors notified." case Model.Status.Review => "Review in progress." case Model.Status.Added => "Added to the AFP." case Model.Status.Rejected => "Submission rejected." } getOrElse "Draft saved. Check the logs for errors and warnings, " + "and submit to editors once successfully built." val archive_url = if (handler.get_archive(submission.id).exists(_.get_ext == "zip")) API_SUBMISSION_DOWNLOAD_ZIP else API_SUBMISSION_DOWNLOAD_TGZ val resubmit = mode match { case Mode.EDIT => "Update" case Mode.SUBMISSION => "Resubmit" } List(submit_form(api.api_url(SUBMISSION, List(ID -> submission.id)), render_if(mode == Mode.SUBMISSION, download_link(api.api_url(archive_url, List(ID -> submission.id)), text("archive")) :: download_link(api.api_url(API_SUBMISSION_DOWNLOAD, List(ID -> submission.id)), text("metadata patch")) :: text(" (apply with: 'patch -p0 < FILE')")) ::: render_if(mode == Mode.SUBMISSION, par( hidden(MESSAGE, submission.message) :: text("Comment: " + submission.message))) ::: section("Metadata") :: render_metadata(submission.meta) ::: section("Status") :: span(text(status_text(submission.status))) :: render_if(submission.build != Model.Build.Running, action_button(api.api_url(API_RESUBMIT), resubmit, submission.id)) ::: render_if(submission.build == Model.Build.Running, action_button(api.api_url(API_BUILD_ABORT), "Abort build", submission.id)) ::: render_if(submission.build == Model.Build.Success && submission.status.isEmpty, action_button(api.api_url(API_SUBMIT), "Send submission to AFP editors", submission.id)) ::: render_if(mode == Mode.SUBMISSION, fieldset(legend("Build") :: bold(text(submission.build.toString)) :: par(text("Isabelle log:") ::: source(submission.log) :: Nil) :: Nil)))) } def render_upload(upload: Model.Upload): XML.Body = { val submit = mode match { case Mode.EDIT => "save" case Mode.SUBMISSION => "submit" } List(submit_form(api.api_url(API_SUBMISSION), List( fieldset(legend("Overview") :: render_metadata(upload.meta)), fieldset(legend("Submit") :: api_button(api.api_url(API_SUBMISSION), "< edit metadata") :: render_if(mode == Mode.SUBMISSION, List( par(List( fieldlabel(MESSAGE, "Message for the editors (optional)"), textfield(MESSAGE, "", upload.message), explanation( MESSAGE, "Note: Anything special or noteworthy about your submission can be covered here. " + "It will not become part of your entry. " + "You're also welcome to leave suggestions for our AFP submission service here. ") )), par(List( fieldlabel(ARCHIVE, "Archive file (.tar.gz or .zip)"), browse(ARCHIVE, List(".zip", ".tar.gz")), explanation(ARCHIVE, "Note: Your zip or tar file must contain exactly one folder for each entry with your theories, ROOT, etc. " + "The folder name must be the short/folder name used in the submission form. " + "Hidden files and folders (e.g., __MACOSX) are not allowed."))))) ::: api_button(api.api_url(API_SUBMISSION_CREATE), submit) :: render_error(ARCHIVE, Val.err((), upload.error)))))) } def render_submission_list(submission_list: Model.Submission_List): XML.Body = { def render_overview(overview: Model.Overview, key: Params.Key): XML.Elem = item( hidden(Nest_Key(key, ID), overview.id) :: hidden(Nest_Key(key, DATE), overview.date.toString) :: hidden(Nest_Key(key, NAME), overview.name) :: span(text(overview.date.toString)) :: span(List(frontend_link(SUBMISSION, List(ID -> overview.id), text(overview.name)))) :: render_if(mode == Mode.SUBMISSION, class_("right")(span(List( selection(Nest_Key(key, STATUS), Some(overview.status.toString), Model.Status.values.toList.map(v => option(v.toString, v.toString))), action_button(api.api_url(API_SUBMISSION_STATUS), "update", key)))))) def list1(ls: List[XML.Elem]): XML.Elem = if (ls.isEmpty) par(Nil) else list(ls) val ls = indexed(submission_list.submissions, Params.empty, ENTRY, (o, k) => (o, k)) val finished = ls.filter(t => Set(Model.Status.Added, Model.Status.Rejected).contains(t._1.status)) List(submit_form(api.api_url(API_SUBMISSION_STATUS), render_if(mode == Mode.SUBMISSION, text("Open") ::: list1(ls.filter(_._1.status == Model.Status.Submitted).map(render_overview)) :: text("In Progress") ::: list1(ls.filter(_._1.status == Model.Status.Review).map(render_overview)) :: text("Finished")) ::: list1(finished.map(render_overview)) :: Nil)) } def render_created(created: Model.Created): XML.Body = { val status = mode match { case Mode.SUBMISSION => "View your submission status: " case Mode.EDIT => "View the entry at: " } List(div( span(text("Entry successfully saved. " + status)) :: break ::: frontend_link(SUBMISSION, List(ID -> created.id), text(api.app_url(SUBMISSION, List(ID -> created.id)))) :: break ::: render_if(mode == Mode.SUBMISSION, span(text("(keep that url!)."))))) } def render_invalid: XML.Body = text("Invalid request") def render(model: Model.T): XML.Body = model match { case create: Model.Create => render_create(create) case upload: Model.Upload => render_upload(upload) case submission: Model.Submission => render_submission(submission) case submissions: Model.Submission_List => render_submission_list(submissions) case created: Model.Created => render_created(created) case Model.Invalid => render_invalid } /* validation */ def validate_topic(id: String, selected: List[Topic]): Val_Opt[Topic] = if (id.isEmpty) Val_Opt.user_err("Select topic first") else { topics.find(_.id == id) match { case Some(topic) => if (selected.contains(topic)) Val_Opt.user_err("Topic already selected") else Val_Opt.ok(topic) case _ => Val_Opt.error } } def validate_new_author( id: String, name: String, orcid_str: String, authors: Map[Author.ID, Author] ): Val_Opt[Author] = { val Id = """[a-zA-Z][a-zA-Z0-9]*""".r id match { case Id() if !authors.contains(id) => if (name.isEmpty || name.trim != name) Val_Opt.user_err("Name must not be empty") else if (orcid_str.isEmpty) Val_Opt.ok(Author(id, name)) else Exn.capture(Orcid(orcid_str)) match { case Exn.Res(orcid) => if (authors.values.exists(_.orcid.exists(_ == orcid))) Val_Opt.user_err("Author with that orcid already exists") else Val_Opt.ok(Author(id, name, orcid = Some(orcid))) case _ => Val_Opt.user_err("Invalid orcid") } case _ => Val_Opt.error } } def validate_new_affil(id: String, address: String, author: Author): Val_Opt[Affiliation] = { if (address.isBlank) Val_Opt.user_err("Address must not be empty") else if (address.contains("@")) { val Id = (author.id + """_email\d*""").r id match { case Id() if !author.emails.map(_.id).contains(id) => val Address = """([^@\s]+)@([^@\s]+)""".r address match { case Address(user, host) => Val_Opt.ok(Email(author.id, id, user, host)) case _ => Val_Opt.user_err("Invalid email address") } case _ => Val_Opt.error } } else { val Id = (author.id + """_homepage\d*""").r id match { case Id() if !author.homepages.map(_.id).contains(id) => - Exn.capture(new URL(address)) match { - case Exn.Res(url) => Val_Opt.ok(Homepage(author.id, id, url)) - case _ => Val_Opt.user_err("Invalid url") - } + if (Url.is_wellformed(address)) Val_Opt.ok(Homepage(author.id, id, Url(address))) + else Val_Opt.user_err("Invalid url") case _ => Val_Opt.error } } } def validate_related( kind: Model.Related.Value, related: String, references: List[Reference] ): Val_Opt[Reference] = if (related.isBlank) Val_Opt.user_err("Reference must not be empty") else { kind match { case Model.Related.DOI => Exn.capture(DOI(related)) match { case Exn.Res(doi) => if (references.contains(doi)) Val_Opt.user_err("Already present") else Val_Opt.ok(doi) case _ => Val_Opt.user_err("Invalid DOI format") } case Model.Related.Plaintext => val formatted = Formatted(related) if (references.contains(formatted)) Val_Opt.user_err("Already present") else Val_Opt.ok(formatted) } } /* param parsing */ def parse_create(params: Params.Data): Option[Model.Create] = { def parse_topic(topic: Params.Data, topics: List[Topic]): Option[Topic] = validate_topic(topic.get(ID).value, topics).opt def parse_email(email: Params.Data, authors: Map[Author.ID, Author]): Option[Email] = authors.get(email.get(ID).value).flatMap( _.emails.find(_.id == email.get(AFFILIATION).value)) def parse_affil(affil: Params.Data, authors: Map[Author.ID, Author]): Option[Affiliation] = authors.get(affil.get(ID).value).flatMap { author => val id = affil.get(AFFILIATION).value if (id.isEmpty) Some(Unaffiliated(author.id)) else (author.emails ++ author.homepages).collectFirst { case e: Email if e.id == id => e case h: Homepage if h.id == id => h } } def parse_related(related: Params.Data, references: List[Reference]): Option[Reference] = Model.Related.from_string(related.get(KIND).value).flatMap( validate_related(_, related.get(INPUT).value, references).opt) def parse_new_author(author: Params.Data, authors: Map[Author.ID, Author]): Option[Author] = validate_new_author( author.get(ID).value, author.get(NAME).value, author.get(ORCID).value, authors).opt def parse_new_affil(affil: Params.Data, authors: Map[Author.ID, Author]): Option[Affiliation] = authors.get(affil.get(AUTHOR).value).flatMap(author => validate_new_affil(affil.get(ID).value, affil.get(AFFILIATION).value, author).opt) def parse_entry(entry: Params.Data, authors: Map[Author.ID, Author]): Option[Model.Create_Entry] = for { topics <- fold(entry.list(TOPIC), List.empty[Topic]) { case (topic, topics) => parse_topic(topic, topics).map(topics :+ _) } affils <- fold(entry.list(AUTHOR), List.empty[Affiliation]) { case (affil, affils) => parse_affil(affil, authors).map(affils :+ _) } notifies <- fold(entry.list(NOTIFY), List.empty[Email]) { case (email, emails) => parse_email(email, authors).map(emails :+ _) } related <- fold(entry.list(RELATED), List.empty[Reference]) { case (related, references) => parse_related(related, references).map(references :+ _) } license <- licenses.get(entry.get(LICENSE).value) } yield Model.Create_Entry( name = Val.ok(entry.get(NAME).value), title = Val.ok(entry.get(TITLE).value), topics = Val.ok(topics), topic_input = parse_topic(entry.get(TOPIC), Nil), license = license, `abstract` = Val.ok(entry.get(ABSTRACT).value), author_input = authors.get(entry.get(AUTHOR).value), notify_input = authors.get(entry.get(NOTIFY).value), affils = Val.ok(affils), notifies = Val.ok(notifies), related = related, related_kind = Model.Related.from_string(entry.get(RELATED).get(KIND).value), related_input = Val.ok(entry.get(RELATED).get(INPUT).value)) for { (new_author_ids, all_authors) <- fold(params.list(AUTHOR), (List.empty[Author.ID], authors)) { case (author, (new_authors, authors)) => parse_new_author(author, authors).map(author => (new_authors :+ author.id, authors.updated(author.id, author))) } (new_affils, all_authors) <- fold(params.list(AFFILIATION), (List.empty[Affiliation], all_authors)) { case (affil, (new_affils, authors)) => parse_new_affil(affil, authors).map { affil => val author = authors(affil.author) (new_affils :+ affil, affil match { case _: Unaffiliated => authors case e: Email => authors.updated(author.id, author.copy(emails = author.emails :+ e)) case h: Homepage => authors.updated(author.id, author.copy(homepages = author.homepages :+ h)) }) } } new_authors = new_author_ids.map(all_authors) entries <- fold(params.list(ENTRY), List.empty[Model.Create_Entry]) { case (entry, entries) => parse_entry(entry, all_authors).map(entries :+ _) } } yield Model.Create( entries = Val.ok(entries), new_authors = Val.ok(new_authors), new_author_input = params.get(AUTHOR).get(NAME).value, new_author_orcid = params.get(AUTHOR).get(ORCID).value, new_affils = Val.ok(new_affils), new_affils_author = all_authors.get(params.get(AFFILIATION).value), new_affils_input = params.get(AFFILIATION).get(ADDRESS).value) } def parse_submission_list(params: Params.Data): Option[Model.Submission_List] = { def parse_overview(entry: Params.Data): Option[Model.Overview] = for { date <- Exn.capture(LocalDate.parse(entry.get(DATE).value)) match { case Exn.Res(date) => Some(date) case Exn.Exn(_) => None } status <- Model.Status.from_string(entry.get(STATUS).value) } yield Model.Overview(entry.get(ID).value, date, entry.get(NAME).value, status) val submissions = fold(params.list(ENTRY), List.empty[Model.Overview]) { case (entry, overviews) => parse_overview(entry).map(overviews :+ _) } submissions.map(Model.Submission_List.apply) } /* control */ def add_entry(params: Params.Data): Option[Model.T] = parse_create(params).map { model => model.copy(entries = Val.ok(model.entries.v :+ Model.Create_Entry(license = licenses.head._2))) } def remove_entry(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) num_entry <- List_Key.num(ENTRY, params.get(Web_App.ACTION).value) } yield model.copy(entries = Val.ok(Utils.remove_at(num_entry, model.entries.v))) def add_author(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) num_entry <- List_Key.num(ENTRY, params.get(Web_App.ACTION).value) entry <- model.entries.v.unapply(num_entry) } yield entry.author_input match { case None => model.update_entry(num_entry, entry.copy( affils = entry.affils.with_err("Select author first"))) case Some(author) => val default_affil = author.emails.headOption.orElse( author.homepages.headOption).getOrElse(Unaffiliated(author.id)) model.update_entry(num_entry, entry.copy( author_input = None, affils = Val.ok(entry.affils.v :+ default_affil))) } def remove_author(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) (action, num_affil) <- List_Key.split(AUTHOR, params.get(Web_App.ACTION).value) num_entry <- List_Key.num(ENTRY, action) entry <- model.entries.v.unapply(num_entry) } yield model.update_entry(num_entry, entry.copy(affils = Val.ok(Utils.remove_at(num_affil, entry.affils.v)))) def add_notify(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) num_entry <- List_Key.num(ENTRY, params.get(Web_App.ACTION).value) entry <- model.entries.v.unapply(num_entry) entry1 <- entry.notify_input match { case Some(author) => author.emails.headOption.map(email => entry.copy( notify_input = None, notifies = Val.ok(entry.notifies.v :+ email))) case None => Some(entry.copy(notifies = entry.notifies.with_err("Select author first"))) } } yield model.update_entry(num_entry, entry1) def remove_notify(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) (action, num_notify) <- List_Key.split(NOTIFY, params.get(Web_App.ACTION).value) num_entry <- List_Key.num(ENTRY, action) entry <- model.entries.v.unapply(num_entry) } yield model.update_entry(num_entry, entry.copy(notifies = Val.ok(Utils.remove_at(num_notify, entry.notifies.v)))) def add_topic(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) num_entry <- List_Key.num(ENTRY, params.get(Web_App.ACTION).value) entry <- model.entries.v.unapply(num_entry) entry_params <- params.list(ENTRY).unapply(num_entry) } yield { val topic = validate_topic(entry_params.get(TOPIC).value, entry.topics.v) val topic_input = if (topic.is_empty) entry.topic_input else None model.update_entry(num_entry, entry.copy( topic_input = topic_input, topics = Val.ok(entry.topics.v ++ topic.opt).perhaps_err(topic))) } def remove_topic(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) (action, num_topic) <- List_Key.split(TOPIC, params.get(Web_App.ACTION).value) num_entry <- List_Key.num(ENTRY, action) entry <- model.entries.v.unapply(num_entry) } yield { val entry1 = entry.copy(topics = Val.ok(Utils.remove_at(num_topic, entry.topics.v))) model.copy(entries = Val.ok(model.entries.v.updated(num_entry, entry1))) } def add_related(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) num_entry <- List_Key.num(ENTRY, params.get(Web_App.ACTION).value) entry <- model.entries.v.unapply(num_entry) } yield { val entry1 = entry.related_kind match { case None => entry.copy(related_input = entry.related_input.with_err("Select reference kind first")) case Some(kind) => val reference = validate_related(kind, entry.related_input.v, entry.related) val related_input = if (reference.is_empty) entry.related_input.v else "" entry.copy( related = entry.related ++ reference.opt, related_input = Val.ok(related_input).perhaps_err(reference)) } model.update_entry(num_entry, entry1) } def remove_related(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) (action, num_related) <- List_Key.split(RELATED, params.get(Web_App.ACTION).value) num_entry <- List_Key.num(ENTRY, action) entry <- model.entries.v.unapply(num_entry) } yield { val entry1 = entry.copy(related = Utils.remove_at(num_related, entry.related)) model.copy(entries = Val.ok(model.entries.v.updated(num_entry, entry1))) } def add_new_author(params: Params.Data): Option[Model.T] = parse_create(params).map { model => val name = model.new_author_input.trim if (name.isEmpty) model.copy(new_authors = model.new_authors.with_err("Name must not be empty")) else { def as_ascii(str: String) = { var res: String = str List("ö" -> "oe", "ü" -> "ue", "ä" -> "ae", "ß" -> "ss").foreach { case (c, rep) => res = res.replace(c, rep) } Normalizer.normalize(res, Normalizer.Form.NFD).replaceAll("[^\\x00-\\x7F]", "") } def make_author_id(name: String): String = { val normalized = as_ascii(name) val Suffix = """^.*?([a-zA-Z]*)$""".r val suffix = normalized match { case Suffix(suffix) => suffix case _ => "" } val Prefix = """^([a-zA-Z]*).*$""".r val prefix = normalized.stripSuffix(suffix) match { case Prefix(prefix) => prefix case _ => "" } val updated_authors = model.updated_authors(authors) var ident = suffix.toLowerCase for { c <- prefix.toLowerCase if updated_authors.contains(ident) } ident += c.toString Utils.make_unique(ident, updated_authors.keySet) } val id = make_author_id(name) val author = validate_new_author(id, model.new_author_input, model.new_author_orcid, model.updated_authors(authors)) val new_author_input = if (author.is_empty) model.new_author_input else "" val new_author_orcid = if (author.is_empty) model.new_author_orcid else "" model.copy( new_author_input = new_author_input, new_author_orcid = new_author_orcid, new_authors = Val.ok(model.new_authors.v ++ author.opt).perhaps_err(author)) } } def remove_new_author(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) num_author <- List_Key.num(AUTHOR, params.get(Web_App.ACTION).value) author <- model.new_authors.v.unapply(num_author) if !model.used_authors.contains(author.id) } yield model.copy(new_authors = Val.ok(Utils.remove_at(num_author, model.new_authors.v))) def add_new_affil(params: Params.Data): Option[Model.T] = parse_create(params).map { model => model.new_affils_author match { case Some(author) => val address = model.new_affils_input.trim if (address.isEmpty) model.copy(new_affils = model.new_affils.with_err("Must not be empty")) else { val id = if (address.contains("@")) Utils.make_unique(author.id + "_email", author.emails.map(_.id).toSet) else Utils.make_unique(author.id + "_homepage", author.homepages.map(_.id).toSet) val affil = validate_new_affil(id, address, author) val new_affils_input = if (affil.is_empty) model.new_affils_input else "" model.copy( new_affils_input = new_affils_input, new_affils = Val.ok(model.new_affils.v ++ affil.opt).perhaps_err(affil)) } case None => model.copy(new_affils = model.new_affils.with_err("Select author first")) } } def remove_new_affil(params: Params.Data): Option[Model.T] = for { model <- parse_create(params) num_affil <- List_Key.num(AFFILIATION, params.get(Web_App.ACTION).value) affil <- model.new_affils.v.unapply(num_affil) if !model.used_affils.contains(affil) } yield model.copy(new_affils = Val.ok(Utils.remove_at(num_affil, model.new_affils.v))) def upload(params: Params.Data): Option[Model.T] = parse_create(params).map { create => var ok = true def validate[A](validator: A => Val[A], value: A): Val[A] = { val res = validator(value) if (res.err.nonEmpty) ok = false res } def title(title: String): Val[String] = if (title.isBlank) Val.err(title, "Title must not be blank") else if (title.trim != title) Val.err(title, "Title must not contain extra spaces") else Val.ok(title) def name(name: String): Val[String] = if (name.isBlank) Val.err(name, "Name must not be blank") else if (!"[a-zA-Z0-9_-]+".r.matches(name)) Val.err(name, "Invalid character in name") else mode match { case Mode.EDIT => if (Server.this.entries.contains(name)) Val.ok(name) else Val.err(name, "Entry does not exist") case Mode.SUBMISSION => if (Server.this.entries.contains(name)) Val.err(name, "Entry already exists") else Val.ok(name) } def entries(entries: List[Model.Create_Entry]): Val[List[Model.Create_Entry]] = if (entries.isEmpty) Val.err(entries, "Must contain at least one entry") else if (Library.duplicates(entries.map(_.name)).nonEmpty) Val.err(entries, "Entries must have different names") else Val.ok(entries) def new_authors(authors: List[Author]): Val[List[Author]] = if (!authors.forall(author => create.used_authors.contains(author.id))) Val.err(authors, "Unused authors") else Val.ok(authors) def new_affils(affils: List[Affiliation]): Val[List[Affiliation]] = if (!affils.forall(affil => create.used_affils.contains(affil))) Val.err(affils, "Unused affils") else Val.ok(affils) def entry_authors(authors: List[Affiliation]): Val[List[Affiliation]] = if (authors.isEmpty) Val.err(authors, "Must contain at least one author") else if (!Utils.is_distinct(authors)) Val.err(authors, "Duplicate affiliations") else Val.ok(authors) def notifies(notifies: List[Email]): Val[List[Email]] = if (notifies.isEmpty) Val.err(notifies, "Must contain at least one maintainer") else if (!Utils.is_distinct(notifies)) Val.err(notifies, "Duplicate emails") else Val.ok(notifies) def topics(topics: List[Topic]): Val[List[Topic]] = if (topics.isEmpty) Val.err(topics, "Must contain at least one topic") else Val.ok(topics) def `abstract`(txt: String): Val[String] = if (txt.isBlank) Val.err(txt, "Entry must contain an abstract") else if (List("\\cite", "\\emph", "\\texttt").exists(txt.contains(_))) Val.err(txt, "LaTeX not allowed, use MathJax for math symbols") else Val.ok(txt) val validated = create.copy( entries = validate( entries, create.entries.v.map(entry => entry.copy( name = validate(name, entry.name.v), title = validate(title, entry.title.v), affils = validate(entry_authors, entry.affils.v), notifies = validate(notifies, entry.notifies.v), topics = validate(topics, entry.topics.v), `abstract` = validate(`abstract`, entry.`abstract`.v) ))), new_authors = validate(new_authors, create.new_authors.v), new_affils = validate(new_affils, create.new_affils.v)) if (ok) { val (updated_authors, entries) = create.create(authors) Model.Upload(Model.Metadata(updated_authors, entries), params.get(MESSAGE).value, "") } else validated } def create(params: Params.Data): Option[Model.T] = { upload(params) match { case Some(upload: Model.Upload) => mode match { case Mode.EDIT => handler.create(Date.now(), upload.meta, upload.message, Bytes.empty, "") Some(Model.Created(upload.meta.entries.head.name)) case Mode.SUBMISSION => val archive = Bytes.decode_base64(params.get(ARCHIVE).get(FILE).value) val file_name = params.get(ARCHIVE).value if ((archive.is_empty || file_name.isEmpty)) { Some(upload.copy(error = "Select a file")) } else if (!file_name.endsWith(".zip") && !file_name.endsWith(".tar.gz")) { Some(upload.copy(error = "Only .zip and .tar.gz archives allowed")) } else { val ext = if (file_name.endsWith(".zip")) ".zip" else ".tar.gz" val id = handler.create(Date.now(), upload.meta, upload.message, archive, ext) Some(Model.Created(id)) } } case _ => None } } def empty_submission: Option[Model.T] = Some(Model.Create(entries = Val.ok(List(Model.Create_Entry(license = licenses.head._2))))) def get_submission(props: Properties.T): Option[Model.Submission] = Properties.get(props, ID).flatMap(handler.get(_, all_topics, licenses)) def resubmit(params: Params.Data): Option[Model.T] = handler.get(params.get(ACTION).value, all_topics, licenses).map(submission => Model.Upload(submission.meta, submission.message, "")) def submit(params: Params.Data): Option[Model.T] = handler.get(params.get(ACTION).value, all_topics, licenses).flatMap { submission => if (submission.status.nonEmpty) None else { handler.submit(submission.id) Some(submission.copy(status = Some(Model.Status.Submitted))) } } def abort_build(params: Params.Data): Option[Model.T] = handler.get(params.get(ACTION).value, all_topics, licenses).flatMap { submission => if (submission.build != Model.Build.Running) None else { handler.abort_build(submission.id) Some(submission.copy(build = Model.Build.Aborted)) } } def set_status(params: Params.Data): Option[Model.T] = { for { list <- parse_submission_list(params) num_entry <- List_Key.num(ENTRY, params.get(ACTION).value) changed <- list.submissions.unapply(num_entry) } yield { if (changed.status == Model.Status.Added && !devel) { progress.echo_if(verbose, "Updating server data...") val id_before = repo.id() repo.pull() repo.update() val id_after = repo.id() if (id_before != id_after) { load() progress.echo("Updated repo from " + id_before + " to " + id_after) } } handler.set_status(changed.id, changed.status) list } } def submission_list: Option[Model.T] = Some(handler.list()) def download(props: Properties.T): Option[Path] = Properties.get(props, ID).flatMap(handler.get_patch) def download_archive(props: Properties.T): Option[Path] = Properties.get(props, ID).flatMap(handler.get_archive) def style_sheet: Option[Path] = Some(afp_structure.base_dir + Path.make(List("tools", "main.css"))) val error = Model.Invalid val endpoints = List( Get(SUBMIT, "empty submission form", _ => empty_submission), Get(SUBMISSION, "get submission", get_submission), Get(SUBMISSIONS, "list submissions", _ => submission_list), Get_File(API_SUBMISSION_DOWNLOAD, "download patch", download), Get_File(API_SUBMISSION_DOWNLOAD_ZIP, "download archive", download_archive), Get_File(API_SUBMISSION_DOWNLOAD_TGZ, "download archive", download_archive), Get_File(API_CSS, "download css", _ => style_sheet), Post(API_RESUBMIT, "get form for resubmit", resubmit), Post(API_SUBMIT, "submit to editors", submit), Post(API_BUILD_ABORT, "abort the build", abort_build), Post(API_SUBMISSION, "get submission form", parse_create), Post(API_SUBMISSION_AUTHORS_ADD, "add author", add_new_author), Post(API_SUBMISSION_AUTHORS_REMOVE, "remove author", remove_new_author), Post(API_SUBMISSION_AFFILIATIONS_ADD, "add affil", add_new_affil), Post(API_SUBMISSION_AFFILIATIONS_REMOVE, "remove affil", remove_new_affil), Post(API_SUBMISSION_ENTRIES_ADD, "add entry", add_entry), Post(API_SUBMISSION_ENTRIES_REMOVE, "remove entry", remove_entry), Post(API_SUBMISSION_ENTRY_AUTHORS_ADD, "add entry author", add_author), Post(API_SUBMISSION_ENTRY_AUTHORS_REMOVE, "remove entry author", remove_author), Post(API_SUBMISSION_ENTRY_NOTIFIES_ADD, "add notify", add_notify), Post(API_SUBMISSION_ENTRY_NOTIFIES_REMOVE, "remove notify", remove_notify), Post(API_SUBMISSION_ENTRY_TOPICS_ADD, "add topic", add_topic), Post(API_SUBMISSION_ENTRY_TOPICS_REMOVE, "remove topic", remove_topic), Post(API_SUBMISSION_ENTRY_RELATED_ADD, "add related", add_related), Post(API_SUBMISSION_ENTRY_RELATED_REMOVE, "remove related", remove_related), Post(API_SUBMISSION_UPLOAD, "upload archive", upload), Post(API_SUBMISSION_CREATE, "create submission", create), Post(API_SUBMISSION_STATUS, "set submission status", set_status)) val head = List( XML.Elem(Markup("script", List( "type" -> "text/javascript", "id" -> "MathJax-script", "async" -> "async", "src" -> "https://cdn.jsdelivr.net/npm/mathjax@3/es5/tex-svg.js")), text("\n")), script( "MathJax={tex:{inlineMath:[['$','$'],['\\\\(','\\\\)']]},processEscapes:true,svg:{fontCache:'global'}}"), style_file(api.api_url(API_CSS))) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("afp_submit", "start afp submission server", Scala_Project.here, { args => var backend_path = Path.current - var frontend_url = new URL("http://localhost:8080") + var frontend_url = Url("http://localhost:8080") var devel = false var verbose = false var port = 8080 var dir: Option[Path] = None val getopts = Getopts(""" Usage: isabelle afp_submit [OPTIONS] Options are: -a PATH backend path (if endpoint is not server root) -b URL application frontend url. Default: """ + frontend_url + """" -d devel mode (serves frontend and skips automatic AFP repository updates) -p PORT server port. Default: """ + port + """ -v verbose -D DIR submission directory Start afp submission server. Server is in "edit" mode unless directory to store submissions in is specified. """, "a:" -> (arg => backend_path = Path.explode(arg)), - "b:" -> (arg => frontend_url = new URL(arg)), + "b:" -> (arg => frontend_url = Url(arg)), "d" -> (_ => devel = true), "p:" -> (arg => port = Value.Int.parse(arg)), "v" -> (_ => verbose = true), "D:" -> (arg => dir = Some(Path.explode(arg)))) getopts(args) val afp_structure = AFP_Structure() val progress = new Console_Progress(verbose = verbose) val (handler, mode) = dir match { case Some(dir) => (Handler.Adapter(dir, afp_structure), Mode.SUBMISSION) case None => (Handler.Edit(afp_structure), Mode.EDIT) } val api = new API(frontend_url, backend_path, devel = devel) val server = new Server(api = api, afp_structure = afp_structure, mode = mode, handler = handler, devel = devel, verbose = verbose, progress = progress, port = port) server.run() }) } diff --git a/tools/metadata.scala b/tools/metadata.scala --- a/tools/metadata.scala +++ b/tools/metadata.scala @@ -1,427 +1,418 @@ /* Author: Fabian Huch, TU Muenchen AFP metadata model and TOML serialization. */ package afp import isabelle.* import java.time.LocalDate -import java.net.{URI, URL} +import java.net.URI object Metadata { /* affiliations */ sealed trait Affiliation { def author: Author.ID } case class Unaffiliated(override val author: Author.ID) extends Affiliation case class Email(override val author: Author.ID, id: Email.ID, address: String) extends Affiliation { private val Address = "([^@]+)@(.+)".r val (user, host) = address match { case Address(user, host) => (user, host) case _ => error("Invalid address: " + address) } } object Email { type ID = String def apply(author: Author.ID, id: Email.ID, user: String, host: String): Email = Email(author, id, user + "@" + host) } - case class Homepage(override val author: Author.ID, id: Homepage.ID, url: URL) - extends Affiliation { - override def equals(that: Any): Boolean = - that match { - case other: Homepage => - other.author == author && other.id == id && other.url.toString == url.toString - case _ => false - } - } + case class Homepage(override val author: Author.ID, id: Homepage.ID, url: Url) extends Affiliation object Homepage { type ID = String } /* authors */ case class Orcid(identifier: String) { require( "^([0-9]{4})-([0-9]{4})-([0-9]{4})-([0-9]{3}[0-9X])$".r.matches(identifier), "Invalid format for orcid: " + quote(identifier)) - def url: URL = new URL("https", "orcid.org", "/" + identifier) + def url: Url = Url("https://orcid.org/" + identifier) } case class Author( id: Author.ID, name: String, emails: List[Email] = Nil, homepages: List[Homepage] = Nil, orcid: Option[Orcid] = None ) object Author { type ID = String } /* topics */ sealed trait Classification { def desc: String - def url: URL + def url: Url } case class ACM(id: String, override val desc: String) extends Classification { - val url = new URL("https", "dl.acm.org", "/topic/ccs2012/" + id) + val url = Url("https://dl.acm.org/topic/ccs2012/" + id) } case class AMS(id: String, hierarchy: List[String]) extends Classification { val code: String = id.length match { case 2 => id + "-XX" case 3 => id + "xx" case 5 => id case _ => error("Invalid ams id:" + id) } override val desc: String = hierarchy.mkString(" / ") - override val url: URL = - new URL("https", "mathscinet.ams.org", "/mathscinet/msc/msc2020.html?t=" + code) + override val url: Url = Url("https://mathscinet.ams.org/mathscinet/msc/msc2020.html?t=" + code) } case class Topic( id: Topic.ID, name: String, classification: List[Classification] = Nil, sub_topics: List[Topic] = Nil ) { def all_topics: List[Topic] = this :: sub_topics.flatMap(_.all_topics) } object Topic { type ID = String } /* releases */ type Date = LocalDate object Isabelle { type Version = String } case class Release(entry: Entry.Name, date: Date, isabelle: Isabelle.Version) /* license */ case class License(id: License.ID, name: String) object License { type ID = String } /* references */ sealed trait Reference case class DOI(identifier: String) extends Reference { require("^10.([1-9][0-9]{3,})/(.+)".r.matches(identifier), "invalid format for DOI: " + quote(identifier)) def uri: URI = new URI("doi:" + identifier) - def url: URL = new URL("https", "doi.org", "/" + identifier) + def url: Url = Url("https://doi.org/" + identifier) def formatted(style: String = "apa"): String = Utils.fetch_text(url, Map("Accept" -> ("text/x-bibliography; style=" + style))) } case class Formatted(rep: String) extends Reference /* misc */ type Change_History = Map[Date, String] type Extra = Map[String, String] /* entry */ case class Entry( name: Entry.Name, title: String, authors: List[Affiliation], date: Date, topics: List[Topic], `abstract`: String, notifies: List[Email], license: License, note: String, contributors: List[Affiliation] = Nil, change_history: Change_History = Map.empty, extra: Extra = Map.empty, releases: List[Release] = Nil, statistics_ignore: Boolean = false, related: List[Reference] = Nil) object Entry { type Name = String } /* toml */ private def by_id[A](elems: Map[String, A], id: String): A = elems.getOrElse(id, error("Elem " + quote(id) + " not found in " + commas_quote(elems.keys))) object TOML { import isabelle.TOML.{Array, Boolean, Key, Local_Date, String, Table} /* affils */ def from_email(email: Email): Table = Table( "user" -> Array(email.user.split('.').map(String(_))), "host" -> Array(email.host.split('.').map(String(_)))) def to_email(author_id: Author.ID, email_id: Email.ID, email: Table): Email = { val user = email.array("user").string.values.map(_.rep) val host = email.array("host").string.values.map(_.rep) Email(author_id, email_id, user.mkString("."), host.mkString(".")) } /* author */ def from_author(author: Author): Table = Table( "name" -> String(author.name), "emails" -> Table(author.emails.map(e => e.id -> from_email(e))), "homepages" -> Table(author.homepages.map(h => h.id -> String(h.url.toString)))) ++ author.orcid.map(orcid => Table("orcid" -> String(orcid.identifier))).getOrElse(Table()) def to_author(author_id: Author.ID, author: Table): Author = { val emails = author.table("emails").table.values.map { case (id, email) => to_email(author_id, id, email) } val homepages = author.table("homepages").string.values.map { - case (id, url) => Homepage(author = author_id, id = id, url = new URL(url.rep)) + case (id, url) => Homepage(author = author_id, id = id, url = Url(url.rep)) } val orcid = author.string.get("orcid").map(_.rep).map(Orcid(_)) Author( id = author_id, name = author.string("name").rep, orcid = orcid, emails = emails, homepages = homepages) } def from_authors(authors: List[Author]): Table = Table(authors.map(author => author.id -> from_author(author))) def to_authors(authors: Table): List[Author] = authors.table.values.map(to_author) /* topics */ def from_acm(acm: ACM): Table = Table("id" -> String(acm.id), "desc" -> String(acm.desc)) def to_acm(acm: Table): ACM = ACM(acm.string("id").rep, acm.string("desc").rep) def from_ams(ams: AMS): Table = Table("id" -> String(ams.id), "hierarchy" -> Array(ams.hierarchy.map(String(_)))) def to_ams(ams: Table): AMS = AMS(ams.string("id").rep, ams.array("hierarchy").string.values.map(_.rep)) def from_classifications(classifications: List[Classification]): Table = Table(classifications.map { case acm: ACM => "acm" -> from_acm(acm) case ams: AMS => "ams" -> from_ams(ams) }) def to_classifications(classifications: Table): List[Classification] = classifications.table.values.map { case ("ams", ams) => to_ams(ams) case ("acm", acm) => to_acm(acm) case (c, _) => error("Unknown topic classification: " + quote(c)) } def from_topics(root_topics: List[Topic]): Table = Table(root_topics.map(t => t.name -> ( Table("classification" -> from_classifications(t.classification)) ++ from_topics(t.sub_topics)))) def to_topics(root_topics: Table): List[Topic] = { def to_topics_rec(topics: List[(Key, Table)], root: Topic.ID): List[Topic] = { topics.map { case (name, data) => val id = (if (root.nonEmpty) root + "/" else "") + name val classifications = to_classifications(data.table("classification")) val sub_topics = data.table.values.filterNot(_._1 == "classification") Topic(id, name, classifications, to_topics_rec(sub_topics, id)) } } to_topics_rec(root_topics.table.values, "") } /* releases */ def from_releases(releases: List[Release]): Table = Table(Utils.group_sorted(releases, (r: Release) => r.entry).view.mapValues { entry_releases => Table(entry_releases.map(r => r.date.toString -> String(r.isabelle))) }.toList) def to_releases(map: Table): List[Release] = map.table.values.flatMap { case (entry, releases) => releases.string.values.map { case (date, version) => Release(entry = entry, date = LocalDate.parse(date), isabelle = version.rep) } } /* affiliation */ def from_affiliations(affiliations: List[Affiliation]): Table = Table(Utils.group_sorted(affiliations, (a: Affiliation) => a.author).view.mapValues(vs => Table(vs.collect { case Email(_, id, _) => "email" -> String(id) case Homepage(_, id, _) => "homepage" -> String(id) })).toList) def to_affiliations(affiliations: Table, authors: Map[Author.ID, Author]): List[Affiliation] = { def to_affiliation(affiliation: (Key, String), author: Author): Affiliation = { affiliation match { case ("email", id) => author.emails.find(_.id == id.rep) getOrElse error("Email not found: " + quote(id.rep)) case ("homepage", id) => author.homepages.find(_.id == id.rep) getOrElse error("Homepage not found: " + quote(id.rep)) case e => error("Unknown affiliation type: " + e) } } affiliations.table.values.flatMap { case (id, author_affiliations) => val author = by_id(authors, id) if (author_affiliations.is_empty) List(Unaffiliated(author.id)) else author_affiliations.string.values.map(to_affiliation(_, author)) } } def from_emails(emails: List[Email]): Table = Table(emails.map(email => email.author -> String(email.id))) def to_emails(emails: Table, authors: Map[Author.ID, Author]): List[Email] = emails.string.values.map { case (author, id) => by_id(authors, author).emails.find(_.id == id.rep) getOrElse error("Email not found: " + quote(id.rep)) } /* license */ def from_licenses(licenses: List[License]): Table = Table(licenses.map(license => license.id -> Table("name" -> String(license.name)))) def to_licenses(licenses: Table): List[License] = { licenses.table.values.map { case (id, license) => License(id, license.string("name").rep) } } def to_license(license: String, licenses: Map[License.ID, License]): License = licenses.getOrElse(license.rep, error("No such license: " + quote(license.rep))) /* history */ def from_change_history(change_history: Change_History): Table = Table(change_history.map { case (date, str) => date.toString -> String(str) }) def to_change_history(change_history: Table): Change_History = change_history.string.values.map { case (date, entry) => LocalDate.parse(date) -> entry.rep }.toMap /* references */ def from_related(references: List[Reference]): Table = { val dois = references.collect { case d: DOI => d } val formatted = references.collect { case f: Formatted => f } Table( "dois" -> Array(dois.map(_.identifier).map(String(_))), "pubs" -> Array(formatted.map(_.rep).map(String(_)))) } def to_related(references: Table): List[Reference] = { val dois = references.array.get("dois").toList.flatMap(_.string.values.map(_.rep)) val pubs = references.array.get("pubs").toList.flatMap(_.string.values.map(_.rep)) dois.map(DOI(_)) ++ pubs.map(Formatted(_)) } /* entry */ def from_entry(entry: Entry): Table = { Table( "title" -> String(entry.title), "authors" -> from_affiliations(entry.authors), "contributors" -> from_affiliations(entry.contributors), "date" -> Local_Date(entry.date), "topics" -> Array(entry.topics.map(_.id).map(String(_))), "abstract" -> String(entry.`abstract`), "notify" -> from_emails(entry.notifies), "license" -> String(entry.license.id), "note" -> String(entry.note), "history" -> from_change_history(entry.change_history), "extra" -> Table(entry.extra.view.mapValues(String(_)).toList), "related" -> from_related(entry.related)) ++ (if (entry.statistics_ignore) Table("statistics_ignore" -> Boolean(true)) else Table()) } def to_entry( name: Entry.Name, entry: Table, authors: Map[Author.ID, Author], topics: Map[Topic.ID, Topic], licenses: Map[License.ID, License], releases: List[Release] ): Entry = Entry( name = name, title = entry.string("title").rep, authors = to_affiliations(entry.table("authors"), authors), date = entry.local_date("date").rep, topics = entry.array("topics").string.values.map(_.rep).map(by_id(topics, _)), `abstract` = entry.string("abstract").rep, notifies = to_emails(entry.table("notify"), authors), license = to_license(entry.string("license"), licenses), note = entry.string("note").rep, contributors = to_affiliations(entry.table("contributors"), authors), change_history = to_change_history(entry.table("history")), extra = entry.table("extra").string.values.map((k, v) => (k, v.rep)).toMap, releases = releases, statistics_ignore = entry.boolean.get("statistics_ignore").map(_.rep).getOrElse(false), related = to_related(entry.table("related"))) } } diff --git a/tools/utils.scala b/tools/utils.scala --- a/tools/utils.scala +++ b/tools/utils.scala @@ -1,59 +1,58 @@ /* Author: Fabian Huch, TU Muenchen Utilities. */ package afp import isabelle.* -import java.net.URL import java.io.{BufferedReader, InputStreamReader, IOException} import scala.collection.immutable.ListMap object Utils { val TIMEOUT = 30*1000 def group_sorted[A, K](l: List[A], f: A => K): ListMap[K, List[A]] = l.foldLeft(ListMap.empty[K, List[A]]) { case (m, a) => m.updatedWith(f(a)) { case Some(as) => Some(as :+ a) case None => Some(List(a)) } } def grouped_sorted[A, K](l: List[A], f: A => K): ListMap[K, A] = group_sorted(l, f).map { case (k, v :: Nil) => k -> v case (k, vs) => error("Not distinct for " + quote(k.toString) + ": " + commas_quote(vs.map(_.toString))) } def the_entry[K, V](m: Map[K, V], k: K): V = m.getOrElse(k, error("Expected key " + quote(k.toString))) - def fetch_text(url: URL, params: Map[String, String]): String = + def fetch_text(url: Url, params: Map[String, String]): String = try { - val conn = url.openConnection() + val conn = url.open_connection() conn.setConnectTimeout(TIMEOUT) conn.setReadTimeout(TIMEOUT) params.foreach { case (param, value) => conn.setRequestProperty(param, value) } File.read_stream(conn.getInputStream) } catch { case _: IOException => error("Could not read from " + quote(url.toString)) } def remove_at[A](i: Int, l: List[A]): List[A] = l.take(i) ++ l.drop(i + 1) def make_unique(prefix: String, elems: Set[String]): String = { if (!elems.contains(prefix)) prefix else { var num = 1 while (elems.contains(prefix + num)) { num += 1 } prefix + num } } def is_distinct[A](it: List[A]): Boolean = it.size == it.distinct.size } \ No newline at end of file diff --git a/tools/web_app.scala b/tools/web_app.scala --- a/tools/web_app.scala +++ b/tools/web_app.scala @@ -1,517 +1,515 @@ /* Author: Fabian Huch, TU Muenchen Technical layer for web apps using server-side rendering with HTML forms. */ package afp import isabelle.* -import java.net.URL -import java.nio.file.Files import scala.annotation.tailrec object Web_App { val FILE = "file" val ACTION = "action" /* form html elements */ object HTML { import isabelle.HTML._ def css(s: String): Attribute = new Attribute("style", s) def name(n: String): Attribute = new Attribute("name", n) def value(v: String): Attribute = new Attribute("value", v) def placeholder(p: String): Attribute = new Attribute("placeholder", p) val italic = new Operator("i") val fieldset = new Operator("fieldset") val button = new Operator("button") def legend(txt: String): XML.Elem = XML.Elem(Markup("legend", Nil), text(txt)) def input(typ: String): XML.Elem = XML.Elem(Markup("input", List("type" -> typ)), Nil) def hidden(k: String, v: String): XML.Elem = id(k)(name(k)(value(v)(input("hidden")))) def textfield(i: String, p: String, v: String): XML.Elem = id(i)(name(i)(value(v)(placeholder(p)(input("text"))))) def browse(i: String, accept: List[String]): XML.Elem = id(i)(name(i)(input("file"))) + ("accept" -> accept.mkString(",")) def label(`for`: String, txt: String): XML.Elem = XML.Elem(Markup("label", List("for" -> `for`)), text(txt)) def fieldlabel(for_elem: String, txt: String): XML.Elem = label(for_elem, " " + txt + ": ") def explanation(for_elem: String, txt: String): XML.Elem = par(List(italic(List(label(for_elem, txt))))) def option(k: String, v: String): XML.Elem = XML.Elem(Markup("option", List("value" -> k)), text(v)) def optgroup(txt: String, opts: XML.Body): XML.Elem = XML.Elem(Markup("optgroup", List("label" -> txt)), opts) def select(i: String, opts: XML.Body): XML.Elem = XML.Elem(Markup("select", List("id" -> i, "name" -> i)), opts) def textarea(i: String, v: String): XML.Elem = XML.Elem(Markup("textarea", List("id" -> i, "name" -> i, "value" -> v)), text(v + "\n")) def radio(i: String, v: String, values: List[(String, String)]): XML.Elem = { def to_option(k: String): XML.Elem = { val elem = id(i + k)(name(i)(value(k)(input("radio")))) if (v == k) elem + ("checked" -> "checked") else elem } span(values.map { case (k, v) => span(List(label(i + k, v), to_option(k))) }) } def selection(i: String, selected: Option[String], opts: XML.Body): XML.Elem = { def sel(elem: XML.Tree): XML.Tree = { selected match { case Some(value) => val Value = new Properties.String("value") elem match { case XML.Elem(Markup("optgroup", props), body) => XML.Elem(Markup("optgroup", props), body.map(sel)) case e@XML.Elem(Markup("option", Value(v)), _) if v == value => e + ("selected" -> "selected") case e => e } case None => elem } } def is_empty_group(elem: XML.Tree): Boolean = elem match { case XML.Elem(Markup("optgroup", _), body) if body.isEmpty => true case _ => false } val default = if (selected.isEmpty) List(option("", "") + ("hidden" -> "hidden")) else Nil select(i, default ::: opts.filterNot(is_empty_group).map(sel)) } def api_button(call: String, label: String): XML.Elem = button(text(label)) + ("formaction" -> call) + ("type" -> "submit") def action_button(call: String, label: String, action: String): XML.Elem = name(ACTION)(value(action)(api_button(call, label))) def submit_form(endpoint: String, body: XML.Body): XML.Elem = { val default_button = css("display: none")(input("submit") + ("formaction" -> endpoint)) val attrs = List("method" -> "post", "target" -> "iframe", "enctype" -> "multipart/form-data") XML.Elem(Markup("form", attrs), default_button :: body) } val UNESCAPE = "unescape" def unescape(html: String): XML.Body = List(XML.Elem(Markup(UNESCAPE, Nil), text(html))) } /* request parameters */ object Params { type Key = String val empty: Key = "" object Nest_Key { def apply(k: Key, field: String): Key = if (k == empty) field else k + "_" + field def unapply(k: Key): Option[(Key, String)] = k.split('_').filterNot(_.isEmpty).toList.reverse match { case k :: ks => Some(ks.reverse.mkString("_"), k) case _ => None } } object List_Key { def apply(k: Key, field: String, i: Int): Key = Nest_Key(k, field + "_" + i.toString) def unapply(k: Key): Option[(Key, (String, Int))] = k.split('_').filterNot(_.isEmpty).toList.reverse match { case Value.Int(i) :: k :: ks => Some(ks.reverse.mkString("_"), (k, i)) case _ => None } def num(field: String, k: Key): Option[Int] = k match { case List_Key(_, (f, i)) if f == field => Some(i) case _ => None } def split(field: String, k: Key): Option[(Key, Int)] = k match { case List_Key(key, (f, i)) if f == field => Some(key, i) case _ => None } } /* strucutred data */ class Data private[Params]( v: Option[String] = None, elem: Map[String, Data] = Map.empty, elems: Map[String, List[Data]] = Map.empty ) { def is_empty: Boolean = v.isEmpty && elem.isEmpty && elems.isEmpty override def toString: String = { val parts = v.map(v => if (v.length <= 100) quote(v) else quote(v.take(100) + "...")).toList ++ elem.toList.map { case (k, v) => k + " -> " + v.toString } ++ elems.toList.map { case (k, v) => k + " -> (" + v.mkString(",") + ")" } "{" + parts.mkString(", ") + "}" } def value: String = v.getOrElse("") def get(field: String): Data = elem.getOrElse(field, new Data()) def list(field: String): List[Data] = elems.getOrElse(field, Nil) } object Data { def from_multipart(parts: List[Multi_Part.Data]): Data = { sealed trait E case class Value(rep: String) extends E case class Index(i: Int, to: E) extends E case class Nest(field: String, to: E) extends E def group_map[A, B, C](v: List[(C, A)], agg: List[A] => B): Map[C, B] = v.groupBy(_._1).view.mapValues(_.map(_._2)).mapValues(agg).toMap def to_list(l: List[(Int, E)]): List[Data] = { val t: List[(Int, Data)] = group_map(l, parse).toList t.sortBy(_._1).map(_._2) } def parse(e: List[E]): Data = { val value = e.collectFirst { case Value(rep) => rep } val nest_by_key = e.collect { case Nest(field, v: Value) => field -> v case Nest(field, v: Nest) => field -> v } val elem = group_map(nest_by_key, parse) val list_by_key = e.collect { case Nest(field, Index(i, to)) => field -> (i -> to) } val elems = group_map(list_by_key, to_list) new Data(value, elem, elems) } @tailrec def expand(key: Key, to: E): E = key match { case List_Key(key, (field, i)) => expand(key, Nest(field, Index(i, to))) case Nest_Key(key, field) => expand(key, Nest(field, to)) case _ => to } val params = parts.flatMap { case Multi_Part.Param(name, value) => List(name -> value) case Multi_Part.File(name, file_name, content) => List(name -> file_name, Nest_Key(name, Web_App.FILE) -> content.encode_base64) } parse(params.map { case (k, v) => expand(k, Value(v)) }) } } } /* form http handling */ object Multi_Part { abstract class Data(name: String) case class Param(name: String, value: String) extends Data(name) case class File(name: String, file_name: String, content: Bytes) extends Data(name) def parse(body: Bytes): List[Data] = { /* Seq for text with embedded bytes */ case class Seq(text: String, bytes: Bytes) { def split_one(sep: String): Option[(Seq, Seq)] = { val text_i = text.indexOf(sep) if (text_i >= 0 && sep.nonEmpty) { val (before_text, at_text) = text.splitAt(text_i) val after_text = at_text.substring(sep.length) // text might be shorter than bytes because of misinterpreted characters var found = false var bytes_i = 0 while (!found && bytes_i < bytes.length) { var sep_ok = true var sep_i = 0 while (sep_ok && sep_i < sep.length) { if (bytes.charAt(bytes_i + sep_i) == sep.charAt(sep_i)) { sep_i += 1 } else { bytes_i += 1 sep_ok = false } } if (sep_ok) found = true } val before_bytes = bytes.subSequence(0, bytes_i) val after_bytes = bytes.subSequence(bytes_i + sep.length, bytes.length) Some(Seq(before_text, before_bytes), Seq(after_text, after_bytes)) } else None } } val s = Seq(body.text, body) def perhaps_unprefix(pfx: String, s: Seq): Seq = { Library.try_unprefix(pfx, s.text) match { case Some(text) => Seq(text, s.bytes.subSequence(pfx.length, s.bytes.length)) case None => s } } val Separator = """--(.*)""".r s.split_one(HTTP.NEWLINE) match { case Some(Seq(Separator(sep), _), rest) => val Param = """Content-Disposition: form-data; name="([^"]+)"""".r val File = """Content-Disposition: form-data; name="([^"]+)"; filename="([^"]+)"""".r def parts(body: Seq): Option[List[Data]] = for { (first_line, more) <- body.split_one(HTTP.NEWLINE) more1 = perhaps_unprefix(HTTP.NEWLINE, more) (value, rest) <- more1.split_one(HTTP.NEWLINE + "--" + sep) rest1 = perhaps_unprefix(HTTP.NEWLINE, rest) } yield first_line.text match { case Param(name) => Multi_Part.Param(name, Line.normalize(value.text)) :: parts(rest1).getOrElse(Nil) case File(name, file_name) => value.split_one(HTTP.NEWLINE + HTTP.NEWLINE) match { case Some(_, content) => Multi_Part.File(name, file_name, content.bytes) :: parts(rest1).getOrElse(Nil) case _ => parts(rest1).getOrElse(Nil) } case _ => Nil } parts(rest).getOrElse(Nil) case _ => Nil } } } /* API with backend base path, and application url (e.g. for frontend links). */ - class API(val app: URL, val base_path: Path, val devel: Boolean = false) { + class API(val app: Url, val base_path: Path, val devel: Boolean = false) { def url(path: Path, params: Properties.T = Nil): String = { def param(p: Properties.Entry): String = Url.encode(p._1) + "=" + Url.encode(p._2) if (params.isEmpty) path.implode else path.implode + "?" + params.map(param).mkString("&") } def api_path(path: Path, external: Boolean = true): Path = (if (devel) Path.explode("backend") else Path.current) + (if (external) base_path else Path.current) + path def api_url(path: Path, params: Properties.T = Nil, external: Boolean = true): String = "/" + url(api_path(path, external = external), params) def app_url(path: Path, params: Properties.T = Nil): String = app.toString + "/" + url(path, params) } abstract class Server[A]( api: API, port: Int = 0, verbose: Boolean = false, progress: Progress = new Progress() ) { def render(model: A): XML.Body val error: A val endpoints: List[Endpoint] val head: XML.Body def output(tree: XML.Tree): String = { def out(body: XML.Body): String = isabelle.HTML.output(body, hidden = true, structural = true) def collect(t: XML.Tree): List[String] = t match { case XML.Elem(Markup(HTML.UNESCAPE, _), List(XML.Text(escaped))) => List(out(HTML.unescape(escaped))) case XML.Elem(_, body) => body.flatMap(collect) case XML.Text(_) => Nil } collect(tree).foldLeft(out(List(tree))) { case (escaped, html) => escaped.replace(html, isabelle.HTML.input(html)) } } def output_document(content: XML.Body, post_height: Boolean = true): String = { val attrs = if (post_height) List("onLoad" -> "parent.postMessage(document.body.scrollHeight, '*')") else Nil cat_lines( List( isabelle.HTML.header, output(XML.elem("head", isabelle.HTML.head_meta :: head)), output(XML.Elem(Markup("body", attrs), content)), isabelle.HTML.footer)) } class UI(path: Path) extends HTTP.Service(path.implode, "GET") { def apply(request: HTTP.Request): Option[HTTP.Response] = { progress.echo_if(verbose, "GET ui") val on_load = """ (function() { window.addEventListener('message', (event) => { if (Number.isInteger(event.data)) { this.style.height=event.data + 32 + 'px' } }) }).call(this)""" val set_src = """ const base = '""" + api.app.toString.replace("/", "\\/") + """' document.getElementById('iframe').src = base + '""" + api.api_url(path).replace("/", "\\/") + """' + window.location.search""" Some(HTTP.Response.html(output_document( List( XML.Elem( Markup( "iframe", List( "id" -> "iframe", "name" -> "iframe", "style" -> "border-style: none; width: 100%", "onload" -> on_load)), isabelle.HTML.text("content")), isabelle.HTML.script(set_src)), post_height = false))) } } sealed abstract class Endpoint(path: Path, method: String = "GET") extends HTTP.Service(api.api_path(path, external = false).implode, method) { def reply(request: HTTP.Request): HTTP.Response final def apply(request: HTTP.Request): Option[HTTP.Response] = Exn.capture(reply(request)) match { case Exn.Res(res) => Some(res) case Exn.Exn(exn) => val id = UUID.random_string() progress.echo_error_message("Internal error <" + id + ">: " + exn) isabelle.error("Internal server error. ID: " + id) } } private def query_params(request: HTTP.Request): Properties.T = { def decode(s: String): Option[Properties.Entry] = s match { case Properties.Eq(k, v) => Some(Url.decode(k) -> Url.decode(v)) case _ => None } Library.try_unprefix(request.query, request.uri_name).toList.flatMap(params => space_explode('&', params).flatMap(decode)) } /* endpoint types */ class Get(val path: Path, description: String, get: Properties.T => Option[A]) extends Endpoint(path) { def reply(request: HTTP.Request): HTTP.Response = { progress.echo_if(verbose, "GET " + description) val params = query_params(request) progress.echo_if(verbose, "params: " + params.toString()) val model = get(params) match { case Some(model) => model case None => progress.echo_if(verbose, "Parsing failed") error } HTTP.Response.html(output_document(render(model))) } } class Get_File(path: Path, description: String, download: Properties.T => Option[Path]) extends Endpoint(path) { def reply(request: HTTP.Request): HTTP.Response = { progress.echo_if(verbose, "DOWNLOAD " + description) val params = query_params(request) progress.echo_if(verbose, "params: " + params.toString()) download(params) match { case Some(path) => HTTP.Response.content(HTTP.Content.read(path)) case None => progress.echo_if(verbose, "Fetching file path failed") HTTP.Response.html(output_document(render(error))) } } } class Post(path: Path, description: String, post: Params.Data => Option[A]) extends Endpoint(path, method = "POST") { def reply(request: HTTP.Request): HTTP.Response = { progress.echo_if(verbose, "POST " + description) val parts = Multi_Part.parse(request.input) val params = Params.Data.from_multipart(parts) progress.echo_if(verbose, "params: " + params.toString) val model = post(params) match { case Some(model) => model case None => progress.echo_if(verbose, "Parsing failed") error } HTTP.Response.html(output_document(render(model))) } } /* server */ private lazy val services = endpoints ::: (if (api.devel) endpoints.collect { case g: Get => new UI(g.path) } else Nil) private lazy val server = HTTP.server(port = port, name = "", services = services) def run(): Unit = { start() @tailrec def loop(): Unit = { Thread.sleep(Long.MaxValue) loop() } Isabelle_Thread.interrupt_handler(_ => server.stop()) { loop() } } def start(): Unit = { server.start() progress.echo("Server started on port " + server.http_server.getAddress.getPort) } def stop(): Unit = { server.stop() progress.echo("Server stopped") } } } \ No newline at end of file