diff --git a/thys/Green/Green.thy b/thys/Green/Green.thy --- a/thys/Green/Green.thy +++ b/thys/Green/Green.thy @@ -1,3117 +1,3117 @@ theory Green imports Paths Derivs Integrals General_Utils begin lemma frontier_Un_subset_Un_frontier: "frontier (s \ t) \ (frontier s) \ (frontier t)" by (simp add: frontier_def Un_Diff) (auto simp add: closure_def interior_def islimpt_def) definition has_partial_derivative:: "(('a::euclidean_space) \ 'b::euclidean_space) \ 'a \ ('a \ 'b) \ ('a) \ bool" where "has_partial_derivative F base_vec F' a \ ((\x::'a::euclidean_space. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + (x \ base_vec) *\<^sub>R base_vec )) has_derivative F') (at a)" definition has_partial_vector_derivative:: "(('a::euclidean_space) \ 'b::euclidean_space) \ 'a \ ( 'b) \ ('a) \ bool" where "has_partial_vector_derivative F base_vec F' a \ ((\x. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + x *\<^sub>R base_vec )) has_vector_derivative F') (at (a \ base_vec))" definition partially_vector_differentiable where "partially_vector_differentiable F base_vec p \ (\F'. has_partial_vector_derivative F base_vec F' p)" definition partial_vector_derivative:: "(('a::euclidean_space) \ 'b::euclidean_space) \ 'a \ 'a \ 'b" where "partial_vector_derivative F base_vec a \ (vector_derivative (\x. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + x *\<^sub>R base_vec)) (at (a \ base_vec)))" lemma partial_vector_derivative_works: assumes "partially_vector_differentiable F base_vec a" shows "has_partial_vector_derivative F base_vec (partial_vector_derivative F base_vec a) a" proof - obtain F' where F'_prop: "((\x. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + x *\<^sub>R base_vec )) has_vector_derivative F') (at (a \ base_vec))" using assms partially_vector_differentiable_def has_partial_vector_derivative_def by blast show ?thesis using Derivative.differentiableI_vector[OF F'_prop] by(simp add: vector_derivative_works partial_vector_derivative_def[symmetric] has_partial_vector_derivative_def[symmetric]) qed lemma fundamental_theorem_of_calculus_partial_vector: fixes a b:: "real" and F:: "('a::euclidean_space \ 'b::euclidean_space)" and i:: "'a" and j:: "'b" and F_j_i:: "('a::euclidean_space \ real)" assumes a_leq_b: "a \ b" and Base_vecs: "i \ Basis" "j \ Basis" and no_i_component: "c \ i = 0 " and has_partial_deriv: "\p \ D. has_partial_vector_derivative (\x. (F x) \ j) i (F_j_i p) p" and domain_subset_of_D: "{x *\<^sub>R i + c |x. a \ x \ x \ b} \ D" shows "((\x. F_j_i( x *\<^sub>R i + c)) has_integral F(b *\<^sub>R i + c) \ j - F(a *\<^sub>R i + c) \ j) (cbox a b)" proof - let ?domain = "{v. \x. a \ x \ x \ b \ v = x *\<^sub>R i + c}" have "\x\ ?domain. has_partial_vector_derivative (\x. (F x) \ j) i (F_j_i x) x" using has_partial_deriv domain_subset_of_D by auto then have "\x\ (cbox a b). ((\x. F(x *\<^sub>R i + c) \ j) has_vector_derivative (F_j_i( x *\<^sub>R i + c))) (at x)" proof(clarsimp simp add: has_partial_vector_derivative_def) fix x::real assume range_of_x: "a \ x" "x \ b" assume ass2: "\x. (\z\a. z \ b \ x = z *\<^sub>R i + c) \ ((\z. F(x - (x \ i) *\<^sub>R i + z *\<^sub>R i) \ j) has_vector_derivative F_j_i x) (at (x \ i))" have "((\z. F((x *\<^sub>R i + c) - ((x *\<^sub>R i + c) \ i) *\<^sub>R i + z *\<^sub>R i) \ j) has_vector_derivative F_j_i (x *\<^sub>R i + c)) (at ((x *\<^sub>R i + c) \ i)) " using range_of_x ass2 by auto then have 0: "((\x. F( c + x *\<^sub>R i) \ j) has_vector_derivative F_j_i (x *\<^sub>R i + c)) (at x)" by (simp add: assms(2) inner_left_distrib no_i_component) have 1: "(\x. F(x *\<^sub>R i + c) \ j) = (\x. F(c + x *\<^sub>R i) \ j)" by (simp add: add.commute) then show "((\x. F(x *\<^sub>R i + c) \ j) has_vector_derivative F_j_i (x *\<^sub>R i + c)) (at x)" using 0 and 1 by auto qed then have "\x\ (cbox a b). ((\x. F(x *\<^sub>R i + c) \ j) has_vector_derivative (F_j_i( x *\<^sub>R i + c))) (at_within x (cbox a b))" using has_vector_derivative_at_within by blast then show "( (\x. F_j_i( x *\<^sub>R i + c)) has_integral F(b *\<^sub>R i + c) \ j - F(a *\<^sub>R i + c) \ j) (cbox a b)" using fundamental_theorem_of_calculus[of "a" "b" "(\x::real. F(x *\<^sub>R i + c) \ j)" "(\x::real. F_j_i( x *\<^sub>R i + c))"] and a_leq_b by auto qed lemma fundamental_theorem_of_calculus_partial_vector_gen: fixes k1 k2:: "real" and F:: "('a::euclidean_space \ 'b::euclidean_space)" and i:: "'a" and F_i:: "('a::euclidean_space \ 'b)" assumes a_leq_b: "k1 \ k2" and unit_len: "i \ i = 1" and no_i_component: "c \ i = 0 " and has_partial_deriv: "\p \ D. has_partial_vector_derivative F i (F_i p) p" and domain_subset_of_D: "{v. \x. k1 \ x \ x \ k2 \ v = x *\<^sub>R i + c} \ D" shows "( (\x. F_i( x *\<^sub>R i + c)) has_integral F(k2 *\<^sub>R i + c) - F(k1 *\<^sub>R i + c)) (cbox k1 k2)" proof - let ?domain = "{v. \x. k1 \ x \ x \ k2 \ v = x *\<^sub>R i + c}" have "\x\ ?domain. has_partial_vector_derivative F i (F_i x) x" using has_partial_deriv domain_subset_of_D by auto then have "\x\ (cbox k1 k2). ((\x. F(x *\<^sub>R i + c)) has_vector_derivative (F_i( x *\<^sub>R i + c))) (at x)" proof (clarsimp simp add: has_partial_vector_derivative_def) fix x::real assume range_of_x: "k1 \ x" "x \ k2" assume ass2: "\x. (\z\k1. z \ k2 \ x = z *\<^sub>R i + c) \ ((\z. F(x - (x \ i) *\<^sub>R i + z *\<^sub>R i)) has_vector_derivative F_i x) (at (x \ i))" have "((\z. F((x *\<^sub>R i + c) - ((x *\<^sub>R i + c) \ i) *\<^sub>R i + z *\<^sub>R i)) has_vector_derivative F_i (x *\<^sub>R i + c)) (at ((x *\<^sub>R i + c) \ i)) " using range_of_x ass2 by auto then have 0: "((\x. F( c + x *\<^sub>R i)) has_vector_derivative F_i (x *\<^sub>R i + c)) (at x)" by (simp add: inner_commute inner_right_distrib no_i_component unit_len) have 1: "(\x. F(x *\<^sub>R i + c)) = (\x. F(c + x *\<^sub>R i))" by (simp add: add.commute) then show "((\x. F(x *\<^sub>R i + c)) has_vector_derivative F_i (x *\<^sub>R i + c)) (at x)" using 0 and 1 by auto qed then have "\x\ (cbox k1 k2). ((\x. F(x *\<^sub>R i + c) ) has_vector_derivative (F_i( x *\<^sub>R i + c))) (at_within x (cbox k1 k2))" using has_vector_derivative_at_within by blast then show "( (\x. F_i( x *\<^sub>R i + c)) has_integral F(k2 *\<^sub>R i + c) - F(k1 *\<^sub>R i + c)) (cbox k1 k2)" using fundamental_theorem_of_calculus[of "k1" "k2" "(\x::real. F(x *\<^sub>R i + c))" "(\x::real. F_i( x *\<^sub>R i + c))"] and a_leq_b by auto qed lemma add_scale_img: assumes "a < b" shows "(\x::real. a + (b - a) * x) ` {0 .. 1} = {a .. b}" using assms apply (auto simp: algebra_simps affine_ineq image_iff) using less_eq_real_def apply force apply (rule_tac x="(x-a)/(b-a)" in bexI) apply (auto simp: field_simps) done lemma add_scale_img': assumes "a \ b" shows "(\x::real. a + (b - a) * x) ` {0 .. 1} = {a .. b}" proof (cases "a = b") case True then show ?thesis by force next case False then show ?thesis using add_scale_img assms by auto qed definition analytically_valid:: "'a::euclidean_space set \ ('a \ 'b::{euclidean_space,times,one}) \ 'a \ bool" where "analytically_valid s F i \ (\a \ s. partially_vector_differentiable F i a) \ continuous_on s F \ \ \TODO: should we replace this with saying that \F\ is partially diffrerentiable on \Dy\,\ \ \i.e. there is a partial derivative on every dimension\ integrable lborel (\p. (partial_vector_derivative F i) p * indicator s p) \ (\x. integral UNIV (\y. (partial_vector_derivative F i (y *\<^sub>R i + x *\<^sub>R (\ b \(Basis - {i}). b))) * (indicator s (y *\<^sub>R i + x *\<^sub>R (\b \ Basis - {i}. b)) ))) \ borel_measurable lborel" (*(\x. integral UNIV (\y. ((partial_vector_derivative F i) (y, x)) * (indicator s (y, x)))) \ borel_measurable lborel)"*) lemma analytically_valid_imp_part_deriv_integrable_on: assumes "analytically_valid (s::(real*real) set) (f::(real*real)\ real) i" shows "(partial_vector_derivative f i) integrable_on s" proof - have "integrable lborel (\p. partial_vector_derivative f i p * indicator s p)" using assms by(simp add: analytically_valid_def indic_ident) then have "integrable lborel (\p::(real*real). if p \ s then partial_vector_derivative f i p else 0)" using indic_ident[of "partial_vector_derivative f i"] by (simp add: indic_ident) then have "(\x. if x \ s then partial_vector_derivative f i x else 0) integrable_on UNIV" using Equivalence_Lebesgue_Henstock_Integration.integrable_on_lborel by auto then show "(partial_vector_derivative f i) integrable_on s" using integrable_restrict_UNIV by auto qed (*******************************************************************************) definition typeII_twoCube :: "((real * real) \ (real * real)) \ bool" where "typeII_twoCube twoC \ \a b g1 g2. a < b \ (\x \ {a..b}. g2 x \ g1 x) \ twoC = (\(y, x). ((1 - y) * (g2 ((1-x)*a + x*b)) + y * (g1 ((1-x)*a + x*b)), (1-x)*a + x*b)) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b}" abbreviation unit_cube where "unit_cube \ cbox (0,0) (1::real,1::real)" definition cubeImage:: "two_cube \ ((real*real) set)" where "cubeImage twoC \ (twoC ` unit_cube)" lemma typeII_twoCubeImg: assumes "typeII_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(y,x). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} " using assms proof (simp add: typeII_twoCube_def cubeImage_def image_def) assume " \a b. a < b \ (\g1 g2. (\x\{a..b}. g2 x \ g1 x) \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a..b} \ g2 piecewise_C1_differentiable_on {a..b})" then obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x\{a..b}. g2 x \ g1 x)" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" by auto have ab1: "a - x * a + x * b \ b" if a1: "0 \ x" "x \ 1" for x using that apply (simp add: algebra_simps) by (metis affine_ineq less_eq_real_def mult.commute twoCisTypeII(1)) have ex: "\z\Green.unit_cube. (d, c) = (case z of (y, x) \ (g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b), a - x * a + x * b))" if c_bounds: "a \ c" "c \ b" and d_bounds: "g2 c \ d" "d \ g1 c" for c d proof - have b_minus_a_nzero: "b - a \ 0" using twoCisTypeII(1) by auto have x_witness: "\k1. c = (1 - k1)*a + k1 * b \ 0 \ k1 \ k1 \ 1" apply (rule_tac x="(c - a)/(b - a)" in exI) using c_bounds \a < b\ apply (simp add: divide_simps algebra_simps) using sum_sqs_eq by blast have y_witness: "\k2. d = (1 - k2)*(g2 c) + k2 * (g1 c) \ 0 \ k2 \ k2 \ 1" proof (cases "g1 c - g2 c = 0") case True with d_bounds show ?thesis by (fastforce simp add: algebra_simps) next case False let ?k2 = "(d - g2 c)/(g1 c - g2 c)" have k2_in_bounds: "0 \ ?k2 \ ?k2 \ 1" using twoCisTypeII(2) c_bounds d_bounds False by simp have "d = (1 - ?k2) * g2 c + ?k2 * g1 c" using False sum_sqs_eq by (fastforce simp add: divide_simps algebra_simps) with k2_in_bounds show ?thesis by fastforce qed show "\x\unit_cube. (d, c) = (case x of (y, x) \ (g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b), a - x * a + x * b))" using x_witness y_witness by (force simp add: left_diff_distrib) qed have "{y. \x\unit_cube. y = twoC x} = {(y, x). a \ x \ x \ b \ g2 x \ y \ y \ g1 x}" apply (auto simp add: twoCisTypeII ab1 left_diff_distrib ex) using order.order_iff_strict twoCisTypeII(1) apply fastforce apply (smt affine_ineq atLeastAtMost_iff mult_left_mono twoCisTypeII)+ done then show "\a b. a < b \ (\g1 g2. (\x\{a..b}. g2 x \ g1 x) \ {y. \x\unit_cube. y = twoC x} = {(y, x). a \ x \ x \ b \ g2 x \ y \ y \ g1 x} \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a..b} \ g2 piecewise_C1_differentiable_on {a..b})" using twoCisTypeII by blast qed definition horizontal_boundary :: "two_cube \ one_chain" where "horizontal_boundary twoC \ {(1, (\x. twoC(x,0))), (-1, (\x. twoC(x,1)))}" definition vertical_boundary :: "two_cube \ one_chain" where "vertical_boundary twoC \ {(-1, (\y. twoC(0,y))), (1, (\y. twoC(1,y)))}" definition boundary :: "two_cube \ one_chain" where "boundary twoC \ horizontal_boundary twoC \ vertical_boundary twoC" definition valid_two_cube where "valid_two_cube twoC \ card (boundary twoC) = 4" definition two_chain_integral:: "two_chain \ ((real*real)\(real)) \ real" where "two_chain_integral twoChain F \ \C\twoChain. (integral (cubeImage C) F)" definition valid_two_chain where "valid_two_chain twoChain \ (\twoCube \ twoChain. valid_two_cube twoCube) \ pairwise (\c1 c2. ((boundary c1) \ (boundary c2)) = {}) twoChain \ inj_on cubeImage twoChain" definition two_chain_boundary:: "two_chain \ one_chain" where "two_chain_boundary twoChain == \(boundary ` twoChain)" definition gen_division where "gen_division s S \ (finite S \ (\S = s) \ pairwise (\X Y. negligible (X \ Y)) S)" definition two_chain_horizontal_boundary:: "two_chain \ one_chain" where "two_chain_horizontal_boundary twoChain \ \(horizontal_boundary ` twoChain)" definition two_chain_vertical_boundary:: "two_chain \ one_chain" where "two_chain_vertical_boundary twoChain \ \(vertical_boundary ` twoChain)" definition only_horizontal_division where "only_horizontal_division one_chain two_chain \ \\ \. finite \ \ finite \ \ (\(k,\) \ \. (\(k', \') \ two_chain_horizontal_boundary two_chain. (\a \ {0..1}. \b \ {0..1}. a \ b \ subpath a b \' = \))) \ (common_sudiv_exists (two_chain_vertical_boundary two_chain) \ \ common_reparam_exists \ (two_chain_vertical_boundary two_chain)) \ boundary_chain \ \ one_chain = \ \ \ \ (\(k,\)\\. valid_path \)" lemma sum_zero_set: assumes "\x \ s. f x = 0" "finite s" "finite t" shows "sum f (s \ t) = sum f t" using assms by (simp add: IntE sum.union_inter_neutral sup_commute) abbreviation "valid_typeII_division s twoChain \ ((\twoCube \ twoChain. typeII_twoCube twoCube) \ (gen_division s (cubeImage ` twoChain)) \ (valid_two_chain twoChain))" lemma two_chain_vertical_boundary_is_boundary_chain: shows "boundary_chain (two_chain_vertical_boundary twoChain)" by(simp add: boundary_chain_def two_chain_vertical_boundary_def vertical_boundary_def) lemma two_chain_horizontal_boundary_is_boundary_chain: shows "boundary_chain (two_chain_horizontal_boundary twoChain)" by(simp add: boundary_chain_def two_chain_horizontal_boundary_def horizontal_boundary_def) definition typeI_twoCube :: "two_cube \ bool" where "typeI_twoCube (twoC::two_cube) \ \a b g1 g2. a < b \ (\x \ {a..b}. g2 x \ g1 x) \ twoC = (\(x,y). ((1-x)*a + x*b, (1 - y) * (g2 ((1-x)*a + x*b)) + y * (g1 ((1-x)*a + x*b)))) \ g1 piecewise_C1_differentiable_on {a..b} \ g2 piecewise_C1_differentiable_on {a..b}" lemma typeI_twoCubeImg: assumes "typeI_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(x,y). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b))) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} " proof - have "\a b. a < b \ (\g1 g2. (\x\{a..b}. g2 x \ g1 x) \ twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))\ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b})" using assms by (simp add: typeI_twoCube_def) then obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x\{a..b}. g2 x \ g1 x)" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" by auto have ex: "\z\Green.unit_cube. (c, d) = (case z of (x, y) \ (a - x * a + x * b, g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b)))" if c_bounds: "a \ c" "c \ b" and d_bounds: "g2 c \ d" "d \ g1 c" for c d proof - have x_witness: "\k1. c = (1 - k1)*a + k1 * b \ 0 \ k1 \ k1 \ 1" proof - let ?k1 = "(c - a)/(b - a)" have k1_in_bounds: "0 \ (c - a)/(b - a) \ (c - a)/(b - a) \ 1" using twoCisTypeI(1) c_bounds by simp have "c = (1 - ?k1)*a + ?k1 * b" using twoCisTypeI(1) sum_sqs_eq by (auto simp add: divide_simps algebra_simps) then show ?thesis using twoCisTypeI k1_in_bounds by fastforce qed have y_witness: "\k2. d = (1 - k2)*(g2 c) + k2 * (g1 c) \ 0 \ k2 \ k2 \ 1" proof (cases "g1 c - g2 c = 0") case True with d_bounds show ?thesis by force next case False let ?k2 = "(d - g2 c)/(g1 c - g2 c)" have k2_in_bounds: "0 \ ?k2 \ ?k2 \ 1" using twoCisTypeI(2) c_bounds d_bounds False by simp have "d = (1 - ?k2) * g2 c + ?k2 * g1 c" using False apply (simp add: divide_simps algebra_simps) using sum_sqs_eq by fastforce then show ?thesis using k2_in_bounds by fastforce qed show "\x\unit_cube. (c, d) = (case x of (x, y) \ (a - x * a + x * b, g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b)))" using x_witness y_witness by (force simp add: left_diff_distrib) qed have "{y. \x\unit_cube. y = twoC x} = {(x, y). a \ x \ x \ b \ g2 x \ y \ y \ g1 x}" apply (auto simp add: twoCisTypeI left_diff_distrib ex) using less_eq_real_def twoCisTypeI(1) apply auto[1] apply (smt affine_ineq twoCisTypeI) apply (smt affine_ineq atLeastAtMost_iff mult_left_mono twoCisTypeI)+ done then show ?thesis unfolding cubeImage_def image_def using twoCisTypeI by auto qed lemma typeI_cube_explicit_spec: assumes "typeI_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(x,y). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b))) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} \ (\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x))) \ (\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b))) \ (\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x))) \ (\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" using assms and typeI_twoCubeImg[of"twoC"] by auto have bottom_edge_explicit: "?bottom_edge = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" by (simp add: twoCisTypeI(4) algebra_simps) have right_edge_explicit: "?right_edge = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" by (simp add: twoCisTypeI(4) algebra_simps) have top_edge_explicit: "?top_edge = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" by (simp add: twoCisTypeI(4) algebra_simps) have left_edge_explicit: "?left_edge = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" by (simp add: twoCisTypeI(4) algebra_simps) show ?thesis using bottom_edge_explicit right_edge_explicit top_edge_explicit left_edge_explicit twoCisTypeI by auto qed lemma typeI_twoCube_smooth_edges: assumes "typeI_twoCube twoC" "(k,\) \ boundary twoC" shows "\ piecewise_C1_differentiable_on {0..1}" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" "(\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" "(\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" "(\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" using assms and typeI_cube_explicit_spec[of"twoC"] by auto have bottom_edge_smooth: "(\x. twoC (x, 0)) piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeI(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_smth: "(\x. (a + (b - a) * x)) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by blast have g2_smooth: "g2 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeI(1)] twoCisTypeI(6) by auto have "(\x. g2 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using Cauchy_Integral_Theorem.piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (a + (b - a) * x, g2 (a + (b - a) * x))) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (a + (b - a) * x, g2 (a + (b - a) * x)))"] apply (simp only: real_pair_basis) by fastforce then show ?thesis using twoCisTypeI(7) by auto qed have top_edge_smooth: "?top_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeI(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_smth: "(\x. (a + (b - a) * x)) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by blast have g1_smooth: "g1 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeI(1)] twoCisTypeI(5) by auto have "(\x. g1 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using Cauchy_Integral_Theorem.piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] by (auto simp add: o_def) then have "(\x. (a + (b - a) * x, g1 (a + (b - a) * x))) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x. (a + (b - a) * x, g1 (a + (b - a) * x)))"] apply (simp only: real_pair_basis) by fastforce then show ?thesis using twoCisTypeI(9) by auto qed have right_edge_smooth: "?right_edge piecewise_C1_differentiable_on {0..1}" proof - have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b))) C1_differentiable_on {0..1}" using scale_shift_smooth C1_differentiable_imp_piecewise by auto then have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b))) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by fastforce then have "(\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b))) piecewise_C1_differentiable_on {0..1}" using pair_prod_smooth_pw_smooth by auto then show ?thesis using twoCisTypeI(8) by auto qed have left_edge_smooth: "?left_edge piecewise_C1_differentiable_on {0..1}" proof - have "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a))) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a))) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by fastforce then have "(\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a))) piecewise_C1_differentiable_on {0..1}" using pair_prod_smooth_pw_smooth by auto then show ?thesis using twoCisTypeI(10) by auto qed have "\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge" using assms by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) then show ?thesis using left_edge_smooth right_edge_smooth top_edge_smooth bottom_edge_smooth by auto qed lemma two_chain_integral_eq_integral_divisable: assumes f_integrable: "\twoCube \ twoChain. F integrable_on cubeImage twoCube" and gen_division: "gen_division s (cubeImage ` twoChain)" and valid_two_chain: "valid_two_chain twoChain" shows "integral s F = two_chain_integral twoChain F" proof - show "integral s F = two_chain_integral twoChain F" proof (simp add: two_chain_integral_def) have partial_deriv_integrable: "\twoCube \ twoChain. ((F) has_integral (integral (cubeImage twoCube) (F))) (cubeImage twoCube)" using f_integrable by auto then have partial_deriv_integrable: "\twoCubeImg. twoCubeImg \ cubeImage ` twoChain \ (F has_integral (integral (twoCubeImg) F)) (twoCubeImg)" using Henstock_Kurzweil_Integration.integrable_neg by force have finite_images: "finite (cubeImage ` twoChain)" using gen_division gen_division_def by auto have negligible_images: "pairwise (\S S'. negligible (S \ S')) (cubeImage ` twoChain)" using gen_division by (auto simp add: gen_division_def pairwise_def) have inj: "inj_on cubeImage twoChain" using valid_two_chain by (simp add: inj_on_def valid_two_chain_def) have "integral s F = (\twoCubeImg\cubeImage ` twoChain. integral twoCubeImg F)" using has_integral_Union[OF finite_images partial_deriv_integrable negligible_images] gen_division by (auto simp add: gen_division_def) also have "\ = (\C\twoChain. integral (cubeImage C) F)" using sum.reindex inj by auto finally show "integral s F = (\C\twoChain. integral (cubeImage C) F)" . qed qed definition only_vertical_division where "only_vertical_division one_chain two_chain \ \\ \. finite \ \ finite \ \ (\(k,\) \ \. (\(k',\') \ two_chain_vertical_boundary two_chain. (\a \ {0..1}. \b \ {0..1}. a \ b \ subpath a b \' = \))) \ (common_sudiv_exists (two_chain_horizontal_boundary two_chain) \ \ common_reparam_exists \ (two_chain_horizontal_boundary two_chain)) \ boundary_chain \ \ one_chain = \ \ \ \ (\(k,\)\\. valid_path \)" abbreviation "valid_typeI_division s twoChain \ (\twoCube \ twoChain. typeI_twoCube twoCube) \ gen_division s (cubeImage ` twoChain) \ valid_two_chain twoChain" lemma field_cont_on_typeI_region_cont_on_edges: assumes typeI_twoC: "typeI_twoCube twoC" and field_cont: "continuous_on (cubeImage twoC) F" and member_of_boundary: "(k,\) \ boundary twoC" shows "continuous_on (\ ` {0 .. 1}) F" proof - obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" "(\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" "(\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" "(\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" using typeI_twoC and typeI_cube_explicit_spec[of"twoC"] by auto let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" let ?Dg1 = "{p. \x. x \ cbox a b \ p = (x, g1(x))}" have line_is_pair_img: "?Dg1 = (\x. (x, g1(x))) ` (cbox a b)" using image_def by auto have field_cont_on_top_edge_image: "continuous_on ?Dg1 F" by (rule continuous_on_subset [OF field_cont]) (auto simp: twoCisTypeI(2) twoCisTypeI(3)) have top_edge_is_compos_of_scal_and_g1: "(\x. twoC(x, 1)) = (\x. (x, g1(x))) \ (\x. a + (b - a) * x)" using twoCisTypeI by auto have Dg1_is_bot_edge_pathimg: "path_image (\x. twoC(x, 1)) = ?Dg1" using line_is_pair_img and top_edge_is_compos_of_scal_and_g1 image_comp path_image_def add_scale_img and twoCisTypeI(1) by (metis (no_types, lifting) cbox_interval) then have cont_on_top: "continuous_on (path_image ?top_edge) F" using field_cont_on_top_edge_image by auto let ?Dg2 = "{p. \x. x \ cbox a b \ p = (x, g2(x))}" have line_is_pair_img: "?Dg2 = (\x. (x, g2(x))) ` (cbox a b)" using image_def by auto have field_cont_on_bot_edge_image: "continuous_on ?Dg2 F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeI(2) twoCisTypeI(3) by auto have bot_edge_is_compos_of_scal_and_g2: "(\x. twoC(x, 0)) = (\x. (x, g2(x))) \ (\x. a + (b - a) * x)" using twoCisTypeI by auto have Dg2_is_bot_edge_pathimg: "path_image (\x. twoC(x, 0)) = ?Dg2" using line_is_pair_img and bot_edge_is_compos_of_scal_and_g2 image_comp path_image_def add_scale_img and twoCisTypeI(1) by (metis (no_types, lifting) cbox_interval) then have cont_on_bot: "continuous_on (path_image ?bottom_edge) F" using field_cont_on_bot_edge_image by auto let ?D_left_edge = "{p. \y. y \ cbox (g2 a) (g1 a) \ p = (a, y)}" have field_cont_on_left_edge_image: "continuous_on ?D_left_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeI(1) twoCisTypeI(3) by auto have "g2 a \ g1 a" using twoCisTypeI(1) twoCisTypeI(2) by auto then have "(\x. g2 a + (g1 a - g2 a) * x) ` {(0::real)..1} = {g2 a .. g1 a}" using add_scale_img'[of "g2 a" "g1 a"] by blast then have left_eq: "?D_left_edge = ?left_edge ` {0..1}" unfolding twoCisTypeI(10) by(auto simp add: subset_iff image_def set_eq_iff Semiring_Normalization.comm_semiring_1_class.semiring_normalization_rules(7)) then have cont_on_left: "continuous_on (path_image ?left_edge) F" using field_cont_on_left_edge_image path_image_def by (metis left_eq field_cont_on_left_edge_image path_image_def) let ?D_right_edge = "{p. \y. y \ cbox (g2 b) (g1 b) \ p = (b, y)}" have field_cont_on_left_edge_image: "continuous_on ?D_right_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeI(1) twoCisTypeI(3) by auto have "g2 b \ g1 b" using twoCisTypeI(1) twoCisTypeI(2) by auto then have "(\x. g2 b + (g1 b - g2 b) * x) ` {(0::real)..1} = {g2 b .. g1 b}" using add_scale_img'[of "g2 b" "g1 b"] by blast then have right_eq: "?D_right_edge = ?right_edge ` {0..1}" unfolding twoCisTypeI(8) by(auto simp add: subset_iff image_def set_eq_iff Semiring_Normalization.comm_semiring_1_class.semiring_normalization_rules(7)) then have cont_on_right: "continuous_on (path_image ?right_edge) F" using field_cont_on_left_edge_image path_image_def by (metis right_eq field_cont_on_left_edge_image path_image_def) have all_edge_cases: "(\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge)" using assms by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) show ?thesis apply (simp add: path_image_def[symmetric]) using cont_on_top cont_on_bot cont_on_right cont_on_left all_edge_cases by blast qed lemma typeII_cube_explicit_spec: assumes "typeII_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(y, x). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} \ (\x. twoC(0, x)) = (\x. (g2 (a + (b - a) * x), a + (b - a) * x)) \ (\y. twoC(y, 1)) = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b)) \ (\x. twoC(1, x)) = (\x. (g1 (a + (b - a) * x), a + (b - a) * x)) \ (\y. twoC(y, 0)) = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" proof - let ?bottom_edge = "(\x. twoC(0, x))" let ?right_edge = "(\y. twoC(y, 1))" let ?top_edge = "(\x. twoC(1, x))" let ?left_edge = "(\y. twoC(y, 0))" obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" using assms and typeII_twoCubeImg[of"twoC"] by auto have bottom_edge_explicit: "?bottom_edge = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" by (simp add: twoCisTypeII(4) algebra_simps) have right_edge_explicit: "?right_edge = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" by (simp add: twoCisTypeII(4) algebra_simps) have top_edge_explicit: "?top_edge = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" by (simp add: twoCisTypeII(4) algebra_simps) have left_edge_explicit: "?left_edge = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" by (simp add: twoCisTypeII(4) algebra_simps) show ?thesis using bottom_edge_explicit right_edge_explicit top_edge_explicit left_edge_explicit twoCisTypeII by auto qed lemma typeII_twoCube_smooth_edges: assumes "typeII_twoCube twoC" "(k,\) \ boundary twoC" shows "\ piecewise_C1_differentiable_on {0..1}" proof - let ?bottom_edge = "(\x. twoC(0, x))" let ?right_edge = "(\y. twoC(y, 1))" let ?top_edge = "(\x. twoC(1, x))" let ?left_edge = "(\y. twoC(y, 0))" obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(0, x)) = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 1)) = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" "(\x. twoC(1, x)) = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 0)) = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" using assms and typeII_cube_explicit_spec[of"twoC"] by auto have bottom_edge_smooth: "?bottom_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x)) -` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by auto then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using scale_shift_smooth C1_differentiable_imp_piecewise by blast have g2_smooth: "g2 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(6) by auto have "(\x. g2 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using Cauchy_Integral_Theorem.piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using twoCisTypeII(7) by auto qed have top_edge_smooth: "?top_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by auto then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using scale_shift_smooth C1_differentiable_imp_piecewise by blast have g1_smooth: "g1 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(5) by auto have "(\x. g1 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using twoCisTypeII(9) by auto qed have right_edge_smooth: "?right_edge piecewise_C1_differentiable_on {0..1}" proof - have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b))) piecewise_C1_differentiable_on {0..1}" by (simp add: C1_differentiable_imp_piecewise) then have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[of "(1,0)" "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))"] by (auto simp add: real_pair_basis) then show ?thesis using twoCisTypeII(8) by auto qed have left_edge_smooth: "?left_edge piecewise_C1_differentiable_on {0..1}" proof - have 0: "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a))) C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by fastforce have "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise[OF C1_differentiable_on_pair[OF 0 C1_differentiable_on_const[of "a" "{0..1}"]]] by force then show ?thesis using twoCisTypeII(10) by auto qed have "\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge" using assms by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) then show ?thesis using left_edge_smooth right_edge_smooth top_edge_smooth bottom_edge_smooth by auto qed lemma field_cont_on_typeII_region_cont_on_edges: assumes typeII_twoC: "typeII_twoCube twoC" and field_cont: "continuous_on (cubeImage twoC) F" and member_of_boundary: "(k,\) \ boundary twoC" shows "continuous_on (\ ` {0 .. 1}) F" proof - obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(0, x)) = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 1)) = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" "(\x. twoC(1, x)) = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 0)) = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" using typeII_twoC and typeII_cube_explicit_spec[of"twoC"] by auto let ?bottom_edge = "(\x. twoC(0, x))" let ?right_edge = "(\y. twoC(y, 1))" let ?top_edge = "(\x. twoC(1, x))" let ?left_edge = "(\y. twoC(y, 0))" let ?Dg1 = "{p. \x. x \ cbox a b \ p = (g1(x), x)}" have line_is_pair_img: "?Dg1 = (\x. (g1(x), x)) ` (cbox a b)" using image_def by auto have field_cont_on_top_edge_image: "continuous_on ?Dg1 F" by (rule continuous_on_subset [OF field_cont]) (auto simp: twoCisTypeII(2) twoCisTypeII(3)) have top_edge_is_compos_of_scal_and_g1: "(\x. twoC(1, x)) = (\x. (g1(x), x)) \ (\x. a + (b - a) * x)" using twoCisTypeII by auto have Dg1_is_bot_edge_pathimg: "path_image (\x. twoC(1, x)) = ?Dg1" using line_is_pair_img and top_edge_is_compos_of_scal_and_g1 image_comp path_image_def add_scale_img and twoCisTypeII(1) by (metis (no_types, lifting) cbox_interval) then have cont_on_top: "continuous_on (path_image ?top_edge) F" using field_cont_on_top_edge_image by auto let ?Dg2 = "{p. \x. x \ cbox a b \ p = (g2(x), x)}" have line_is_pair_img: "?Dg2 = (\x. (g2(x), x)) ` (cbox a b)" using image_def by auto have field_cont_on_bot_edge_image: "continuous_on ?Dg2 F" by (rule continuous_on_subset [OF field_cont]) (auto simp add: twoCisTypeII(2) twoCisTypeII(3)) have bot_edge_is_compos_of_scal_and_g2: "(\x. twoC(0, x)) = (\x. (g2(x), x)) \ (\x. a + (b - a) * x)" using twoCisTypeII by auto have Dg2_is_bot_edge_pathimg: "path_image (\x. twoC(0, x)) = ?Dg2" unfolding path_image_def using line_is_pair_img and bot_edge_is_compos_of_scal_and_g2 image_comp add_scale_img [OF \a < b\] by (metis (no_types, lifting) box_real(2)) then have cont_on_bot: "continuous_on (path_image ?bottom_edge) F" using field_cont_on_bot_edge_image by auto let ?D_left_edge = "{p. \y. y \ cbox (g2 a) (g1 a) \ p = (y, a)}" have field_cont_on_left_edge_image: "continuous_on ?D_left_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeII(1) twoCisTypeII(3) by auto have "g2 a \ g1 a" using twoCisTypeII(1) twoCisTypeII(2) by auto then have "(\x. g2 a + x * (g1 a - g2 a)) ` {(0::real)..1} = {g2 a .. g1 a}" using add_scale_img'[of "g2 a" "g1 a"] by (auto simp add: ac_simps) with \g2 a \ g1 a\ have left_eq: "?D_left_edge = ?left_edge ` {0..1}" by (simp only: twoCisTypeII(10)) auto then have cont_on_left: "continuous_on (path_image ?left_edge) F" using field_cont_on_left_edge_image path_image_def by (metis left_eq field_cont_on_left_edge_image path_image_def) let ?D_right_edge = "{p. \y. y \ cbox (g2 b) (g1 b) \ p = (y, b)}" have field_cont_on_left_edge_image: "continuous_on ?D_right_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeII(1) twoCisTypeII(3) by auto have "g2 b \ g1 b" using twoCisTypeII(1) twoCisTypeII(2) by auto then have "(\x. g2 b + x * (g1 b - g2 b)) ` {(0::real)..1} = {g2 b .. g1 b}" using add_scale_img'[of "g2 b" "g1 b"] by (auto simp add: ac_simps) with \g2 b \ g1 b\ have right_eq: "?D_right_edge = ?right_edge ` {0..1}" by (simp only: twoCisTypeII(8)) auto then have cont_on_right: "continuous_on (path_image ?right_edge) F" using field_cont_on_left_edge_image path_image_def by (metis right_eq field_cont_on_left_edge_image path_image_def) have all_edge_cases: "(\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge)" using assms unfolding boundary_def horizontal_boundary_def vertical_boundary_def by blast show ?thesis apply (simp add: path_image_def[symmetric]) using cont_on_top cont_on_bot cont_on_right cont_on_left all_edge_cases by blast qed lemma two_cube_boundary_is_boundary: "boundary_chain (boundary C)" by (auto simp add: boundary_chain_def boundary_def horizontal_boundary_def vertical_boundary_def) lemma common_boundary_subdiv_exists_refl: assumes "\(k,\)\boundary twoC. valid_path \" shows "common_boundary_sudivision_exists (boundary twoC) (boundary twoC)" using assms chain_subdiv_chain_refl common_boundary_sudivision_exists_def two_cube_boundary_is_boundary by blast lemma common_boundary_subdiv_exists_refl': assumes "\(k,\)\C. valid_path \" "boundary_chain (C::(int \ (real \ real \ real)) set)" shows "common_boundary_sudivision_exists (C) (C)" using assms chain_subdiv_chain_refl common_boundary_sudivision_exists_def by blast lemma gen_common_boundary_subdiv_exists_refl_twochain_boundary: assumes "\(k,\)\C. valid_path \" "boundary_chain (C::(int \ (real \ real \ real)) set)" shows "common_sudiv_exists (C) (C)" using assms chain_subdiv_chain_refl common_boundary_sudivision_exists_def common_subdiv_imp_gen_common_subdiv by blast lemma two_chain_boundary_is_boundary_chain: shows "boundary_chain (two_chain_boundary twoChain)" by (simp add: boundary_chain_def two_chain_boundary_def boundary_def horizontal_boundary_def vertical_boundary_def) lemma typeI_edges_are_valid_paths: assumes "typeI_twoCube twoC" "(k,\) \ boundary twoC" shows "valid_path \" using typeI_twoCube_smooth_edges[OF assms] C1_differentiable_imp_piecewise by (auto simp: valid_path_def) lemma typeII_edges_are_valid_paths: assumes "typeII_twoCube twoC" "(k,\) \ boundary twoC" shows "valid_path \" using typeII_twoCube_smooth_edges[OF assms] C1_differentiable_imp_piecewise by (auto simp: valid_path_def) lemma finite_two_chain_vertical_boundary: assumes "finite two_chain" shows "finite (two_chain_vertical_boundary two_chain)" using assms by (simp add: two_chain_vertical_boundary_def vertical_boundary_def) lemma finite_two_chain_horizontal_boundary: assumes "finite two_chain" shows "finite (two_chain_horizontal_boundary two_chain)" using assms by (simp add: two_chain_horizontal_boundary_def horizontal_boundary_def) locale R2 = fixes i j assumes i_is_x_axis: "i = (1::real,0::real)" and j_is_y_axis: "j = (0::real, 1::real)" begin lemma analytically_valid_y: assumes "analytically_valid s F i" shows "(\x. integral UNIV (\y. (partial_vector_derivative F i) (y, x) * (indicator s (y, x)))) \ borel_measurable lborel" proof - have "{(1::real, 0::real), (0, 1)} - {(1, 0)} = {(0::real,1::real)}" by force with assms show ?thesis using assms by(simp add: real_pair_basis analytically_valid_def i_is_x_axis j_is_y_axis) qed lemma analytically_valid_x: assumes "analytically_valid s F j" shows "(\x. integral UNIV (\y. ((partial_vector_derivative F j) (x, y)) * (indicator s (x, y)))) \ borel_measurable lborel" proof - have "{(1::real, 0::real), (0, 1)} - {(0, 1)} = {(1::real, 0::real)}" by force with assms show ?thesis by(simp add: real_pair_basis analytically_valid_def i_is_x_axis j_is_y_axis) qed lemma Greens_thm_type_I: fixes F:: "((real*real) \ (real * real))" and gamma1 gamma2 gamma3 gamma4 :: "(real \ (real * real))" and a:: "real" and b:: "real" and g1:: "(real \ real)" and g2:: "(real \ real)" assumes Dy_def: "Dy_pair = {(x::real,y) . x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" and gamma1_def: "gamma1 = (\x. (a + (b - a) * x, g2(a + (b - a) * x)))" and gamma1_smooth: "gamma1 piecewise_C1_differentiable_on {0..1}" and (*TODO: This should be piecewise smooth*) gamma2_def: "gamma2 = (\x. (b, g2(b) + x *\<^sub>R (g1(b) - g2(b))))" and gamma3_def: "gamma3 = (\x. (a + (b - a) * x, g1(a + (b - a) * x)))" and gamma3_smooth: "gamma3 piecewise_C1_differentiable_on {0..1}" and gamma4_def: "gamma4 = (\x. (a, g2(a) + x *\<^sub>R (g1(a) - g2(a))))" and F_i_analytically_valid: "analytically_valid Dy_pair (\p. F(p) \ i) j" and g2_leq_g1: "\x \ cbox a b. (g2 x) \ (g1 x)" and (*This is needed otherwise what would Dy be?*) a_lt_b: "a < b" shows "(line_integral F {i} gamma1) + (line_integral F {i} gamma2) - (line_integral F {i} gamma3) - (line_integral F {i} gamma4) = (integral Dy_pair (\a. - (partial_vector_derivative (\p. F(p) \ i) j a)))" "line_integral_exists F {i} gamma4" "line_integral_exists F {i} gamma3" "line_integral_exists F {i} gamma2" "line_integral_exists F {i} gamma1" proof - let ?F_b' = "partial_vector_derivative (\a. (F a) \ i) j" have F_first_is_continuous: "continuous_on Dy_pair (\a. F(a) \ i)" using F_i_analytically_valid by (auto simp add: analytically_valid_def) let ?f = "(\x. if x \ (Dy_pair) then (partial_vector_derivative (\a. (F a) \ i) j) x else 0)" have f_lesbegue_integrable: "integrable lborel ?f" using F_i_analytically_valid by (auto simp add: analytically_valid_def indic_ident) have partially_vec_diff: "\a \ Dy_pair. partially_vector_differentiable (\a. (F a) \ i) j a" using F_i_analytically_valid by (auto simp add: analytically_valid_def indicator_def) have x_axis_integral_measurable: "(\x. integral UNIV (\y. ?f(x, y))) \ borel_measurable lborel" proof - have "(\p. (?F_b' p) * indicator (Dy_pair) p) = (\x. if x \ (Dy_pair) then (?F_b') x else 0)" using indic_ident[of "?F_b'"] by auto then have "\x y. ?F_b' (x,y) * indicator (Dy_pair) (x,y) = (\x. if x \ (Dy_pair) then (?F_b') x else 0) (x,y)" by auto then show ?thesis using analytically_valid_x[OF F_i_analytically_valid] by (auto simp add: indicator_def) qed have F_partially_differentiable: "\a\Dy_pair. has_partial_vector_derivative (\x. (F x) \ i) j (?F_b' a) a" using partial_vector_derivative_works partially_vec_diff by fastforce have g1_g2_continuous: "continuous_on (cbox a b) g1" "continuous_on (cbox a b) g2" proof - have shift_scale_cont: "continuous_on {a..b} (\x. (x - a)*(1/(b-a)))" by (intro continuous_intros) have shift_scale_inv: "(\x. a + (b - a) * x) \ (\x. (x - a)*(1/(b-a))) = id" using a_lt_b by (auto simp add: o_def) have img_shift_scale: "(\x. (x - a)*(1/(b-a)))`{a..b} = {0..1}" using a_lt_b apply (auto simp: divide_simps image_iff) apply (rule_tac x="x * (b - a) + a" in bexI) using le_diff_eq by fastforce+ have gamma1_y_component: "(\x. g2(a + (b - a) * x)) = g2 \ (\x.(a + (b - a) * x))" by auto have "continuous_on {0..1} (\x. g2(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma1_smooth], of "(\x. j)", OF continuous_on_const] by (simp add: gamma1_def j_is_y_axis) then have "continuous_on {a..b} ((\x. g2(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then have "continuous_on {a..b} (g2 \ (\x.(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using gamma1_y_component by auto then show "continuous_on (cbox a b) g2" using a_lt_b by (simp add: shift_scale_inv) have gamma3_y_component: "(\x. g1(a + (b - a) * x)) = g1 \ (\x.(a + (b - a) * x))" by auto have "continuous_on {0..1} (\x. g1(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma3_smooth], of "(\x. j)", OF continuous_on_const] by (simp add: gamma3_def j_is_y_axis) then have "continuous_on {a..b} ((\x. g1(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then have "continuous_on {a..b} (g1 \ (\x.(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using gamma3_y_component by auto then show "continuous_on (cbox a b) g1" using a_lt_b by (simp add: shift_scale_inv) qed have g2_scale_j_contin: "continuous_on (cbox a b) (\x. (0, g2 x))" by (intro continuous_intros g1_g2_continuous) let ?Dg2 = "{p. \x. x \ cbox a b \ p = (x, g2(x))}" have line_is_pair_img: "?Dg2 = (\x. (x, g2(x))) ` (cbox a b)" using image_def by auto have g2_path_continuous: "continuous_on (cbox a b) (\x. (x, g2(x)))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma1_image: "continuous_on ?Dg2 (\a. F(a) \ i)" apply (rule continuous_on_subset [OF F_first_is_continuous]) by (auto simp add: Dy_def g2_leq_g1) have gamma1_is_compos_of_scal_and_g2: "gamma1 = (\x. (x, g2(x))) \ (\x. a + (b - a) * x)" using gamma1_def by auto have add_scale_img: "(\x. a + (b - a) * x) ` {0 .. 1} = {a .. b}" using add_scale_img and a_lt_b by auto then have Dg2_is_gamma1_pathimg: "path_image gamma1 = ?Dg2" by (metis (no_types, lifting) box_real(2) gamma1_is_compos_of_scal_and_g2 image_comp line_is_pair_img path_image_def) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma1_as_euclid_space_fun: "gamma1 = (\x. (a + (b - a) * x) *\<^sub>R i + (0, g2 (a + (b - a) * x)))" using i_is_x_axis gamma1_def by auto have 0: "line_integral F {i} gamma1 = integral (cbox a b) (\x. F(x, g2(x)) \ i )" "line_integral_exists F {i} gamma1" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(1)] _ gamma1_as_euclid_space_fun, of "F"] gamma1_def gamma1_smooth g2_scale_j_contin a_lt_b add_scale_img Dg2_is_gamma1_pathimg and field_cont_on_gamma1_image by (auto simp: pathstart_def pathfinish_def i_is_x_axis) then show "(line_integral_exists F {i} gamma1)" by metis have gamma2_x_const: "\x. gamma2 x \ i = b" by (simp add: i_is_x_axis gamma2_def) have 1: "(line_integral F {i} gamma2) = 0" "(line_integral_exists F {i} gamma2)" using line_integral_on_pair_straight_path[OF gamma2_x_const] straight_path_diffrentiable_x gamma2_def by (auto simp add: mult.commute) then show "(line_integral_exists F {i} gamma2)" by metis have "continuous_on (cbox a b) (\x. F(x, g2(x)) \ i)" using line_is_pair_img and g2_path_continuous and field_cont_on_gamma1_image Topological_Spaces.continuous_on_compose i_is_x_axis j_is_y_axis by auto then have 6: "(\x. F(x, g2(x)) \ i) integrable_on (cbox a b)" using integrable_continuous [of "a" "b" "(\x. F(x, g2(x)) \ i)"] by auto have g1_scale_j_contin: "continuous_on (cbox a b) (\x. (0, g1 x))" by (intro continuous_intros g1_g2_continuous) let ?Dg1 = "{p. \x. x \ cbox a b \ p = (x, g1(x))}" have line_is_pair_img: "?Dg1 = (\x. (x, g1(x)) ) ` (cbox a b)" using image_def by auto have g1_path_continuous: "continuous_on (cbox a b) (\x. (x, g1(x)))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma3_image: "continuous_on ?Dg1 (\a. F(a) \ i)" apply (rule continuous_on_subset [OF F_first_is_continuous]) by (auto simp add: Dy_def g2_leq_g1) have gamma3_is_compos_of_scal_and_g1: "gamma3 = (\x. (x, g1(x))) \ (\x. a + (b - a) * x)" using gamma3_def by auto then have Dg1_is_gamma3_pathimg: "path_image gamma3 = ?Dg1" by (metis (no_types, lifting) box_real(2) image_comp line_is_pair_img local.add_scale_img path_image_def) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma3_as_euclid_space_fun: "gamma3 = (\x. (a + (b - a) * x) *\<^sub>R i + (0, g1 (a + (b - a) * x)))" using i_is_x_axis gamma3_def by auto have 2: "line_integral F {i} gamma3 = integral (cbox a b) (\x. F(x, g1(x)) \ i )" "line_integral_exists F {i} gamma3" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(1)] _ gamma3_as_euclid_space_fun, of "F"] gamma3_def and gamma3_smooth and g1_scale_j_contin and a_lt_b and add_scale_img Dg1_is_gamma3_pathimg and field_cont_on_gamma3_image by (auto simp: pathstart_def pathfinish_def i_is_x_axis) then show "(line_integral_exists F {i} gamma3)" by metis have gamma4_x_const: "\x. gamma4 x \ i = a" using gamma4_def by (auto simp add: real_inner_class.inner_add_left inner_not_same_Basis i_is_x_axis) have 3: "(line_integral F {i} gamma4) = 0" "(line_integral_exists F {i} gamma4)" using line_integral_on_pair_straight_path[OF gamma4_x_const] straight_path_diffrentiable_x gamma4_def by (auto simp add: mult.commute) then show "(line_integral_exists F {i} gamma4)" by metis have "continuous_on (cbox a b) (\x. F(x, g1(x)) \ i)" using line_is_pair_img and g1_path_continuous and field_cont_on_gamma3_image continuous_on_compose i_is_x_axis j_is_y_axis by auto then have 7: "(\x. F(x, g1(x)) \ i) integrable_on (cbox a b)" using integrable_continuous [of "a" "b" "(\x. F(x, g1(x)) \ i)"] by auto have partial_deriv_one_d_integrable: "((\y. ?F_b'(xc, y)) has_integral F(xc,g1(xc)) \ i - F(xc, g2(xc)) \ i) (cbox (g2 xc) (g1 xc))" if "xc \ cbox a b" for xc proof - have "{(xc', y). y \ cbox (g2 xc) (g1 xc) \ xc' = xc} \ Dy_pair" using that by (auto simp add: Dy_def) then show "((\y. ?F_b' (xc, y)) has_integral F(xc, g1 xc) \ i - F(xc, g2 xc) \ i) (cbox (g2 xc) (g1 xc))" using that and Base_vecs and F_partially_differentiable and Dy_def [symmetric] and g2_leq_g1 and fundamental_theorem_of_calculus_partial_vector [of "g2 xc" "g1 xc" "j" "i" "xc *\<^sub>R i" "Dy_pair" "F" "?F_b'" ] by (auto simp add: Groups.ab_semigroup_add_class.add.commute i_is_x_axis j_is_y_axis) qed have partial_deriv_integrable: "(?F_b') integrable_on Dy_pair" by (simp add: F_i_analytically_valid analytically_valid_imp_part_deriv_integrable_on) have 4: "integral Dy_pair ?F_b' = integral (cbox a b) (\x. integral (cbox (g2 x) (g1 x)) (\y. ?F_b' (x, y)))" proof - have x_axis_gauge_integrable: "\x. (\y. ?f(x,y)) integrable_on UNIV" proof - fix x::real have "\x. x \ cbox a b \ (\y. ?f (x, y)) = (\y. 0)" by (auto simp add: Dy_def) then have f_integrable_x_not_in_range: "\x. x \ cbox a b \ (\y. ?f (x, y)) integrable_on UNIV" by (simp add: integrable_0) let ?F_b'_oneD = "(\x. (\y. if y \ (cbox (g2 x) ( g1 x)) then ?F_b' (x,y) else 0))" have f_value_x_in_range: "\x \ cbox a b. ?F_b'_oneD x = (\y. ?f(x,y))" by (auto simp add: Dy_def) have "\x \ cbox a b. ?F_b'_oneD x integrable_on UNIV" using has_integral_integrable integrable_restrict_UNIV partial_deriv_one_d_integrable by blast then have f_integrable_x_in_range: "\x. x \ cbox a b \ (\y. ?f (x, y)) integrable_on UNIV" using f_value_x_in_range by auto show "(\y. ?f (x, y)) integrable_on UNIV" using f_integrable_x_not_in_range and f_integrable_x_in_range by auto qed have arg: "(\a. if a \ Dy_pair then partial_vector_derivative (\a. F a \ i) j a else 0) = (\x. if x \ Dy_pair then if x \ Dy_pair then partial_vector_derivative (\a. F a \ i) j x else 0 else 0)" by auto have arg2: "Dy_pair = {(x, y). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. g2 x \ i \ y \ i \ y \ i \ g1 x \ i)}" using Dy_def by auto have arg3: "\ x. x \ Dy_pair \ (\x. if x \ Dy_pair then partial_vector_derivative (\a. F a \ i) j x else 0) x = (\x. partial_vector_derivative (\a. F a \ i) j x) x" by auto have arg4: "\x. x \ (cbox a b) \ (\x. integral (cbox (g2 x) (g1 x)) (\y. if (x, y) \ Dy_pair then partial_vector_derivative (\a. F a \ i) j (x, y) else 0)) x = (\x. integral (cbox (g2 x) (g1 x)) (\y. partial_vector_derivative (\a. F a \ i) j (x, y))) x" apply (simp add: Dy_def) by (smt Henstock_Kurzweil_Integration.integral_cong atLeastAtMost_iff) show ?thesis using gauge_integral_Fubini_curve_bounded_region_x [OF f_lesbegue_integrable x_axis_gauge_integrable x_axis_integral_measurable arg arg2] Henstock_Kurzweil_Integration.integral_cong[OF arg3, of "Dy_pair" "(\x. x)"] Henstock_Kurzweil_Integration.integral_cong[OF arg4, of "cbox a b" "(\x. x)"] by auto qed have 5: "(integral Dy_pair (\a. (?F_b' a)) = integral (cbox a b) (\x. F(x, g1(x)) \ i - F(x, g2(x)) \ i))" using 4 Henstock_Kurzweil_Integration.integral_cong partial_deriv_one_d_integrable integrable_def by (smt integral_unique) show "(line_integral F {i} gamma1) + (line_integral F {i} gamma2) - (line_integral F {i} gamma3) - (line_integral F {i} gamma4) = (integral Dy_pair (\a. - (?F_b' a)))" using "0"(1) "1"(1) "2"(1) "3"(1) 5 "6" "7" by (simp add: Henstock_Kurzweil_Integration.integral_diff) qed theorem Greens_thm_type_II: fixes F:: "((real*real) \ (real * real))" and gamma4 gamma3 gamma2 gamma1 :: "(real \ (real * real))" and a:: "real" and b:: "real" and g1:: "(real \ real)" and g2:: "(real \ real)" assumes Dx_def: "Dx_pair = {(x::real,y) . y \ cbox a b \ x \ cbox (g2 y) (g1 y)}" and gamma4_def: "gamma4 = (\x. (g2(a + (b - a) * x), a + (b - a) * x))" and gamma4_smooth: "gamma4 piecewise_C1_differentiable_on {0..1}" and (*TODO: This should be piecewise smooth*) gamma3_def: "gamma3 = (\x. (g2(b) + x *\<^sub>R (g1(b) - g2(b)), b))" and gamma2_def: "gamma2 = (\x. (g1(a + (b - a) * x), a + (b - a) * x))" and gamma2_smooth: "gamma2 piecewise_C1_differentiable_on {0..1}" and gamma1_def: "gamma1 = (\x. (g2(a) + x *\<^sub>R (g1(a) - g2(a)), a))" and F_j_analytically_valid: "analytically_valid Dx_pair (\p. F(p) \ j) i" and g2_leq_g1: "\x \ cbox a b. (g2 x) \ (g1 x)" and (*This is needed otherwise what would Dy be?*) a_lt_b: "a < b" shows "-(line_integral F {j} gamma4) - (line_integral F {j} gamma3) + (line_integral F {j} gamma2) + (line_integral F {j} gamma1) = (integral Dx_pair (\a. (partial_vector_derivative (\a. (F a) \ j) i a)))" "line_integral_exists F {j} gamma4" "line_integral_exists F {j} gamma3" "line_integral_exists F {j} gamma2" "line_integral_exists F {j} gamma1" proof - let ?F_a' = "partial_vector_derivative (\a. (F a) \ j) i" have F_first_is_continuous: "continuous_on Dx_pair (\a. F(a) \ j)" using F_j_analytically_valid by (auto simp add: analytically_valid_def) let ?f = "(\x. if x \ (Dx_pair) then (partial_vector_derivative (\a. (F a) \ j) i) x else 0)" have f_lesbegue_integrable: "integrable lborel ?f" using F_j_analytically_valid by (auto simp add: analytically_valid_def indic_ident) have partially_vec_diff: "\a \ Dx_pair. partially_vector_differentiable (\a. (F a) \ j) i a" using F_j_analytically_valid by (auto simp add: analytically_valid_def indicator_def) have "\x y. ?F_a' (x,y) * indicator (Dx_pair) (x,y) = (\x. if x \ (Dx_pair) then ?F_a' x else 0) (x,y)" using indic_ident[of "?F_a'"] by auto then have y_axis_integral_measurable: "(\x. integral UNIV (\y. ?f(y, x))) \ borel_measurable lborel" using analytically_valid_y[OF F_j_analytically_valid] by (auto simp add: indicator_def) have F_partially_differentiable: "\a\Dx_pair. has_partial_vector_derivative (\x. (F x) \ j) i (?F_a' a) a" using partial_vector_derivative_works partially_vec_diff by fastforce have g1_g2_continuous: "continuous_on (cbox a b) g1" "continuous_on (cbox a b) g2" proof - have shift_scale_cont: "continuous_on {a..b} (\x. (x - a)*(1/(b-a)))" by (intro continuous_intros) have shift_scale_inv: "(\x. a + (b - a) * x) \ (\x. (x - a)*(1/(b-a))) = id" using a_lt_b by (auto simp add: o_def) have img_shift_scale: "(\x. (x - a)*(1/(b-a)))`{a..b} = {0..1}" apply (auto simp: divide_simps image_iff) apply (rule_tac x="x * (b - a) + a" in bexI) using a_lt_b by (auto simp: algebra_simps mult_le_cancel_left affine_ineq) have "continuous_on {0..1} (\x. g2(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma4_smooth], of "(\x. i)", OF continuous_on_const] by (simp add: gamma4_def i_is_x_axis) then have "continuous_on {a..b} ((\x. g2(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then show "continuous_on (cbox a b) g2" using a_lt_b by (simp add: shift_scale_inv) have "continuous_on {0..1} (\x. g1(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma2_smooth], of "(\x. i)", OF continuous_on_const] by (simp add: gamma2_def i_is_x_axis) then have "continuous_on {a..b} ((\x. g1(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then show "continuous_on (cbox a b) g1" using a_lt_b by (simp add: shift_scale_inv) qed have g2_scale_i_contin: "continuous_on (cbox a b) (\x. (g2 x, 0))" by (intro continuous_intros g1_g2_continuous) let ?Dg2 = "{p. \x. x \ cbox a b \ p = (g2(x), x)}" have line_is_pair_img: "?Dg2 = (\x. (g2(x), x) ) ` (cbox a b)" using image_def by auto have g2_path_continuous: "continuous_on (cbox a b) (\x. (g2(x), x))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma4_image: "continuous_on ?Dg2 (\a. F(a) \ j)" by (rule continuous_on_subset [OF F_first_is_continuous]) (auto simp: Dx_def g2_leq_g1) have gamma4_is_compos_of_scal_and_g2: "gamma4 = (\x. (g2(x), x)) \ (\x. a + (b - a) * x)" using gamma4_def by auto have add_scale_img: "(\x. a + (b - a) * x) ` {0 .. 1} = {a .. b}" using add_scale_img and a_lt_b by auto then have Dg2_is_gamma4_pathimg: "path_image gamma4 = ?Dg2" using line_is_pair_img and gamma4_is_compos_of_scal_and_g2 image_comp path_image_def by (metis (no_types, lifting) cbox_interval) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma4_as_euclid_space_fun: "gamma4 = (\x. (a + (b - a) * x) *\<^sub>R j + (g2 (a + (b - a) * x), 0))" using j_is_y_axis gamma4_def by auto have 0: "(line_integral F {j} gamma4) = integral (cbox a b) (\x. F(g2(x), x) \ j)" "line_integral_exists F {j} gamma4" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(2)] _ gamma4_as_euclid_space_fun] gamma4_def gamma4_smooth g2_scale_i_contin a_lt_b add_scale_img Dg2_is_gamma4_pathimg and field_cont_on_gamma4_image by (auto simp: pathstart_def pathfinish_def j_is_y_axis) then show "line_integral_exists F {j} gamma4" by metis have gamma3_y_const: "\x. gamma3 x \ j = b" by (simp add: gamma3_def j_is_y_axis) have 1: "(line_integral F {j} gamma3) = 0" "(line_integral_exists F {j} gamma3)" using line_integral_on_pair_straight_path[OF gamma3_y_const] straight_path_diffrentiable_y gamma3_def by (auto simp add: mult.commute) then show "line_integral_exists F {j} gamma3" by auto have "continuous_on (cbox a b) (\x. F(g2(x), x) \ j)" by (smt Collect_mono_iff continuous_on_compose2 continuous_on_eq field_cont_on_gamma4_image g2_path_continuous line_is_pair_img) then have 6: "(\x. F(g2(x), x) \ j) integrable_on (cbox a b)" using integrable_continuous by blast have g1_scale_i_contin: "continuous_on (cbox a b) (\x. (g1 x, 0))" by (intro continuous_intros g1_g2_continuous) let ?Dg1 = "{p. \x. x \ cbox a b \ p = (g1(x), x)}" have line_is_pair_img: "?Dg1 = (\x. (g1(x), x) ) ` (cbox a b)" using image_def by auto have g1_path_continuous: "continuous_on (cbox a b) (\x. (g1(x), x))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma2_image: "continuous_on ?Dg1 (\a. F(a) \ j)" by (rule continuous_on_subset [OF F_first_is_continuous]) (auto simp: Dx_def g2_leq_g1) have "gamma2 = (\x. (g1(x), x)) \ (\x. a + (b - a) * x)" using gamma2_def by auto then have Dg1_is_gamma2_pathimg: "path_image gamma2 = ?Dg1" using line_is_pair_img image_comp path_image_def add_scale_img by (metis (no_types, lifting) cbox_interval) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma2_as_euclid_space_fun: "gamma2 = (\x. (a + (b - a) * x) *\<^sub>R j + (g1 (a + (b - a) * x), 0))" using j_is_y_axis gamma2_def by auto have 2: "(line_integral F {j} gamma2) = integral (cbox a b) (\x. F(g1(x), x) \ j)" "(line_integral_exists F {j} gamma2)" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(2)] _ gamma2_as_euclid_space_fun] gamma2_def and gamma2_smooth and g1_scale_i_contin and a_lt_b and add_scale_img Dg1_is_gamma2_pathimg and field_cont_on_gamma2_image by (auto simp: pathstart_def pathfinish_def j_is_y_axis) then show "line_integral_exists F {j} gamma2" by metis have gamma1_y_const: "\x. gamma1 x \ j = a" using gamma1_def by (auto simp add: real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis) have 3: "(line_integral F {j} gamma1) = 0" "(line_integral_exists F {j} gamma1)" using line_integral_on_pair_straight_path[OF gamma1_y_const] straight_path_diffrentiable_y gamma1_def by (auto simp add: mult.commute) then show "line_integral_exists F {j} gamma1" by auto have "continuous_on (cbox a b) (\x. F(g1(x), x) \ j)" by (smt Collect_mono_iff continuous_on_compose2 continuous_on_eq field_cont_on_gamma2_image g1_path_continuous line_is_pair_img) then have 7: "(\x. F(g1(x), x) \ j) integrable_on (cbox a b)" using integrable_continuous [of "a" "b" "(\x. F(g1(x), x) \ j)"] by auto have partial_deriv_one_d_integrable: "((\y. ?F_a'(y, xc)) has_integral F(g1(xc), xc) \ j - F(g2(xc), xc) \ j) (cbox (g2 xc) (g1 xc))" if "xc \ cbox a b" for xc::real proof - have "{(y, xc'). y \ cbox (g2 xc) (g1 xc) \ xc' = xc} \ Dx_pair" using that by (auto simp add: Dx_def) then show ?thesis using that and Base_vecs and F_partially_differentiable and Dx_def [symmetric] and g2_leq_g1 and fundamental_theorem_of_calculus_partial_vector [of "g2 xc" "g1 xc" "i" "j" "xc *\<^sub>R j" "Dx_pair" "F" "?F_a'" ] by (auto simp add: Groups.ab_semigroup_add_class.add.commute i_is_x_axis j_is_y_axis) qed have "?f integrable_on UNIV" by (simp add: f_lesbegue_integrable integrable_on_lborel) then have partial_deriv_integrable: "?F_a' integrable_on Dx_pair" using integrable_restrict_UNIV by auto have 4: "integral Dx_pair ?F_a' = integral (cbox a b) (\x. integral (cbox (g2 x) (g1 x)) (\y. ?F_a' (y, x)))" proof - have y_axis_gauge_integrable: "(\y. ?f(y, x)) integrable_on UNIV" for x proof - let ?F_a'_oneD = "(\x. (\y. if y \ (cbox (g2 x) ( g1 x)) then ?F_a' (y, x) else 0))" have "\x. x \ cbox a b \ (\y. ?f (y, x)) = (\y. 0)" by (auto simp add: Dx_def) then have f_integrable_x_not_in_range: "\x. x \ cbox a b \ (\y. ?f (y, x)) integrable_on UNIV" by (simp add: integrable_0) have "\x \ cbox a b. ?F_a'_oneD x = (\y. ?f(y, x))" using g2_leq_g1 by (auto simp add: Dx_def) moreover have "\x \ cbox a b. ?F_a'_oneD x integrable_on UNIV" using has_integral_integrable integrable_restrict_UNIV partial_deriv_one_d_integrable by blast ultimately have "\x. x \ cbox a b \ (\y. ?f (y, x)) integrable_on UNIV" by auto then show "(\y. ?f (y, x)) integrable_on UNIV" using f_integrable_x_not_in_range by auto qed have arg: "(\a. if a \ Dx_pair then partial_vector_derivative (\a. F a \ j) i a else 0) = (\x. if x \ Dx_pair then if x \ Dx_pair then partial_vector_derivative (\a. F a \ j) i x else 0 else 0)" by auto have arg2: "Dx_pair = {(y, x). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. g2 x \ i \ y \ i \ y \ i \ g1 x \ i)}" using Dx_def by auto have arg3: "\ x. x \ Dx_pair \ (\x. if x \ Dx_pair then partial_vector_derivative (\a. F a \ j) i x else 0) x = (\x. partial_vector_derivative (\a. F a \ j) i x) x" by auto have arg4: "\x. x \ (cbox a b) \ (\x. integral (cbox (g2 x) (g1 x)) (\y. if (y, x) \ Dx_pair then partial_vector_derivative (\a. F a \ j) i (y, x) else 0)) x = (\x. integral (cbox (g2 x) (g1 x)) (\y. partial_vector_derivative (\a. F a \ j) i (y, x))) x" apply (clarsimp simp: Dx_def) by (smt Henstock_Kurzweil_Integration.integral_cong atLeastAtMost_iff) show ?thesis using gauge_integral_Fubini_curve_bounded_region_y [OF f_lesbegue_integrable y_axis_gauge_integrable y_axis_integral_measurable arg arg2] Henstock_Kurzweil_Integration.integral_cong[OF arg3, of "Dx_pair" "(\x. x)"] Henstock_Kurzweil_Integration.integral_cong[OF arg4, of "cbox a b" "(\x. x)"] by auto qed have "((integral Dx_pair (\a. (?F_a' a))) = integral (cbox a b) (\x. F(g1(x), x) \ j - F(g2(x), x) \ j))" using 4 Henstock_Kurzweil_Integration.integral_cong partial_deriv_one_d_integrable integrable_def by (smt integral_unique) then have "integral Dx_pair (\a. - (?F_a' a)) = - integral (cbox a b) (\x. F(g1(x), x) \ j - F(g2(x), x) \ j)" using partial_deriv_integrable and integral_neg by auto then have 5: "integral Dx_pair (\a. - (?F_a' a)) = integral (cbox a b) (\x. ( F(g2(x), x) \ j - F(g1(x), x) \ j))" using 6 7 and integral_neg [of _ "(\x. F(g1 x, x) \ j - F(g2 x, x) \ j)"] by auto have "(line_integral F {j} gamma4) + (line_integral F {j} gamma3) - (line_integral F {j} gamma2) - (line_integral F {j} gamma1) = (integral Dx_pair (\a. -(?F_a' a)))" using 0 1 2 3 5 6 7 Henstock_Kurzweil_Integration.integral_diff[of "(\x. F(g2(x), x) \ j)" "cbox a b" "(\x. F(g1(x), x) \ j)"] by (auto) then show "-(line_integral F {j} gamma4) - (line_integral F {j} gamma3) + (line_integral F {j} gamma2) + (line_integral F {j} gamma1) = (integral Dx_pair (\a. (?F_a' a)))" by (simp add: \integral Dx_pair (\a. - ?F_a' a) = - integral (cbox a b) (\x. F(g1 x, x) \ j - F(g2 x, x) \ j)\ \integral Dx_pair ?F_a' = integral (cbox a b) (\x. F(g1 x, x) \ j - F(g2 x, x) \ j)\) qed end locale green_typeII_cube = R2 + fixes twoC F assumes two_cube: "typeII_twoCube twoC" and valid_two_cube: "valid_two_cube twoC" and f_analytically_valid: "analytically_valid (cubeImage twoC) (\x. (F x) \ j) i" begin lemma GreenThm_typeII_twoCube: shows "integral (cubeImage twoC) (\a. partial_vector_derivative (\x. (F x) \ j) i a) = one_chain_line_integral F {j} (boundary twoC)" "\(k,\) \ boundary twoC. line_integral_exists F {j} \" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" have line_integral_around_boundary: "one_chain_line_integral F {j} (boundary twoC) = line_integral F {j} ?bottom_edge + line_integral F {j} ?right_edge - line_integral F {j} ?top_edge - line_integral F {j} ?left_edge" proof (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite1: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}" by auto then have sum_step1: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. k * line_integral F {j} g) = line_integral F {j} (\x. twoC (x, 0)) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}. k * line_integral F {j} g)" using sum.insert_remove [OF finite1] using valid_two_cube apply (simp only: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def) by (auto simp: card_insert_if split: if_split_asm) have three_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))} = 3" using valid_two_cube apply (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def) by (auto simp: card_insert_if split: if_split_asm) have finite2: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))}" by auto then have sum_step2: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (-1, \x. twoC (x, 1))}. k * line_integral F {j} g) = (- line_integral F {j} (\x. twoC (x, 1))) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {j} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) have two_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))} = 2" using three_distinct_edges by (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite3: "finite {(- 1::int, \y. twoC (0, y))}" by auto then have sum_step3: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {j} g) = ( line_integral F {j} (\y. twoC (1, y))) + (\(k, g)\{(- (1::real), \y. twoC (0, y))}. k * line_integral F {j} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) show "(\x\{(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. case x of (k, g) \ k * line_integral F {j} g) = line_integral F {j} (\x. twoC (x, 0)) + line_integral F {j} (\y. twoC (1, y)) - line_integral F {j} (\x. twoC (x, 1)) - line_integral F {j} (\y. twoC (0, y))" using sum_step1 sum_step2 sum_step3 by auto qed obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" using two_cube and typeII_twoCubeImg[of"twoC"] by auto have left_edge_explicit: "?left_edge = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" by (simp add: twoCisTypeII(4) algebra_simps) have left_edge_smooth: "?left_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_smth: "(\x. (a + (b - a) * x)) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by blast have g2_smooth: "g2 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(6) by auto have "(\x. g2 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using Cauchy_Integral_Theorem.piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using left_edge_explicit by auto qed have top_edge_explicit: "?top_edge = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" and right_edge_explicit: "?right_edge = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" by (simp_all add: twoCisTypeII(4) algebra_simps) have right_edge_smooth: "?right_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise [OF scale_shift_smooth] by auto have g1_smooth: "g1 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(5) by auto have "(\x. g1 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using right_edge_explicit by auto qed have bottom_edge_explicit: "?bottom_edge = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" by (simp add: twoCisTypeII(4) algebra_simps) show "integral (cubeImage twoC) (\a. partial_vector_derivative (\x. (F x) \ j) i a) = one_chain_line_integral F {j} (boundary twoC)" using Greens_thm_type_II[OF twoCisTypeII(3) left_edge_explicit left_edge_smooth top_edge_explicit right_edge_explicit right_edge_smooth bottom_edge_explicit f_analytically_valid twoCisTypeII(2) twoCisTypeII(1)] line_integral_around_boundary by auto have "line_integral_exists F {j} \" if "(k,\) \ boundary twoC" for k \ proof - have "line_integral_exists F {j} (\y. twoC (0, y))" "line_integral_exists F {j} (\x. twoC (x, 1))" "line_integral_exists F {j} (\y. twoC (1, y))" "line_integral_exists F {j} (\x. twoC (x, 0))" using Greens_thm_type_II[OF twoCisTypeII(3) left_edge_explicit left_edge_smooth top_edge_explicit right_edge_explicit right_edge_smooth bottom_edge_explicit f_analytically_valid twoCisTypeII(2) twoCisTypeII(1)] line_integral_around_boundary by auto then show "line_integral_exists F {j} \" using that by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then show "\(k,\) \ boundary twoC. line_integral_exists F {j} \" by auto qed lemma line_integral_exists_on_typeII_Cube_boundaries': assumes "(k,\) \ boundary twoC" shows "line_integral_exists F {j} \" using assms GreenThm_typeII_twoCube(2) by blast end locale green_typeII_chain = R2 + fixes F two_chain s assumes valid_typeII_div: "valid_typeII_division s two_chain" and F_anal_valid: "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\x. (F x) \ j) i" begin lemma two_chain_valid_valid_cubes: "\two_cube \ two_chain. valid_two_cube two_cube" using valid_typeII_div by (auto simp add: valid_two_chain_def) lemma typeII_chain_line_integral_exists_boundary': shows "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using green_typeII_cube.line_integral_exists_on_typeII_Cube_boundaries'[of i j] F_anal_valid valid_typeII_div apply(auto simp add: two_chain_boundary_def) using R2_axioms green_typeII_cube_axioms_def green_typeII_cube_def two_chain_valid_valid_cubes by blast then show integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" by (simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) qed lemma typeII_chain_line_integral_exists_boundary'': "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {j} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using green_typeII_cube.line_integral_exists_on_typeII_Cube_boundaries'[of i j] valid_typeII_div apply (simp add: two_chain_boundary_def boundary_def) using F_anal_valid R2_axioms green_typeII_cube_axioms_def green_typeII_cube_def two_chain_valid_valid_cubes by fastforce then show integ_exis_vert: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {j} \" by (simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) qed lemma typeII_cube_line_integral_exists_boundary: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using valid_typeII_div typeII_chain_line_integral_exists_boundary' typeII_chain_line_integral_exists_boundary'' apply (auto simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def) using boundary_def by auto lemma type_II_chain_horiz_bound_valid: "\(k,\) \ two_chain_horizontal_boundary two_chain. valid_path \" using valid_typeII_div typeII_edges_are_valid_paths by (force simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) lemma type_II_chain_vert_bound_valid: (*This and the previous one need to be used in all proofs*) "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using typeII_edges_are_valid_paths valid_typeII_div by (force simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) lemma members_of_only_horiz_div_line_integrable': assumes "only_horizontal_division one_chain two_chain" "(k::int, \)\one_chain" "(k::int, \)\one_chain" "finite two_chain" "\two_cube \ two_chain. valid_two_cube two_cube" shows "line_integral_exists F {j} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using typeII_cube_line_integral_exists_boundary by blast have integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" using typeII_chain_line_integral_exists_boundary' assms by auto have integ_exis_horiz: "(\k \. (\(k', \') \ two_chain_horizontal_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {j} \)" using integ_exis type_II_chain_horiz_bound_valid using line_integral_exists_subpath[of "F" "{j}"] by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def boundary_def) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_horizontal_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "one_chain = \ \ \" "((common_sudiv_exists (two_chain_vertical_boundary two_chain) \) \ common_reparam_exists \ (two_chain_vertical_boundary two_chain))" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using assms(1) unfolding only_horizontal_division_def by blast have finite_j: "finite {j}" by auto show "line_integral_exists F {j} \" proof(cases "common_sudiv_exists (two_chain_vertical_boundary two_chain) \") case True show ?thesis using gen_common_subdivision_imp_eq_line_integral(2)[OF True two_chain_vertical_boundary_is_boundary_chain hv_props(6) integ_exis_vert finite_two_chain_vertical_boundary[OF assms(4)] hv_props(5) finite_j] integ_exis_horiz[of "\"] assms(3) case_prod_conv hv_props(2) hv_props(3) by fastforce next case False have i: "{j} \ Basis" using j_is_y_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_vertical_boundary two_chain. \b\{j}. continuous_on (path_image \2) (\x. F x \ b)" using F_anal_valid field_cont_on_typeII_region_cont_on_edges valid_typeII_div by (fastforce simp add: analytically_valid_def two_chain_vertical_boundary_def boundary_def path_image_def) show "line_integral_exists F {j} \" using common_reparam_exists_imp_eq_line_integral(2)[OF finite_j hv_props(5) finite_two_chain_vertical_boundary[OF assms(4)] hv_props(6) two_chain_vertical_boundary_is_boundary_chain ii _ hv_props(7) type_II_chain_vert_bound_valid] integ_exis_horiz[of "\"] assms(3) hv_props False by fastforce qed qed lemma GreenThm_typeII_twoChain: shows "two_chain_integral two_chain (partial_vector_derivative (\a. (F a) \ j) i) = one_chain_line_integral F {j} (two_chain_boundary two_chain)" proof (simp add: two_chain_boundary_def one_chain_line_integral_def two_chain_integral_def) let ?F_a' = "partial_vector_derivative (\a. (F a) \ j) i" have "(\(k,g)\ boundary twoCube. k * line_integral F {j} g) = integral (cubeImage twoCube) (\a. ?F_a' a)" if "twoCube \ two_chain" for twoCube using green_typeII_cube.GreenThm_typeII_twoCube(1) valid_typeII_div F_anal_valid one_chain_line_integral_def valid_two_chain_def by (simp add: R2_axioms green_typeII_cube_axioms_def green_typeII_cube_def that) then have double_sum_eq_sum: "(\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {j} g)) = (\twoCube\(two_chain). integral (cubeImage twoCube) (\a. ?F_a' a))" using Finite_Cartesian_Product.sum_cong_aux by auto have pairwise_disjoint_boundaries: "\x\ (boundary ` two_chain). (\y\ (boundary ` two_chain). (x \ y \ (x \ y = {})))" using valid_typeII_div by (fastforce simp add: image_def valid_two_chain_def pairwise_def) have finite_boundaries: "\B \ (boundary` two_chain). finite B" using valid_typeII_div image_iff by (fastforce simp add: valid_two_cube_def valid_two_chain_def) have boundary_inj: "inj_on boundary two_chain" using valid_typeII_div by (force simp add: valid_two_cube_def valid_two_chain_def pairwise_def inj_on_def) have "(\x\(\x\two_chain. boundary x). case x of (k, g) \ k * line_integral F {j} g) = (\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {j} g))" using sum.reindex[OF boundary_inj, of "\x. (\(k,g)\ x. k * line_integral F {j} g)"] using sum.Union_disjoint[OF finite_boundaries pairwise_disjoint_boundaries, of "\x. case x of (k, g) \ (k::int) * line_integral F {j} g"] by auto then show "(\C\two_chain. integral (cubeImage C) (\a. ?F_a' a)) = (\x\(\x\two_chain. boundary x). case x of (k, g) \ k * line_integral F {j} g)" using double_sum_eq_sum by auto qed lemma GreenThm_typeII_divisible: assumes gen_division: "gen_division s (cubeImage ` two_chain)" (*This should follow from the assumption that images are not negligible*) shows "integral s (partial_vector_derivative (\x. (F x) \ j) i) = one_chain_line_integral F {j} (two_chain_boundary two_chain)" proof - let ?F_a' = "(partial_vector_derivative (\x. (F x) \ j) i)" have "integral s (\x. ?F_a' x) = two_chain_integral two_chain (\a. ?F_a' a)" proof (simp add: two_chain_integral_def) have partial_deriv_integrable: "(?F_a' has_integral (integral (cubeImage twoCube) ?F_a')) (cubeImage twoCube)" if "twoCube \ two_chain" for twoCube by (simp add: analytically_valid_imp_part_deriv_integrable_on F_anal_valid integrable_integral that) then have partial_deriv_integrable: "\twoCubeImg. twoCubeImg \ cubeImage ` two_chain \ ((\x. ?F_a' x) has_integral (integral (twoCubeImg) (\x. ?F_a' x))) (twoCubeImg)" using integrable_neg by force have finite_images: "finite (cubeImage ` two_chain)" using gen_division gen_division_def by auto have negligible_images: "pairwise (\S S'. negligible (S \ S')) (cubeImage ` two_chain)" using gen_division by (auto simp add: gen_division_def pairwise_def) have "inj_on cubeImage two_chain" using valid_typeII_div valid_two_chain_def by auto then have "(\twoCubeImg\cubeImage ` two_chain. integral twoCubeImg (\x. ?F_a' x)) = (\C\two_chain. integral (cubeImage C) (\a. ?F_a' a))" using sum.reindex by auto then show "integral s (\x. ?F_a' x) = (\C\two_chain. integral (cubeImage C) (\a. ?F_a' a))" using has_integral_Union[OF finite_images partial_deriv_integrable negligible_images] gen_division by (auto simp add: gen_division_def) qed then show ?thesis using GreenThm_typeII_twoChain F_anal_valid by auto qed lemma GreenThm_typeII_divisible_region_boundary_gen: assumes only_horizontal_division: "only_horizontal_division \ two_chain" shows "integral s (partial_vector_derivative (\x. (F x) \ j) i) = one_chain_line_integral F {j} \" proof - let ?F_a' = "(partial_vector_derivative (\x. (F x) \ j) i)" (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have horiz_line_integral_zero: "one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if "oneCube \ two_chain_horizontal_boundary(two_chain)" for oneCube proof - from that obtain x y1 y2 k where horiz_edge_def: "oneCube = (k, (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) let ?horiz_edge = "(snd oneCube)" have horiz_edge_y_const: "\t. (?horiz_edge t) \ j = x" by (auto simp add: horiz_edge_def real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis) have horiz_edge_is_straight_path: "?horiz_edge = (\t. (y2 + t * (y1 - y2), x))" by (auto simp: horiz_edge_def algebra_simps) have "\x. ?horiz_edge differentiable at x" using horiz_edge_is_straight_path straight_path_diffrentiable_y by (metis mult_commute_abs) then show "line_integral F {j} (snd oneCube) = 0" using line_integral_on_pair_straight_path(1) j_is_y_axis real_pair_basis horiz_edge_y_const by blast qed then show "(\x\two_chain_horizontal_boundary two_chain. case x of (k, g) \ k * line_integral F {j} g) = 0" by (force intro: sum.neutral) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain" using valid_typeII_div finite_image_iff gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed have boundary_is_vert_hor: "two_chain_boundary two_chain = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by(auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \z. x (z, 0)), (- 1, \z. x (z, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \z. twoCube2 (z, 0)), (- 1, \z. twoCube2 (z, 1))} = {}" if "twoCube\two_chain" "twoCube2\two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have card_4: "card {(- 1, \y. twoCube2 (0::real, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using valid_typeII_div valid_two_chain_def that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def valid_two_cube_def) show ?thesis using card_4 by (auto simp: True card_insert_if split: if_split_asm) next case False show ?thesis using valid_typeII_div valid_two_chain_def by (simp add: boundary_def vertical_boundary_def horizontal_boundary_def pairwise_def \twoCube \ twoCube2\ that) qed then have "\ ((\twoCube. {(- 1, \y. twoCube (0::real, y)), (1::real, \y. twoCube (1::real, y))}) ` two_chain) \ \ ((\twoCube. {(1::int, \z. twoCube (z, 0::real)), (- 1, \z. twoCube (z, 1::real))}) ` two_chain) = {}" by (fastforce simp add: Union_disjoint) then show ?thesis by force qed qed have "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (simp add: hor_vert_finite sum.union_disjoint one_chain_line_integral_def) then have y_axis_line_integral_is_only_vertical: "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using horiz_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>H(2-\) such that \ = \ \ \\<^sub>v(2-\)*) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_horizontal_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "\ = \ \ \" "(common_sudiv_exists (two_chain_vertical_boundary two_chain) \ \ common_reparam_exists \ (two_chain_vertical_boundary two_chain))" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using only_horizontal_division by(fastforce simp add: only_horizontal_division_def) have "finite {j}" by auto then have eq_integrals: " one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" proof(cases "common_sudiv_exists (two_chain_vertical_boundary two_chain) \") case True then show ?thesis using gen_common_subdivision_imp_eq_line_integral(1)[OF True two_chain_vertical_boundary_is_boundary_chain hv_props(6) _ hor_vert_finite(1) hv_props(5)] typeII_chain_line_integral_exists_boundary' by force next case False have integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" using typeII_chain_line_integral_exists_boundary' assms by (fastforce simp add: valid_two_chain_def) have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using typeII_cube_line_integral_exists_boundary by blast have valid_paths: "\(k,\) \ two_chain_horizontal_boundary two_chain. valid_path \" using typeII_edges_are_valid_paths valid_typeII_div by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) have integ_exis_horiz: "(\k \. (\(k', \')\two_chain_horizontal_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {j} \)" using integ_exis valid_paths line_integral_exists_subpath[of "F" "{j}"] by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def boundary_def) have finite_j: "finite {j}" by auto have i: "{j} \ Basis" using j_is_y_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_vertical_boundary two_chain. \b\{j}. continuous_on (path_image \2) (\x. F x \ b)" using valid_typeII_div field_cont_on_typeII_region_cont_on_edges F_anal_valid by (fastforce simp add: analytically_valid_def two_chain_vertical_boundary_def boundary_def path_image_def) show "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using hv_props(4) False common_reparam_exists_imp_eq_line_integral(1)[OF finite_j hv_props(5) hor_vert_finite(1) hv_props(6) two_chain_vertical_boundary_is_boundary_chain ii _ hv_props(7) type_II_chain_vert_bound_valid] by fastforce qed (*Since \ \ \\<^sub>H(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have line_integral_on_path: "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" proof (simp only: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if oneCube: "oneCube \ \" for oneCube proof - obtain k \ where k_gamma: "(k,\) = oneCube" using prod.collapse by blast then obtain k' \' a b where kp_gammap: "(k'::int, \') \ two_chain_horizontal_boundary two_chain" "a \ {0 .. 1}" "b \ {0..1}" "subpath a b \' = \" using hv_props oneCube by (smt case_prodE split_conv) obtain x y1 y2 where horiz_edge_def: "(k',\') = (k', (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div kp_gammap by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) have horiz_edge_y_const: "\t. \ (t) \ j = x" using horiz_edge_def kp_gammap(4) by (auto simp add: real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis subpath_def) have horiz_edge_is_straight_path: "\ = (\t::real. ((1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t, x::real))" proof - fix t::real have "(1 - (b - a)*t + a) * (y2) + ((b-a)*t + a) * y1 = (1 - (b - a)*t + a) * (y2) + ((b-a)*t + a) * y1" by auto then have "\ = (\t::real. ((1 - (b - a)*t - a) * (y2) + ((b-a)*t + a) * y1, x::real))" using horiz_edge_def Product_Type.snd_conv Product_Type.fst_conv kp_gammap(4) by (simp add: subpath_def diff_diff_eq[symmetric]) also have "\ = (\t::real. ((1*y2 - (b - a)*y2*t - a*y2) + ((b-a)*y1*t + a*y1), x::real))" by(auto simp add: ring_class.ring_distribs(2) Groups.diff_diff_eq left_diff_distrib) also have "... = (\t::real. ((1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t, x::real))" by (force simp add: left_diff_distrib) finally show "\ = (\t::real. ((1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t, x::real))" . qed show "line_integral F {j} (snd oneCube) = 0" proof - have "\x. \ differentiable at x" by (simp add: horiz_edge_is_straight_path straight_path_diffrentiable_y) then have "line_integral F {j} \ = 0" by (simp add: horiz_edge_y_const line_integral_on_pair_straight_path(1)) then show ?thesis using Product_Type.snd_conv k_gamma by auto qed qed then have "\x\\. (case x of (k, g) \ (k::int) * line_integral F {j} g) = 0" by auto then show "(\x\\. case x of (k, g) \ real_of_int k * line_integral F {j} g) = (\x\two_chain_vertical_boundary two_chain. case x of (k, g) \ real_of_int k * line_integral F {j} g)" using hv_props(1) hv_props(3) hv_props(5) sum_zero_set hor_vert_finite(1) eq_integrals by (clarsimp simp add: one_chain_line_integral_def sum_zero_set) qed then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using line_integral_on_path by auto then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_boundary two_chain)" using y_axis_line_integral_is_only_vertical by auto then show ?thesis using valid_typeII_div GreenThm_typeII_divisible by auto qed lemma GreenThm_typeII_divisible_region_boundary: assumes two_cubes_trace_vertical_boundaries: "two_chain_vertical_boundary two_chain \ \" and boundary_of_region_is_subset_of_partition_boundary: "\ \ two_chain_boundary two_chain" shows "integral s (partial_vector_derivative (\x. (F x) \ j) i) = one_chain_line_integral F {j} \" proof - let ?F_a' = "(partial_vector_derivative (\x. (F x) \ j) i)" (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have horiz_line_integral_zero: "one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if one: "oneCube \ two_chain_horizontal_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where horiz_edge_def: "oneCube = (k, (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div one by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) let ?horiz_edge = "(snd oneCube)" have horiz_edge_y_const: "\t. (?horiz_edge t) \ j = x" using horiz_edge_def by (auto simp add: real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis) have horiz_edge_is_straight_path: "?horiz_edge = (\t. (y2 + t * (y1 - y2), x))" by (simp add: add_diff_eq diff_add_eq mult.commute right_diff_distrib horiz_edge_def) show "line_integral F {j} (snd oneCube) = 0" by (metis horiz_edge_is_straight_path horiz_edge_y_const line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_y) qed then show "(\x\two_chain_horizontal_boundary two_chain. case x of (k, g) \ k * line_integral F {j} g) = 0" by (simp add: prod.case_eq_if) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain" using valid_typeII_div finite_image_iff gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed have boundary_is_vert_hor: "two_chain_boundary two_chain = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by(auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \z. x (z, 0)), (- 1, \z. x (z, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \y. twoCube2 (y, 0)), (- 1, \y. twoCube2 (y, 1))} = {}" if "twoCube\two_chain" "twoCube2\two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have "card {(- 1, \y. twoCube2 (0::real, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using valid_typeII_div valid_two_chain_def that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def valid_two_cube_def) then show ?thesis by (auto simp: True card_insert_if split: if_split_asm) next case False show ?thesis using valid_typeII_div valid_two_chain_def by (simp add: boundary_def vertical_boundary_def horizontal_boundary_def pairwise_def \twoCube \ twoCube2\ that(1) that(2)) qed then have "\ ((\twoCube. {(- 1, \y. twoCube (0::real, y)), (1::real, \y. twoCube (1::real, y))}) ` two_chain) \ \ ((\twoCube. {(1::int, \z. twoCube (z, 0::real)), (- 1, \z. twoCube (z, 1::real))}) ` two_chain) = {}" by (force simp add: Complete_Lattices.Union_disjoint) then show ?thesis by force qed qed have "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (auto simp add: one_chain_line_integral_def hor_vert_finite sum.union_disjoint) then have y_axis_line_integral_is_only_vertical: "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using horiz_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>H(2-\) such that \ = \ \ \\<^sub>v(2-\)*) have "\\. \ \ (two_chain_horizontal_boundary two_chain) \ \ = \ \ (two_chain_vertical_boundary two_chain)" proof let ?\ = "\ - (two_chain_vertical_boundary two_chain)" show "?\ \ two_chain_horizontal_boundary two_chain \ \ = ?\ \ two_chain_vertical_boundary two_chain" using two_cubes_trace_vertical_boundaries boundary_of_region_is_subset_of_partition_boundary boundary_is_vert_hor by blast qed then obtain \ where h_props: "\ \ (two_chain_horizontal_boundary two_chain)" "\ = \ \ (two_chain_vertical_boundary two_chain)" by auto have h_vert_disj: "\ \ (two_chain_vertical_boundary two_chain) = {}" using horiz_verti_disjoint h_props(1) by auto have h_finite: "finite \" using hor_vert_finite h_props(1) Finite_Set.rev_finite_subset by force have line_integral_on_path: "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} \ + one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" by (auto simp add: one_chain_line_integral_def h_props sum.union_disjoint[OF h_finite hor_vert_finite(1) h_vert_disj]) (*Since \ \ \\<^sub>H(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have "one_chain_line_integral F {j} \ = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_horizontal_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where horiz_edge_def: "oneCube = (k, (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div oneCube by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) let ?horiz_edge = "(snd oneCube)" have horiz_edge_y_const: "\t. (?horiz_edge t) \ j = x" by (simp add: j_is_y_axis horiz_edge_def) have horiz_edge_is_straight_path: "?horiz_edge = (\t. (y2 + t * (y1 - y2), x))" by (simp add: add_diff_eq diff_add_eq mult.commute right_diff_distrib horiz_edge_def) show "line_integral F {j} (snd oneCube) = 0" by (simp add: horiz_edge_is_straight_path j_is_y_axis line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_y) qed then have "\oneCube \ \. line_integral F {j} (snd oneCube) = 0" using h_props by auto then have "\x\\. (case x of (k, g) \ (k::int) * line_integral F {j} g) = 0" by auto then show "(\x\\. case x of (k, g) \ k * line_integral F {j} g) = 0" - using sum_not_0 by blast + by simp qed then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using line_integral_on_path by auto then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_boundary two_chain)" using y_axis_line_integral_is_only_vertical by auto then show ?thesis using valid_typeII_div GreenThm_typeII_divisible by auto qed end locale green_typeI_cube = R2 + fixes twoC F assumes two_cube: "typeI_twoCube twoC" and valid_two_cube: "valid_two_cube twoC" and f_analytically_valid: "analytically_valid (cubeImage twoC) (\x. (F x) \ i) j" begin lemma GreenThm_typeI_twoCube: shows "integral (cubeImage twoC) (\a. - partial_vector_derivative (\p. F p \ i) j a) = one_chain_line_integral F {i} (boundary twoC)" "\(k,\) \ boundary twoC. line_integral_exists F {i} \" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" have line_integral_around_boundary: "one_chain_line_integral F {i} (boundary twoC) = line_integral F {i} ?bottom_edge + line_integral F {i} ?right_edge - line_integral F {i} ?top_edge - line_integral F {i} ?left_edge" proof (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite1: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}" by auto have sum_step1: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. k * line_integral F {i} g) = line_integral F {i} (\x. twoC (x, 0)) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}. k * line_integral F {i} g)" using sum.insert_remove [OF finite1] valid_two_cube by (auto simp: horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def card_insert_if split: if_split_asm) have three_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))} = 3" using valid_two_cube apply (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def) by (auto simp: card_insert_if split: if_split_asm) have finite2: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))}" by auto have sum_step2: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (-1, \x. twoC (x, 1))}. k * line_integral F {i} g) = (- line_integral F {i} (\x. twoC (x, 1))) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {i} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) have two_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))} = 2" using three_distinct_edges by (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite3: "finite {(- 1::int, \y. twoC (0, y))}" by auto have sum_step3: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {i} g) = (line_integral F {i} (\y. twoC (1, y))) + (\(k, g)\{(- (1::real), \y. twoC (0, y))}. k * line_integral F {i} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) show "(\x\{(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. case x of (k, g) \ k * line_integral F {i} g) = line_integral F {i} (\x. twoC (x, 0)) + line_integral F {i} (\y. twoC (1, y)) - line_integral F {i} (\x. twoC (x, 1)) - line_integral F {i} (\y. twoC (0, y))" using sum_step1 sum_step2 sum_step3 by auto qed obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" "(\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" "(\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" "(\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" using two_cube and typeI_cube_explicit_spec[of"twoC"] by auto have bottom_edge_smooth: "(\x. twoC (x, 0)) piecewise_C1_differentiable_on {0..1}" using typeI_twoCube_smooth_edges two_cube boundary_def vertical_boundary_def horizontal_boundary_def by auto have top_edge_smooth: "?top_edge piecewise_C1_differentiable_on {0..1}" using typeI_twoCube_smooth_edges two_cube boundary_def vertical_boundary_def horizontal_boundary_def by auto show "integral (cubeImage twoC) (\a. - partial_vector_derivative (\p. F p \ i) j a) = one_chain_line_integral F {i} (boundary twoC)" using Greens_thm_type_I[OF twoCisTypeI(3) twoCisTypeI(7) bottom_edge_smooth twoCisTypeI(8) twoCisTypeI(9) top_edge_smooth twoCisTypeI(10) f_analytically_valid twoCisTypeI(2) twoCisTypeI(1)] line_integral_around_boundary by auto have "line_integral_exists F {i} (\y. twoC (0, y))" "line_integral_exists F {i} (\x. twoC (x, 1))" "line_integral_exists F {i} (\y. twoC (1, y))" "line_integral_exists F {i} (\x. twoC (x, 0))" using Greens_thm_type_I[OF twoCisTypeI(3) twoCisTypeI(7) bottom_edge_smooth twoCisTypeI(8) twoCisTypeI(9) top_edge_smooth twoCisTypeI(10) f_analytically_valid twoCisTypeI(2) twoCisTypeI(1)] line_integral_around_boundary by auto then have "line_integral_exists F {i} \" if "(k,\) \ boundary twoC" for k \ using that by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) then show "\(k,\) \ boundary twoC. line_integral_exists F {i} \" by auto qed lemma line_integral_exists_on_typeI_Cube_boundaries': assumes "(k,\) \ boundary twoC" shows "line_integral_exists F {i} \" using assms two_cube valid_two_cube f_analytically_valid GreenThm_typeI_twoCube(2) by blast end locale green_typeI_chain = R2 + fixes F two_chain s assumes valid_typeI_div: "valid_typeI_division s two_chain" and F_anal_valid: "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\x. (F x) \ i) j" begin lemma two_chain_valid_valid_cubes: "\two_cube \ two_chain. valid_two_cube two_cube" using valid_typeI_div by (auto simp add: valid_two_chain_def) lemma typeI_cube_line_integral_exists_boundary': assumes "\two_cube \ two_chain. typeI_twoCube two_cube" assumes "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\x. (F x) \ i) j" assumes "\two_cube \ two_chain. valid_two_cube two_cube" shows "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {i} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using green_typeI_cube.line_integral_exists_on_typeI_Cube_boundaries'[of i j] assms using R2_axioms green_typeI_cube_axioms_def green_typeI_cube_def two_chain_boundary_def by fastforce then show integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {i} \" by (simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) qed lemma typeI_cube_line_integral_exists_boundary'': "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using green_typeI_cube.line_integral_exists_on_typeI_Cube_boundaries'[of i j] valid_typeI_div apply (simp add: two_chain_boundary_def boundary_def) using F_anal_valid R2_axioms green_typeI_cube_axioms_def green_typeI_cube_def two_chain_valid_valid_cubes by fastforce then show integ_exis_horiz: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" by (simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) qed lemma typeI_cube_line_integral_exists_boundary: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary' typeI_cube_line_integral_exists_boundary'' apply (auto simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def) by (meson R2_axioms green_typeI_chain.F_anal_valid green_typeI_chain_axioms green_typeI_cube.line_integral_exists_on_typeI_Cube_boundaries' green_typeI_cube_axioms_def green_typeI_cube_def two_chain_valid_valid_cubes valid_typeI_div) lemma type_I_chain_horiz_bound_valid: "\(k,\) \ two_chain_horizontal_boundary two_chain. valid_path \" using typeI_edges_are_valid_paths valid_typeI_div by (force simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) lemma type_I_chain_vert_bound_valid: (*This and the previous one need to be used in all proofs*) assumes "\two_cube \ two_chain. typeI_twoCube two_cube" shows "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using typeI_edges_are_valid_paths assms(1) by (force simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) lemma members_of_only_vertical_div_line_integrable': assumes "only_vertical_division one_chain two_chain" "(k::int, \)\one_chain" "(k::int, \)\one_chain" "finite two_chain" shows "line_integral_exists F {i} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary by blast have integ_exis_horiz: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary'' assms by auto have valid_paths: "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using type_I_chain_vert_bound_valid valid_typeI_div by linarith have integ_exis_vert: "(\k \. (\(k', \')\two_chain_vertical_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {i} \)" using integ_exis valid_paths line_integral_exists_subpath[of "F" "{i}"] by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def boundary_def) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_vertical_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "one_chain = \ \ \" "(common_sudiv_exists (two_chain_horizontal_boundary two_chain) \) \ common_reparam_exists \ (two_chain_horizontal_boundary two_chain)" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using assms(1) unfolding only_vertical_division_def by blast have finite_i: "finite {i}" by auto show "line_integral_exists F {i} \" proof(cases "common_sudiv_exists (two_chain_horizontal_boundary two_chain) \") case True show ?thesis using gen_common_subdivision_imp_eq_line_integral(2)[OF True two_chain_horizontal_boundary_is_boundary_chain hv_props(6) integ_exis_horiz finite_two_chain_horizontal_boundary[OF assms(4)] hv_props(5) finite_i] integ_exis_vert[of "\"] assms(3) case_prod_conv hv_props(2) hv_props(3) by fastforce next case False have i: "{i} \ Basis" using i_is_x_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_horizontal_boundary two_chain. \b\{i}. continuous_on (path_image \2) (\x. F x \ b)" using assms field_cont_on_typeI_region_cont_on_edges F_anal_valid valid_typeI_div by (fastforce simp add: analytically_valid_def two_chain_horizontal_boundary_def boundary_def path_image_def) show "line_integral_exists F {i} \" using common_reparam_exists_imp_eq_line_integral(2)[OF finite_i hv_props(5) finite_two_chain_horizontal_boundary[OF assms(4)] hv_props(6) two_chain_horizontal_boundary_is_boundary_chain ii _ hv_props(7) type_I_chain_horiz_bound_valid] integ_exis_vert[of "\"] False assms(3) hv_props(2-4) by fastforce qed qed lemma GreenThm_typeI_two_chain: "two_chain_integral two_chain (\a. - partial_vector_derivative (\x. (F x) \ i) j a) = one_chain_line_integral F {i} (two_chain_boundary two_chain)" proof (simp add: two_chain_boundary_def one_chain_line_integral_def two_chain_integral_def) let ?F_b' = "partial_vector_derivative (\x. (F x) \ i) j" have all_two_cubes_have_four_distict_edges: "\twoCube \ two_chain. card (boundary twoCube) = 4" using valid_typeI_div valid_two_chain_def valid_two_cube_def by auto have no_shared_edges_have_similar_orientations: "\twoCube1 \ two_chain. \twoCube2 \ two_chain. twoCube1 \ twoCube2 \ boundary twoCube1 \ boundary twoCube2 = {}" using valid_typeI_div valid_two_chain_def by (auto simp add: pairwise_def) have "(\(k,g)\ boundary twoCube. k * line_integral F {i} g) = integral (cubeImage twoCube) (\a. - ?F_b' a)" if "twoCube \ two_chain" for twoCube proof - have analytically_val: " analytically_valid (cubeImage twoCube) (\x. F x \ i) j" using that F_anal_valid by auto show "(\(k, g)\boundary twoCube. k * line_integral F {i} g) = integral (cubeImage twoCube) (\a. - ?F_b' a)" using green_typeI_cube.GreenThm_typeI_twoCube apply (simp add: one_chain_line_integral_def) by (simp add: R2_axioms analytically_val green_typeI_cube_axioms_def green_typeI_cube_def that two_chain_valid_valid_cubes valid_typeI_div) qed then have double_sum_eq_sum: "(\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {i} g)) = (\twoCube\(two_chain). integral (cubeImage twoCube) (\a. - ?F_b' a))" using Finite_Cartesian_Product.sum_cong_aux by auto have pairwise_disjoint_boundaries: "\x\ (boundary ` two_chain). (\y\ (boundary ` two_chain). (x \ y \ (x \ y = {})))" using no_shared_edges_have_similar_orientations by (force simp add: image_def disjoint_iff_not_equal) have finite_boundaries: "\B \ (boundary` two_chain). finite B" using all_two_cubes_have_four_distict_edges using image_iff by fastforce have boundary_inj: "inj_on boundary two_chain" using all_two_cubes_have_four_distict_edges and no_shared_edges_have_similar_orientations by (force simp add: inj_on_def) have "(\x\(\(boundary` two_chain)). case x of (k, g) \ k * line_integral F {i} g) = (sum \ sum) (\x. case x of (k, g) \ (k::int) * line_integral F {i} g) (boundary ` two_chain)" using sum.Union_disjoint[OF finite_boundaries pairwise_disjoint_boundaries] by simp also have "... = (\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {i} g))" using sum.reindex[OF boundary_inj, of "\x. (\(k,g)\ x. k * line_integral F {i} g)"] by auto finally show "(\C\two_chain. - integral (cubeImage C) (partial_vector_derivative (\x. F x \ i) j)) = (\x\(\x\two_chain. boundary x). case x of (k, g) \ real_of_int k * line_integral F {i} g)" using double_sum_eq_sum by auto qed lemma GreenThm_typeI_divisible: assumes gen_division: "gen_division s (cubeImage ` two_chain)" shows "integral s (\x. - partial_vector_derivative (\a. F(a) \ i) j x) = one_chain_line_integral F {i} (two_chain_boundary two_chain)" proof - let ?F_b' = "partial_vector_derivative (\a. F(a) \ i) j" have "integral s (\x. - ?F_b' x) = two_chain_integral two_chain (\a. - ?F_b' a)" proof (simp add: two_chain_integral_def) have "(\f\two_chain. integral (cubeImage f) (partial_vector_derivative (\p. F p \ i) j)) = integral s (partial_vector_derivative (\p. F p \ i) j)" by (metis analytically_valid_imp_part_deriv_integrable_on F_anal_valid gen_division two_chain_integral_def two_chain_integral_eq_integral_divisable valid_typeI_div) then show "- integral s (partial_vector_derivative (\a. F a \ i) j) = (\C\two_chain. - integral (cubeImage C) (partial_vector_derivative (\a. F a \ i) j))" by (simp add: sum_negf) qed then show ?thesis using GreenThm_typeI_two_chain assms by auto qed lemma GreenThm_typeI_divisible_region_boundary: assumes gen_division: "gen_division s (cubeImage ` two_chain)" and two_cubes_trace_horizontal_boundaries: "two_chain_horizontal_boundary two_chain \ \" and boundary_of_region_is_subset_of_partition_boundary: "\ \ two_chain_boundary two_chain" shows "integral s (\x. - partial_vector_derivative (\a. F(a) \ i) j x) = one_chain_line_integral F {i} \" proof - let ?F_b' = "partial_vector_derivative (\a. F(a) \ i)" have all_two_cubes_have_four_distict_edges: "\twoCube \ two_chain. card (boundary twoCube) = 4" using valid_typeI_div valid_two_chain_def valid_two_cube_def by auto have no_shared_edges_have_similar_orientations: "\twoCube1 \ two_chain. \twoCube2 \ two_chain. twoCube1 \ twoCube2 \ boundary twoCube1 \ boundary twoCube2 = {}" using valid_typeI_div valid_two_chain_def by (auto simp add: pairwise_def) (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have vert_line_integral_zero: "one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_vertical_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where vert_edge_def: "oneCube = (k, (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div oneCube by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) let ?vert_edge = "(snd oneCube)" have vert_edge_x_const: "\t. (?vert_edge t) \ i = x" by (simp add: i_is_x_axis vert_edge_def) have vert_edge_is_straight_path: "?vert_edge = (\t. (x, y2 + t * (y1 - y2)))" using vert_edge_def Product_Type.snd_conv by (auto simp add: mult.commute right_diff_distrib') show ?thesis by (simp add: i_is_x_axis line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_x vert_edge_is_straight_path) qed then show "(\x\two_chain_vertical_boundary two_chain. case x of (k, g) \ k * line_integral F {i} g) = 0" using comm_monoid_add_class.sum.neutral by (simp add: prod.case_eq_if) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def by (metis all_two_cubes_have_four_distict_edges card.infinite finite_UN_I finite_imageD gen_division gen_division_def zero_neq_numeral valid_typeI_div valid_two_chain_def) have boundary_is_vert_hor: "(two_chain_boundary two_chain) = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by(auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \z. x (z, 0)), (- 1, \z. x (z, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \z. twoCube2 (z, 0)), (- 1, \z. twoCube2 (z, 1))} = {}" if "twoCube\two_chain" "twoCube2\two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have "card {(- 1::int, \y. twoCube2 (0::real, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using all_two_cubes_have_four_distict_edges that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def) then show ?thesis by (auto simp: True card_insert_if split: if_split_asm) next case False then show ?thesis using no_shared_edges_have_similar_orientations by (simp add: that boundary_def vertical_boundary_def horizontal_boundary_def) qed then have "\ ((\twoCube. {(-1::int, \y. twoCube (0,y)), (1, \y. twoCube (1, y))}) ` two_chain) \ \ ((\twoCube. {(1, \y. twoCube (y, 0)), (-1, \z. twoCube (z, 1))}) ` two_chain) = {}" using Complete_Lattices.Union_disjoint by force then show ?thesis by force qed qed have "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (auto simp add: one_chain_line_integral_def hor_vert_finite sum.union_disjoint) then have x_axis_line_integral_is_only_horizontal: "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using vert_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>V(2-\) such that \ = \ \ \\<^sub>H(2-\)*) have "\\. \ \ (two_chain_vertical_boundary two_chain) \ \ = \ \ (two_chain_horizontal_boundary two_chain)" proof let ?\ = "\ - (two_chain_horizontal_boundary two_chain)" show "?\ \ two_chain_vertical_boundary two_chain \ \ = ?\ \ two_chain_horizontal_boundary two_chain" using two_cubes_trace_horizontal_boundaries boundary_of_region_is_subset_of_partition_boundary boundary_is_vert_hor by blast qed then obtain \ where v_props: "\ \ (two_chain_vertical_boundary two_chain)" "\ = \ \ (two_chain_horizontal_boundary two_chain)" by auto have v_horiz_disj: "\ \ (two_chain_horizontal_boundary two_chain) = {}" using horiz_verti_disjoint v_props(1) by auto have v_finite: "finite \" using hor_vert_finite v_props(1) Finite_Set.rev_finite_subset by force have line_integral_on_path: "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} \ + one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" by(auto simp add: one_chain_line_integral_def v_props sum.union_disjoint[OF v_finite hor_vert_finite(2) v_horiz_disj]) (*Since \ \ \\<^sub>V(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have "one_chain_line_integral F {i} \ = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_vertical_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where vert_edge_def: "oneCube = (k, (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div oneCube by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) let ?vert_edge = "(snd oneCube)" have vert_edge_x_const: "\t. (?vert_edge t) \ i = x" by (simp add: i_is_x_axis vert_edge_def) have vert_edge_is_straight_path: "?vert_edge = (\t. (x, y2 + t * (y1 - y2)))" by (auto simp: vert_edge_def algebra_simps) have "\x. ?vert_edge differentiable at x" by (metis mult.commute vert_edge_is_straight_path straight_path_diffrentiable_x) then show "line_integral F {i} (snd oneCube) = 0" using line_integral_on_pair_straight_path(1) vert_edge_x_const by blast qed then have "\oneCube \ \. line_integral F {i} (snd oneCube) = 0" using v_props by auto then show "(\x\\. case x of (k, g) \ k * line_integral F {i} g) = 0" using comm_monoid_add_class.sum.neutral by (simp add: prod.case_eq_if) qed then have "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_boundary two_chain)" using x_axis_line_integral_is_only_horizontal by (simp add: line_integral_on_path) then show ?thesis using assms and GreenThm_typeI_divisible by auto qed lemma GreenThm_typeI_divisible_region_boundary_gen: assumes valid_typeI_div: "valid_typeI_division s two_chain" and f_analytically_valid: "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\a. F(a) \ i) j" and only_vertical_division: "only_vertical_division \ two_chain" shows "integral s (\x. - partial_vector_derivative (\a. F(a) \ i) j x) = one_chain_line_integral F {i} \" proof - let ?F_b' = "partial_vector_derivative (\a. F(a) \ i)" have all_two_cubes_have_four_distict_edges: "\twoCube \ two_chain. card (boundary twoCube) = 4" using valid_typeI_div valid_two_chain_def valid_two_cube_def by auto have no_shared_edges_have_similar_orientations: "\twoCube1 \ two_chain. \twoCube2 \ two_chain. twoCube1 \ twoCube2 \ boundary twoCube1 \ boundary twoCube2 = {}" using valid_typeI_div valid_two_chain_def by (auto simp add: pairwise_def) (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have vert_line_integral_zero: "one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_vertical_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where vert_edge_def: "oneCube = (k, (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div oneCube by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) let ?vert_edge = "(snd oneCube)" have vert_edge_x_const: "\t. (?vert_edge t) \ i = x" by (simp add: i_is_x_axis vert_edge_def) have vert_edge_is_straight_path: "?vert_edge = (\t. (x, y2 + t * (y1 - y2)))" by (auto simp: vert_edge_def algebra_simps) show ?thesis by (simp add: i_is_x_axis line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_x vert_edge_is_straight_path) qed then show "(\x\two_chain_vertical_boundary two_chain. case x of (k, g) \ k * line_integral F {i} g) = 0" using comm_monoid_add_class.sum.neutral by (simp add: prod.case_eq_if) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain" using assms(1) finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed have boundary_is_vert_hor: "two_chain_boundary two_chain = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by (auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \y. x (y, 0)), (- 1, \y. x (y, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \y. twoCube2 (y, 0)), (- 1, \y. twoCube2 (y, 1))} = {}" if "twoCube \ two_chain" "twoCube2 \ two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have "card {(- 1::int, \y. twoCube2 (0, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using all_two_cubes_have_four_distict_edges that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def) then show ?thesis by (auto simp: card_insert_if True split: if_split_asm) next case False then show ?thesis using no_shared_edges_have_similar_orientations by (simp add: that boundary_def vertical_boundary_def horizontal_boundary_def) qed then have "\ ((\twoCube. {(- 1, \y. twoCube (0, y)), (1, \y. twoCube (1, y))}) ` two_chain) \ \ ((\twoCube. {(1::int, \y. twoCube (y, 0)), (- 1, \y. twoCube (y, 1))}) ` two_chain) = {}" using Complete_Lattices.Union_disjoint by force then show ?thesis by force qed qed have "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (auto simp add: one_chain_line_integral_def hor_vert_finite sum.union_disjoint) then have x_axis_line_integral_is_only_horizontal: "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using vert_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>V(2-\) such that \ = \ \ \\<^sub>H(2-\)*) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_vertical_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "\ = \ \ \" "(common_sudiv_exists (two_chain_horizontal_boundary two_chain) \) \ common_reparam_exists \ (two_chain_horizontal_boundary two_chain)" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using only_vertical_division by (auto simp add: only_vertical_division_def) have "finite {i}" by auto then have eq_integrals: " one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" proof(cases "common_sudiv_exists (two_chain_horizontal_boundary two_chain) \") case True then show ?thesis using gen_common_subdivision_imp_eq_line_integral(1)[OF True two_chain_horizontal_boundary_is_boundary_chain hv_props(6) _ hor_vert_finite(2) hv_props(5)] typeI_cube_line_integral_exists_boundary'' by force next case False have integ_exis_horiz: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary'' assms by (fastforce simp add: valid_two_chain_def) have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary by blast have valid_paths: "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using typeI_edges_are_valid_paths assms by (fastforce simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) have integ_exis_vert: "(\k \. (\(k', \') \ two_chain_vertical_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {i} \)" using integ_exis valid_paths line_integral_exists_subpath[of "F" "{i}"] by (fastforce simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) have finite_i: "finite {i}" by auto have i: "{i} \ Basis" using i_is_x_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_horizontal_boundary two_chain. \b\{i}. continuous_on (path_image \2) (\x. F x \ b)" using assms(1) field_cont_on_typeI_region_cont_on_edges assms(2) by (fastforce simp add: analytically_valid_def two_chain_horizontal_boundary_def boundary_def path_image_def) have *: "common_reparam_exists \ (two_chain_horizontal_boundary two_chain)" using hv_props(4) False by auto show "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using common_reparam_exists_imp_eq_line_integral(1)[OF finite_i hv_props(5) hor_vert_finite(2) hv_props(6) two_chain_horizontal_boundary_is_boundary_chain ii * hv_props(7) type_I_chain_horiz_bound_valid] by fastforce qed (*Since \ \ \\<^sub>H(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have line_integral_on_path: "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" proof (auto simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ \" for oneCube proof - obtain k \ where k_gamma: "(k,\) = oneCube" by (metis coeff_cube_to_path.cases) then obtain k' \' a b where kp_gammap: "(k'::int, \') \ two_chain_vertical_boundary two_chain" "a \ {0 .. 1}" "b \ {0..1}" "subpath a b \' = \" using hv_props oneCube by (smt case_prodE split_conv) obtain x y1 y2 where vert_edge_def: "(k',\') = (k', (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div kp_gammap by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) have vert_edge_x_const: "\t. \ (t) \ i = x" by (metis (no_types, lifting) Pair_inject fstI i_is_x_axis inner_Pair_0(2) kp_gammap(4) real_inner_1_right subpath_def vert_edge_def) have "\ = (\t::real. (x::real, (1 - (b - a)*t - a) * (y2) + ((b-a)*t + a) * y1))" using vert_edge_def Product_Type.snd_conv Product_Type.fst_conv kp_gammap(4) by (simp add: subpath_def diff_diff_eq[symmetric]) also have "... = (\t::real. (x::real, (1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t))" by (simp add: algebra_simps) finally have vert_edge_is_straight_path: "\ = (\t::real. (x::real, (1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t))" . show "line_integral F {i} (snd oneCube) = 0" proof - have "\x. \ differentiable at x" by (simp add: straight_path_diffrentiable_x vert_edge_is_straight_path) then have "line_integral F {i} \ = 0" using line_integral_on_pair_straight_path(1) vert_edge_x_const by blast then show ?thesis using Product_Type.snd_conv k_gamma by auto qed qed then have "\x\\. (case x of (k, g) \ (k::int) * line_integral F {i} g) = 0" by auto then show "(\x\\. case x of (k, g) \ real_of_int k * line_integral F {i} g) = (\x\two_chain_horizontal_boundary two_chain. case x of (k, g) \ of_int k * line_integral F {i} g)" using hv_props(1) hv_props(3) hv_props(5) sum_zero_set hor_vert_finite(2) eq_integrals apply(auto simp add: one_chain_line_integral_def) by (smt Un_commute sum_zero_set) qed then have "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_boundary two_chain)" using x_axis_line_integral_is_only_horizontal line_integral_on_path by auto then show ?thesis using assms GreenThm_typeI_divisible by auto qed end locale green_typeI_typeII_chain = R2: R2 i j + T1: green_typeI_chain i j F two_chain_typeI + T2: green_typeII_chain i j F two_chain_typeII for i j F two_chain_typeI two_chain_typeII begin lemma GreenThm_typeI_typeII_divisible_region_boundary: assumes gen_divisions: "gen_division s (cubeImage ` two_chain_typeI)" "gen_division s (cubeImage ` two_chain_typeII)" and typeI_two_cubes_trace_horizontal_boundaries: "two_chain_horizontal_boundary two_chain_typeI \ \" and typeII_two_cubes_trace_vertical_boundaries: "two_chain_vertical_boundary two_chain_typeII \ \" and boundary_of_region_is_subset_of_partition_boundaries: "\ \ two_chain_boundary two_chain_typeI" "\ \ two_chain_boundary two_chain_typeII" shows "integral s (\x. partial_vector_derivative (\a. F a \ j) i x - partial_vector_derivative (\a. F a \ i) j x) = one_chain_line_integral F {i, j} \" proof - let ?F_b' = "partial_vector_derivative (\a. F a \ i) j" let ?F_a' = "partial_vector_derivative (\a. F a \ j) i" have typeI_regions_integral: "integral s (\x. - partial_vector_derivative (\a. F a \ i) j x) = one_chain_line_integral F {i} \" using T1.GreenThm_typeI_divisible_region_boundary gen_divisions(1) typeI_two_cubes_trace_horizontal_boundaries boundary_of_region_is_subset_of_partition_boundaries(1) by blast have typeII_regions_integral: "integral s (partial_vector_derivative (\x. F x \ j) i) = one_chain_line_integral F {j} \" using T2.GreenThm_typeII_divisible_region_boundary gen_divisions(2) typeII_two_cubes_trace_vertical_boundaries boundary_of_region_is_subset_of_partition_boundaries(2) by auto have integral_dis: "integral s (\x. ?F_a' x - ?F_b' x) = integral s (\x. ?F_a' x) + integral s (\x. - ?F_b' x)" proof - have "\twoCube \ two_chain_typeII. (?F_a' has_integral integral (cubeImage twoCube) ?F_a') (cubeImage twoCube)" by (simp add: analytically_valid_imp_part_deriv_integrable_on T2.F_anal_valid has_integral_iff) then have "\u. u \ (cubeImage ` two_chain_typeII) \ (?F_a' has_integral integral u ?F_a') u" by auto then have "(?F_a' has_integral (\img\cubeImage ` two_chain_typeII. integral img ?F_a')) s" using gen_divisions(2) unfolding gen_division_def by (metis has_integral_Union) then have F_a'_integrable: "(?F_a' integrable_on s)" by auto have "\twoCube \ two_chain_typeI. (?F_b' has_integral integral (cubeImage twoCube) ?F_b') (cubeImage twoCube)" using analytically_valid_imp_part_deriv_integrable_on T1.F_anal_valid by blast then have "\u. u \ (cubeImage ` two_chain_typeI) \ (?F_b' has_integral integral u ?F_b') u" by auto then have "(?F_b' has_integral (\img\cubeImage ` two_chain_typeI. integral img ?F_b')) s" using gen_divisions(1) unfolding gen_division_def by (metis has_integral_Union) then show ?thesis by (simp add: F_a'_integrable Henstock_Kurzweil_Integration.integral_diff has_integral_iff) qed have line_integral_dist: "one_chain_line_integral F {i, j} \ = one_chain_line_integral F {i} \ + one_chain_line_integral F {j} \" proof (simp add: one_chain_line_integral_def) have "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if kg: "(k,g) \ \" for k g proof - obtain twoCube_typeI where twoCube_typeI_props: "twoCube_typeI \ two_chain_typeI" "(k, g) \ boundary twoCube_typeI" "typeI_twoCube twoCube_typeI" "continuous_on (cubeImage twoCube_typeI) (\x. F(x) \ i)" using boundary_of_region_is_subset_of_partition_boundaries(1) two_chain_boundary_def T1.valid_typeI_div T1.F_anal_valid kg by (auto simp add: analytically_valid_def) obtain twoCube_typeII where twoCube_typeII_props: "twoCube_typeII \ two_chain_typeII" "(k, g) \ boundary twoCube_typeII" "typeII_twoCube twoCube_typeII" "continuous_on (cubeImage twoCube_typeII) (\x. F(x) \ j)" using boundary_of_region_is_subset_of_partition_boundaries(2) two_chain_boundary_def T2.valid_typeII_div kg T2.F_anal_valid by (auto simp add: analytically_valid_def) have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" proof - have int_exists_i: "line_integral_exists F {i} g" using T1.typeI_cube_line_integral_exists_boundary assms kg by (auto simp add: valid_two_chain_def) have int_exists_j: "line_integral_exists F {j} g" using T2.typeII_cube_line_integral_exists_boundary assms kg by (auto simp add: valid_two_chain_def) have finite: "finite {i, j}" by auto show ?thesis using line_integral_sum_gen[OF finite int_exists_i int_exists_j] R2.i_is_x_axis R2.j_is_y_axis by auto qed then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed then have line_integral_distrib: "(\(k,g)\\. k * line_integral F {i, j} g) = (\(k,g)\\. k * line_integral F {i} g + k * line_integral F {j} g)" by (force intro: sum.cong split_cong) have "(\x. (case x of (k, g) \ (k::int) * line_integral F {i} g) + (case x of (k, g) \ (k::int) * line_integral F {j} g)) = (\x. (case x of (k, g) \ (k * line_integral F {i} g) + (k::int) * line_integral F {j} g))" using comm_monoid_add_class.sum.distrib by auto then show "(\(k, g)\\. k * line_integral F {i, j} g) = (\(k, g)\\. (k::int) * line_integral F {i} g) + (\(k, g)\\. (k::int) * line_integral F {j} g)" using comm_monoid_add_class.sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" \] line_integral_distrib by presburger qed show ?thesis using integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto qed lemma GreenThm_typeI_typeII_divisible_region': assumes only_vertical_division: "only_vertical_division one_chain_typeI two_chain_typeI" "boundary_chain one_chain_typeI" and only_horizontal_division: "only_horizontal_division one_chain_typeII two_chain_typeII" "boundary_chain one_chain_typeII" and typeI_and_typII_one_chains_have_gen_common_subdiv: "common_sudiv_exists one_chain_typeI one_chain_typeII" shows "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeI" "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeII" proof - let ?F_b' = "partial_vector_derivative (\x. (F x) \ i) j" let ?F_a' = "partial_vector_derivative (\x. (F x) \ j) i" have one_chain_i_integrals: "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII \ (\(k,\)\one_chain_typeI. line_integral_exists F {i} \) \ (\(k,\)\one_chain_typeII. line_integral_exists F {i} \)" proof (intro conjI) have "finite two_chain_typeI" using T1.valid_typeI_div finite_image_iff by (auto simp add: gen_division_def valid_two_chain_def) then show ii: "\(k, \)\one_chain_typeI. line_integral_exists F {i} \" using T1.members_of_only_vertical_div_line_integrable' assms by fastforce have "finite (two_chain_horizontal_boundary two_chain_typeI)" by (meson T1.valid_typeI_div finite_imageD finite_two_chain_horizontal_boundary gen_division_def valid_two_chain_def) then have "finite one_chain_typeI" using only_vertical_division(1) only_vertical_division_def by auto moreover have "finite one_chain_typeII" using only_horizontal_division(1) only_horizontal_division_def by auto ultimately show "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII" and "\(k, \)\one_chain_typeII. line_integral_exists F {i} \" using gen_common_subdivision_imp_eq_line_integral[OF typeI_and_typII_one_chains_have_gen_common_subdiv only_vertical_division(2) only_horizontal_division(2)] ii by auto qed have one_chain_j_integrals: "one_chain_line_integral F {j} one_chain_typeII = one_chain_line_integral F {j} one_chain_typeI \ (\(k,\)\one_chain_typeII. line_integral_exists F {j} \) \ (\(k,\)\one_chain_typeI. line_integral_exists F {j} \)" proof (intro conjI) have "finite two_chain_typeII" using T2.valid_typeII_div finite_image_iff by (auto simp add: gen_division_def valid_two_chain_def) then show ii: "\(k,\)\one_chain_typeII. line_integral_exists F {j} \" using T2.members_of_only_horiz_div_line_integrable' assms T2.two_chain_valid_valid_cubes by blast have typeII_and_typI_one_chains_have_common_subdiv: "common_sudiv_exists one_chain_typeII one_chain_typeI" by (simp add: common_sudiv_exists_comm typeI_and_typII_one_chains_have_gen_common_subdiv) have iv: "finite one_chain_typeI" using only_vertical_division(1) only_vertical_division_def by auto moreover have iv': "finite one_chain_typeII" using only_horizontal_division(1) only_horizontal_division_def by auto ultimately show "one_chain_line_integral F {j} one_chain_typeII = one_chain_line_integral F {j} one_chain_typeI" "\(k, \)\one_chain_typeI. line_integral_exists F {j} \" using gen_common_subdivision_imp_eq_line_integral[OF typeII_and_typI_one_chains_have_common_subdiv only_horizontal_division(2) only_vertical_division(2) ii] ii by auto qed have typeI_regions_integral: "integral s (\x. - ?F_b' x) = one_chain_line_integral F {i} one_chain_typeI" using T1.GreenThm_typeI_divisible_region_boundary_gen T1.valid_typeI_div T1.F_anal_valid only_vertical_division(1) by auto have typeII_regions_integral: "integral s ?F_a' = one_chain_line_integral F {j} one_chain_typeII" using T2.GreenThm_typeII_divisible_region_boundary_gen T2.valid_typeII_div T2.F_anal_valid only_horizontal_division(1) by auto have line_integral_dist: "one_chain_line_integral F {i, j} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeI + one_chain_line_integral F {j} one_chain_typeI \ one_chain_line_integral F {i, j} one_chain_typeII = one_chain_line_integral F {i} one_chain_typeII + one_chain_line_integral F {j} one_chain_typeII" proof (simp add: one_chain_line_integral_def) have line_integral_distrib: "(\(k,g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeI. k * line_integral F {i} g + k * line_integral F {j} g) \ (\(k,g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeII. k * line_integral F {i} g + k * line_integral F {j} g)" proof - have 0: "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeII" for k g proof - have "line_integral_exists F {i} g" "line_integral_exists F {j} g" "finite {i, j}" using one_chain_i_integrals one_chain_j_integrals that by fastforce+ moreover have "{i} \ {j} = {}" by (simp add: R2.i_is_x_axis R2.j_is_y_axis) ultimately have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" by (metis insert_is_Un line_integral_sum_gen(1)) then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed have "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeI" for k g proof - have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" by (smt that disjoint_insert(2) finite.emptyI finite.insertI R2.i_is_x_axis inf_bot_right insert_absorb insert_commute insert_is_Un R2.j_is_y_axis line_integral_sum_gen(1) one_chain_i_integrals one_chain_j_integrals prod.case_eq_if singleton_inject snd_conv) then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed then show ?thesis using 0 by (smt sum.cong split_cong) qed show "(\(k::int, g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeI. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeI. k * line_integral F {j} g) \ (\(k::int, g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeII. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeII. k * line_integral F {j} g)" proof - have 0: "(\x. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x. (case x of (k::int, g) \ (k * line_integral F {i} g) + k * line_integral F {j} g))" using comm_monoid_add_class.sum.distrib by auto then have 1: "(\x\one_chain_typeI. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeI. (case x of (k::int, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" by presburger have "(\x\one_chain_typeII. (case x of (k, g) \ k * line_integral F {i} g) + (case x of (k, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeII. (case x of (k, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" using 0 by presburger then show ?thesis using sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeI"] sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeII"] line_integral_distrib 1 by auto qed qed have integral_dis: "integral s (\x. ?F_a' x - ?F_b' x) = integral s (\x. ?F_a' x) + integral s (\x. - ?F_b' x)" proof - have "(?F_a' has_integral integral (cubeImage twoCube) ?F_a') (cubeImage twoCube)" if "twoCube \ two_chain_typeII" for twoCube by (simp add: analytically_valid_imp_part_deriv_integrable_on T2.F_anal_valid has_integral_integrable_integral that) then have "\u. u \ (cubeImage ` two_chain_typeII) \ (?F_a' has_integral integral u ?F_a') u" by auto then have "(?F_a' has_integral (\img\cubeImage ` two_chain_typeII. integral img ?F_a')) s" using T2.valid_typeII_div unfolding gen_division_def by (metis has_integral_Union) then have F_a'_integrable: "(?F_a' integrable_on s)" by auto have "\twoCube \ two_chain_typeI. (?F_b' has_integral integral (cubeImage twoCube) ?F_b') (cubeImage twoCube)" using analytically_valid_imp_part_deriv_integrable_on T1.F_anal_valid by blast then have "\u. u \ (cubeImage ` two_chain_typeI) \ (?F_b' has_integral integral u ?F_b') u" by auto then have "(?F_b' has_integral (\img\cubeImage ` two_chain_typeI. integral img ?F_b')) s" using T1.valid_typeI_div unfolding gen_division_def by (metis has_integral_Union) then show ?thesis by (simp add: F_a'_integrable Henstock_Kurzweil_Integration.integral_diff has_integral_iff) qed show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeI" using one_chain_j_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeII" using one_chain_i_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto qed lemma GreenThm_typeI_typeII_divisible_region: assumes only_vertical_division: "only_vertical_division one_chain_typeI two_chain_typeI" "boundary_chain one_chain_typeI" and only_horizontal_division: "only_horizontal_division one_chain_typeII two_chain_typeII" "boundary_chain one_chain_typeII" and typeI_and_typII_one_chains_have_common_subdiv: "common_boundary_sudivision_exists one_chain_typeI one_chain_typeII" shows "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeI" "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeII" using GreenThm_typeI_typeII_divisible_region' only_vertical_division only_horizontal_division common_subdiv_imp_gen_common_subdiv[OF typeI_and_typII_one_chains_have_common_subdiv] by auto lemma GreenThm_typeI_typeII_divisible_region_finite_holes: assumes valid_cube_boundary: "\(k,\)\boundary C. valid_path \" and only_vertical_division: "only_vertical_division (boundary C) two_chain_typeI" and only_horizontal_division: "only_horizontal_division (boundary C) two_chain_typeII" and s_is_oneCube: "s = cubeImage C" shows "integral (cubeImage C) (\x. partial_vector_derivative (\x. F x \ j) i x - partial_vector_derivative (\x. F x \ i) j x) = one_chain_line_integral F {i, j} (boundary C)" using GreenThm_typeI_typeII_divisible_region[OF only_vertical_division two_cube_boundary_is_boundary only_horizontal_division two_cube_boundary_is_boundary common_boundary_subdiv_exists_refl[OF assms(1)]] s_is_oneCube by auto lemma GreenThm_typeI_typeII_divisible_region_equivallent_boundary: assumes gen_divisions: "gen_division s (cubeImage ` two_chain_typeI)" "gen_division s (cubeImage ` two_chain_typeII)" and typeI_two_cubes_trace_horizontal_boundaries: "two_chain_horizontal_boundary two_chain_typeI \ one_chain_typeI" and typeII_two_cubes_trace_vertical_boundaries: "two_chain_vertical_boundary two_chain_typeII \ one_chain_typeII" and boundary_of_region_is_subset_of_partition_boundaries: "one_chain_typeI \ two_chain_boundary two_chain_typeI" "one_chain_typeII \ two_chain_boundary two_chain_typeII" and typeI_and_typII_one_chains_have_common_subdiv: "common_boundary_sudivision_exists one_chain_typeI one_chain_typeII" shows "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeI" "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeII" proof - let ?F_b' = "partial_vector_derivative (\x. (F x) \ i) j" let ?F_a' = "partial_vector_derivative (\x. (F x) \ j) i" have one_chain_i_integrals: "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII \ (\(k,\)\one_chain_typeI. line_integral_exists F {i} \) \ (\(k,\)\one_chain_typeII. line_integral_exists F {i} \)" proof (intro conjI) have i: "boundary_chain one_chain_typeI" using two_chain_boundary_is_boundary_chain boundary_chain_def boundary_of_region_is_subset_of_partition_boundaries(1) by blast have i': "boundary_chain one_chain_typeII" using two_chain_boundary_is_boundary_chain boundary_chain_def boundary_of_region_is_subset_of_partition_boundaries(2) by blast have "\k \. (k,\)\one_chain_typeI \ line_integral_exists F {i} \" using T1.typeI_cube_line_integral_exists_boundary assms by (fastforce simp add: valid_two_chain_def) then show ii: "\(k,\)\one_chain_typeI. line_integral_exists F {i} \" by auto have "finite (two_chain_boundary two_chain_typeI)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeI" using T1.valid_typeI_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeI \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have "finite one_chain_typeI" using boundary_of_region_is_subset_of_partition_boundaries(1) finite_subset by fastforce moreover have "finite (two_chain_boundary two_chain_typeII)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeII" using T2.valid_typeII_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeII \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have "finite one_chain_typeII" using boundary_of_region_is_subset_of_partition_boundaries(2) finite_subset by fastforce ultimately show "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII" "\(k, \)\one_chain_typeII. line_integral_exists F {i} \" using ii common_subdivision_imp_eq_line_integral[OF typeI_and_typII_one_chains_have_common_subdiv i i' ii] by auto qed have one_chain_j_integrals: "one_chain_line_integral F {j} one_chain_typeI = one_chain_line_integral F {j} one_chain_typeII \ (\(k,\)\one_chain_typeI. line_integral_exists F {j} \) \ (\(k,\)\one_chain_typeII. line_integral_exists F {j} \)" proof (intro conjI) have i: "boundary_chain one_chain_typeI" and i': "boundary_chain one_chain_typeII" using two_chain_boundary_is_boundary_chain boundary_of_region_is_subset_of_partition_boundaries unfolding boundary_chain_def by blast+ have "line_integral_exists F {j} \" if "(k,\)\one_chain_typeII" for k \ proof - have F_is_continuous: "\twoC \ two_chain_typeII. continuous_on (cubeImage twoC) (\a. F(a) \ j)" using T2.F_anal_valid by(simp add: analytically_valid_def) show "line_integral_exists F {j} \" using that T2.valid_typeII_div boundary_of_region_is_subset_of_partition_boundaries(2) using green_typeII_cube.line_integral_exists_on_typeII_Cube_boundaries' assms valid_two_chain_def apply (simp add: two_chain_boundary_def) by (metis T2.typeII_cube_line_integral_exists_boundary case_prodD subset_iff that two_chain_boundary_def) qed then show ii: " \(k,\)\one_chain_typeII. line_integral_exists F {j} \" by auto have "finite (two_chain_boundary two_chain_typeI)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeI" using T1.valid_typeI_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeI \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have iv: "finite one_chain_typeI" using boundary_of_region_is_subset_of_partition_boundaries(1) finite_subset by fastforce have "finite (two_chain_boundary two_chain_typeII)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeII" using T2.valid_typeII_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeII \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have iv': "finite one_chain_typeII" using boundary_of_region_is_subset_of_partition_boundaries(2) finite_subset by fastforce have typeII_and_typI_one_chains_have_common_subdiv: "common_boundary_sudivision_exists one_chain_typeII one_chain_typeI" using typeI_and_typII_one_chains_have_common_subdiv common_boundary_sudivision_commutative by auto show "one_chain_line_integral F {j} one_chain_typeI = one_chain_line_integral F {j} one_chain_typeII" "\(k, \)\one_chain_typeI. line_integral_exists F {j} \" using common_subdivision_imp_eq_line_integral[OF typeII_and_typI_one_chains_have_common_subdiv i' i ii iv' iv] ii by auto qed have typeI_regions_integral: "integral s (\x. - ?F_b' x) = one_chain_line_integral F {i} one_chain_typeI" using T1.GreenThm_typeI_divisible_region_boundary gen_divisions(1) typeI_two_cubes_trace_horizontal_boundaries boundary_of_region_is_subset_of_partition_boundaries(1) by auto have typeII_regions_integral: "integral s ?F_a' = one_chain_line_integral F {j} one_chain_typeII" using T2.GreenThm_typeII_divisible_region_boundary gen_divisions(2) typeII_two_cubes_trace_vertical_boundaries boundary_of_region_is_subset_of_partition_boundaries(2) by auto have line_integral_dist: "one_chain_line_integral F {i, j} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeI + one_chain_line_integral F {j} one_chain_typeI \ one_chain_line_integral F {i, j} one_chain_typeII = one_chain_line_integral F {i} one_chain_typeII + one_chain_line_integral F {j} one_chain_typeII" proof (simp add: one_chain_line_integral_def) have line_integral_distrib: "(\(k,g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeI. k * line_integral F {i} g + k * line_integral F {j} g) \ (\(k,g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeII. k * line_integral F {i} g + k * line_integral F {j} g)" proof - have 0: "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeII" for k g proof - have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" proof - have finite: "finite {i, j}" by auto have line_integral_all: "\i\{i, j}. line_integral_exists F {i} g" using one_chain_i_integrals one_chain_j_integrals that by auto show ?thesis using line_integral_sum_gen[OF finite] R2.i_is_x_axis R2.j_is_y_axis line_integral_all by auto qed then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed have "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeI" for k g proof - have finite: "finite {i, j}" by auto have line_integral_all: "\i\{i, j}. line_integral_exists F {i} g" using one_chain_i_integrals one_chain_j_integrals that by auto have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" using line_integral_sum_gen[OF finite] R2.i_is_x_axis R2.j_is_y_axis line_integral_all by auto then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed then show ?thesis using 0 by (smt sum.cong split_cong) qed show "(\(k::int, g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeI. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeI. k * line_integral F {j} g) \ (\(k::int, g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeII. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeII. k * line_integral F {j} g)" proof - have 0: "(\x. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x. (case x of (k::int, g) \ (k * line_integral F {i} g) + k * line_integral F {j} g))" using comm_monoid_add_class.sum.distrib by auto then have 1: "(\x\one_chain_typeI. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeI. (case x of (k::int, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" by presburger have "(\x\one_chain_typeII. (case x of (k, g) \ k * line_integral F {i} g) + (case x of (k, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeII. (case x of (k, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" using 0 by presburger then show ?thesis using sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeI"] sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeII"] line_integral_distrib 1 by auto qed qed have integral_dis: "integral s (\x. ?F_a' x - ?F_b' x) = integral s (\x. ?F_a' x) + integral s (\x. - ?F_b' x)" proof - have "(?F_a' has_integral (\img\cubeImage ` two_chain_typeII. integral img ?F_a')) s" proof - have "(?F_a' has_integral integral (cubeImage twoCube) ?F_a') (cubeImage twoCube)" if "twoCube \ two_chain_typeII" for twoCube by (simp add: analytically_valid_imp_part_deriv_integrable_on T2.F_anal_valid has_integral_integrable_integral that) then have "\u. u \ (cubeImage ` two_chain_typeII) \ (?F_a' has_integral integral u ?F_a') u" by auto then show ?thesis using gen_divisions(2) unfolding gen_division_def by (metis has_integral_Union) qed then have F_a'_integrable: "(?F_a' integrable_on s)" by auto have "(?F_b' has_integral (\img\cubeImage ` two_chain_typeI. integral img ?F_b')) s" proof - have "\twoCube \ two_chain_typeI. (?F_b' has_integral integral (cubeImage twoCube) ?F_b') (cubeImage twoCube)" by (simp add: analytically_valid_imp_part_deriv_integrable_on T1.F_anal_valid has_integral_integrable_integral) then have "\u. u \ (cubeImage ` two_chain_typeI) \ (?F_b' has_integral integral u ?F_b') u" by auto then show ?thesis using gen_divisions(1) unfolding gen_division_def by (metis has_integral_Union) qed then show ?thesis using F_a'_integrable Henstock_Kurzweil_Integration.integral_diff by auto qed show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeI" using one_chain_j_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeII" using one_chain_i_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto qed end end diff --git a/thys/Probabilistic_Noninterference/Language_Semantics.thy b/thys/Probabilistic_Noninterference/Language_Semantics.thy --- a/thys/Probabilistic_Noninterference/Language_Semantics.thy +++ b/thys/Probabilistic_Noninterference/Language_Semantics.thy @@ -1,1608 +1,1608 @@ section \Simple While Language with probabilistic choice and parallel execution\ theory Language_Semantics imports Interface begin subsection \Preliminaries\ text\Trivia\ declare zero_le_mult_iff[simp] declare split_mult_pos_le[simp] declare zero_le_divide_iff[simp] lemma in_minus_Un[simp]: assumes "i \ I" shows "I - {i} Un {i} = I" and "{i} Un (I - {i}) = I" apply(metis Un_commute assms insert_Diff_single insert_absorb insert_is_Un) by (metis assms insert_Diff_single insert_absorb insert_is_Un) lemma less_plus_cases[case_names Left Right]: assumes *: "(i::nat) < n1 \ phi" and **: "\ i2. i = n1 + i2 \ phi" shows phi proof(cases "i < n1") case True thus ?thesis using * by simp next case False hence "n1 \ i" by simp then obtain i2 where "i = n1 + i2" by (metis le_iff_add) thus ?thesis using ** by blast qed lemma less_plus_elim[elim!, consumes 1, case_names Left Right]: assumes i: "(i:: nat) < n1 + n2" and *: "i < n1 \ phi" and **: "\ i2. \i2 < n2; i = n1 + i2\ \ phi" shows phi apply(rule less_plus_cases[of i n1]) using assms by auto lemma nth_append_singl[simp]: "i < length al \ (al @ [a]) ! i = al ! i" by (auto simp add: nth_append) lemma take_append_singl[simp]: assumes "n < length al" shows "take n (al @ [a]) = take n al" using assms by (induct al rule: rev_induct) auto lemma length_unique_prefix: "al1 \ al \ al2 \ al \ length al1 = length al2 \ al1 = al2" by (metis not_equal_is_parallel parallelE prefix_same_cases less_eq_list_def) text\take:\ lemma take_length[simp]: "take (length al) al = al" using take_all by auto lemma take_le: assumes "n < length al" shows "take n al @ [al ! n] \ al" by (metis assms take_Suc_conv_app_nth take_is_prefix less_eq_list_def) lemma list_less_iff_prefix: "a < b \ strict_prefix a b" by (metis le_less less_eq_list_def less_irrefl prefix_order.le_less prefix_order.less_irrefl) lemma take_lt: "n < length al \ take n al < al" unfolding list_less_iff_prefix using prefix_order.le_less[of "take n al" al] by (simp add: take_is_prefix) (metis length_take min_absorb2 nat_le_linear not_less) lemma le_take: assumes "n1 \ n2" shows "take n1 al \ take n2 al" using assms proof(induct al arbitrary: n1 n2) case (Cons a al) thus ?case by (cases n1 n2 rule: nat.exhaust[case_product nat.exhaust]) auto qed auto lemma inj_take: assumes "n1 \ length al" and "n2 \ length al" shows "take n1 al = take n2 al \ n1 = n2" proof- {assume "take n1 al = take n2 al" hence "n1 = n2" using assms proof(induct al arbitrary: n1 n2) case (Cons a al) thus ?case by (cases n1 n2 rule: nat.exhaust[case_product nat.exhaust]) auto qed auto } thus ?thesis by auto qed lemma lt_take: "n1 < n2 \ n2 \ length al \ take n1 al < take n2 al" by (metis inj_take le_less_trans le_take not_less_iff_gr_or_eq order.not_eq_order_implies_strict order.strict_implies_order) text\lsum:\ definition lsum :: "('a \ nat) \ 'a list \ nat" where "lsum f al \ sum_list (map f al)" lemma lsum_simps[simp]: "lsum f [] = 0" "lsum f (al @ [a]) = lsum f al + f a" unfolding lsum_def by auto lemma lsum_append: "lsum f (al @ bl) = lsum f al + lsum f bl" unfolding lsum_def by auto lemma lsum_cong[fundef_cong]: assumes "\ a. a \ set al \ f a = g a" shows "lsum f al = lsum g al" using assms unfolding lsum_def by (metis map_eq_conv) lemma lsum_gt_0[simp]: assumes "al \ []" and "\ a. a \ set al \ 0 < f a" shows "0 < lsum f al" using assms by (induct rule: rev_induct) auto lemma lsum_mono[simp]: assumes "al \ bl" shows "lsum f al \ lsum f bl" proof- obtain cl where bl: "bl = al @ cl" using assms unfolding prefix_def less_eq_list_def by blast thus ?thesis unfolding bl lsum_append by simp qed lemma lsum_mono2[simp]: assumes f: "\ b. b \ set bl \ f b > 0" and le: "al < bl" shows "lsum f al < lsum f bl" proof- obtain cl where bl: "bl = al @ cl" and cl: "cl \ []" using assms unfolding list_less_iff_prefix prefix_def strict_prefix_def by blast have "lsum f al < lsum f al + lsum f cl" using f cl unfolding bl by simp also have "... = lsum f bl" unfolding bl lsum_append by simp finally show ?thesis . qed lemma lsum_take[simp]: "lsum f (take n al) \ lsum f al" by (metis lsum_mono take_is_prefix less_eq_list_def) lemma less_lsum_nchotomy: assumes f: "\ a. a \ set al \ 0 < f a" and i: "(i::nat) < lsum f al" shows "\ n j. n < length al \ j < f (al ! n) \ i = lsum f (take n al) + j" using assms proof(induct rule: rev_induct) case (snoc a al) hence i: "i < lsum f al + f a" and pos: "0 < f a" "\a'. a' \ set al \ 0 < f a'" by auto from i show ?case proof(cases rule: less_plus_elim) case Left then obtain n j where "n < length al \ j < f (al ! n) \ i = lsum f (take n al) + j" using pos snoc by auto thus ?thesis apply - apply(rule exI[of _ n]) apply(rule exI[of _ j]) by auto next case (Right j) thus ?thesis apply - apply(rule exI[of _ "length al"]) apply(rule exI[of _ j]) by simp qed qed auto lemma less_lsum_unique: assumes "\ a. a \ set al \ (0::nat) < f a" and "n1 < length al \ j1 < f (al ! n1)" and "n2 < length al \ j2 < f (al ! n2)" and "lsum f (take n1 al) + j1 = lsum f (take n2 al) + j2" shows "n1 = n2 \ j1 = j2" using assms proof(induct al arbitrary: n1 n2 j1 j2 rule: rev_induct) case (snoc a al) hence pos: "0 < f a" "\ a'. a' \ set al \ 0 < f a'" and n1: "n1 < length al + 1" and n2: "n2 < length al + 1" by auto from n1 show ?case proof(cases rule: less_plus_elim) case Left note n1 = Left hence j1: "j1 < f (al ! n1)" using snoc by auto obtain al' where al: "al = (take n1 al) @ ((al ! n1) # al')" using n1 by (metis append_take_drop_id Cons_nth_drop_Suc) have "j1 < lsum f ((al ! n1) # al')" using pos j1 unfolding lsum_def by simp hence "lsum f (take n1 al) + j1 < lsum f (take n1 al) + lsum f ((al ! n1) # al')" by simp also have "... = lsum f al" unfolding lsum_append[THEN sym] using al by simp finally have lsum1: "lsum f (take n1 al) + j1 < lsum f al" . from n2 show ?thesis proof(cases rule: less_plus_elim) case Left note n2 = Left hence j2: "j2 < f (al ! n2)" using snoc by auto show ?thesis apply(rule snoc(1)) using snoc using pos n1 j1 n2 j2 by auto next case Right hence n2: "n2 = length al" by simp hence j2: "j2 < f a" using snoc by simp have "lsum f (take n1 al) + j1 = lsum f al + j2" using n1 n2 snoc by simp hence False using lsum1 by auto thus ?thesis by simp qed next case Right hence n1: "n1 = length al" by simp hence j1: "j1 < f a" using snoc by simp from n2 show ?thesis proof(cases rule: less_plus_elim) case Left note n2 = Left hence j2: "j2 < f (al ! n2)" using snoc by auto obtain al' where al: "al = (take n2 al) @ ((al ! n2) # al')" using n2 by (metis append_take_drop_id Cons_nth_drop_Suc) have "j2 < lsum f ((al ! n2) # al')" using pos j2 unfolding lsum_def by simp hence "lsum f (take n2 al) + j2 < lsum f (take n2 al) + lsum f ((al ! n2) # al')" by simp also have "... = lsum f al" unfolding lsum_append[THEN sym] using al by simp finally have lsum2: "lsum f (take n2 al) + j2 < lsum f al" . have "lsum f al + j1 = lsum f (take n2 al) + j2" using n1 n2 snoc by simp hence False using lsum2 by auto thus ?thesis by simp next case Right hence n2: "n2 = length al" by simp have "j1 = j2" using n1 n2 snoc by simp thus ?thesis using n1 n2 by simp qed qed qed auto definition locate_pred where "locate_pred f al (i::nat) n_j \ fst n_j < length al \ snd n_j < f (al ! (fst n_j)) \ i = lsum f (take (fst n_j) al) + (snd n_j)" definition locate where "locate f al i \ SOME n_j. locate_pred f al i n_j" definition locate1 where "locate1 f al i \ fst (locate f al i)" definition locate2 where "locate2 f al i \ snd (locate f al i)" lemma locate_pred_ex: assumes "\ a. a \ set al \ 0 < f a" and "i < lsum f al" shows "\ n_j. locate_pred f al i n_j" using assms less_lsum_nchotomy unfolding locate_pred_def by force lemma locate_pred_unique: assumes "\ a. a \ set al \ 0 < f a" and "locate_pred f al i n1_j1" "locate_pred f al i n2_j2" shows "n1_j1 = n2_j2" using assms less_lsum_unique unfolding locate_pred_def apply(cases n1_j1, cases n2_j2) apply simp by blast lemma locate_locate_pred: assumes "\ a. a \ set al \ (0::nat) < f a" and "i < lsum f al" shows "locate_pred f al i (locate f al i)" proof- obtain n_j where "locate_pred f al i n_j" using assms locate_pred_ex[of al f i] by auto thus ?thesis unfolding locate_def apply(intro someI[of "locate_pred f al i"]) by blast qed lemma locate_locate_pred_unique: assumes "\ a. a \ set al \ (0::nat) < f a" and "locate_pred f al i n_j" shows "n_j = locate f al i" unfolding locate_def apply(rule sym, rule some_equality) using assms locate_locate_pred apply force using assms locate_pred_unique by blast lemma locate: assumes "\ a. a \ set al \ 0 < f a" and "i < lsum f al" shows "locate1 f al i < length al \ locate2 f al i < f (al ! (locate1 f al i)) \ i = lsum f (take (locate1 f al i) al) + (locate2 f al i)" using assms locate_locate_pred unfolding locate1_def locate2_def locate_pred_def by auto lemma locate_unique: assumes "\ a. a \ set al \ 0 < f a" and "n < length al" and "j < f (al ! n)" and "i = lsum f (take n al) + j" shows "n = locate1 f al i \ j = locate2 f al i" proof- define n_j where "n_j = (n,j)" have "locate_pred f al i n_j" using assms unfolding n_j_def locate_pred_def by auto hence "n_j = locate f al i" using assms locate_locate_pred_unique by blast thus ?thesis unfolding n_j_def locate1_def locate2_def by (metis fst_conv n_j_def snd_conv) qed text\sum:\ lemma sum_2[simp]: "sum f {..< 2} = f 0 + f (Suc 0)" proof- have "{..< 2} = {0, Suc 0}" by auto thus ?thesis by force qed lemma inj_Plus[simp]: "inj ((+) (a::nat))" unfolding inj_on_def by auto lemma inj_on_Plus[simp]: "inj_on ((+) (a::nat)) A" unfolding inj_on_def by auto lemma Plus_int[simp]: fixes a :: nat shows "(+) b ` {..< a} = {b ..< b + a}" proof safe fix x::nat assume "x \ {b..< b + a}" hence "b \ x" and x: "x < a + b" by auto then obtain y where xb: "x = b + y" by (metis le_iff_add) hence "y < a" using x by simp thus "x \ (+) b ` {.. n. n \ N \ finite (B1 n)" and int: "\ m n. {m, n} \ N \ m \ n \ B1 m \ B1 n = {}" and ss: "\ n. n \ N \ sum f1 (B1 n) = b2 n" shows "sum f1 A1 = a2" (is "?L = a2") proof- have "?L = sum (%n. sum f1 (B1 n)) N" unfolding A1 using sum.UNION_disjoint[of N B1 f1] fin int by simp also have "... = sum b2 N" using ss fin by auto also have "... = a2" using a2 by simp finally show ?thesis . qed lemma sum_UN_intro: assumes A1: "A1 = (UN n : N. B1 n)" and A2: "A2 = (UN n : N. B2 n)" and fin: "finite N" "\ n. n \ N \ finite (B1 n) \ finite (B2 n)" and int: "\ m n. {m, n} \ N \ m \ n \ B1 m \ B1 n = {}" "\ m n. {m, n} \ N \ B2 m \ B2 n = {}" and ss: "\ n. n \ N \ sum f1 (B1 n) = sum f2 (B2 n)" shows "sum f1 A1 = sum f2 A2" (is "?L = ?R") proof- have "?L = sum (%n. sum f1 (B1 n)) N" unfolding A1 using sum.UNION_disjoint[of N B1 f1] fin int by simp also have "... = sum (%n. sum f2 (B2 n)) N" using ss fin sum.mono_neutral_left by auto also have "... = ?R" unfolding A2 using sum.UNION_disjoint[of N B2 f2] fin int by simp finally show ?thesis . qed lemma sum_Minus_intro: fixes f1 :: "'a1 \ real" and f2 :: "'a2 \ real" assumes "B1 = A1 - {a1}" and "B2 = A2 - {a2}" and "a1 : A1" and "a2 : A2" and "finite A1" and "finite A2" "sum f1 A1 = sum f2 A2" and "f1 a1 = f2 a2" shows "sum f1 B1 = sum f2 B2" proof- have 1: "A1 = B1 Un {a1}" and 2: "A2 = B2 Un {a2}" using assms by blast+ from assms have "a1 \ B1" by simp with 1 \finite A1\ have "sum f1 A1 = sum f1 B1 + f1 a1" by simp hence 3: "sum f1 B1 = sum f1 A1 - f1 a1" by simp from assms have "a2 \ B2" by simp with 2 \finite A2 \have "sum f2 A2 = sum f2 B2 + f2 a2" by simp hence "sum f2 B2 = sum f2 A2 - f2 a2" by simp thus ?thesis using 3 assms by simp qed lemma sum_singl_intro: assumes "b = f a" and "finite A" and "a \ A" and "\ a'. \a' \ A; a' \ a\ \ f a' = 0" shows "sum f A = b" proof- define B where "B = A - {a}" have "A = B Un {a}" unfolding B_def using assms by blast moreover have "B Int {a} = {}" unfolding B_def by blast ultimately have "sum f A = sum f B + sum f {a}" using assms sum.union_disjoint by blast moreover have "\ b \ B. f b = 0" using assms unfolding B_def by auto ultimately show ?thesis using assms by auto qed lemma sum_all0_intro: assumes "b = 0" and "\ a. a \ A \ f a = 0" shows "sum f A = b" apply(cases "finite A") by(metis assms sum.neutral sum.infinite)+ lemma sum_1: assumes I: "finite I" and ss: "sum f I = 1" and i: "i \ I - I1" and I1: "I1 \ I" and f: "\i. i \ I \ (0::real) \ f i" shows "f i \ 1 - sum f I1" proof- have "sum f I = sum f ({i} Un (I - {i}))" using i by (metis DiffE insert_Diff_single insert_absorb insert_is_Un) also have "... = sum f {i} + sum f (I - {i})" apply(rule sum.union_disjoint) using I by auto finally have "1 = f i + sum f (I - {i})" unfolding ss[THEN sym] by simp moreover have "sum f (I - {i}) \ sum f I1" apply(rule sum_mono2) using assms by auto ultimately have "1 \ f i + sum f I1" by auto thus ?thesis by auto qed subsection \Syntax\ datatype ('test, 'atom, 'choice) cmd = Done | Atm "'atom" | Seq "('test, 'atom, 'choice) cmd" "('test, 'atom, 'choice) cmd" ("_ ;; _" [60, 61] 60) | While "'test" "('test, 'atom, 'choice) cmd" | Ch 'choice "('test, 'atom, 'choice) cmd" "('test, 'atom, 'choice) cmd" | Par "('test, 'atom, 'choice) cmd list" | ParT "('test, 'atom, 'choice) cmd list" (* Commands containing no while loops: *) fun noWhile where "noWhile Done \ True" | "noWhile (Atm atm) \ True" | "noWhile (c1 ;; c2) \ noWhile c1 \ noWhile c2" | "noWhile (While tst c) \ False" | "noWhile (Ch ch c1 c2) \ noWhile c1 \ noWhile c2" | "noWhile (Par cl) \ (\ c \ set cl. noWhile c)" | "noWhile (ParT cl) \ (\ c \ set cl. noWhile c)" (* ``Finished" commands: *) fun finished where "finished Done \ True" | "finished (Atm atm) \ False" | "finished (c1 ;; c2) \ False" | "finished (While tst c) \ False" | "finished (Ch ch c1 c2) \ False" | "finished (Par cl) \ (\ c \ set cl. finished c)" | "finished (ParT cl) \ (\ c \ set cl. finished c)" definition noWhileL where "noWhileL cl \ \ c \ set cl. noWhile c" lemma fin_Par_noWhileL[simp]: "noWhile (Par cl) \ noWhileL cl" unfolding noWhileL_def by simp lemma fin_ParT_noWhileL[simp]: "noWhile (ParT cl) \ noWhileL cl" unfolding noWhileL_def by simp declare noWhile.simps(6) [simp del] declare noWhile.simps(7) [simp del] lemma noWhileL_intro[intro]: assumes "\ c. c \ set cl \ noWhile c" shows "noWhileL cl" using assms unfolding noWhileL_def by auto lemma noWhileL_fin[simp]: assumes "noWhileL cl" and "c \ set cl" shows "noWhile c" using assms unfolding noWhileL_def by simp lemma noWhileL_update[simp]: assumes cl: "noWhileL cl" and c': "noWhile c'" shows "noWhileL (cl[n := c'])" proof(cases "n < length cl") case True show ?thesis unfolding noWhileL_def proof safe fix c assume "c \ set (cl[n := c'])" hence "c \ insert c' (set cl)" using set_update_subset_insert by fastforce thus "noWhile c" using assms by (cases "c = c'") auto qed qed (insert cl, auto) definition finishedL where "finishedL cl \ \ c \ set cl. finished c" lemma finished_Par_finishedL[simp]: "finished (Par cl) \ finishedL cl" unfolding finishedL_def by simp lemma finished_ParT_finishedL[simp]: "finished (ParT cl) \ finishedL cl" unfolding finishedL_def by simp declare finished.simps(6) [simp del] declare finished.simps(7) [simp del] lemma finishedL_intro[intro]: assumes "\ c. c \ set cl \ finished c" shows "finishedL cl" using assms unfolding finishedL_def by auto lemma finishedL_finished[simp]: assumes "finishedL cl" and "c \ set cl" shows "finished c" using assms unfolding finishedL_def by simp lemma finishedL_update[simp]: assumes cl: "finishedL cl" and c': "finished c'" shows "finishedL (cl[n := c'])" proof(cases "n < length cl") case True show ?thesis unfolding finishedL_def proof safe fix c assume "c \ set (cl[n := c'])" hence "c \ insert c' (set cl)" using set_update_subset_insert by fastforce thus "finished c" using assms by (cases "c = c'") auto qed qed (insert cl, auto) lemma finished_fin[simp]: "finished c \ noWhile c" by(induct c) auto lemma finishedL_noWhileL[simp]: "finishedL cl \ noWhileL cl" unfolding finishedL_def noWhileL_def by auto locale PL = fixes aval :: "'atom \ 'state \ 'state" and tval :: "'test \ 'state \ bool" and cval :: "'choice \ 'state \ real" assumes properCh: "\ ch s. 0 \ cval ch s \ cval ch s \ 1" begin lemma [simp]: "(n::nat) < N \ 0 \ 1 / N" by auto lemma [simp]: "(n::nat) < N \ 1 / N \ 1" by auto lemma [simp]: "(n::nat) < N \ 0 \ 1 - 1 / N" by (simp add: divide_simps) lemma sum_equal: "0 < (N::nat) \ sum (\ n. 1/N) {..< N} = 1" unfolding sum_constant by auto fun proper where "proper Done \ True" | "proper (Atm x) \ True" | "proper (Seq c1 c2) \ proper c1 \ proper c2" | "proper (While tst c) \ proper c" | "proper (Ch ch c1 c2) \ proper c1 \ proper c2" | "proper (Par cl) \ cl \ [] \ (\ c \ set cl. proper c)" | "proper (ParT cl) \ cl \ [] \ (\ c \ set cl. proper c)" definition properL where "properL cl \ cl \ [] \ (\ c \ set cl. proper c)" lemma proper_Par_properL[simp]: "proper (Par cl) \ properL cl" unfolding properL_def by simp lemma proper_ParT_properL[simp]: "proper (ParT cl) \ properL cl" unfolding properL_def by simp declare proper.simps(6) [simp del] declare proper.simps(7) [simp del] lemma properL_intro[intro]: "\cl \ []; \ c. c \ set cl \ proper c\ \ properL cl" unfolding properL_def by auto lemma properL_notEmp[simp]: "properL cl \ cl \ []" unfolding properL_def by simp lemma properL_proper[simp]: "\properL cl; c \ set cl\ \ proper c" unfolding properL_def by simp lemma properL_update[simp]: assumes cl: "properL cl" and c': "proper c'" shows "properL (cl[n := c'])" proof(cases "n < length cl") case True show ?thesis unfolding properL_def proof safe fix c assume "c \ set (cl[n := c'])" hence "c \ insert c' (set cl)" using set_update_subset_insert by fastforce thus "proper c" using assms by (cases "c = c'") auto qed (insert cl, auto) qed (insert cl, auto) lemma proper_induct[consumes 1, case_names Done Atm Seq While Ch Par ParT]: assumes *: "proper c" and Done: "phi Done" and Atm: "\ atm. phi (Atm atm)" and Seq: "\ c1 c2. \phi c1; phi c2\ \ phi (c1 ;; c2)" and While: "\ tst c. phi c \ phi (While tst c)" and Ch: "\ ch c1 c2. \phi c1; phi c2\ \ phi (Ch ch c1 c2)" and Par: "\ cl. \properL cl; \ c. c \ set cl \ phi c\ \ phi (Par cl)" and ParT: "\ cl. \properL cl; \ c. c \ set cl \ phi c\ \ phi (ParT cl)" shows "phi c" using * apply(induct c) using assms unfolding properL_def by auto subsubsection \Operational Small-Step Semantics\ (* "The Finished Threads": The sublist of finished threads from a list of threads *) definition "theFT cl \ {n. n < length cl \ finished (cl!n)}" (* "The NopnFinished Threads": *) definition "theNFT cl \ {n. n < length cl \ \ finished (cl!n)}" lemma finite_theFT[simp]: "finite (theFT cl)" unfolding theFT_def by simp lemma theFT_length[simp]: "n \ theFT cl \ n < length cl" unfolding theFT_def by simp lemma theFT_finished[simp]: "n \ theFT cl \ finished (cl!n)" unfolding theFT_def by simp lemma finite_theNFT[simp]: "finite (theNFT cl)" unfolding theNFT_def by simp lemma theNFT_length[simp]: "n \ theNFT cl \ n < length cl" unfolding theNFT_def by simp lemma theNFT_notFinished[simp]: "n \ theNFT cl \ \ finished (cl!n)" unfolding theNFT_def by simp lemma theFT_Int_theNFT[simp]: "theFT cl Int theNFT cl = {}" and "theNFT cl Int theFT cl = {}" unfolding theFT_def theNFT_def by auto lemma theFT_Un_theNFT[simp]: "theFT cl Un theNFT cl = {..< length cl}" and "theNFT cl Un theFT cl = {..< length cl}" unfolding theFT_def theNFT_def by auto lemma in_theFT_theNFT[simp]: assumes "n1 \ theFT cl" and "n2 \ theNFT cl" shows "n1 \ n2" and "n2 \ n1" using assms theFT_Int_theNFT by blast+ (* The cumulated weight of the finished threads: *) definition "WtFT cl \ sum (\ (n::nat). 1/(length cl)) (theFT cl)" (* The cumulated weight of the non-finished threads: *) definition "WtNFT cl \ sum (\ (n::nat). 1/(length cl)) (theNFT cl)" lemma WtFT_WtNFT[simp]: assumes "0 < length cl" shows "WtFT cl + WtNFT cl = 1" (is "?A = 1") proof- let ?w = "\ n. 1 / length cl" let ?L = "theFT cl" let ?Lnot = "theNFT cl" have 1: "{..< length cl} = ?L Un ?Lnot" by auto have "?A = sum ?w ?L + sum ?w ?Lnot" unfolding WtFT_def WtNFT_def by simp also have "... = sum ?w {..< length cl}" unfolding 1 apply(rule sum.union_disjoint[THEN sym]) by auto also have "... = 1" unfolding sum_equal[OF assms] by auto finally show ?thesis . qed lemma WtNFT_1_WtFT: "0 < length cl \ WtNFT cl = 1 - WtFT cl" by (simp add: algebra_simps) lemma WtNFT_WtFT_1[simp]: assumes "0 < length cl" and "WtFT cl \ 1" shows "WtNFT cl / (1 - WtFT cl) = 1" (is "?A / ?B = 1") proof- have A: "?A = ?B" using assms WtNFT_1_WtFT by auto show ?thesis unfolding A using assms by auto qed lemma WtFT_ge_0[simp]: "WtFT cl \ 0" unfolding WtFT_def by (rule sum_nonneg) simp lemma WtFT_le_1[simp]: "WtFT cl \ 1" (is "?L \ 1") proof- let ?N = "length cl" have "?L \ sum (\ n::nat. 1/?N) {..< ?N}" unfolding WtFT_def apply(rule sum_mono2) by auto also have "... \ 1" - by (metis div_by_0 le_cases neq0_conv not_one_le_zero of_nat_0 sum_not_0 sum_equal) + by (metis div_by_0 le_cases neq0_conv not_one_le_zero of_nat_0 sum.neutral sum_equal) finally show ?thesis . qed lemma le_1_WtFT[simp]: "0 \ 1 - WtFT cl" (is "0 \ ?R") proof- have a: "-1 \ - WtFT cl" by simp have "(0 :: real) = 1 + (-1)" by simp also have "1 + (-1) \ 1 + (- WtFT cl)" using a by arith also have "... = ?R" by simp finally show ?thesis . qed lemma WtFT_lt_1[simp]: "WtFT cl \ 1 \ WtFT cl < 1" using WtFT_le_1 by (auto simp add: le_less) lemma lt_1_WtFT[simp]: "WtFT cl \ 1 \ 0 < 1 - WtFT cl" using le_1_WtFT by (metis le_1_WtFT eq_iff_diff_eq_0 less_eq_real_def) lemma notFinished_WtFT[simp]: assumes "n < length cl" and "\ finished (cl ! n)" shows "1 / length cl \ 1 - WtFT cl" proof- have "0 < length cl" using assms by auto thus ?thesis unfolding WtFT_def apply(intro sum_1[of "{..< length cl}"]) using assms by auto qed (* The branching of a command: *) fun brn :: "('test, 'atom, 'choice) cmd \ nat" where "brn Done = 1" | "brn (Atm atm) = 1" | "brn (c1 ;; c2) = brn c1" | "brn (While tst c) = 1" | "brn (Ch ch c1 c2) = 2" | "brn (Par cl) = lsum brn cl" | "brn (ParT cl) = lsum brn cl" lemma brn_gt_0: "proper c \ 0 < brn c" by (induct rule: proper_induct) auto lemma brn_gt_0_L: "\properL cl; c \ set cl\ \ 0 < brn c" by (metis brn_gt_0 properL_def) (* The locate-thread and locate-index operators. Given a thread list cl with n = length cl and i < (\ l < length cl . brn cl), (locateT cl i, locateI cl i) are (k, j) from the paper's Figure 1. *) definition "locateT \ locate1 brn" definition "locateI \ locate2 brn" definition "brnL cl n \ lsum brn (take n cl)" lemma brnL_lsum: "brnL cl (length cl) = lsum brn cl" unfolding brnL_def by simp lemma brnL_unique: assumes "properL cl" and "n1 < length cl \ j1 < brn (cl ! n1)" and "n2 < length cl \ j2 < brn (cl ! n2)" and "brnL cl n1 + j1 = brnL cl n2 + j2" shows "n1 = n2 \ j1 = j2" apply (rule less_lsum_unique) using assms brn_gt_0 unfolding brnL_def properL_def by auto lemma brn_Par_simp[simp]: "brn (Par cl) = brnL cl (length cl)" unfolding brnL_lsum by simp lemma brn_ParT_simp[simp]: "brn (ParT cl) = brnL cl (length cl)" unfolding brnL_lsum by simp declare brn.simps(6)[simp del] declare brn.simps(7)[simp del] lemma brnL_0[simp]: "brnL cl 0 = 0" unfolding brnL_def by auto lemma brnL_Suc[simp]: "n < length cl \ brnL cl (Suc n) = brnL cl n + brn (cl ! n)" unfolding brnL_def using take_Suc_conv_app_nth[of n cl] by simp lemma brnL_mono[simp]: "n1 \ n2 \ brnL cl n1 \ brnL cl n2" using le_take[of n1 n2 cl] unfolding brnL_def by simp lemma brnL_mono2[simp]: assumes p: "properL cl" and n: "n1 < n2" and l: "n2 \ length cl" shows "brnL cl n1 < brnL cl n2" (is "?L < ?R") proof- have 1: "\c. c \ set (take n2 cl) \ 0 < brn c" using p by (metis brn_gt_0 in_set_takeD properL_proper) have "take n1 cl < take n2 cl" using n l lt_take by auto hence "lsum brn (take n1 cl) < lsum brn (take n2 cl)" using lsum_mono2[of "take n2 cl" "%c. brn c" "take n1 cl"] 1 by simp thus ?thesis unfolding brnL_def . qed lemma brn_index[simp]: assumes n: "n < length cl" and i: "i < brn (cl ! n)" shows "brnL cl n + i < brnL cl (length cl)" (is "?L < ?R") proof- have "?L < brnL cl (Suc n)" using assms by simp also have "... \ ?R" using n brnL_mono[of "Suc n" "length cl" cl] by simp finally show ?thesis . qed lemma brnL_gt_0[simp]: "\properL cl; 0 < n\ \ 0 < brnL cl n" by (metis properL_def brnL_mono brnL_mono2 le_0_eq length_greater_0_conv nat_le_linear neq0_conv) lemma locateTI: assumes "properL cl" and "ii < brn (Par cl)" shows "locateT cl ii < length cl \ locateI cl ii < brn (cl ! (locateT cl ii)) \ ii = brnL cl (locateT cl ii) + locateI cl ii" using assms locate[of cl brn] brn_gt_0 unfolding locateT_def locateI_def brnL_def unfolding brnL_lsum[THEN sym] by auto lemma locateTI_unique: assumes "properL cl" and "n < length cl" and "i < brn (cl ! n)" and "ii = brnL cl n + i" shows "n = locateT cl ii \ i = locateI cl ii" using assms locate_unique[of cl brn] brn_gt_0 unfolding locateT_def locateI_def brnL_def unfolding brnL_lsum[THEN sym] by auto (* pickFT picks a finished thread (if there is any). It will be used to perform a dummy transition in case the cumulated weight of the finished threads is 0; specifically, one will assign probability 1 to that picked fresh. (Obviously, the particular choice does not matter.) *) definition pickFT_pred where "pickFT_pred cl n \ n < length cl \ finished (cl!n)" definition pickFT where "pickFT cl \ SOME n. pickFT_pred cl n" lemma pickFT_pred: assumes "WtFT cl = 1" shows "\ n. pickFT_pred cl n" proof(rule ccontr, unfold not_ex) assume "\n. \ pickFT_pred cl n" hence "\ n. n < length cl \ \ finished (cl!n)" unfolding pickFT_pred_def by auto hence "theFT cl = {}" unfolding theFT_def by auto hence "WtFT cl = 0" unfolding WtFT_def by simp thus False using assms by simp qed lemma pickFT_pred_pickFT: "WtFT cl = 1 \ pickFT_pred cl (pickFT cl)" unfolding pickFT_def by(auto intro: someI_ex pickFT_pred) lemma pickFT_length[simp]: "WtFT cl = 1 \ pickFT cl < length cl" using pickFT_pred_pickFT unfolding pickFT_pred_def by auto lemma pickFT_finished[simp]: "WtFT cl = 1 \ finished (cl ! (pickFT cl))" using pickFT_pred_pickFT unfolding pickFT_pred_def by auto lemma pickFT_theFT[simp]: "WtFT cl = 1 \ pickFT cl \ theFT cl" unfolding theFT_def by auto (* The weight, continuation and effect defined as a single operator: *) fun wt_cont_eff where "wt_cont_eff Done s i = (1, Done, s)" | "wt_cont_eff (Atm atm) s i = (1, Done, aval atm s)" | "wt_cont_eff (c1 ;; c2) s i = (case wt_cont_eff c1 s i of (x, c1', s') \ if finished c1' then (x, c2, s') else (x, c1' ;; c2, s'))" | "wt_cont_eff (While tst c) s i = (if tval tst s then (1, c ;; (While tst c), s) else (1, Done, s))" | "wt_cont_eff (Ch ch c1 c2) s i = (if i = 0 then cval ch s else 1 - cval ch s, if i = 0 then c1 else c2, s)" | "wt_cont_eff (Par cl) s ii = (if cl ! (locateT cl ii) \ set cl then (case wt_cont_eff (cl ! (locateT cl ii)) s (locateI cl ii) of (w, c', s') \ ((1 / (length cl)) * w, Par (cl [(locateT cl ii) := c']), s')) else undefined)" | "wt_cont_eff (ParT cl) s ii = (if cl ! (locateT cl ii) \ set cl then (case wt_cont_eff (cl ! (locateT cl ii)) s (locateI cl ii) of (w, c', s') \ (if WtFT cl = 1 then (if locateT cl ii = pickFT cl \ locateI cl ii = 0 then 1 else 0) else if finished (cl ! (locateT cl ii)) then 0 else (1 / (length cl)) / (1 - WtFT cl) * w, ParT (cl [(locateT cl ii) := c']), s')) else undefined)" (* weight, cont (transition) and effect: *) definition wt where "wt c s i = fst (wt_cont_eff c s i)" definition cont where "cont c s i = fst (snd (wt_cont_eff c s i))" definition eff where "eff c s i = snd (snd (wt_cont_eff c s i))" (* Their individual equations (corresponding to the paper's Figure 1): *) lemma wt_Done[simp]: "wt Done s i = 1" unfolding wt_def by simp lemma wt_Atm[simp]: "wt (Atm atm) s i = 1" unfolding wt_def by simp lemma wt_Seq[simp]: "wt (c1 ;; c2) s = wt c1 s" proof- {fix i have "wt (c1 ;; c2) s i = wt c1 s i" proof(cases "wt_cont_eff c1 s i") case (fields _ c1' _) thus ?thesis unfolding wt_def by(cases c1', auto) qed } thus ?thesis by auto qed lemma wt_While[simp]: "wt (While tst c) s i = 1" unfolding wt_def by simp lemma wt_Ch_L[simp]: "wt (Ch ch c1 c2) s 0 = cval ch s" unfolding wt_def by simp lemma wt_Ch_R[simp]: "wt (Ch ch c1 c2) s (Suc n) = 1 - cval ch s" unfolding wt_def by simp (* *) lemma cont_Done[simp]: "cont Done s i = Done" unfolding cont_def by simp lemma cont_Atm[simp]: "cont (Atm atm) s i = Done" unfolding cont_def by simp lemma cont_Seq_finished[simp]: "finished (cont c1 s i) \ cont (c1 ;; c2) s i = c2" unfolding cont_def by(cases "wt_cont_eff c1 s i") auto lemma cont_Seq_notFinished[simp]: assumes "\ finished (cont c1 s i)" shows "cont (c1 ;; c2) s i = (cont c1 s i) ;; c2" proof(cases "wt_cont_eff c1 s i") case (fields _ c1' _) thus ?thesis using assms unfolding cont_def by(cases c1') auto qed lemma cont_Seq_not_eq_finished[simp]: "\ finished c2 \ \ finished (cont (Seq c1 c2) s i)" by (cases "finished (cont c1 s i)") auto lemma cont_While_False[simp]: "tval tst s = False \ cont (While tst c) s i = Done" unfolding cont_def by simp lemma cont_While_True[simp]: "tval tst s = True \ cont (While tst c) s i = c ;; (While tst c)" unfolding cont_def by simp lemma cont_Ch_L[simp]: "cont (Ch ch c1 c2) s 0 = c1" unfolding cont_def by simp lemma cont_Ch_R[simp]: "cont (Ch ch c1 c2) s (Suc n) = c2" unfolding cont_def by simp (* *) lemma eff_Done[simp]: "eff Done s i = s" unfolding eff_def by simp lemma eff_Atm[simp]: "eff (Atm atm) s i = aval atm s" unfolding eff_def by simp lemma eff_Seq[simp]: "eff (c1 ;; c2) s = eff c1 s" proof- {fix i have "eff (c1 ;; c2) s i = eff c1 s i" proof(cases "wt_cont_eff c1 s i") case (fields _ c1' _) thus ?thesis unfolding eff_def by(cases c1') auto qed } thus ?thesis by auto qed lemma eff_While[simp]: "eff (While tst c) s i = s" unfolding eff_def by simp lemma eff_Ch[simp]: "eff (Ch ch c1 c2) s i = s" unfolding eff_def by simp (* wt-cont-eff simps for parallel composition *) lemma brnL_nchotomy: assumes "properL cl" and "ii < brnL cl (length cl)" shows "\ n i. n < length cl \ i < brn (cl ! n) \ ii = brnL cl n + i" unfolding brnL_def apply(rule less_lsum_nchotomy) using assms brn_gt_0 unfolding brnL_lsum[THEN sym] by auto corollary brnL_cases[consumes 2, case_names Local, elim]: assumes "properL cl" and "ii < brnL cl (length cl)" and "\ n i. \n < length cl; i < brn (cl ! n); ii = brnL cl n + i\ \ phi" shows phi using assms brnL_nchotomy by blast lemma wt_cont_eff_Par[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" shows "wt (Par cl) s (brnL cl n + i) = 1 / (length cl) * wt (cl ! n) s i" (is "?wL = ?wR") (* *) "cont (Par cl) s (brnL cl n + i) = Par (cl [n := cont (cl ! n) s i])" (is "?mL = ?mR") (* *) "eff (Par cl) s (brnL cl n + i) = eff (cl ! n) s i" (is "?eL = ?eR") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have "?wL = 1 / (length cl)* wt (cl ! n1) s i1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding wt_def by simp also have "... = ?wR" unfolding n_i by simp finally show "?wL = ?wR" . (* *) have "?mL = Par (cl [n1 := cont (cl ! n1) s i1])" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding cont_def by simp also have "... = ?mR" unfolding n_i by simp finally show "?mL = ?mR" . (* *) have "?eL = eff (cl ! n1) s i1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding eff_def by simp also have "eff (cl ! n1) s i1 = ?eR" unfolding n_i by simp finally show "?eL = ?eR" . qed lemma cont_eff_ParT[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" shows "cont (ParT cl) s (brnL cl n + i) = ParT (cl [n := cont (cl ! n) s i])" (is "?mL = ?mR") (* *) "eff (ParT cl) s (brnL cl n + i) = eff (cl ! n) s i" (is "?eL = ?eR") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have "?mL = ParT (cl [n1 := cont (cl ! n1) s i1])" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding cont_def by simp also have "... = ?mR" unfolding n_i by simp finally show "?mL = ?mR" . (* *) have "?eL = eff (cl ! n1) s i1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding eff_def by simp also have "eff (cl ! n1) s i1 = ?eR" unfolding n_i by simp finally show "?eL = ?eR" . qed lemma wt_ParT_WtFT_pickFT_0[simp]: assumes p: "properL cl" and WtFT: "WtFT cl = 1" shows "wt (ParT cl) s (brnL cl (pickFT cl)) = 1" (is "?wL = 1") proof- define ii where "ii = brnL cl (pickFT cl)" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have ni: "pickFT cl < length cl" "0 < brn (cl! (pickFT cl))" using WtFT p brn_gt_0 by auto hence n_i: "pickFT cl = n1" "0 = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique[of cl "pickFT cl" 0 ii] by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using ni by (metis add.comm_neutral brn_index) hence "n1 < length cl" unfolding n1_def using ni p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have n1i1: "n1 = pickFT cl \ i1 = 0" using WtFT ni unfolding n_i by auto show "?wL = 1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using WtFT n1i1 set unfolding n1_def i1_def unfolding wt_def by simp qed lemma wt_ParT_WtFT_notPickFT_0[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" and WtFT: "WtFT cl = 1" and ni: "n = pickFT cl \ i \ 0" shows "wt (ParT cl) s (brnL cl n + i) = 0" (is "?wL = 0") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have n1i1: "n1 \ pickFT cl \ i1 \ 0" using WtFT ni unfolding n_i by auto show "?wL = 0" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using WtFT n1i1 set unfolding n1_def i1_def unfolding wt_def by force qed lemma wt_ParT_notWtFT_finished[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" and WtFT: "WtFT cl \ 1" and f: "finished (cl ! n)" shows "wt (ParT cl) s (brnL cl n + i) = 0" (is "?wL = 0") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have f: "finished (cl ! n1)" using f unfolding n_i by auto show "?wL = 0" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using WtFT f set unfolding n1_def i1_def unfolding wt_def by simp qed lemma wt_cont_eff_ParT_notWtFT_notFinished[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" and WtFT: "WtFT cl \ 1" and nf: "\ finished (cl ! n)" shows "wt (ParT cl) s (brnL cl n + i) = (1 / (length cl)) / (1 - WtFT cl) * wt (cl ! n) s i" (is "?wL = ?wR") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have nf: "\ finished (cl ! n1)" using nf unfolding n_i by auto have "?wL = (1 / (length cl)) / (1 - WtFT cl) * wt (cl ! n1) s i1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using WtFT nf set unfolding n1_def i1_def unfolding wt_def by simp also have "... = ?wR" unfolding n_i by simp finally show "?wL = ?wR" . qed (* *) lemma wt_ge_0[simp]: assumes "proper c" and "i < brn c" shows "0 \ wt c s i" using assms proof (induct c arbitrary: i s rule: proper_induct) case (Ch ch c1 c2) thus ?case using properCh by (cases i) (auto simp: algebra_simps) next case (Par cl ii) have "properL cl" and "ii < brnL cl (length cl)" using Par by auto thus ?case apply (cases rule: brnL_cases) using Par by simp next case (ParT cl ii) have "properL cl" and "ii < brnL cl (length cl)" using ParT by auto thus ?case proof(cases rule: brnL_cases) case (Local n i) show ?thesis proof (cases "WtFT cl = 1") case True thus ?thesis using Local ParT by (cases "n = pickFT cl \ i = 0") auto next case False thus ?thesis using Local ParT by (cases "finished (cl ! n)") auto qed qed qed auto lemma wt_le_1[simp]: assumes "proper c" and "i < brn c" shows "wt c s i \ 1" using assms proof (induct c arbitrary: i s rule: proper_induct) case (Ch ch c1 c2) thus ?case using properCh by (cases i) auto next case (Par cl ii) hence "properL cl" and "ii < brnL cl (length cl)" by auto thus ?case apply (cases rule: brnL_cases) apply simp using Par apply auto by (metis add_increasing2 diff_is_0_eq gr0_conv_Suc less_imp_diff_less less_or_eq_imp_le nth_mem of_nat_0_le_iff of_nat_Suc) next case (ParT cl ii) have "properL cl" and "ii < brnL cl (length cl)" using ParT by auto thus ?case proof(cases rule: brnL_cases) case (Local n i) show ?thesis proof (cases "WtFT cl = 1") case True thus ?thesis using Local ParT by (cases "n = pickFT cl \ i = 0", auto) next case False note sch = False thus ?thesis using Local ParT proof (cases "finished (cl ! n)") case False note cln = False let ?L1 = "1 / (length cl)" let ?L2 = "wt (cl ! n) s i" let ?R = "WtFT cl" have "0 \ ?L1" and "0 \ ?L2" using ParT Local by auto moreover have "?L2 \ 1" using ParT Local by auto ultimately have "?L1 * ?L2 \ ?L1" by (metis mult_right_le_one_le) also have "?L1 \ 1 - ?R" using ParT Local cln by auto finally have "?L1 * ?L2 \ 1 - ?R" . thus ?thesis using Local ParT cln sch by (auto simp: pos_divide_le_eq mult.commute) qed (insert sch ParT Local, auto) qed qed qed auto abbreviation fromPlus ("(1{_..<+_})") where "{a ..<+ b} \ {a ..< a + b}" lemma brnL_UN: assumes "properL cl" shows "{..< brnL cl (length cl)} = (\ n < length cl. {brnL cl n ..<+ brn (cl!n)})" (is "?L = (\ n < length cl. ?R n)") proof safe fix ii assume ii: "ii < brnL cl (length cl)" from assms ii show "ii \ (\ n < length cl. ?R n)" proof(cases rule: brnL_cases) case (Local n i) hence "ii \ ?R n" by simp thus ?thesis using Local by force qed qed auto lemma brnL_Int_lt: assumes n12: "n1 < n2" and n2: "n2 < length cl" shows "{brnL cl n1 ..<+ brn (cl!n1)} \ {brnL cl n2 ..<+ brn (cl!n2)} = {}" proof- have "Suc n1 \ n2" using assms by simp hence "brnL cl (Suc n1) \ brnL cl n2" by simp thus ?thesis using assms by simp qed lemma brnL_Int: assumes "n1 \ n2" and "n1 < length cl" and "n2 < length cl" shows "{brnL cl n1 ..<+ brn (cl!n1)} \ {brnL cl n2 ..<+ brn (cl!n2)} = {}" proof(cases "n1 < n2") case True thus ?thesis using assms brnL_Int_lt by auto next case False hence "n2 < n1" using assms by simp thus ?thesis using brnL_Int_lt assms by fastforce qed lemma sum_wt_Par_sub[simp]: assumes cl: "properL cl" and n: "n < length cl" and I: "I \ {..< brn (cl ! n)}" shows "sum (wt (Par cl) s) ((+) (brnL cl n) ` I) = 1 /(length cl) * sum (wt (cl ! n) s) I" (is "?L = ?wSch * ?R") proof- have "?L = sum (%i. ?wSch * wt (cl ! n) s i) I" apply(rule sum.reindex_cong[of "(+) (brnL cl n)"]) using assms by auto also have "... = ?wSch * ?R" unfolding sum_distrib_left by simp finally show ?thesis . qed lemma sum_wt_Par[simp]: assumes cl: "properL cl" and n: "n < length cl" shows "sum (wt (Par cl) s) {brnL cl n ..<+ brn (cl!n)} = 1 /(length cl) * sum (wt (cl ! n) s) {..< brn (cl ! n)}" (is "?L = ?W * ?R") using assms by (simp add: sum_distrib_left) lemma sum_wt_ParT_sub_WtFT_pickFT_0[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" and I: "I \ {..< brn (cl ! (pickFT cl))}" "0 \ I" shows "sum (wt (ParT cl) s) ((+) (brnL cl (pickFT cl)) ` I) = 1" (is "?L = 1") proof- let ?n = "pickFT cl" let ?w = "%i. if i = 0 then (1::real) else 0" have n: "?n < length cl" using nf by simp have 0: "I = {0} Un (I - {0})" using I by auto have finI: "finite I" using I by (metis finite_nat_iff_bounded) have "?L = sum ?w I" proof (rule sum.reindex_cong [of "plus (brnL cl ?n)"]) fix i assume i: "i \ I" have "i < brn (cl ! ?n)" using i I by auto note i = this i show "wt (ParT cl) s (brnL cl ?n + i) = ?w i" using nf n i cl by (cases "i = 0") auto qed (insert assms, auto) also have "... = sum ?w ({0} Un (I - {0}))" using 0 by auto also have "... = sum ?w {0::real} + sum ?w (I - {0})" using sum.union_disjoint[of "{0}" "I - {0}" ?w] finI by auto also have "... = 1" by simp finally show ?thesis . qed lemma sum_wt_ParT_sub_WtFT_pickFT_0_2[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" and II: "II \ {..< brnL cl (length cl)}" "brnL cl (pickFT cl) \ II" shows "sum (wt (ParT cl) s) II = 1" (is "?L = 1") proof- let ?n = "pickFT cl" let ?w = "%ii. if ii = brnL cl (pickFT cl) then (1::real) else 0" have n: "?n < length cl" using nf by simp have 0: "II = {brnL cl ?n} Un (II - {brnL cl ?n})" using II by auto have finI: "finite II" using II by (metis finite_nat_iff_bounded) have "?L = sum ?w II" proof(rule sum.cong) fix ii assume ii: "ii \ II" hence ii: "ii < brnL cl (length cl)" using II by auto from cl ii show "wt (ParT cl) s ii = ?w ii" proof(cases rule: brnL_cases) case (Local n i) show ?thesis proof(cases "ii = brnL cl (pickFT cl)") case True have "n = pickFT cl \ i = 0" apply(intro brnL_unique[of cl]) using Local cl nf brn_gt_0 unfolding True by auto thus ?thesis using cl nf True by simp next case False hence "n = pickFT cl \ i \ 0" unfolding Local by auto thus ?thesis using Local ii nf cl False by auto qed qed qed simp also have "... = sum ?w ({brnL cl ?n} Un (II - {brnL cl ?n}))" using 0 by simp also have "... = sum ?w {brnL cl ?n} + sum ?w (II - {brnL cl ?n})" apply(rule sum.union_disjoint) using finI by auto also have "... = 1" by simp finally show ?thesis . qed lemma sum_wt_ParT_sub_WtFT_notPickFT_0[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" and n: "n < length cl" and I: "I \ {..< brn (cl ! n)}" and nI: "n = pickFT cl \ 0 \ I" shows "sum (wt (ParT cl) s) ((+) (brnL cl n) ` I) = 0" (is "?L = 0") proof- have "?L = sum (%i. 0) I" proof (rule sum.reindex_cong [of "plus (brnL cl n)"]) fix i assume i: "i \ I" hence "n = pickFT cl \ i \ 0" using nI by metis moreover have "i < brn (cl ! n)" using i I by auto ultimately show "wt (ParT cl) s (brnL cl n + i) = 0" using nf n cl by simp qed (insert assms, auto) also have "... = 0" by simp finally show ?thesis . qed lemma sum_wt_ParT_sub_notWtFT_finished[simp]: assumes cl: "properL cl" and nf: "WtFT cl \ 1" and n: "n < length cl" and cln: "finished (cl!n)" and I: "I \ {..< brn (cl ! n)}" shows "sum (wt (ParT cl) s) ((+) (brnL cl n) ` I) = 0" (is "?L = 0") proof- have "?L = sum (%i. 0) I" apply(rule sum.reindex_cong[of "(+) (brnL cl n)"]) using assms by auto also have "... = 0" by simp finally show ?thesis . qed lemma sum_wt_ParT_sub_notWtFT_notFinished[simp]: assumes cl: "properL cl" and nf: "WtFT cl \ 1" and n: "n < length cl" and cln: "\ finished (cl!n)" and I: "I \ {..< brn (cl ! n)}" shows "sum (wt (ParT cl) s) ((+) (brnL cl n) ` I) = (1 / (length cl)) / (1 - WtFT cl) * sum (wt (cl ! n) s) I" (is "?L = ?w / (1 - ?wF) * ?R") proof- have "?L = sum (%i. ?w / (1 - ?wF) * wt (cl ! n) s i) I" apply(rule sum.reindex_cong[of "(+) (brnL cl n)"]) using assms by auto also have "... = ?w / (1 - ?wF) * ?R" unfolding sum_distrib_left by simp finally show ?thesis . qed lemma sum_wt_ParT_WtFT_pickFT_0[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" shows "sum (wt (ParT cl) s) {brnL cl (pickFT cl) ..<+ brn (cl ! (pickFT cl))} = 1" proof- let ?n = "pickFT cl" have 1: "{brnL cl ?n ..<+ brn (cl ! ?n)} = (+) (brnL cl ?n) ` {..< brn (cl ! ?n)}" by simp show ?thesis unfolding 1 apply(rule sum_wt_ParT_sub_WtFT_pickFT_0) using assms apply simp_all by (metis brn_gt_0_L nth_mem pickFT_length) qed lemma sum_wt_ParT_WtFT_notPickFT_0[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" and n: "n < length cl" "n \ pickFT cl" shows "sum (wt (ParT cl) s) {brnL cl n ..<+ brn (cl!n)} = 0" proof- have 1: "{brnL cl n ..<+ brn (cl!n)} = (+) (brnL cl n) ` {..< brn (cl!n)}" by simp show ?thesis unfolding 1 apply(rule sum_wt_ParT_sub_WtFT_notPickFT_0) using assms by auto qed lemma sum_wt_ParT_notWtFT_finished[simp]: assumes cl: "properL cl" and "WtFT cl \ 1" and n: "n < length cl" and cln: "finished (cl!n)" shows "sum (wt (ParT cl) s) {brnL cl n ..<+ brn (cl!n)} = 0" proof- have 1: "{brnL cl n ..<+ brn (cl!n)} = (+) (brnL cl n) ` {..< brn (cl!n)}" by simp show ?thesis unfolding 1 apply(rule sum_wt_ParT_sub_notWtFT_finished) using assms by auto qed lemma sum_wt_ParT_notWtFT_notFinished[simp]: assumes cl: "properL cl" and nf: "WtFT cl \ 1" and n: "n < length cl" and cln: "\ finished (cl!n)" shows "sum (wt (ParT cl) s) {brnL cl n ..<+ brn (cl!n)} = (1 / (length cl)) / (1 - WtFT cl) * sum (wt (cl ! n) s) {..< brn (cl ! n)}" proof- have 1: "{brnL cl n ..<+ brn (cl!n)} = (+) (brnL cl n) ` {..< brn (cl!n)}" by simp show ?thesis unfolding 1 apply(rule sum_wt_ParT_sub_notWtFT_notFinished) using assms by auto qed lemma sum_wt[simp]: assumes "proper c" shows "sum (wt c s) {..< brn c} = 1" using assms proof (induct c arbitrary: s rule: proper_induct) case (Par cl) let ?w = "\ n. 1 / (length cl) * sum (wt (cl ! n) s) {..< brn (cl ! n)}" show ?case proof (rule sum_UN_introL [of _ "%n. {brnL cl n ..<+ brn (cl!n)}" "{..< length cl}" _ ?w]) have "1 = sum (\ n. 1 / (length cl)) {..< length cl}" using Par by simp also have "... = sum ?w {..< length cl}" using Par by simp finally show "1 = sum ?w {..< length cl}" . next fix m n assume "{m, n} \ {.. m \ n" thus "{brnL cl m ..<+ brn (cl!m)} \ {brnL cl n ..<+ brn (cl!n)} = {}" using brnL_Int by auto qed(insert Par brnL_UN sum_wt_Par, auto) next case (ParT cl) let ?v = "1/(length cl)" let ?wtF = "WtFT cl" let ?wtNF = "WtNFT cl" let ?w = "\n. if ?wtF = 1 then (if n = pickFT cl then 1 else 0) else (if finished (cl!n) then 0 else ?v / (1 - ?wtF) * sum (wt (cl ! n) s) {..< brn (cl ! n)})" define w where "w = ?w" have w: "\ n. ?wtF \ 1 \ n < length cl \ \ finished (cl!n) \ w n = ?v / (1 - ?wtF)" unfolding w_def using ParT by simp show ?case proof(cases "WtFT cl = 1") case True with ParT show ?thesis by simp next case False note nf = False show ?thesis proof (rule sum_UN_introL [of _ "%n. {brnL cl n ..<+ brn (cl!n)}" "{..< length cl}" _ w]) show "1 = sum w {..< length cl}" proof(cases "?wtF = 1") case True note sch = True let ?n = "pickFT cl" let ?L = "{?n}" let ?Lnot = "{.. {..iii {..< brn c}" shows "sum (wt c s) I \ 1" proof- define J where "J = {..< brn c}" have "I \ J" and "finite J" using ** unfolding J_def by auto moreover have "\ j \ J. wt c s j \ 0" unfolding J_def using * by simp ultimately have "sum (wt c s) I \ sum (wt c s) J" using sum_mono2[of J I "wt c s"] by auto also have "... = 1" using * unfolding J_def by simp finally show "sum (wt c s) I \ 1" unfolding J_def by simp qed lemma sum_le_1[simp]: assumes *: "proper c" and **: "i < brn c" shows "sum (wt c s) {..i} \ 1" proof- have "{..i} \ {..< brn c}" using ** by auto thus ?thesis using assms sum_subset_le_1[of c "{..i}" s] by blast qed subsubsection \Operations on configurations\ definition "cont_eff cf b = snd (wt_cont_eff (fst cf) (snd cf) b)" lemma cont_eff: "cont_eff cf b = (cont (fst cf) (snd cf) b, eff (fst cf) (snd cf) b)" unfolding cont_eff_def cont_def eff_def by simp end (* context PL *) end diff --git a/thys/Smooth_Manifolds/Analysis_More.thy b/thys/Smooth_Manifolds/Analysis_More.thy --- a/thys/Smooth_Manifolds/Analysis_More.thy +++ b/thys/Smooth_Manifolds/Analysis_More.thy @@ -1,1394 +1,1394 @@ section \Library Additions\ theory Analysis_More imports "HOL-Analysis.Equivalence_Lebesgue_Henstock_Integration" "HOL-Library.Function_Algebras" "HOL-Types_To_Sets.Linear_Algebra_On" begin lemma openin_open_Int'[intro]: "open S \ openin (top_of_set U) (S \ U)" by (auto simp: openin_open) subsection \Parametricity rules for topology\ text \TODO: also check with theory \Transfer_Euclidean_Space_Vector\ in AFP/ODE...\ context includes lifting_syntax begin lemma Sigma_transfer[transfer_rule]: "(rel_set A ===> (A ===> rel_set B) ===> rel_set (rel_prod A B)) Sigma Sigma" unfolding Sigma_def by transfer_prover lemma filterlim_transfer[transfer_rule]: "((A ===> B) ===> rel_filter B ===> rel_filter A ===> (=)) filterlim filterlim" if [transfer_rule]: "bi_unique B" unfolding filterlim_iff by transfer_prover lemma nhds_transfer[transfer_rule]: "(A ===> rel_filter A) nhds nhds" if [transfer_rule]: "bi_unique A" "bi_total A" "(rel_set A ===> (=)) open open" unfolding nhds_def by transfer_prover lemma at_within_transfer[transfer_rule]: "(A ===> rel_set A ===> rel_filter A) at_within at_within" if [transfer_rule]: "bi_unique A" "bi_total A" "(rel_set A ===> (=)) open open" unfolding at_within_def by transfer_prover lemma continuous_on_transfer[transfer_rule]: "(rel_set A ===> (A ===> B) ===> (=)) continuous_on continuous_on" if [transfer_rule]: "bi_unique A" "bi_total A" "(rel_set A ===> (=)) open open" "bi_unique B" "bi_total B" "(rel_set B ===> (=)) open open" unfolding continuous_on_def by transfer_prover lemma continuous_on_transfer_right_total[transfer_rule]: "(rel_set A ===> (A ===> B) ===> (=)) (\X::'a::t2_space set. continuous_on (X \ Collect AP)) (\Y::'b::t2_space set. continuous_on Y)" if DomainA: "Domainp A = AP" and [folded DomainA, transfer_rule]: "bi_unique A" "right_total A" "(rel_set A ===> (=)) (openin (top_of_set (Collect AP))) open" "bi_unique B" "bi_total B" "(rel_set B ===> (=)) open open" unfolding DomainA[symmetric] proof (intro rel_funI) fix X Y f g assume H[transfer_rule]: "rel_set A X Y" "(A ===> B) f g" from H(1) have XA: "x \ X \ Domainp A x" for x by (auto simp: rel_set_def) then have *: "X \ Collect (Domainp A) = X" by auto have "openin (top_of_set (Collect (Domainp A))) (Collect (Domainp A))" by auto show " continuous_on (X \ Collect (Domainp A)) f = continuous_on Y g" unfolding continuous_on_eq_continuous_within continuous_within_topological * apply transfer apply safe subgoal for x B apply (drule bspec, assumption, drule spec, drule mp, assumption, drule mp, assumption) apply clarsimp subgoal for AA apply (rule exI[where x="AA \ Collect (Domainp A)"]) by (auto intro: XA) done subgoal using XA by (force simp: openin_subtopology) done qed lemma continuous_on_transfer_right_total2[transfer_rule]: "(rel_set A ===> (A ===> B) ===> (=)) (\X::'a::t2_space set. continuous_on X) (\Y::'b::t2_space set. continuous_on Y)" if DomainB: "Domainp B = BP" and [folded DomainB, transfer_rule]: "bi_unique A" "bi_total A" "(rel_set A ===> (=)) open open" "bi_unique B" "right_total B" "(rel_set B ===> (=)) ((openin (top_of_set (Collect BP)))) open" unfolding DomainB[symmetric] proof (intro rel_funI) fix X Y f g assume H[transfer_rule]: "rel_set A X Y" "(A ===> B) f g" show "continuous_on X f = continuous_on Y g" unfolding continuous_on_eq_continuous_within continuous_within_topological apply transfer apply safe subgoal for x C apply (clarsimp simp: openin_subtopology) apply (drule bspec, assumption, drule spec, drule mp, assumption, drule mp, assumption) apply clarsimp by (meson Domainp_applyI H(1) H(2) rel_setD1) subgoal for x C proof - let ?sub = "top_of_set (Collect (Domainp B))" assume cont: "\x\X. \Ba\{A. Ball A (Domainp B)}. openin (top_of_set (Collect (Domainp B))) Ba \ f x \ Ba \ (\Aa. open Aa \ x \ Aa \ (\y\X. y \ Aa \ f y \ Ba))" and x: "x \ X" "open C" "f x \ C" let ?B = "C \ Collect (Domainp B)" have "?B \ {A. Ball A (Domainp B)}" by auto have "openin ?sub (Collect (Domainp B))" by auto then have "openin ?sub ?B" using \open C\ by auto moreover have "f x \ ?B" using x apply transfer apply auto by (meson Domainp_applyI H(1) H(2) rel_setD1) ultimately obtain D where "open D \ x \ D \ (\y\X. y \ D \ f y \ ?B)" using cont x by blast then show "\A. open A \ x \ A \ (\y\X. y \ A \ f y \ C)" by auto qed done qed lemma generate_topology_transfer[transfer_rule]: includes lifting_syntax assumes [transfer_rule]: "right_total A" "bi_unique A" shows "(rel_set (rel_set A) ===> rel_set A ===> (=)) (generate_topology o (insert (Collect (Domainp A)))) generate_topology" proof (intro rel_funI) fix B C X Y assume t[transfer_rule]: "rel_set (rel_set A) B C" "rel_set A X Y" then have "X \ Collect (Domainp A)" by (auto simp: rel_set_def) with t have rI: "rel_set A (X \ Collect (Domainp A)) Y" by (auto simp: inf_absorb1) have eq_UNIV_I: "Z = UNIV" if [transfer_rule]: "rel_set A {a. Domainp A a} Z" for Z using that assms apply (auto simp: right_total_def rel_set_def) using bi_uniqueDr by fastforce show "(generate_topology \ insert (Collect (Domainp A))) B X = generate_topology C Y" unfolding o_def proof (rule iffI) fix x assume "generate_topology (insert (Collect (Domainp A)) B) X" then show "generate_topology C Y" unfolding o_def using rI proof (induction X arbitrary: Y) case [transfer_rule]: UNIV with eq_UNIV_I[of Y] show ?case by (simp add: generate_topology.UNIV) next case (Int a b) note [transfer_rule] = Int(5) obtain a' where a'[transfer_rule]: "rel_set A (a \ Collect (Domainp A)) a'" by (metis Domainp_iff Domainp_set Int_Collect) obtain b' where b'[transfer_rule]: "rel_set A (b \ Collect (Domainp A)) b'" by (metis Domainp_iff Domainp_set Int_Collect) from Int.IH(1)[OF a'] Int.IH(2)[OF b'] have "generate_topology C a'" "generate_topology C b'" by auto from generate_topology.Int[OF this] have "generate_topology C (a' \ b')" . also have "a' \ b' = Y" by transfer auto finally show ?case by (simp add: generate_topology.Int) next case (UN K) note [transfer_rule] = UN(3) have "\K'. \k. rel_set A (k \ Collect (Domainp A)) (K' k)" by (rule choice) (metis Domainp_iff Domainp_set Int_Collect) then obtain K' where K': "\k. rel_set A (k \ Collect (Domainp A)) (K' k)" by metis from UN.IH[OF _ this] have "generate_topology C k'" if "k' \ K'`K" for k' using that by auto from generate_topology.UN[OF this] have "generate_topology C (\(K' ` K))" . also from K' have [transfer_rule]: "(rel_set (=) ===> rel_set A) (\x. x \ Collect (Domainp A)) K'" by (fastforce simp: rel_fun_def rel_set_def) have "\(K' ` K) = Y" by transfer auto finally show ?case by (simp add: generate_topology.UN) next case (Basis s) from this(1) show ?case proof assume "s = Collect (Domainp A)" with eq_UNIV_I[of Y] Basis(2) show ?case by (simp add: generate_topology.UNIV) next assume "s \ B" with Basis(2) obtain t where [transfer_rule]: "rel_set A (s \ Collect (Domainp A)) t" by auto from Basis(1) t(1) have s: "s \ Collect (Domainp A) = s" by (force simp: rel_set_def) have "t \ C" using \s \ B\ s by transfer auto also note [transfer_rule] = Basis(2) have "t = Y" by transfer auto finally show ?case by (rule generate_topology.Basis) qed qed next assume "generate_topology C Y" then show "generate_topology (insert (Collect (Domainp A)) B) X" using \rel_set A X Y\ proof (induction arbitrary: X) case [transfer_rule]: UNIV have "UNIV = (UNIV::'b set)" by auto then have "X = {a. Domainp A a}" by transfer then show ?case by (intro generate_topology.Basis) auto next case (Int a b) obtain a' b' where [transfer_rule]: "rel_set A a' a" "rel_set A b' b" by (meson assms(1) right_total_def right_total_rel_set) from generate_topology.Int[OF Int.IH(1)[OF this(1)] Int.IH(2)[OF this(2)]] have "generate_topology (insert {a. Domainp A a} B) (a' \ b')" by simp also define I where "I = a \ b" from \rel_set A X (a \ b)\ have [transfer_rule]: "rel_set A X I" by (simp add: I_def) from I_def have "a' \ b' = X" by transfer simp finally show ?case . next case (UN K) have "\K'. \k. rel_set A (K' k) k" by (rule choice) (meson assms(1) right_total_def right_total_rel_set) then obtain K' where K': "\k. rel_set A (K' k) k" by metis from UN.IH[OF _ this] have "generate_topology (insert {a. Domainp A a} B) k" if "k \ K'`K" for k using that by auto from generate_topology.UN[OF this] have "generate_topology (insert {a. Domainp A a} B) (\(K'`K))" by auto also from K' have [transfer_rule]: "(rel_set (=) ===> rel_set A) K' id" by (fastforce simp: rel_fun_def rel_set_def) define U where "U = (\(id ` K))" from \rel_set A X _\ have [transfer_rule]: "rel_set A X U" by (simp add: U_def) from U_def have "\(K' ` K) = X" by transfer simp finally show ?case . next case (Basis s) note [transfer_rule] = \rel_set A X s\ from \s \ C\ have "X \ B" by transfer then show ?case by (intro generate_topology.Basis) auto qed qed qed end subsection \Miscellaneous\ lemmas [simp del] = mem_ball lemma in_closureI[intro, simp]: "x \ X \ x \ closure X" using closure_subset by auto lemmas open_continuous_vimage = continuous_on_open_vimage[THEN iffD1, rule_format] lemma open_continuous_vimage': "open s \ continuous_on s f \ open B \ open (s \ f -` B)" using open_continuous_vimage[of s f B] by (auto simp: Int_commute) lemma support_on_mono: "support_on carrier f \ support_on carrier g" if "\x. x \ carrier \ f x \ 0 \ g x \ 0" using that by (auto simp: support_on_def) lemma image_prod: "(\(x, y). (f x, g y)) ` (A \ B) = f ` A \ g ` B" by auto subsection \Closed support\ definition "csupport_on X S = closure (support_on X S)" lemma closed_csupport_on[intro, simp]: "closed (csupport_on carrier \)" by (auto simp: csupport_on_def) lemma not_in_csupportD: "x \ csupport_on carrier \ \ x \ carrier \ \ x = 0" by (auto simp: csupport_on_def support_on_def) lemma csupport_on_mono: "csupport_on carrier f \ csupport_on carrier g" if "\x. x \ carrier \ f x \ 0 \ g x \ 0" unfolding csupport_on_def apply (rule closure_mono) using that by (rule support_on_mono) subsection \Homeomorphism\ lemma homeomorphism_empty[simp]: "homeomorphism {} t f f' \ t = {}" "homeomorphism s {} f f' \ s = {}" by (auto simp: homeomorphism_def) lemma homeomorphism_add: "homeomorphism UNIV UNIV (\x. x + c) (\x. x - c)" for c::"_::real_normed_vector" unfolding homeomorphism_def by (auto simp: algebra_simps continuous_intros intro!: image_eqI[where x="x - c" for x]) lemma in_range_scaleR_iff: "x \ range ((*\<^sub>R) c) \ c = 0 \ x = 0" for x::"_::real_vector" by (auto simp: intro!: image_eqI[where x="x /\<^sub>R c"]) lemma homeomorphism_scaleR: "homeomorphism UNIV UNIV (\x. c *\<^sub>R x::_::real_normed_vector) (\x. x /\<^sub>R c)" if "c \ 0" using that unfolding homeomorphism_def by (auto simp: in_range_scaleR_iff algebra_simps intro!: continuous_intros) lemma homeomorphism_prod: "homeomorphism (a \ b) (c \ d) (\(x, y). (f x, g y)) (\(x, y). (f' x, g' y))" if "homeomorphism a c f f'" "homeomorphism b d g g'" using that by (simp add: homeomorphism_def image_prod) (auto simp add: split_beta intro!: continuous_intros elim: continuous_on_compose2) subsection \Generalizations\ lemma openin_subtopology_eq_generate_topology: "openin (top_of_set S) x = generate_topology (insert S ((\B. B \ S) ` BB)) x" if open_gen: "open = generate_topology BB" and subset: "x \ S" proof - have "generate_topology (insert S ((\B. B \ S) ` BB)) (T \ S)" if "generate_topology BB T" for T using that proof (induction) case UNIV then show ?case by (auto intro!: generate_topology.Basis) next case (Int a b) have "generate_topology (insert S ((\B. B \ S) ` BB)) (a \ S \ (b \ S))" by (rule generate_topology.Int) (use Int in auto) then show ?case by (simp add: ac_simps) next case (UN K) then have "generate_topology (insert S ((\B. B \ S) ` BB)) (\k\K. k \ S)" by (intro generate_topology.UN) auto then show ?case by simp next case (Basis s) then show ?case by (intro generate_topology.Basis) auto qed moreover have "\T. generate_topology BB T \ x = T \ S" if "generate_topology (insert S ((\B. B \ S) ` BB)) x" "x \ UNIV" using that proof (induction) case UNIV then show ?case by simp next case (Int a b) then show ?case using generate_topology.Int by auto next case (UN K) from UN.IH have "\k\K-{UNIV}. \T. generate_topology BB T \ k = T \ S" by auto from this[THEN bchoice] obtain T where T: "\k. k \ T ` (K - {UNIV}) \ generate_topology BB k" "\k. k \ K - {UNIV} \ k = (T k) \ S" by auto from generate_topology.UN[OF T(1)] have "generate_topology BB (\(T ` (K - {UNIV})))" by auto moreover have "\K = (\(T ` (K - {UNIV}))) \ S" if "UNIV \ K" using T(2) UN that by auto ultimately show ?case apply (cases "UNIV \ K") subgoal using UN by auto subgoal by auto done next case (Basis s) then show ?case using generate_topology.UNIV generate_topology.Basis by blast qed moreover have "\T. generate_topology BB T \ UNIV = T \ S" if "generate_topology (insert S ((\B. B \ S) ` BB)) x" "x = UNIV" proof - have "S = UNIV" using that \x \ S\ by auto then show ?thesis by (simp add: generate_topology.UNIV) qed ultimately show ?thesis by (metis open_gen open_openin openin_open_Int' openin_subtopology) qed subsection \Equal topologies\ lemma topology_eq_iff: "t = s \ (topspace t = topspace s \ (\x\topspace t. openin t x = openin s x))" by (metis (full_types) openin_subset topology_eq) subsection \Finer topologies\ definition finer_than (infix "(finer'_than)" 50) where "T1 finer_than T2 \ continuous_map T1 T2 (\x. x)" lemma finer_than_iff_nhds: "T1 finer_than T2 \ (\X. openin T2 X \ openin T1 (X \ topspace T1)) \ (topspace T1 \ topspace T2)" by (auto simp: finer_than_def continuous_map_alt) lemma continuous_on_finer_topo: "continuous_map s t f" if "continuous_map s' t f" "s finer_than s'" using that by (auto simp: finer_than_def o_def dest: continuous_map_compose) lemma continuous_on_finer_topo2: "continuous_map s t f" if "continuous_map s t' f" "t' finer_than t" using that by (auto simp: finer_than_def o_def dest: continuous_map_compose) lemma antisym_finer_than: "S = T" if "S finer_than T" "T finer_than S" using that apply (auto simp: finer_than_def topology_eq_iff continuous_map_alt) apply (metis inf.orderE)+ done lemma subtopology_finer_than[simp]: "top_of_set X finer_than euclidean" by (auto simp: finer_than_iff_nhds openin_subtopology) subsection \Support\ lemma support_on_nonneg_sum: "support_on X (\x. \i\S. f i x) = (\i\S. support_on X (f i))" if "finite S" "\x i . x \ X \ i \ S \ f i x \ 0" for f::"_\_\_::ordered_comm_monoid_add" using that by (auto simp: support_on_def sum_nonneg_eq_0_iff) lemma support_on_nonneg_sum_subset: "support_on X (\x. \i\S. f i x) \ (\i\S. support_on X (f i))" for f::"_\_\_::ordered_comm_monoid_add" - by (cases "finite S") (auto simp: support_on_def dest!: sum_not_0) +by (cases "finite S") (auto simp: support_on_def, meson sum.neutral) lemma support_on_nonneg_sum_subset': "support_on X (\x. \i\S x. f i x) \ (\x\X. (\i\S x. support_on X (f i)))" for f::"_\_\_::ordered_comm_monoid_add" - by (auto simp: support_on_def dest!: sum_not_0) + by (auto simp: support_on_def, meson sum.neutral) subsection \Final topology (Bourbaki, General Topology I, 4.)\ definition "final_topology X Y f = topology (\U. U \ X \ (\i. openin (Y i) (f i -` U \ topspace (Y i))))" lemma openin_final_topology: "openin (final_topology X Y f) = (\U. U \ X \ (\i. openin (Y i) (f i -` U \ topspace (Y i))))" unfolding final_topology_def apply (rule topology_inverse') unfolding istopology_def proof safe fix S T i assume "\i. openin (Y i) (f i -` S \ topspace (Y i))" "\i. openin (Y i) (f i -` T \ topspace (Y i))" then have "openin (Y i) (f i -` S \ topspace (Y i) \ (f i -` T \ topspace (Y i)))" (is "openin _ ?I") by auto also have "?I = f i -` (S \ T) \ topspace (Y i)" (is "_ = ?R") by auto finally show "openin (Y i) ?R" . next fix K i assume "\U\K. U \ X \ (\i. openin (Y i) (f i -` U \ topspace (Y i)))" then have "openin (Y i) (\X\K. f i -` X \ topspace (Y i))" by (intro openin_Union) auto then show "openin (Y i) (f i -` \K \ topspace (Y i))" by (auto simp: vimage_Union) qed force+ lemma topspace_final_topology: "topspace (final_topology X Y f) = X" if "\i. f i \ topspace (Y i) \ X" proof - have *: "f i -` X \ topspace (Y i) = topspace (Y i)" for i using that by auto show ?thesis unfolding topspace_def unfolding openin_final_topology apply (rule antisym) apply force apply (rule subsetI) apply (rule UnionI[where X=X]) using that by (auto simp: *) qed lemma continuous_on_final_topologyI2: "continuous_map (Y i) (final_topology X Y f) (f i)" if "\i. f i \ topspace (Y i) \ X" using that by (auto simp: openin_final_topology continuous_map_alt topspace_final_topology) lemma continuous_on_final_topologyI1: "continuous_map (final_topology X Y f) Z g" if hyp: "\i. continuous_map (Y i) Z (g o f i)" and that: "\i. f i \ topspace (Y i) \ X" "g \ X \ topspace Z" unfolding continuous_map_alt proof safe fix V assume V: "openin Z V" have oV: "openin (Y i) (f i -` g -` V \ topspace (Y i))" for i using hyp[rule_format, of i] V by (auto simp: continuous_map_alt vimage_comp dest!: spec[where x=V]) have *: "f i -` g -` V \ f i -` X \ topspace (Y i) = f i -` g -` V \ topspace (Y i)" (is "_ = ?rhs i") for i using that by auto show "openin (final_topology X Y f) (g -` V \ topspace (final_topology X Y f))" by (auto simp: openin_final_topology oV topspace_final_topology that *) qed (use that in \auto simp: topspace_final_topology\) lemma continuous_on_final_topology_iff: "continuous_map (final_topology X Y f) Z g \ (\i. continuous_map (Y i) Z (g o f i))" if "\i. f i \ topspace (Y i) \ X" "g \ X \ topspace Z" using that by (auto intro!: continuous_on_final_topologyI1[OF _ that] intro: continuous_map_compose[OF continuous_on_final_topologyI2[OF that(1)]]) subsection \Quotient topology\ definition map_topology :: "('a \ 'b) \ 'a topology \ 'b topology" where "map_topology p X = final_topology (p ` topspace X) (\_. X) (\(_::unit). p)" lemma openin_map_topology: "openin (map_topology p X) = (\U. U \ p ` topspace X \ openin X (p -` U \ topspace X))" by (auto simp: map_topology_def openin_final_topology) lemma topspace_map_topology[simp]: "topspace (map_topology f T) = f ` topspace T" unfolding map_topology_def by (subst topspace_final_topology) auto lemma continuous_on_map_topology: "continuous_map T (map_topology f T) f" unfolding continuous_map_alt openin_map_topology by auto lemma continuous_map_composeD: "continuous_map T X (g \ f) \ g \ f ` topspace T \ topspace X" by (auto simp: continuous_map_def) lemma continuous_on_map_topology2: "continuous_map T X (g \ f) \ continuous_map (map_topology f T) X g" unfolding map_topology_def apply safe subgoal apply (rule continuous_on_final_topologyI1) subgoal by assumption subgoal by force subgoal by (rule continuous_map_composeD) done subgoal apply (erule continuous_map_compose[rotated]) apply (rule continuous_on_final_topologyI2) by force done lemma map_sub_finer_than_commute: "map_topology f (subtopology T (f -` X)) finer_than subtopology (map_topology f T) X" by (auto simp: finer_than_def continuous_map_def openin_subtopology openin_map_topology topspace_subtopology) lemma sub_map_finer_than_commute: "subtopology (map_topology f T) X finer_than map_topology f (subtopology T (f -` X))" if "openin T (f -` X)"\ \this is more or less the condition from \<^url>\https://math.stackexchange.com/questions/705840/quotient-topology-vs-subspace-topology\\ unfolding finer_than_def continuous_map_alt proof (rule conjI, clarsimp) fix U assume "openin (map_topology f (subtopology T (f -` X))) U" then obtain W where W: "U \ f ` (topspace T \ f -` X)" "openin T W" "f -` U \ (topspace T \ f -` X) = W \ f -` X" by (auto simp: topspace_subtopology openin_subtopology openin_map_topology) have "(f -` f ` W \ f -` X) \ topspace T = W \ topspace T \ f -` X" apply auto by (metis Int_iff W(3) vimage_eq) also have "openin T \" by (auto intro!: W that) finally show "openin (subtopology (map_topology f T) X) (U \ (f ` topspace T \ X))" using W unfolding topspace_subtopology topspace_map_topology openin_subtopology openin_map_topology by (intro exI[where x="(f ` W \ X)"]) auto qed auto lemma subtopology_map_topology: "subtopology (map_topology f T) X = map_topology f (subtopology T (f -` X))" if "openin T (f -` X)" apply (rule antisym_finer_than) using sub_map_finer_than_commute[OF that] map_sub_finer_than_commute[of f T X] by auto lemma quotient_map_map_topology: "quotient_map X (map_topology f X) f" by (auto simp: quotient_map_def openin_map_topology ac_simps) (simp_all add: vimage_def Int_def) lemma topological_space_quotient: "class.topological_space (openin (map_topology f euclidean))" if "surj f" apply standard apply (auto simp: ) using that by (auto simp: openin_map_topology) lemma t2_space_quotient: "class.t2_space (open::'b set \ bool)" if open_def: "open = (openin (map_topology (p::'a::t2_space\'b::topological_space) euclidean))" "surj p" and open_p: "\X. open X \ open (p ` X)" and "closed {(x, y). p x = p y}" (is "closed ?R") apply (rule class.t2_space.intro) subgoal by (unfold open_def, rule topological_space_quotient; fact) proof standard fix a b::'b obtain x y where a_def: "a = p x" and b_def: "b = p y" using \surj p\ by fastforce assume "a \ b" with \closed ?R\ have "open (-?R)" "(x, y) \ -?R" by (auto simp add: a_def b_def) from open_prod_elim[OF this] obtain N\<^sub>x N\<^sub>y where "open N\<^sub>x" "open N\<^sub>y" "(x, y) \ N\<^sub>x \ N\<^sub>y" "N\<^sub>x \ N\<^sub>y \ -?R" . then have "p ` N\<^sub>x \ p ` N\<^sub>y = {}" by auto moreover from \open N\<^sub>x\ \open N\<^sub>y\ have "open (p ` N\<^sub>x)" "open (p ` N\<^sub>y)" using open_p by blast+ moreover have "a \ p ` N\<^sub>x" "b \ p ` N\<^sub>y" using \(x, y) \ _ \ _\ by (auto simp: a_def b_def) ultimately show "\U V. open U \ open V \ a \ U \ b \ V \ U \ V = {}" by blast qed lemma second_countable_topology_quotient: "class.second_countable_topology (open::'b set \ bool)" if open_def: "open = (openin (map_topology (p::'a::second_countable_topology\'b::topological_space) euclidean))" "surj p" and open_p: "\X. open X \ open (p ` X)" apply (rule class.second_countable_topology.intro) subgoal by (unfold open_def, rule topological_space_quotient; fact) proof standard have euclidean_def: "euclidean = map_topology p euclidean" by (simp add: openin_inverse open_def) have continuous_on: "continuous_on UNIV p" using continuous_map_iff_continuous2 continuous_on_map_topology euclidean_def by fastforce from ex_countable_basis[where 'a='a] obtain A::"'a set set" where "countable A" "topological_basis A" by auto define B where "B = (\X. p ` X) ` A" have "countable (B::'b set set)" by (auto simp: B_def intro!: \countable A\) moreover have "topological_basis B" proof (rule topological_basisI) fix B' assume "B' \ B" then show "open B'" using \topological_basis A\ by (auto simp: B_def topological_basis_open intro!: open_p) next fix x::'b and O' assume "open O'" "x \ O'" have "open (p -` O')" using \open O'\ by (rule open_vimage) (auto simp: continuous_on) obtain y where y: "y \ p -` {x}" using \x \ O'\ by auto (metis UNIV_I open_def(2) rangeE) then have "y \ p -` O'" using \x \ O'\ by auto from topological_basisE[OF \topological_basis A\ \open (p -` O')\ this] obtain C where "C \ A" "y \ C" "C \ p -` O'" . let ?B' = "p ` C" have "?B' \ B" using \C \ A\ by (auto simp: B_def) moreover have "x \ ?B'" using y \y \ C\ \x \ O'\ by auto moreover have "?B' \ O'" using \C \ _\ by auto ultimately show "\B'\B. x \ B' \ B' \ O'" by metis qed ultimately show "\B::'b set set. countable B \ open = generate_topology B" by (auto simp: topological_basis_imp_subbasis) qed subsection \Closure\ lemma closure_Union: "closure (\X) = (\x\X. closure x)" if "finite X" using that by (induction X) auto subsection \Compactness\ lemma compact_if_closed_subset_of_compact: "compact S" if "closed S" "compact T" "S \ T" proof (rule compactI) fix UU assume UU: "\t\UU. open t" "S \ \UU" have "T \ \(insert (- S) (UU))" "\B. B \ insert (- S) UU \ open B" using UU \S \ T\ by (auto simp: open_Compl \closed S\) from compactE[OF \compact T\ this] obtain \' where \: "\' \ insert (- S) UU" "finite \'" "T \ \\'" by metis show "\C'\UU. finite C' \ S \ \C'" apply (rule exI[where x="\' - {-S}"]) using \ UU apply auto proof - fix x assume "x \ S" with \ \S \ T\ obtain U where "x \ U" "U \ \'" using \ by auto then show "\X\\' - {- S}. x \ X" using \ UU \x \ S\ apply - apply (rule bexI[where x=U]) by auto qed qed subsection \Locally finite\ definition "locally_finite_on X I U \ (\p\X. \N. p\N \ open N \ finite {i\I. U i \ N \ {}})" lemmas locally_finite_onI = locally_finite_on_def[THEN iffD2, rule_format] lemma locally_finite_onE: assumes "locally_finite_on X I U" assumes "p \ X" obtains N where "p \ N" "open N" "finite {i\I. U i \ N \ {}}" using assms by (auto simp: locally_finite_on_def) lemma locally_finite_onD: assumes "locally_finite_on X I U" assumes "p \ X" shows "finite {i\I. p \ U i}" apply (rule locally_finite_onE[OF assms]) apply (rule finite_subset) by auto lemma locally_finite_on_open_coverI: "locally_finite_on X I U" if fin: "\j. j \ I \ finite {i\I. U i \ U j \ {}}" and open_cover: "X \ (\i\I. U i)" "\i. i \ I \ open (U i)" proof (rule locally_finite_onI) fix p assume "p \ X" then obtain i where i: "i \ I" "p \ U i" "open (U i)" using open_cover by blast show "\N. p \ N \ open N \ finite {i \ I. U i \ N \ {}}" by (intro exI[where x="U i"] conjI i fin) qed lemma locally_finite_compactD: "finite {i\I. U i \ V \ {}}" if lf: "locally_finite_on X I U" and compact: "compact V" and subset: "V \ X" proof - have "\N. \p \ X. p \ N p \ open (N p) \ finite {i\I. U i \ N p \ {}}" by (rule bchoice) (auto elim!: locally_finite_onE[OF lf, rule_format]) then obtain N where N: "\p. p \ X \ p \ N p" "\p. p \ X \ open (N p)" "\p. p \ X \ finite {i\I. U i \ N p \ {}}" by blast have "V \ (\p\X. N p)" "\B. B \ N ` X \ open B" using N subset by force+ from compactE[OF compact this] obtain C where C: "C \ X" "finite C" "V \ \(N ` C)" by (metis finite_subset_image) then have "{i\I. U i \ V \ {}} \ {i\I. U i \ \(N ` C) \ {}}" by force also have "\ \ (\c\C. {i\I. U i \ N c \ {}})" by force also have "finite \" apply (rule finite_Union) using C by (auto intro!: C N) finally (finite_subset) show ?thesis . qed lemma closure_Int_open_eq_empty: "open S \ (closure T \ S) = {} \ T \ S = {}" by (auto simp: open_Int_closure_eq_empty ac_simps) lemma locally_finite_on_subset: assumes "locally_finite_on X J U" assumes "\i. i \ I \ V i \ U i" "I \ J" shows "locally_finite_on X I V" proof (rule locally_finite_onI) fix p assume "p \ X" from locally_finite_onE[OF assms(1) this] obtain N where "p \ N" "open N" "finite {i \ J. U i \ N \ {}}" . then show "\N. p \ N \ open N \ finite {i \ I. V i \ N \ {}}" apply (intro exI[where x=N]) using assms by (auto elim!: finite_subset[rotated]) qed lemma locally_finite_on_closure: "locally_finite_on X I (\x. closure (U x))" if "locally_finite_on X I U" proof (rule locally_finite_onI) fix p assume "p \ X" from locally_finite_onE[OF that this] obtain N where "p \ N" "open N" "finite {i \ I. U i \ N \ {}}" . then show "\N. p \ N \ open N \ finite {i \ I. closure (U i) \ N \ {}}" by (auto intro!: exI[where x=N] simp: closure_Int_open_eq_empty) qed lemma locally_finite_on_closedin_Union_closure: "closedin (top_of_set X) (\i\I. closure (U i))" if "locally_finite_on X I U" "\i. i \ I \ closure (U i) \ X" unfolding closedin_def apply safe subgoal using that(2) by auto subgoal apply (subst openin_subopen) proof clarsimp fix x assume x: "x \ X" "\i\I. x \ closure (U i)" from locally_finite_onE[OF that(1) \x \ X\] obtain N where N: "x \ N" "open N" "finite {i \ I. U i \ N \ {}}" (is "finite ?I"). define N' where "N' = N - (\i \ ?I. closure (U i))" have "open N'" by (auto simp: N'_def intro!: N) then have "openin (top_of_set X) (X \ N')" by (rule openin_open_Int) moreover have "x \ X \ N'" using x by (auto simp: N'_def N) moreover have "X \ N' \ X - (\i\I. closure (U i))" using x that(2) apply (auto simp: N'_def) by (meson N(2) closure_iff_nhds_not_empty dual_order.refl) ultimately show "\T. openin (top_of_set X) T \ x \ T \ T \ X - (\i\I. closure (U i))" by auto qed done lemma closure_subtopology_minimal: "S \ T \ closedin (top_of_set X) T \ closure S \ X \ T" apply (auto simp: closedin_closed) using closure_minimal by blast lemma locally_finite_on_closure_Union: "(\i\I. closure (U i)) = closure (\i\I. (U i)) \ X" if "locally_finite_on X I U" "\i. i \ I \ closure (U i) \ X" proof (rule antisym) show "(\i\I. closure (U i)) \ closure (\i\I. U i) \ X" using that apply auto by (metis (no_types, lifting) SUP_le_iff closed_closure closure_minimal closure_subset subsetCE) show "closure (\i\I. U i) \ X \ (\i\I. closure (U i))" apply (rule closure_subtopology_minimal) apply auto using that by (auto intro!: locally_finite_on_closedin_Union_closure) qed subsection \Refinement of cover\ definition refines :: "'a set set \ 'a set set \ bool" (infix "refines" 50) where "A refines B \ (\s\A. (\t. t \ B \ s \ t))" lemma refines_subset: "x refines y" if "z refines y" "x \ z" using that by (auto simp: refines_def) subsection \Functions as vector space\ instantiation "fun" :: (type, scaleR) scaleR begin definition scaleR_fun :: "real \ ('a \ 'b) \ 'a \ 'b" where "scaleR_fun r f = (\x. r *\<^sub>R f x)" lemma scaleR_fun_beta[simp]: "(r *\<^sub>R f) x = r *\<^sub>R f x" by (simp add: scaleR_fun_def) instance .. end instance "fun" :: (type, real_vector) real_vector by standard (auto simp: scaleR_fun_def algebra_simps) subsection \Additional lemmas\ lemmas [simp del] = vimage_Un vimage_Int lemma finite_Collect_imageI: "finite {U \ f ` X. P U}" if "finite {x\X. P (f x)}" proof - have "{d \ f ` X. P d} \ f ` {c \ X. P (f c)}" by blast then show ?thesis using finite_surj that by blast qed lemma plus_compose: "(x + y) \ f = (x \ f) + (y \ f)" by auto lemma mult_compose: "(x * y) \ f = (x \ f) * (y \ f)" by auto lemma scaleR_compose: "(c *\<^sub>R x) \ f = c *\<^sub>R (x \ f)" by (auto simp:) lemma image_scaleR_ball: fixes a :: "'a::real_normed_vector" shows "c \ 0 \ (*\<^sub>R) c ` ball a r = ball (c *\<^sub>R a) (abs c *\<^sub>R r)" proof (auto simp: mem_ball dist_norm, goal_cases) case (1 b) have "norm (c *\<^sub>R a - c *\<^sub>R b) = abs c * norm (a - b)" by (auto simp: norm_scaleR[symmetric] algebra_simps simp del: norm_scaleR) also have "\ < abs c * r" apply (rule mult_strict_left_mono) using 1 by auto finally show ?case . next case (2 x) have "norm (a - x /\<^sub>R c) < r" proof - have "norm (a - x /\<^sub>R c) = abs c *\<^sub>R norm (a - x /\<^sub>R c) /\<^sub>R abs c" using 2 by auto also have "abs c *\<^sub>R norm (a - x /\<^sub>R c) = norm (c *\<^sub>R a - x)" using 2 by (auto simp: norm_scaleR[symmetric] algebra_simps simp del: norm_scaleR) also have "\ < \c\ * r" by fact also have "\c\ * r /\<^sub>R \c\ = r" using 2 by auto finally show ?thesis using 2 by auto qed then have xdc: "x /\<^sub>R c \ ball a r" by (auto simp: mem_ball dist_norm) show ?case apply (rule image_eqI[OF _ xdc]) using 2 by simp qed subsection \Continuity\ lemma continuous_within_topologicalE: assumes "continuous (at x within s) f" "open B" "f x \ B" obtains A where "open A" "x \ A" "\y. y \ s \ y \ A \ f y \ B" using assms continuous_within_topological by metis lemma continuous_within_topologicalE': assumes "continuous (at x) f" "open B" "f x \ B" obtains A where "open A" "x \ A" "f ` A \ B" using assms continuous_within_topologicalE[OF assms] by (metis UNIV_I image_subsetI) lemma continuous_on_inverse: "continuous_on S f \ 0 \ f ` S \ continuous_on S (\x. inverse (f x))" for f::"_\_::real_normed_div_algebra" by (auto simp: continuous_on_def intro!: tendsto_inverse) subsection \@{term "(has_derivative)"}\ lemma has_derivative_plus_fun[derivative_intros]: "(x + y has_derivative x' + y') (at a within A)" if [derivative_intros]: "(x has_derivative x') (at a within A)" "(y has_derivative y') (at a within A)" by (auto simp: plus_fun_def intro!: derivative_eq_intros) lemma has_derivative_scaleR_fun[derivative_intros]: "(x *\<^sub>R y has_derivative x *\<^sub>R y') (at a within A)" if [derivative_intros]: "(y has_derivative y') (at a within A)" by (auto simp: scaleR_fun_def intro!: derivative_eq_intros) lemma has_derivative_times_fun[derivative_intros]: "(x * y has_derivative (\h. x a * y' h + x' h * y a)) (at a within A)" if [derivative_intros]: "(x has_derivative x') (at a within A)" "(y has_derivative y') (at a within A)" for x y::"_\'a::real_normed_algebra" by (auto simp: times_fun_def intro!: derivative_eq_intros) lemma real_sqrt_has_derivative_generic: "x \ 0 \ (sqrt has_derivative (*) ((if x > 0 then 1 else -1) * inverse (sqrt x) / 2)) (at x within S)" apply (rule has_derivative_at_withinI) using DERIV_real_sqrt_generic[of x "(if x > 0 then 1 else -1) * inverse (sqrt x) / 2"] at_within_open[of x "UNIV - {0}"] by (auto simp: has_field_derivative_def open_delete ac_simps split: if_splits) lemma sqrt_has_derivative: "((\x. sqrt (f x)) has_derivative (\xa. (if 0 < f x then 1 else - 1) / (2 * sqrt (f x)) * f' xa)) (at x within S)" if "(f has_derivative f') (at x within S)" "f x \ 0" by (rule has_derivative_eq_rhs[OF has_derivative_compose[OF that(1) real_sqrt_has_derivative_generic, OF that(2)]]) (auto simp: divide_simps) lemmas has_derivative_norm_compose[derivative_intros] = has_derivative_compose[OF _ has_derivative_norm] subsection \Differentiable\ lemmas differentiable_on_empty[simp] lemma differentiable_transform_eventually: "f differentiable (at x within X)" if "g differentiable (at x within X)" "f x = g x" "\\<^sub>F x in (at x within X). f x = g x" using that apply (auto simp: differentiable_def) subgoal for D apply (rule exI[where x=D]) apply (auto simp: has_derivative_within) by (simp add: eventually_mono Lim_transform_eventually) done lemma differentiable_within_eqI: "f differentiable at x within X" if "g differentiable at x within X" "\x. x \ X \ f x = g x" "x \ X" "open X" apply (rule differentiable_transform_eventually) apply (rule that) apply (auto simp: that) proof - have "\\<^sub>F x in at x within X. x \ X" using \open X\ using eventually_at_topological by blast then show " \\<^sub>F x in at x within X. f x = g x" by eventually_elim (auto simp: that) qed lemma differentiable_eqI: "f differentiable at x" if "g differentiable at x" "\x. x \ X \ f x = g x" "x \ X" "open X" using that unfolding at_within_open[OF that(3,4), symmetric] by (rule differentiable_within_eqI) lemma differentiable_on_eqI: "f differentiable_on S" if "g differentiable_on S" "\x. x \ S \ f x = g x" "open S" using that differentiable_eqI[of g _ S f] by (auto simp: differentiable_on_eq_differentiable_at) lemma differentiable_on_comp: "(f o g) differentiable_on S" if "g differentiable_on S" "f differentiable_on (g ` S)" using that by (auto simp: differentiable_on_def intro: differentiable_chain_within) lemma differentiable_on_comp2: "(f o g) differentiable_on S" if "f differentiable_on T" "g differentiable_on S" "g ` S \ T" apply (rule differentiable_on_comp) apply (rule that) apply (rule differentiable_on_subset) apply (rule that) apply (rule that) done lemmas differentiable_on_compose2 = differentiable_on_comp2[unfolded o_def] lemma differentiable_on_openD: "f differentiable at x" if "f differentiable_on X" "open X" "x \ X" using differentiable_on_eq_differentiable_at that by blast lemma differentiable_on_add_fun[intro, simp]: "x differentiable_on UNIV \ y differentiable_on UNIV \ x + y differentiable_on UNIV" by (auto simp: plus_fun_def) lemma differentiable_on_mult_fun[intro, simp]: "x differentiable_on UNIV \ y differentiable_on UNIV \ x * y differentiable_on UNIV" for x y::"_\'a::real_normed_algebra" by (auto simp: times_fun_def) lemma differentiable_on_scaleR_fun[intro, simp]: "y differentiable_on UNIV \ x *\<^sub>R y differentiable_on UNIV" by (auto simp: scaleR_fun_def) lemma sqrt_differentiable: "(\x. sqrt (f x)) differentiable (at x within S)" if "f differentiable (at x within S)" "f x \ 0" using that using sqrt_has_derivative[of f _ x S] by (auto simp: differentiable_def) lemma sqrt_differentiable_on: "(\x. sqrt (f x)) differentiable_on S" if "f differentiable_on S" "0 \ f ` S" using sqrt_differentiable[of f _ S] that by (force simp: differentiable_on_def) lemma differentiable_on_inverse: "f differentiable_on S \ 0 \ f ` S \ (\x. inverse (f x)) differentiable_on S" for f::"_\_::real_normed_field" by (auto simp: differentiable_on_def intro!: differentiable_inverse) lemma differentiable_on_openI: "f differentiable_on S" if "open S" "\x. x \ S \ \f'. (f has_derivative f') (at x)" using that by (auto simp: differentiable_on_def at_within_open[where S=S] differentiable_def) lemmas differentiable_norm_compose_at = differentiable_compose[OF differentiable_norm_at] lemma differentiable_on_Pair: "f differentiable_on S \ g differentiable_on S \ (\x. (f x, g x)) differentiable_on S" unfolding differentiable_on_def using differentiable_Pair[of f _ S g] by auto lemma differentiable_at_fst: "(\x. fst (f x)) differentiable at x within X" if "f differentiable at x within X" using that by (auto simp: differentiable_def dest!: has_derivative_fst) lemma differentiable_at_snd: "(\x. snd (f x)) differentiable at x within X" if "f differentiable at x within X" using that by (auto simp: differentiable_def dest!: has_derivative_snd) lemmas frechet_derivative_worksI = frechet_derivative_works[THEN iffD1] lemma sin_differentiable_at: "(\x. sin (f x::real)) differentiable at x within X" if "f differentiable at x within X" using differentiable_def has_derivative_sin that by blast lemma cos_differentiable_at: "(\x. cos (f x::real)) differentiable at x within X" if "f differentiable at x within X" using differentiable_def has_derivative_cos that by blast subsection \Frechet derivative\ lemmas frechet_derivative_transform_within_open_ext = fun_cong[OF frechet_derivative_transform_within_open] lemmas frechet_derivative_at' = frechet_derivative_at[symmetric] lemma frechet_derivative_plus_fun: "x differentiable at a \ y differentiable at a \ frechet_derivative (x + y) (at a) = frechet_derivative x (at a) + frechet_derivative y (at a)" by (rule frechet_derivative_at') (auto intro!: derivative_eq_intros frechet_derivative_worksI) lemmas frechet_derivative_plus = frechet_derivative_plus_fun[unfolded plus_fun_def] lemma frechet_derivative_zero_fun: "frechet_derivative 0 (at a) = 0" by (auto simp: frechet_derivative_const zero_fun_def) lemma frechet_derivative_sin: "frechet_derivative (\x. sin (f x)) (at x) = (\xa. frechet_derivative f (at x) xa * cos (f x))" if "f differentiable (at x)" for f::"_\real" by (rule frechet_derivative_at'[OF has_derivative_sin[OF frechet_derivative_worksI[OF that]]]) lemma frechet_derivative_cos: "frechet_derivative (\x. cos (f x)) (at x) = (\xa. frechet_derivative f (at x) xa * - sin (f x))" if "f differentiable (at x)" for f::"_\real" by (rule frechet_derivative_at'[OF has_derivative_cos[OF frechet_derivative_worksI[OF that]]]) lemma differentiable_sum_fun: "(\i. i \ I \ (f i differentiable at a)) \ sum f I differentiable at a" by (induction I rule: infinite_finite_induct) (auto simp: zero_fun_def plus_fun_def) lemma frechet_derivative_sum_fun: "(\i. i \ I \ (f i differentiable at a)) \ frechet_derivative (\i\I. f i) (at a) = (\i\I. frechet_derivative (f i) (at a))" by (induction I rule: infinite_finite_induct) (auto simp: frechet_derivative_zero_fun frechet_derivative_plus_fun differentiable_sum_fun) lemma sum_fun_def: "(\i\I. f i) = (\x. \i\I. f i x)" by (induction I rule: infinite_finite_induct) auto lemmas frechet_derivative_sum = frechet_derivative_sum_fun[unfolded sum_fun_def] lemma frechet_derivative_times_fun: "f differentiable at a \ g differentiable at a \ frechet_derivative (f * g) (at a) = (\x. f a * frechet_derivative g (at a) x + frechet_derivative f (at a) x * g a)" for f g::"_\'a::real_normed_algebra" by (rule frechet_derivative_at') (auto intro!: derivative_eq_intros frechet_derivative_worksI) lemmas frechet_derivative_times = frechet_derivative_times_fun[unfolded times_fun_def] lemma frechet_derivative_scaleR_fun: "y differentiable at a \ frechet_derivative (x *\<^sub>R y) (at a) = x *\<^sub>R frechet_derivative y (at a)" by (rule frechet_derivative_at') (auto intro!: derivative_eq_intros frechet_derivative_worksI) lemmas frechet_derivative_scaleR = frechet_derivative_scaleR_fun[unfolded scaleR_fun_def] lemma frechet_derivative_compose: "frechet_derivative (f o g) (at x) = frechet_derivative (f) (at (g x)) o frechet_derivative g (at x)" if "g differentiable at x" "f differentiable at (g x)" by (meson diff_chain_at frechet_derivative_at' frechet_derivative_works that) lemma frechet_derivative_compose_eucl: "frechet_derivative (f o g) (at x) = (\v. \i\Basis. ((frechet_derivative g (at x) v) \ i) *\<^sub>R frechet_derivative f (at (g x)) i)" (is "?l = ?r") if "g differentiable at x" "f differentiable at (g x)" proof (rule ext) fix v interpret g: linear "frechet_derivative g (at x)" using that(1) by (rule linear_frechet_derivative) interpret f: linear "frechet_derivative f (at (g x))" using that(2) by (rule linear_frechet_derivative) have "frechet_derivative (f o g) (at x) v = frechet_derivative f (at (g x)) (\i\Basis. (frechet_derivative g (at x) v \ i) *\<^sub>R i)" unfolding frechet_derivative_compose[OF that] o_apply by (simp add: euclidean_representation) also have "\ = ?r v" by (auto simp: g.sum g.scaleR f.sum f.scaleR) finally show "?l v = ?r v" . qed lemma frechet_derivative_works_on_open: "f differentiable_on X \ open X \ x \ X \ (f has_derivative frechet_derivative f (at x)) (at x)" and frechet_derivative_works_on: "f differentiable_on X \ x \ X \ (f has_derivative frechet_derivative f (at x within X)) (at x within X)" by (auto simp: differentiable_onD differentiable_on_openD frechet_derivative_worksI) lemma frechet_derivative_inverse: "frechet_derivative (\x. inverse (f x)) (at x) = (\h. - 1 / (f x)\<^sup>2 * frechet_derivative f (at x) h)" if "f differentiable at x" "f x \ 0" for f::"_\_::real_normed_field" apply (rule frechet_derivative_at') using that by (auto intro!: derivative_eq_intros frechet_derivative_worksI simp: divide_simps algebra_simps power2_eq_square) lemma frechet_derivative_sqrt: "frechet_derivative (\x. sqrt (f x)) (at x) = (\v. (if f x > 0 then 1 else -1) / (2 * sqrt (f x)) * frechet_derivative f (at x) v)" if "f differentiable at x" "f x \ 0" apply (rule frechet_derivative_at') apply (rule sqrt_has_derivative[THEN has_derivative_eq_rhs]) by (auto intro!: frechet_derivative_worksI that simp: divide_simps) lemma frechet_derivative_norm: "frechet_derivative (\x. norm (f x)) (at x) = (\v. frechet_derivative f (at x) v \ sgn (f x))" if "f differentiable at x" "f x \ 0" for f::"_\_::real_inner" apply (rule frechet_derivative_at') by (auto intro!: derivative_eq_intros frechet_derivative_worksI that simp: divide_simps) lemma (in bounded_linear) frechet_derivative: "frechet_derivative f (at x) = f" apply (rule frechet_derivative_at') apply (rule has_derivative_eq_rhs) apply (rule has_derivative) by (auto intro!: derivative_eq_intros) bundle no_matrix_mult begin no_notation matrix_matrix_mult (infixl "**" 70) end lemma (in bounded_bilinear) frechet_derivative: includes no_matrix_mult shows "x differentiable at a \ y differentiable at a \ frechet_derivative (\a. x a ** y a) (at a) = (\h. x a ** frechet_derivative y (at a) h + frechet_derivative x (at a) h ** y a)" by (rule frechet_derivative_at') (auto intro!: FDERIV frechet_derivative_worksI) lemma frechet_derivative_divide: "frechet_derivative (\x. f x / g x) (at x) = (\h. frechet_derivative f (at x) h / (g x) -frechet_derivative g (at x) h * f x / (g x)\<^sup>2)" if "f differentiable at x" "g differentiable at x" "g x \ 0" for f::"_\_::real_normed_field" using that by (auto simp: divide_inverse_commute bounded_bilinear.frechet_derivative[OF bounded_bilinear_mult] frechet_derivative_inverse) lemma frechet_derivative_pair: "frechet_derivative (\x. (f x, g x)) (at x) = (\v. (frechet_derivative f (at x) v, frechet_derivative g (at x) v))" if "f differentiable (at x)" "g differentiable (at x)" apply (rule frechet_derivative_at') apply (rule derivative_eq_intros) apply (rule frechet_derivative_worksI) apply fact apply (rule frechet_derivative_worksI) apply fact .. lemma frechet_derivative_fst: "frechet_derivative (\x. fst (f x)) (at x) = (\xa. fst (frechet_derivative f (at x) xa))" if "(f differentiable at x)" for f::"_\(_::real_normed_vector \ _::real_normed_vector)" apply (rule frechet_derivative_at') using that by (auto intro!: derivative_eq_intros frechet_derivative_worksI) lemma frechet_derivative_snd: "frechet_derivative (\x. snd (f x)) (at x) = (\xa. snd (frechet_derivative f (at x) xa))" if "(f differentiable at x)" for f::"_\(_::real_normed_vector \ _::real_normed_vector)" apply (rule frechet_derivative_at') using that by (auto intro!: derivative_eq_intros frechet_derivative_worksI) lemma frechet_derivative_eq_vector_derivative_1: assumes "f differentiable at t" shows "frechet_derivative f (at t) 1 = vector_derivative f (at t)" apply (subst frechet_derivative_eq_vector_derivative) apply (rule assms) by auto subsection \Linear algebra\ lemma (in vector_space) dim_pos_finite_dimensional_vector_spaceE: assumes "dim (UNIV::'b set) > 0" obtains basis where "finite_dimensional_vector_space scale basis" proof - from assms obtain b where b: "local.span b = local.span UNIV" "local.independent b" by (auto simp: dim_def split: if_splits) then have "dim UNIV = card b" by (rule dim_eq_card) with assms have "finite b" by (auto simp: card_ge_0_finite) then have "finite_dimensional_vector_space scale b" by unfold_locales (auto simp: b) then show ?thesis .. qed context vector_space_on begin context includes lifting_syntax assumes "\(Rep::'s \ 'b) (Abs::'b \ 's). type_definition Rep Abs S" begin interpretation local_typedef_vector_space_on S scale "TYPE('s)" by unfold_locales fact lemmas_with [var_simplified explicit_ab_group_add, unoverload_type 'd, OF type.ab_group_add_axioms type_vector_space_on_with, folded dim_S_def, untransferred, var_simplified implicit_ab_group_add]: lt_dim_pos_finite_dimensional_vector_spaceE = vector_space.dim_pos_finite_dimensional_vector_spaceE end lemmas_with [cancel_type_definition, OF S_ne, folded subset_iff', simplified pred_fun_def, folded finite_dimensional_vector_space_on_with, simplified\\too much?\]: dim_pos_finite_dimensional_vector_spaceE = lt_dim_pos_finite_dimensional_vector_spaceE end subsection \Extensional function space\ text \f is zero outside A. We use such functions to canonically represent functions whose domain is A\ definition extensional0 :: "'a set \ ('a \ 'b::zero) \ bool" where "extensional0 A f = (\x. x \ A \ f x = 0)" lemma extensional0_0[intro, simp]: "extensional0 X 0" by (auto simp: extensional0_def) lemma extensional0_UNIV[intro, simp]: "extensional0 UNIV f" by (auto simp: extensional0_def) lemma ext_extensional0: "f = g" if "extensional0 S f" "extensional0 S g" "\x. x \ S \ f x = g x" using that by (force simp: extensional0_def fun_eq_iff) lemma extensional0_add[intro, simp]: "extensional0 S f \ extensional0 S g \ extensional0 S (f + g::_\'a::comm_monoid_add)" by (auto simp: extensional0_def) lemma extensinoal0_mult[intro, simp]: "extensional0 S x \ extensional0 S y \ extensional0 S (x * y)" for x y::"_\'a::mult_zero" by (auto simp: extensional0_def) lemma extensional0_scaleR[intro, simp]: "extensional0 S f \ extensional0 S (c *\<^sub>R f::_\'a::real_vector)" by (auto simp: extensional0_def) lemma extensional0_outside: "x \ S \ extensional0 S f \ f x = 0" by (auto simp: extensional0_def) lemma subspace_extensional0: "subspace (Collect (extensional0 X))" by (auto simp: subspace_def) text \Send the function f to its canonical representative as a function with domain A\ definition restrict0 :: "'a set \ ('a \ 'b::zero) \ 'a \ 'b" where "restrict0 A f x = (if x \ A then f x else 0)" lemma restrict0_UNIV[simp]: "restrict0 UNIV = (\x. x)" by (intro ext) (auto simp: restrict0_def) lemma extensional0_restrict0[intro, simp]: "extensional0 A (restrict0 A f)" by (auto simp: extensional0_def restrict0_def) lemma restrict0_times: "restrict0 A (x * y) = restrict0 A x * restrict0 A y" for x::"'a\'b::mult_zero" by (auto simp: restrict0_def[abs_def]) lemma restrict0_apply_in[simp]: "x \ A \ restrict0 A f x = f x" by (auto simp: restrict0_def) lemma restrict0_apply_out[simp]: "x \ A \ restrict0 A f x = 0" by (auto simp: restrict0_def) lemma restrict0_scaleR: "restrict0 A (c *\<^sub>R f::_\'a::real_vector) = c *\<^sub>R restrict0 A f" by (auto simp: restrict0_def[abs_def]) lemma restrict0_add: "restrict0 A (f + g::_\'a::real_vector) = restrict0 A f + restrict0 A g" by (auto simp: restrict0_def[abs_def]) lemma restrict0_restrict0: "restrict0 X (restrict0 Y f) = restrict0 (X \ Y) f" by (auto simp: restrict0_def) end