diff --git a/thys/Green/Derivs.thy b/thys/Green/Derivs.thy --- a/thys/Green/Derivs.thy +++ b/thys/Green/Derivs.thy @@ -1,654 +1,652 @@ theory Derivs imports General_Utils begin lemma field_simp_has_vector_derivative [derivative_intros]: "(f has_field_derivative y) F \ (f has_vector_derivative y) F" by (simp add: has_field_derivative_iff_has_vector_derivative) lemma continuous_on_cases_empty [continuous_intros]: "\closed S; continuous_on S f; \x. \x \ S; \ P x\ \ f x = g x\ \ continuous_on S (\x. if P x then f x else g x)" using continuous_on_cases [of _ "{}"] by force lemma inj_on_cases: assumes "inj_on f (Collect P \ S)" "inj_on g (Collect (Not \ P) \ S)" "f ` (Collect P \ S) \ g ` (Collect (Not \ P) \ S) = {}" shows "inj_on (\x. if P x then f x else g x) S" using assms by (force simp: inj_on_def) lemma inj_on_arccos: "S \ {-1..1} \ inj_on arccos S" by (metis atLeastAtMost_iff cos_arccos inj_onI subsetCE) lemma has_vector_derivative_componentwise_within: "(f has_vector_derivative f') (at a within S) \ (\i \ Basis. ((\x. f x \ i) has_vector_derivative (f' \ i)) (at a within S))" apply (simp add: has_vector_derivative_def) apply (subst has_derivative_componentwise_within) apply simp done lemma has_vector_derivative_pair_within: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "\u. u \ Basis \ ((\x. f x \ u) has_vector_derivative f' \ u) (at x within S)" "\u. u \ Basis \ ((\x. g x \ u) has_vector_derivative g' \ u) (at x within S)" shows "((\x. (f x, g x)) has_vector_derivative (f',g')) (at x within S)" apply (subst has_vector_derivative_componentwise_within) apply (auto simp: assms Basis_prod_def) done lemma piecewise_C1_differentiable_const: shows "(\x. c) piecewise_C1_differentiable_on s" using continuous_on_const by (auto simp add: piecewise_C1_differentiable_on_def) declare piecewise_C1_differentiable_const [simp, derivative_intros] declare piecewise_C1_differentiable_neg [simp, derivative_intros] declare piecewise_C1_differentiable_add [simp, derivative_intros] declare piecewise_C1_differentiable_diff [simp, derivative_intros] (*move to Cauchy_Integral_Theorem*) lemma piecewise_C1_differentiable_on_ident [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_vector" shows "(\x. x) piecewise_C1_differentiable_on S" unfolding piecewise_C1_differentiable_on_def using C1_differentiable_on_ident by (blast intro: continuous_on_id C1_differentiable_on_ident) lemma piecewise_C1_differentiable_on_mult [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_algebra" assumes "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on S" shows "(\x. f x * g x) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def apply safe apply (blast intro: continuous_intros) apply (rename_tac A B) apply (rule_tac x="A \ B" in exI) apply (auto intro: C1_differentiable_on_mult C1_differentiable_on_subset) done lemma C1_differentiable_on_cdiv [simp, derivative_intros]: fixes f :: "real \ 'a :: real_normed_field" shows "f C1_differentiable_on S \ (\x. f x / c) C1_differentiable_on S" by (simp add: divide_inverse) lemma piecewise_C1_differentiable_on_cdiv [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_field" assumes "f piecewise_C1_differentiable_on S" shows "(\x. f x / c) piecewise_C1_differentiable_on S" by (simp add: divide_inverse piecewise_C1_differentiable_const piecewise_C1_differentiable_on_mult assms) lemma sqrt_C1_differentiable [simp, derivative_intros]: assumes f: "f C1_differentiable_on S" and fim: "f ` S \ {0<..}" shows "(\x. sqrt (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) show ?thesis using assms unfolding C1_differentiable_on_def has_field_derivative_iff_has_vector_derivative [symmetric] by (fastforce intro!: contf continuous_intros derivative_intros) qed lemma sqrt_piecewise_C1_differentiable [simp, derivative_intros]: assumes f: "f piecewise_C1_differentiable_on S" and fim: "f ` S \ {0<..}" shows "(\x. sqrt (f x)) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def by (fastforce intro!: continuous_intros derivative_intros) lemma fixes f :: "real \ 'a::{banach,real_normed_field}" assumes f: "f C1_differentiable_on S" shows sin_C1_differentiable [simp, derivative_intros]: "(\x. sin (f x)) C1_differentiable_on S" and cos_C1_differentiable [simp, derivative_intros]: "(\x. cos (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) note df_sin = field_vector_diff_chain_at [where g=sin, unfolded o_def] note df_cos = field_vector_diff_chain_at [where g=cos, unfolded o_def] show "(\x. sin (f x)) C1_differentiable_on S" "(\x. cos (f x)) C1_differentiable_on S" using assms unfolding C1_differentiable_on_def has_field_derivative_iff_has_vector_derivative [symmetric] apply auto by (rule contf continuous_intros df_sin df_cos derivative_intros exI conjI ballI | force)+ qed lemma has_derivative_abs: fixes a::real assumes "a \ 0" shows "(abs has_derivative ((*) (sgn a))) (at a)" proof - have [simp]: "norm = abs" using real_norm_def by force show ?thesis using has_derivative_norm [where 'a=real, simplified] assms by (simp add: mult_commute_abs) qed lemma abs_C1_differentiable [simp, derivative_intros]: fixes f :: "real \ real" assumes f: "f C1_differentiable_on S" and "0 \ f ` S" shows "(\x. abs (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) note df = DERIV_chain [where f=abs and g=f, unfolded o_def] show ?thesis using assms unfolding C1_differentiable_on_def has_field_derivative_iff_has_vector_derivative [symmetric] apply clarify apply (rule df exI conjI ballI)+ apply (force simp: has_field_derivative_def intro: has_derivative_abs continuous_intros contf)+ done qed lemma C1_differentiable_on_pair [simp, derivative_intros]: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "f C1_differentiable_on S" "g C1_differentiable_on S" shows "(\x. (f x, g x)) C1_differentiable_on S" using assms unfolding C1_differentiable_on_def apply safe apply (rename_tac A B) apply (intro exI ballI conjI) apply (rule_tac f'="A x" and g'="B x" in has_vector_derivative_pair_within) using has_vector_derivative_componentwise_within by (blast intro: continuous_on_Pair)+ lemma piecewise_C1_differentiable_on_pair [simp, derivative_intros]: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on S" shows "(\x. (f x, g x)) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def by (blast intro!: continuous_intros C1_differentiable_on_pair del: continuous_on_discrete intro: C1_differentiable_on_subset) lemma test2: assumes s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" and "x \ {0..1}" "x \ (\t. (v-u) *\<^sub>R t + u) -` s" shows "vector_derivative (\x. g ((v-u) * x + u)) (at x within {0..1}) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u) within{0..1})" proof - have i:"(g has_vector_derivative vector_derivative g (at ((v - u) * x + u))) (at ((v-u) * x + u))" using assms s [of "(v - u) * x + u"] uv mult_left_le [of x "v-u"] by (auto simp: vector_derivative_works) have ii:"((\x. g ((v - u) * x + u)) has_vector_derivative (v - u) *\<^sub>R vector_derivative g (at ((v - u) * x + u))) (at x)" by (intro vector_diff_chain_at [simplified o_def] derivative_eq_intros | simp add: i)+ have 0: "0 \ (v - u) * x + u" using assms uv by auto have 1: "(v - u) * x + u \ 1" using assms uv by simp (metis add.commute atLeastAtMost_iff atLeastatMost_empty_iff diff_ge_0_iff_ge empty_iff le_diff_eq mult_left_le) have iii: "vector_derivative g (at ((v - u) * x + u) within {0..1}) = vector_derivative g (at ((v - u) * x + u))" using Derivative.vector_derivative_at_within_ivl[OF i, of "0" "1", OF 0 1] by auto have iv: "vector_derivative (\x. g ((v - u) * x + u)) (at x within {0..1}) = (v - u) *\<^sub>R vector_derivative g (at ((v - u) * x + u))" using Derivative.vector_derivative_at_within_ivl[OF ii, of "0" "1"] assms by auto show ?thesis using iii iv by auto qed lemma C1_differentiable_on_components: assumes "\i. i \ Basis \ (\x. f x \ i) C1_differentiable_on s" shows "f C1_differentiable_on s" proof (clarsimp simp add: C1_differentiable_on_def has_vector_derivative_def) have *:"\f i x. x *\<^sub>R (f \ i) = (x *\<^sub>R f) \ i" by auto have "\f'. \i\Basis. \x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R f' x \ i)) (at x) \ continuous_on s f'" using assms lambda_skolem_euclidean[of "\i D. (\x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R D x)) (at x)) \ continuous_on s D"] apply (simp only: C1_differentiable_on_def has_vector_derivative_def *) using continuous_on_componentwise[of "s"] by metis then obtain f' where f': "\i\Basis. \x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R f' x \ i)) (at x) \ continuous_on s f'" by auto then have 0: "(\x\s. (f has_derivative (\z. z *\<^sub>R f' x)) (at x)) \ continuous_on s f'" using f' has_derivative_componentwise_within[of "f", where S= UNIV] by auto then show "\D. (\x\s. (f has_derivative (\z. z *\<^sub>R D x)) (at x)) \ continuous_on s D" by metis qed lemma piecewise_C1_differentiable_on_components: assumes "finite t" "\i. i \ Basis \ (\x. f x \ i) C1_differentiable_on s - t" "\i. i \ Basis \ continuous_on s (\x. f x \ i)" shows "f piecewise_C1_differentiable_on s" using C1_differentiable_on_components assms continuous_on_componentwise piecewise_C1_differentiable_on_def by blast lemma all_components_smooth_one_pw_smooth_is_pw_smooth: assumes "\i. i \ Basis - {j} \ (\x. f x \ i) C1_differentiable_on s" assumes "(\x. f x \ j) piecewise_C1_differentiable_on s" shows "f piecewise_C1_differentiable_on s" proof - have is_cont: "\i\Basis. continuous_on s (\x. f x \ i)" using assms C1_differentiable_imp_continuous_on piecewise_C1_differentiable_on_def by fastforce obtain t where t:"(finite t \ (\x. f x \ j) C1_differentiable_on s - t)" using assms(2) piecewise_C1_differentiable_on_def by auto show ?thesis using piecewise_C1_differentiable_on_components[where ?f = "f"] using assms(2) piecewise_C1_differentiable_on_def C1_differentiable_on_subset[OF assms(1) Diff_subset, where ?B1 ="t"] t is_cont by fastforce qed lemma derivative_component_fun_component: fixes i::"'a::euclidean_space" assumes "f differentiable (at x)" shows "((vector_derivative f (at x)) \ i) = ((vector_derivative (\x. (f x) \ i) (at x)) )" proof - have "((\x. f x \ i) has_vector_derivative vector_derivative f (at x) \ i) (at x)" using assms and bounded_linear.has_vector_derivative[of "(\x. x \ i)" "f" "(vector_derivative f (at x))" "(at x)"] and bounded_linear_inner_left[of "i"] and vector_derivative_works[of "f" "(at x)"] by blast then show "((vector_derivative f (at x)) \ i) = ((vector_derivative (\x. (f x) \ i) (at x)) )" using vector_derivative_works[of "(\x. (f x) \ i)" "(at x)"] and differentiableI_vector[of "(\x. (f x) \ i)" "(vector_derivative f (at x) \ i)" "(at x)"] and Derivative.vector_derivative_at by force qed lemma gamma_deriv_at_within: assumes a_leq_b: "a < b" and x_within_bounds: "x \ {a..b}" and gamma_differentiable: "\x \ {a .. b}. \ differentiable at x" shows "vector_derivative \ (at x within {a..b}) = vector_derivative \ (at x)" using Derivative.vector_derivative_at_within_ivl[of "\" "vector_derivative \ (at x)" "x" "a" "b"] gamma_differentiable x_within_bounds a_leq_b by (auto simp add: vector_derivative_works) lemma islimpt_diff_finite: assumes "finite (t::'a::t1_space set)" shows "x islimpt s - t = x islimpt s" proof- have iii: "s - t = s - (t \ s)" by auto have "(t \ s) \ s" by auto have ii:"finite (t \ s)" using assms(1) by auto have i: "(t \ s) \ (s - (t \ s)) = ( s)" using assms by auto then have "x islimpt s - (t \ s) = x islimpt s" using islimpt_Un_finite[OF ii,where ?t = "s - (t \ s)"] i assms(1) by force then show ?thesis using iii by auto qed lemma ivl_limpt_diff: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "x islimpt {a..b} - s" proof- have "x islimpt {a..b}" proof (cases "x \{a,b}") have i: "finite {a,b}" and ii: "{a, b} \ {a<..{a,b}" then show ?thesis using islimpt_Un_finite[OF i, where ?t= "{a<..{a,b}" then show "x islimpt {a..b}" using assms by auto qed then show "x islimpt {a..b} - s" using islimpt_diff_finite[OF assms(1)] assms by fastforce qed lemma ivl_closure_diff_del: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "x \ closure (({a..b} - s) - {x})" using ivl_limpt_diff islimpt_in_closure assms by blast lemma ivl_not_trivial_limit_within: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "at x within {a..b} - s \ bot" using assms ivl_closure_diff_del not_trivial_limit_within by blast lemma vector_derivative_at_within_non_trivial_limit: "at x within s \ bot \ (f has_vector_derivative f') (at x) \ vector_derivative f (at x within s) = f'" using has_vector_derivative_at_within vector_derivative_within by fastforce lemma vector_derivative_at_within_ivl_diff: "finite s \ a < b \ (x::real) \ {a..b} - s \ (f has_vector_derivative f') (at x) \ vector_derivative f (at x within {a..b} - s) = f'" using vector_derivative_at_within_non_trivial_limit ivl_not_trivial_limit_within by fastforce lemma gamma_deriv_at_within_diff: assumes a_leq_b: "a < b" and x_within_bounds: "x \ {a..b} - s" and gamma_differentiable: "\x \ {a .. b} - s. \ differentiable at x" and s_subset: "s \ {a..b}" and finite_s: "finite s" shows "vector_derivative \ (at x within {a..b} - s) = vector_derivative \ (at x)" using vector_derivative_at_within_ivl_diff [of "s" "a" "b" "x" "\" "vector_derivative \ (at x)"] gamma_differentiable x_within_bounds a_leq_b s_subset finite_s by (auto simp add: vector_derivative_works) lemma gamma_deriv_at_within_gen: assumes a_leq_b: "a < b" and x_within_bounds: "x \ s" and s_subset: "s \ {a..b}" and gamma_differentiable: "\x \ s. \ differentiable at x" shows "vector_derivative \ (at x within ({a..b})) = vector_derivative \ (at x)" using Derivative.vector_derivative_at_within_ivl[of "\" "vector_derivative \ (at x)" "x" "a" "b"] gamma_differentiable x_within_bounds a_leq_b s_subset by (auto simp add: vector_derivative_works) lemma derivative_component_fun_component_at_within_gen: assumes gamma_differentiable: "\x \ s. \ differentiable at x" and s_subset: "s \ {0..1}" shows "\x \ s. vector_derivative (\x. \ x) (at x within {0..1}) \ (i::'a:: euclidean_space) = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof - have gamma_i_component_smooth: "\x \ s. (\x. \ x \ i) differentiable at x" using gamma_differentiable by auto show "\x \ s. vector_derivative (\x. \ x) (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof fix x::real assume x_within_bounds: "x \ s" have gamma_deriv_at_within: "vector_derivative (\x. \ x) (at x within {0..1}) = vector_derivative (\x. \ x) (at x)" using gamma_deriv_at_within_gen[of "0" "1"] x_within_bounds gamma_differentiable s_subset by (auto simp add: vector_derivative_works) then have gamma_component_deriv_at_within: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_deriv_at_within_gen[of "0" "1", where \ = "(\x. \ x \ i)"] x_within_bounds gamma_i_component_smooth s_subset by (auto simp add: vector_derivative_works) have gamma_component_deriv_eq_gamma_deriv_component: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x) (at x) \ i" using derivative_component_fun_component[of "\" "x" "i"] gamma_differentiable x_within_bounds by auto show "vector_derivative \ (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_component_deriv_eq_gamma_deriv_component gamma_component_deriv_at_within gamma_deriv_at_within by auto qed qed lemma derivative_component_fun_component_at_within: assumes gamma_differentiable: "\x \ {0 .. 1}. \ differentiable at x" shows "\x \ {0..1}. vector_derivative (\x. \ x) (at x within {0..1}) \ (i::'a:: euclidean_space) = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof - have gamma_i_component_smooth: "\x \ {0 .. 1}. (\x. \ x \ i) differentiable at x" using gamma_differentiable by auto show "\x \ {0..1}. vector_derivative (\x. \ x) (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof fix x::real assume x_within_bounds: "x \ {0..1}" have gamma_deriv_at_within: "vector_derivative (\x. \ x) (at x within {0..1}) = vector_derivative (\x. \ x) (at x)" using gamma_deriv_at_within[of "0" "1"] x_within_bounds gamma_differentiable by (auto simp add: vector_derivative_works) have gamma_component_deriv_at_within: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x \ i) (at x within {0..1})" using Derivative.vector_derivative_at_within_ivl[of "(\x. (\ x) \ i)" "vector_derivative (\x. (\ x) \ i) (at x)" "x" "0" "1"] has_vector_derivative_at_within[of "(\x. \ x \ i)" "vector_derivative (\x. \ x \ i) (at x)" "x" "{0..1}"] gamma_i_component_smooth x_within_bounds by (simp add: vector_derivative_works) have gamma_component_deriv_eq_gamma_deriv_component: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x) (at x) \ i" using derivative_component_fun_component[of "\" "x" "i"] gamma_differentiable x_within_bounds by auto show "vector_derivative \ (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_component_deriv_eq_gamma_deriv_component gamma_component_deriv_at_within gamma_deriv_at_within by auto qed qed lemma straight_path_diffrentiable_x: fixes b :: "real" and y1 ::"real" assumes gamma_def: "\ = (\x. (b, y2 + y1 * x))" shows "\x. \ differentiable at x" unfolding gamma_def differentiable_def by (fast intro!: derivative_intros) lemma straight_path_diffrentiable_y: fixes b :: "real" and y1 y2 ::"real" assumes gamma_def: "\ = (\x. (y2 + y1 * x , b))" shows "\x. \ differentiable at x" unfolding gamma_def differentiable_def by (fast intro!: derivative_intros) lemma piecewise_C1_differentiable_on_imp_continuous_on: assumes "f piecewise_C1_differentiable_on s" shows "continuous_on s f" using assms by (auto simp add: piecewise_C1_differentiable_on_def) lemma boring_lemma1: fixes f :: "real\real" assumes "(f has_vector_derivative D) (at x)" shows "((\x. (f x, 0)) has_vector_derivative ((D,0::real))) (at x)" proof- have *: "((\x. (f x) *\<^sub>R (1,0)) has_vector_derivative (D *\<^sub>R (1,0))) (at x)" using bounded_linear.has_vector_derivative[OF Real_Vector_Spaces.bounded_linear_scaleR_left assms(1), of "(1,0)"] by auto have "((\x. (f x) *\<^sub>R (1,0)) has_vector_derivative (D,0)) (at x)" proof - have "(D, 0::'a) = D *\<^sub>R (1, 0)" by simp then show ?thesis by (metis (no_types) *) qed then show ?thesis by auto qed lemma boring_lemma2: fixes f :: "real\real" assumes "(f has_vector_derivative D) (at x)" shows "((\x. (0, f x)) has_vector_derivative (0, D)) (at x)" proof- have *: "((\x. (f x) *\<^sub>R (0,1)) has_vector_derivative (D *\<^sub>R (0,1))) (at x)" using bounded_linear.has_vector_derivative[OF Real_Vector_Spaces.bounded_linear_scaleR_left assms(1), of "(0,1)"] by auto then have "((\x. (f x) *\<^sub>R (0,1)) has_vector_derivative ((0,D))) (at x)" using scaleR_Pair Real_Vector_Spaces.real_scaleR_def proof - have "(0::'b, D) = D *\<^sub>R (0, 1)" by auto then show ?thesis by (metis (no_types) *) qed then show ?thesis by auto qed lemma pair_prod_smooth_pw_smooth: assumes "(f::real\real) C1_differentiable_on s" "(g::real\real) piecewise_C1_differentiable_on s" shows "(\x. (f x, g x)) piecewise_C1_differentiable_on s" proof - have f_cont: "continuous_on s f" using assms(1) C1_differentiable_imp_continuous_on by fastforce have g_cont: "continuous_on s g" using assms(2) by (auto simp add: piecewise_C1_differentiable_on_def) obtain t where t:"(finite t \ g C1_differentiable_on s - t)" using assms(2) piecewise_C1_differentiable_on_def by auto show ?thesis using piecewise_C1_differentiable_on_components[where ?f = "(\x. (f x, g x))"] apply (simp add: real_pair_basis) using assms(2) piecewise_C1_differentiable_on_def C1_differentiable_on_subset[OF assms(1) Diff_subset, where ?B1 ="t"] t f_cont g_cont by fastforce qed lemma scale_shift_smooth: shows "(\x. a + b * x) C1_differentiable_on s" proof - show "(\x. a + b * x) C1_differentiable_on s" - using Cauchy_Integral_Theorem.C1_differentiable_on_mult - Cauchy_Integral_Theorem.C1_differentiable_on_add - Cauchy_Integral_Theorem.C1_differentiable_on_const - Cauchy_Integral_Theorem.C1_differentiable_on_ident + using C1_differentiable_on_mult C1_differentiable_on_add + C1_differentiable_on_const C1_differentiable_on_ident by auto qed lemma open_diff: assumes "finite (t::'a::t1_space set)" "open (s::'a set)" shows "open (s -t)" using assms proof(induction "t") show "open s \ open (s - {})" by auto next fix x::"'a::t1_space" fix F::"'a::t1_space set" assume step: "finite F " " x \ F" "open s" then have i: "(s - insert x F) = (s - F) - {x}" by auto assume ind_hyp: " (open s \ open (s - F))" show "open (s - insert x F)" apply (simp only: i) using open_delete[of "s -F"] ind_hyp[OF step(3)] by auto qed lemma has_derivative_transform_within: assumes "0 < d" and "x \ s" and "\x'\s. dist x' x < d \ f x' = g x'" and "(f has_derivative f') (at x within s)" shows "(g has_derivative f') (at x within s)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within) lemma has_derivative_transform_within_ivl: assumes "(0::real) < b" and "\x\{a..b} -s. f x = g x" and "x \ {a..b} -s" and "(f has_derivative f') (at x within {a..b} - s)" shows "(g has_derivative f') (at x within {a..b} - s)" using has_derivative_transform_within[of "b" "x" "{a..b} - s"] assms by auto lemma has_vector_derivative_transform_within_ivl: assumes "(0::real) < b" and "\x\{a..b} -s . f x = g x" and "x \ {a..b} - s" and "(f has_vector_derivative f') (at x within {a..b} - s)" shows "(g has_vector_derivative f') (at x within {a..b} - s)" using assms has_derivative_transform_within_ivl apply (auto simp add: has_vector_derivative_def) by blast lemma has_derivative_transform_at: assumes "0 < d" and "\x'. dist x' x < d \ f x' = g x'" and "(f has_derivative f') (at x)" shows "(g has_derivative f') (at x)" using has_derivative_transform_within [of d x UNIV f g f'] assms by simp lemma has_vector_derivative_transform_at: assumes "0 < d" and "\x'. dist x' x < d \ f x' = g x'" and "(f has_vector_derivative f') (at x)" shows "(g has_vector_derivative f') (at x)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_at) lemma C1_diff_components_2: assumes "b \ Basis" assumes "f C1_differentiable_on s" shows "(\x. f x \ b) C1_differentiable_on s" proof - obtain D where D:"(\x\s. (f has_derivative (\z. z *\<^sub>R D x)) (at x))" "continuous_on s D" using assms(2) by (fastforce simp add: C1_differentiable_on_def has_vector_derivative_def) show ?thesis proof (simp add: C1_differentiable_on_def has_vector_derivative_def, intro exI conjI) show "continuous_on s (\x. D x \ b)" using D continuous_on_componentwise assms(1) by fastforce show "(\x\s. ((\x. f x \ b) has_derivative (\y. y * (\x. D x \ b) x)) (at x))" using has_derivative_inner_left D(1) by fastforce qed qed lemma eq_smooth: assumes "0 < d" "\x\s. \y. dist x y < d \ f y = g y" (*This crappy condition cannot be loosened :( *) "f C1_differentiable_on s" shows "g C1_differentiable_on s" proof (simp add: C1_differentiable_on_def) obtain D where D: "(\x\s. (f has_vector_derivative D x) (at x)) \ continuous_on s D" using assms by (auto simp add: C1_differentiable_on_def) then have f: "(\x\s. (g has_vector_derivative D x) (at x))" using assms(1-2) by (metis dist_commute has_vector_derivative_transform_at) have "(\x\s. (g has_vector_derivative D x) (at x)) \ continuous_on s D" using D f by auto then show "\D. (\x\s. (g has_vector_derivative D x) (at x)) \ continuous_on s D" by metis qed lemma eq_pw_smooth: assumes "0 < d" "\x\s. \y. dist x y < d \ f y = g y" (*This crappy condition cannot be loosened :( *) "\x\s. f x = g x" "f piecewise_C1_differentiable_on s" shows "g piecewise_C1_differentiable_on s" proof (simp add: piecewise_C1_differentiable_on_def) have g_cont: "continuous_on s g" using assms piecewise_C1_differentiable_const by (simp add: piecewise_C1_differentiable_on_def) obtain t where t: "finite t \ f C1_differentiable_on s - t" using assms by (auto simp add: piecewise_C1_differentiable_on_def) then have "g C1_differentiable_on s - t" using assms eq_smooth by blast then show "continuous_on s g \ (\t. finite t \ g C1_differentiable_on s - t)" using t g_cont by metis qed lemma scale_piecewise_C1_differentiable_on: assumes "f piecewise_C1_differentiable_on s" shows "(\x. (c::real) * (f x)) piecewise_C1_differentiable_on s" proof (simp add: piecewise_C1_differentiable_on_def, intro conjI) show "continuous_on s (\x. c * f x)" using assms continuous_on_mult_left by (auto simp add: piecewise_C1_differentiable_on_def) show "\t. finite t \ (\x. c * f x) C1_differentiable_on s - t" using assms continuous_on_mult_left by (auto simp add: piecewise_C1_differentiable_on_def) qed lemma eq_smooth_gen: assumes "f C1_differentiable_on s" "\x. f x = g x" shows "g C1_differentiable_on s" using assms unfolding C1_differentiable_on_def by (metis (no_types, lifting) has_vector_derivative_weaken UNIV_I top_greatest) lemma subpath_compose: shows "(subpath a b \) = \ o (\x. (b - a) * x + a)" by (auto simp add: subpath_def) lemma subpath_smooth: assumes "\ C1_differentiable_on {0..1}" "0 \ a" "a < b" "b \ 1" shows "(subpath a b \) C1_differentiable_on {0..1}" proof- have "\ C1_differentiable_on {a..b}" apply (rule C1_differentiable_on_subset) using assms by auto then have "\ C1_differentiable_on (\x. (b - a) * x + a) ` {0..1}" using \a < b\ closed_segment_eq_real_ivl closed_segment_real_eq by auto moreover have "finite ({0..1} \ (\x. (b - a) * x + a) -` {x})" for x proof - have "((\x. (b - a) * x + a) -` {x}) = {(x -a) /(b-a)}" using assms by (auto simp add: divide_simps) then show ?thesis by auto qed ultimately show ?thesis by (force simp add: subpath_compose intro: C1_differentiable_compose derivative_intros) qed lemma has_vector_derivative_divide[derivative_intros]: fixes a :: "'a::real_normed_field" shows "(f has_vector_derivative x) F \ ((\x. f x / a) has_vector_derivative (x / a)) F" unfolding divide_inverse by (fact has_vector_derivative_mult_left) end 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] + using 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] + using 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] + using 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] + using 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" 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/Green/Paths.thy b/thys/Green/Paths.thy --- a/thys/Green/Paths.thy +++ b/thys/Green/Paths.thy @@ -1,3584 +1,3584 @@ theory Paths imports Derivs General_Utils Integrals begin (*This has everything related to paths purely*) lemma reverse_subpaths_join: shows " subpath 1 (1 / 2) p +++ subpath (1 / 2) 0 p = reversepath p" using reversepath_subpath join_subpaths_middle pathfinish_subpath pathstart_subpath reversepath_joinpaths by (metis (no_types, lifting)) (*Below F cannot be from 'a \ 'b because the dot product won't work. We have that g returns 'a and then F takes the output of g, so F should start from 'a Then we have to compute the dot product of the vector b with both the derivative of g, and F. Since the derivative of g returns the same type as g, accordingly F should return the same type as g, i.e. 'a. *) definition line_integral:: "('a::euclidean_space \ 'a::euclidean_space) \ (('a) set) \ (real \ 'a) \ real" where "line_integral F basis g \ integral {0 .. 1} (\x. \b\basis. (F(g x) \ b) * (vector_derivative g (at x within {0..1}) \ b))" definition line_integral_exists where "line_integral_exists F basis \ \ (\x. \b\basis. F(\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) integrable_on {0..1}" lemma line_integral_on_pair_straight_path: fixes F::"('a::euclidean_space) \ 'a" and g :: "real \ real" and \ assumes gamma_const: "\x. \(x)\ i = a" and gamma_smooth: "\x \ {0 .. 1}. \ differentiable at x" shows "(line_integral F {i} \) = 0" "(line_integral_exists F {i} \)" proof (simp add: line_integral_def) have *: "F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i) = 0" if "0 \ x \ x \ 1" for x proof - have "((\x. \(x)\ i) has_vector_derivative 0) (at x)" using vector_derivative_const_at[of "a" "x"] and gamma_const by auto then have "(vector_derivative \ (at x) \ i) = 0" using derivative_component_fun_component[ of "\" "x" "i"] and gamma_smooth and that by (simp add: vector_derivative_at) then have "(vector_derivative \ (at x within {0 .. 1}) \ i) = 0" using has_vector_derivative_at_within vector_derivative_at_within_ivl that by (metis atLeastAtMost_iff gamma_smooth vector_derivative_works zero_less_one) then show ?thesis by auto qed then have "((\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) has_integral 0) {0..1}" using has_integral_is_0[of "{0 .. 1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))"] by auto then have "((\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) integrable_on {0..1})" by auto then show "line_integral_exists F {i} \" by (auto simp add:line_integral_exists_def) show "integral {0..1} (\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) = 0" using * has_integral_is_0[of "{0 .. 1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))"] by auto qed lemma line_integral_on_pair_path_strong: fixes F::"('a::euclidean_space) \ ('a)" and g::"real \ 'a" and \::"(real \ 'a)" and i::'a assumes i_norm_1: "norm i = 1" and g_orthogonal_to_i: "\x. g(x) \ i = 0" and gamma_is_in_terms_of_i: "\ = (\x. f(x) *\<^sub>R i + g(f(x)))" and gamma_smooth: "\ piecewise_C1_differentiable_on {0 .. 1}" and g_continuous_on_f: "continuous_on (f ` {0..1}) g" and path_start_le_path_end: "(pathstart \) \ i \ (pathfinish \) \ i" and field_i_comp_cont: "continuous_on (path_image \) (\x. F x \ i)" shows "line_integral F {i} \ = integral (cbox ((pathstart \) \ i) ((pathfinish \) \ i)) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i))" "line_integral_exists F {i} \" proof (simp add: line_integral_def) obtain s where gamma_differentiable: "finite s" "(\x \ {0 .. 1} - s. \ differentiable at x)" using gamma_smooth by (auto simp add: C1_differentiable_on_eq piecewise_C1_differentiable_on_def) then have gamma_i_component_smooth: "\x \ {0 .. 1} - s. (\x. \ x \ i) differentiable at x" by auto have field_cont_on_path: "continuous_on ((\x. \ x \ i) ` {0..1}) (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" proof - have 0: "(\x. \ x \ i) = f" proof fix x show "\ x \ i = f x" using g_orthogonal_to_i i_norm_1 by (simp only: gamma_is_in_terms_of_i real_inner_class.inner_add_left g_orthogonal_to_i inner_scaleR_left inner_same_Basis norm_eq_1) qed show ?thesis unfolding 0 apply (rule continuous_on_compose2 [of _ "(\x. F(x) \ i)" "f ` { 0..1}" "(\x. x *\<^sub>R i + g x)"] field_i_comp_cont g_continuous_on_f field_i_comp_cont continuous_intros)+ by (auto simp add: gamma_is_in_terms_of_i path_image_def) qed have path_start_le_path_end': "\ 0 \ i \ \ 1 \ i" using path_start_le_path_end by (auto simp add: pathstart_def pathfinish_def) have gamm_cont: "continuous_on {0..1} (\a. \ a \ i)" apply(rule continuous_on_inner) using gamma_smooth apply (simp add: piecewise_C1_differentiable_on_def) using continuous_on_const by auto then obtain c d where cd: "c \ d" "(\a. \ a \ i) ` {0..1} = {c..d}" by (meson continuous_image_closed_interval zero_le_one) then have subset_cd: "(\a. \ a \ i) ` {0..1} \ {c..d}" by auto have field_cont_on_path_cd: "continuous_on {c..d} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" using field_cont_on_path cd by auto have path_vector_deriv_line_integrals: "\x\{0..1} - s. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x)) (at x)" using gamma_i_component_smooth and derivative_component_fun_component and vector_derivative_works by blast then have "\x\{0..1} - s. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x within ({0..1}))) (at x within ({0..1}))" using has_vector_derivative_at_within vector_derivative_at_within_ivl by fastforce then have has_int:"((\x. vector_derivative (\x. \ x \ i) (at x within {0..1}) *\<^sub>R (F ((\ x \ i) *\<^sub>R i + g (\ x \ i)) \ i)) has_integral integral {\ 0 \ i..\ 1 \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)) {0..1}" using has_integral_substitution_strong[OF gamma_differentiable(1) rel_simps(44) path_start_le_path_end' subset_cd field_cont_on_path_cd gamm_cont, of "(\x. vector_derivative (\x. \(x) \ i) (at x within ({0..1})))"] gamma_is_in_terms_of_i by (auto simp only: has_field_derivative_iff_has_vector_derivative) then have has_int':"((\x. (F(\(x)) \ i)*(vector_derivative (\x. \(x) \ i) (at x within ({0..1})))) has_integral integral {((pathstart \) \ i)..((pathfinish \) \ i)} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)) {0..1}" using gamma_is_in_terms_of_i i_norm_1 apply (auto simp add: pathstart_def pathfinish_def) apply (simp only: real_inner_class.inner_add_left inner_not_same_Basis g_orthogonal_to_i inner_scaleR_left norm_eq_1) by (auto simp add: algebra_simps) have substitute: "integral ({((pathstart \) \ i)..((pathfinish \) \ i)}) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i)) = integral ({0..1}) (\x. (F(\(x)) \ i)*(vector_derivative (\x. \(x) \ i) (at x within ({0..1}))))" using gamma_is_in_terms_of_i integral_unique[OF has_int] i_norm_1 apply (auto simp add: pathstart_def pathfinish_def) apply (simp only: real_inner_class.inner_add_left inner_not_same_Basis g_orthogonal_to_i inner_scaleR_left norm_eq_1) by (auto simp add: algebra_simps) have comp_in_eq_comp_out: "\x \ {0..1} - s. (vector_derivative (\x. \(x) \ i) (at x within {0..1})) = (vector_derivative \ (at x within {0..1})) \ i" proof fix x:: real assume ass:"x \ {0..1} -s" then have x_bounds:"x \ {0..1}" by auto have "\ differentiable at x" using ass gamma_differentiable by auto then have dotprod_in_is_out: "vector_derivative (\x. \(x) \ i) (at x) = (vector_derivative \ (at x)) \ i" using derivative_component_fun_component by force then have 0: "(vector_derivative \ (at x)) \ i = (vector_derivative \ (at x within {0..1})) \ i" proof - have "(\ has_vector_derivative vector_derivative \ (at x)) (at x)" using \\ differentiable at x\ vector_derivative_works by blast moreover have "0 \ x \ x \ 1" using x_bounds by presburger ultimately show ?thesis by (simp add: vector_derivative_at_within_ivl) qed have 1: "vector_derivative (\x. \(x) \ i) (at x) = vector_derivative (\x. \(x) \ i) (at x within {0..1})" using path_vector_deriv_line_integrals and vector_derivative_at_within_ivl and x_bounds by (metis ass atLeastAtMost_iff zero_less_one) show "vector_derivative (\x. \ x \ i) (at x within {0..1}) = vector_derivative \ (at x within {0..1}) \ i" using 0 and 1 and dotprod_in_is_out by auto qed show "integral {0..1} (\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) = integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" using substitute and comp_in_eq_comp_out and negligible_finite Henstock_Kurzweil_Integration.integral_spike [of "s" "{0..1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))" "(\x. F (\ x) \ i * vector_derivative (\x. \ x \ i) (at x within {0..1}))"] by (metis (no_types, lifting) gamma_differentiable(1)) have "((\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) has_integral integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)) {0..1}" using has_int' and comp_in_eq_comp_out and negligible_finite Henstock_Kurzweil_Integration.has_integral_spike [of "s" "{0..1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))" "(\x. F (\ x) \ i * vector_derivative (\x. \ x \ i) (at x within {0..1}))" "integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)"] by (metis (no_types, lifting) gamma_differentiable(1)) then have "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) integrable_on {0..1}" using integrable_on_def by auto then show "line_integral_exists F {i} \" by (auto simp add: line_integral_exists_def) qed lemma line_integral_on_pair_path: fixes F::"('a::euclidean_space) \ ('a)" and g::"real \ 'a" and \::"(real \ 'a)" and i::'a assumes i_norm_1: "norm i = 1" and g_orthogonal_to_i: "\x. g(x) \ i = 0" and gamma_is_in_terms_of_i: "\ = (\x. f(x) *\<^sub>R i + g(f(x)))" and gamma_smooth: "\ C1_differentiable_on {0 .. 1}" and g_continuous_on_f: "continuous_on (f ` {0..1}) g" and path_start_le_path_end: "(pathstart \) \ i \ (pathfinish \) \ i" and field_i_comp_cont: "continuous_on (path_image \) (\x. F x \ i)" shows "(line_integral F {i} \) = integral (cbox ((pathstart \) \ i) ((pathfinish \) \ i)) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i))" proof (simp add: line_integral_def) have gamma_differentiable: "\x \ {0 .. 1}. \ differentiable at x" using gamma_smooth C1_differentiable_on_eq by blast then have gamma_i_component_smooth: "\x \ {0 .. 1}. (\x. \ x \ i) differentiable at x" by auto have vec_deriv_i_comp_cont: "continuous_on {0..1} (\x. vector_derivative (\x. \ x \ i) (at x within {0..1}))" proof - have "continuous_on {0..1} (\x. vector_derivative (\x. \ x) (at x within {0..1}))" using gamma_smooth C1_differentiable_on_eq by (smt C1_differentiable_on_def atLeastAtMost_iff continuous_on_eq vector_derivative_at_within_ivl) then have deriv_comp_cont: "continuous_on {0..1} (\x. vector_derivative (\x. \ x) (at x within {0..1}) \ i)" by (simp add: continuous_intros) show ?thesis using derivative_component_fun_component_at_within[OF gamma_differentiable, of "i"] continuous_on_eq[OF deriv_comp_cont, of "(\x. vector_derivative (\x. \ x \ i) (at x within {0..1}))"] by fastforce qed have field_cont_on_path: "continuous_on ((\x. \ x \ i) ` {0..1}) (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" proof - have 0: "(\x. \ x \ i) = f" proof fix x show "\ x \ i = f x" using g_orthogonal_to_i i_norm_1 by (simp only: gamma_is_in_terms_of_i real_inner_class.inner_add_left g_orthogonal_to_i inner_scaleR_left inner_same_Basis norm_eq_1) qed show ?thesis unfolding 0 apply (rule continuous_on_compose2 [of _ "(\x. F(x) \ i)" "f ` { 0..1}" "(\x. x *\<^sub>R i + g x)"] field_i_comp_cont g_continuous_on_f field_i_comp_cont continuous_intros)+ by (auto simp add: gamma_is_in_terms_of_i path_image_def) qed have path_vector_deriv_line_integrals: "\x\{0..1}. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x)) (at x)" using gamma_i_component_smooth and derivative_component_fun_component and vector_derivative_works by blast then have "\x\{0..1}. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x within {0..1})) (at x within {0..1})" using has_vector_derivative_at_within vector_derivative_at_within_ivl by fastforce then have substitute: "integral (cbox ((pathstart \) \ i) ((pathfinish \) \ i)) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i)) = integral (cbox 0 1) (\x. (F(\(x)) \ i)*(vector_derivative (\x. \(x) \ i) (at x within {0..1})))" using gauge_integral_by_substitution [of "0" "1" "(\x. (\ x) \ i)" "(\x. vector_derivative (\x. \(x) \ i) (at x within {0..1}))" "(\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i))"] and path_start_le_path_end and vec_deriv_i_comp_cont and field_cont_on_path and gamma_is_in_terms_of_i i_norm_1 apply (auto simp add: pathstart_def pathfinish_def) apply (simp only: real_inner_class.inner_add_left inner_not_same_Basis g_orthogonal_to_i inner_scaleR_left norm_eq_1) by (auto) (*integration by substitution*) have comp_in_eq_comp_out: "\x \ {0..1}. (vector_derivative (\x. \(x) \ i) (at x within {0..1})) = (vector_derivative \ (at x within {0..1})) \ i" proof fix x:: real assume x_bounds: "x \ {0..1}" then have "\ differentiable at x" using gamma_differentiable by auto then have dotprod_in_is_out: "vector_derivative (\x. \(x) \ i) (at x) = (vector_derivative \ (at x)) \ i" using derivative_component_fun_component by force then have 0: "(vector_derivative \ (at x)) \ i = (vector_derivative \ (at x within {0..1})) \ i" using has_vector_derivative_at_within and x_bounds and vector_derivative_at_within_ivl by (smt atLeastAtMost_iff gamma_differentiable inner_commute vector_derivative_works) have 1: "vector_derivative (\x. \(x) \ i) (at x) = vector_derivative (\x. \(x) \ i) (at x within {0..1})" using path_vector_deriv_line_integrals and vector_derivative_at_within_ivl and x_bounds by fastforce show "vector_derivative (\x. \ x \ i) (at x within {0..1}) = vector_derivative \ (at x within {0..1}) \ i" using 0 and 1 and dotprod_in_is_out by auto qed show "integral {0..1} (\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) = integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" using substitute and comp_in_eq_comp_out and Henstock_Kurzweil_Integration.integral_cong [of "{0..1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))" "(\x. F (\ x) \ i * vector_derivative (\x. \ x \ i) (at x within {0..1}))"] by auto qed lemma content_box_cases: "content (box a b) = (if \i\Basis. a\i \ b\i then prod (\i. b\i - a\i) Basis else 0)" by (simp add: measure_lborel_box_eq inner_diff) lemma content_box_cbox: shows "content (box a b) = content (cbox a b)" by (simp add: content_box_cases content_cbox_cases) lemma content_eq_0: "content (box a b) = 0 \ (\i\Basis. b\i \ a\i)" by (auto simp: content_box_cases not_le intro: less_imp_le antisym eq_refl) lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" by (auto simp add: content_cbox_cases less_le prod_nonneg) lemma content_lt_nz: "0 < content (box a b) \ content (box a b) \ 0" using Paths.content_eq_0 zero_less_measure_iff by blast lemma content_subset: "cbox a b \ box c d \ content (cbox a b) \ content (box c d)" unfolding measure_def by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq emeasure_lborel_box_eq) lemma sum_content_null: assumes "content (box a b) = 0" and "p tagged_division_of (box a b)" shows "sum (\(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)" proof (rule sum.neutral, rule) fix y assume y: "y \ p" obtain x k where xk: "y = (x, k)" using surj_pair[of y] by blast note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]] from this(2) obtain c d where k: "k = cbox c d" by blast have "(\(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x" unfolding xk by auto also have "\ = 0" using content_subset[OF assm(1)[unfolded k]] content_pos_le[of "cbox c d"] unfolding assms(1) k by auto finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . qed lemma has_integral_null [intro]: "content(box a b) = 0 \ (f has_integral 0) (box a b)" by (simp add: content_box_cbox content_eq_0_interior) lemma line_integral_distrib: assumes "line_integral_exists f basis g1" "line_integral_exists f basis g2" "valid_path g1" "valid_path g2" shows "line_integral f basis (g1 +++ g2) = line_integral f basis g1 + line_integral f basis g2" "line_integral_exists f basis (g1 +++ g2)" proof - obtain s1 s2 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) obtain i1 i2 where 1: "((\x. \b\basis. f (g1 x) \ b * (vector_derivative g1 (at x within {0..1}) \ b)) has_integral i1) {0..1}" and 2: "((\x. \b\basis. f (g2 x) \ b * (vector_derivative g2 (at x within {0..1}) \ b)) has_integral i2) {0..1}" using assms by (auto simp: line_integral_exists_def) have i1: "((\x. 2 * (\b\basis. f (g1 (2 * x)) \ b * (vector_derivative g1 (at (2 * x) within {0..1}) \ b))) has_integral i1) {0..1/2}" and i2: "((\x. 2 * (\b\basis. f (g2 (2 * x - 1)) \ b * (vector_derivative g2 (at ((2 * x) - 1) within {0..1}) \ b))) has_integral i2) {1/2..1}" using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z*2) within {0..1})" for z proof - have i:"\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" for z proof- have g1_at_z:"\0 \ z; z*2 < 1; z*2 \ s1\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at (z*2))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g1(2*x))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s1 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have z_ge: "z\ 1" by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" using Derivative.vector_derivative_at_within_ivl[OF g1_at_z[OF ass] ass(1) z_ge] by auto qed assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have "(g1 has_vector_derivative ((vector_derivative g1 (at (z*2))))) (at (z*2))" using s1 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g1 (at (z*2) within {0..1})) = (vector_derivative g1 (at (z*2)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2) within {0..1})" using i[OF ass] ii by auto qed have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1) within {0..1})" for z proof - have i:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" for z proof- have g2_at_z:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g2 (2*x - 1))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s2 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have z_le: "z\ 1" by auto have z_ge: "0 \ z" using ass by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" using Derivative.vector_derivative_at_within_ivl[OF g2_at_z[OF ass] z_ge z_le] by auto qed assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have "(g2 has_vector_derivative ((vector_derivative g2 (at (z*2 - 1))))) (at (z*2 - 1))" using s2 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g2 (at (z*2 - 1) within {0..1})) = (vector_derivative g2 (at (z*2 - 1)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1) within {0..1})" using i[OF ass] ii by auto qed have lem1: "((\x. \b\basis. f ((g1+++g2) x) \ b * (vector_derivative (g1+++g2) (at x within {0..1}) \ b)) has_integral i1) {0..1/2}" apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) ((*)2 -` s1)"]) using s1 apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) by (simp add: sum_distrib_left) moreover have lem2: "((\x. \b\basis. f ((g1+++g2) x) \ b * (vector_derivative (g1+++g2) (at x within {0..1}) \ b)) has_integral i2) {1/2..1}" apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) using s2 apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) by (simp add: sum_distrib_left) ultimately show "line_integral f basis (g1 +++ g2) = line_integral f basis g1 + line_integral f basis g2" apply (simp add: line_integral_def) apply (rule integral_unique [OF has_integral_combine [where c = "1/2"]]) using 1 2 integral_unique apply auto done show "line_integral_exists f basis (g1 +++ g2)" proof (simp add: line_integral_exists_def integrable_on_def) have "(1::real) \ 1 * 2 \ (0::real) \ 1 / 2" by simp then show "\r. ((\r. \a\basis. f ((g1 +++ g2) r) \ a * (vector_derivative (g1 +++ g2) (at r within {0..1}) \ a)) has_integral r) {0..1}" using has_integral_combine [where c = "1/2"] 1 2 divide_le_eq_numeral1(1) lem1 lem2 by blast qed qed lemma line_integral_exists_joinD1: assumes "line_integral_exists f basis (g1 +++ g2)" "valid_path g1" shows "line_integral_exists f basis g1" proof - obtain s1 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. \b\basis. f ((g1 +++ g2) (x/2)) \ b * (vector_derivative (g1 +++ g2) (at (x/2) within {0..1}) \ b)) integrable_on {0..1}" using assms apply (auto simp: line_integral_exists_def) apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) done then have *:"(\x. \b\basis. ((f ((g1 +++ g2) (x/2)) \ b) / 2)* (vector_derivative (g1 +++ g2) (at (x/2) within {0..1}) \ b)) integrable_on {0..1}" by (auto simp: Groups_Big.sum_distrib_left dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z*2) within {0..1})" for z proof - have i:"\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" for z proof- have g1_at_z:"\0 \ z; z*2 < 1; z*2 \ s1\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at (z*2))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g1(2*x))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s1 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have z_ge: "z\ 1" by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" using Derivative.vector_derivative_at_within_ivl[OF g1_at_z[OF ass] ass(1) z_ge] by auto qed assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have "(g1 has_vector_derivative ((vector_derivative g1 (at (z*2))))) (at (z*2))" using s1 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g1 (at (z*2) within {0..1})) = (vector_derivative g1 (at (z*2)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2) within {0..1})" using i[OF ass] ii by auto qed show ?thesis using s1 apply (auto simp: line_integral_exists_def) apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g1) done qed lemma line_integral_exists_joinD2: assumes "line_integral_exists f basis (g1 +++ g2)" "valid_path g2" shows "line_integral_exists f basis g2" proof - obtain s2 where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. \b\basis. f ((g1 +++ g2) (x/2 + 1/2)) \ b * (vector_derivative (g1 +++ g2) (at (x/2 + 1/2) within {0..1}) \ b)) integrable_on {0..1}" using assms apply (auto simp: line_integral_exists_def) apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) apply (simp add: image_affinity_atLeastAtMost_diff) done then have *:"(\x. \b\basis. ((f ((g1 +++ g2) (x/2 + 1/2)) \ b) / 2)* (vector_derivative (g1 +++ g2) (at (x/2 + 1/2) within {0..1}) \ b)) integrable_on {0..1}" by (auto simp: Groups_Big.sum_distrib_left dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1) within {0..1})" for z proof - have i:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" for z proof- have g2_at_z:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g2 (2*x - 1))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s2 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have z_le: "z\ 1" by auto have z_ge: "0 \ z" using ass by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" using Derivative.vector_derivative_at_within_ivl[OF g2_at_z[OF ass] z_ge z_le] by auto qed assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have "(g2 has_vector_derivative ((vector_derivative g2 (at (z*2 - 1))))) (at (z*2 - 1))" using s2 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g2 (at (z*2 - 1) within {0..1})) = (vector_derivative g2 (at (z*2 - 1)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1) within {0..1})" using i[OF ass] ii by auto qed show ?thesis using s2 apply (auto simp: line_integral_exists_def) apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g2) done qed lemma has_line_integral_on_reverse_path: assumes g: "valid_path g" and int: "((\x. \b\basis. F (g x) \ b * (vector_derivative g (at x within {0..1}) \ b)) has_integral c){0..1}" shows "((\x. \b\basis. F ((reversepath g) x) \ b * (vector_derivative (reversepath g) (at x within {0..1}) \ b)) has_integral -c){0..1}" proof - { fix s x assume xs: "g C1_differentiable_on ({0..1} - s)" "x \ (-) 1 ` s" "0 \ x" "x \ 1" have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = - vector_derivative g (at (1 - x) within {0..1})" proof - obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" using xs by (force simp: has_vector_derivative_def C1_differentiable_on_def) have "(g o (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" apply (rule vector_diff_chain_within) apply (intro vector_diff_chain_within derivative_eq_intros | simp)+ apply (rule has_vector_derivative_at_within [OF f']) done then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" by (simp add: o_def) show ?thesis using xs by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) qed } note * = this obtain S where "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" using g by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) then show ?thesis using has_integral_affinity01 [OF int, where m= "-1" and c=1] unfolding reversepath_def by (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) (auto simp: * has_integral_neg Groups_Big.sum_negf) qed lemma line_integral_on_reverse_path: assumes "valid_path \" "line_integral_exists F basis \" shows "line_integral F basis \ = - (line_integral F basis (reversepath \))" "line_integral_exists F basis (reversepath \)" proof - obtain i where 0: "((\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) has_integral i){0..1}" using assms unfolding integrable_on_def line_integral_exists_def by auto then have 1: "((\x. \b\basis. F ((reversepath \) x) \ b * (vector_derivative (reversepath \) (at x within {0..1}) \ b)) has_integral -i){0..1}" using has_line_integral_on_reverse_path assms by auto then have rev_line_integral:"line_integral F basis (reversepath \) = -i" using line_integral_def Henstock_Kurzweil_Integration.integral_unique by (metis (no_types)) have line_integral: "line_integral F basis \ = i" using line_integral_def 0 Henstock_Kurzweil_Integration.integral_unique by blast show "line_integral F basis \ = - (line_integral F basis (reversepath \))" using line_integral rev_line_integral by auto show "line_integral_exists F basis (reversepath \)" using 1 line_integral_exists_def by auto qed lemma line_integral_exists_on_degenerate_path: assumes "finite basis" shows "line_integral_exists F basis (\x. c)" proof- have every_component_integrable: "\b\basis. (\x. F ((\x. c) x) \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) integrable_on {0..1}" proof fix b assume b_in_basis: "b \ basis" have cont_field_zero_one: "continuous_on {0..1} (\x. F ((\x. c) x) \ b)" using continuous_on_const by fastforce have cont_path_zero_one: "continuous_on {0..1} (\x. (vector_derivative (\x. c) (at x within {0..1})) \ b)" proof - have "((vector_derivative (\x. c) (at x within {0..1})) \ b) = 0" if "x \ {0..1}" for x proof - have "vector_derivative (\x. c) (at x within {0..1}) = 0" using that gamma_deriv_at_within[of "0" "1"] differentiable_const vector_derivative_const_at by fastforce then show "vector_derivative (\x. c) (at x within {0..1}) \ b = 0" by auto qed then show "continuous_on {0..1} (\x. (vector_derivative (\x. c) (at x within {0..1})) \ b)" using continuous_on_const[of "{0..1}" "0"] continuous_on_eq[of "{0..1}" "\x. 0" "(\x. (vector_derivative (\x. c) (at x within {0..1})) \ b)"] by auto qed show "(\x. F (c) \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) integrable_on {0..1}" using cont_field_zero_one cont_path_zero_one continuous_on_mult integrable_continuous_real by blast qed have integrable_sum': "\t f s. finite t \ \a\t. f a integrable_on s \ (\x. \a\t. f a x) integrable_on s" using integrable_sum by metis have field_integrable_on_basis: "(\x. \b\basis. F (c) \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) integrable_on {0..1}" using integrable_sum'[OF assms(1) every_component_integrable] by auto then show ?thesis using line_integral_exists_def by auto qed lemma degenerate_path_is_valid_path: "valid_path (\x. c)" by (auto simp add: valid_path_def piecewise_C1_differentiable_on_def continuous_on_const) lemma line_integral_degenerate_path: assumes "finite basis" shows "line_integral F basis (\x. c) = 0" proof (simp add: line_integral_def) have "((vector_derivative (\x. c) (at x within {0..1})) \ b) = 0" if "x \ {0..1}" for x b proof - have "vector_derivative (\x. c) (at x within {0..1}) = 0" using that gamma_deriv_at_within[of "0" "1"] differentiable_const vector_derivative_const_at by fastforce then show "vector_derivative (\x. c) (at x within {0..1}) \ b = 0" by auto qed then have 0: "\x. x \ {0..1} \ (\x. \b\basis. F c \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) x = (\x. 0) x" by auto then show "integral {0..1} (\x. \b\basis. F c \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) = 0" using integral_cong[of "{0..1}", OF 0] integral_0 by auto qed definition point_path where "point_path \ \ \c. \ = (\x. c)" lemma line_integral_point_path: assumes "point_path \" assumes "finite basis" shows "line_integral F basis \ = 0" using assms(1) point_path_def line_integral_degenerate_path[OF assms(2)] by force lemma line_integral_exists_point_path: assumes "finite basis" "point_path \" shows "line_integral_exists F basis \" using assms apply(simp add: point_path_def) using line_integral_exists_on_degenerate_path by auto lemma line_integral_exists_subpath: assumes f: "line_integral_exists f basis g" and g: "valid_path g" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" shows "(line_integral_exists f basis (subpath u v g))" proof (cases "v=u") case tr: True have zero: "(\b\basis. f (g u) \ b * (vector_derivative (\x. g u) (at x within {0..1}) \ b)) = 0" if "x \ {0..1}" for x::real proof - have "(vector_derivative (\x. g u) (at x within {0..1})) = 0" using Deriv.has_vector_derivative_const that Derivative.vector_derivative_at_within_ivl by fastforce then show "(\b\basis. f (g u) \ b * (vector_derivative (\x. g u) (at x within {0..1}) \ b)) = 0" by auto qed then have "((\x. \b\basis. f (g u) \ b * (vector_derivative (\x. g u) (at x within {0..1}) \ b)) has_integral 0) {0..1}" by (meson has_integral_is_0) then show ?thesis using f tr by (auto simp add: line_integral_def line_integral_exists_def subpath_def) next case False obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast have *: "((\x. \b\basis. f (g ((v - u) * x + u)) \ b * (vector_derivative g (at ((v - u) * x + u) within {0..1}) \ b)) has_integral (1 / (v - u)) * integral {u..v} (\x. \b\basis. f (g (x)) \ b * (vector_derivative g (at x within {0..1}) \ b))) {0..1}" using f uv apply (simp add: line_integral_exists_def subpath_def) apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) apply (simp_all add: has_integral_integral) apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) apply (simp add: divide_simps False) done have vd:"\x. x \ {0..1} \ x \ (\t. (v-u) *\<^sub>R t + u) -` s \ vector_derivative (\x. g ((v-u) * x + u)) (at x within {0..1}) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u) within {0..1})" using test2[OF s fs uv] by auto have arg:"\x. (\n\basis. (v - u) * (f (g ((v - u) * x + u)) \ n) * (vector_derivative g (at ((v - u) * x + u) within {0..1}) \ n)) = (\b\basis. f (g ((v - u) * x + u)) \ b * (v - u) * (vector_derivative g (at ((v - u) * x + u) within {0..1}) \ b))" by (simp add: mult.commute) have"((\x. \b\basis. f (g ((v - u) * x + u)) \ b * (vector_derivative (\x. g ((v - u) * x + u)) (at x within {0..1}) \ b)) has_integral (integral {u..v} (\x. \b\basis. f (g (x)) \ b * (vector_derivative g (at x within {0..1}) \ b)))) {0..1}" apply (cut_tac Henstock_Kurzweil_Integration.has_integral_mult_right [OF *, where c = "v-u"]) using fs assms apply (simp add: False subpath_def line_integral_exists_def) apply (rule_tac S = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) apply (auto simp: inj_on_def False vd finite_vimageI scaleR_conv_of_real Groups_Big.sum_distrib_left mult.assoc[symmetric] arg) done then show "(line_integral_exists f basis (subpath u v g))" by(auto simp add: line_integral_exists_def subpath_def integrable_on_def) qed (*This should have everything that has to do with onecubes*) type_synonym path = "real \ (real * real)" type_synonym one_cube = "(real \ (real * real))" type_synonym one_chain = "(int * path) set" type_synonym two_cube = "(real * real) \ (real * real)" type_synonym two_chain = "two_cube set" definition one_chain_line_integral :: "((real * real) \ (real * real)) => ((real*real) set) \ one_chain \ real" where "one_chain_line_integral F b C \ (\(k,g)\C. k * (line_integral F b g))" definition boundary_chain where "boundary_chain s \ (\(k, \) \ s. k = 1 \ k = -1)" fun coeff_cube_to_path::"(int * one_cube) \ path" where "coeff_cube_to_path (k, \) = (if k = 1 then \ else (reversepath \))" fun rec_join :: "(int*path) list \ path" where "rec_join [] = (\x.0)" | "rec_join [oneC] = coeff_cube_to_path oneC" | "rec_join (oneC # xs) = coeff_cube_to_path oneC +++ (rec_join xs)" fun valid_chain_list where "valid_chain_list [] = True" | "valid_chain_list [oneC] = True" | "valid_chain_list (oneC#l) = (pathfinish (coeff_cube_to_path (oneC)) = pathstart (rec_join l) \ valid_chain_list l)" lemma joined_is_valid: assumes boundary_chain: "boundary_chain (set l)" and valid_path: "\k \. (k, \) \ set l \ valid_path \" and valid_chain_list_ass: "valid_chain_list l" shows "valid_path (rec_join l)" using assms proof(induction l) case Nil then show ?case using C1_differentiable_imp_piecewise[OF C1_differentiable_on_const[of "0" "{0..1}"]] by (auto simp add: valid_path_def) next case (Cons a l) have *: "valid_path (rec_join ((k::int, \) # l))" if "boundary_chain (set (l))" "(\k' \'. (k', \') \ set l \ valid_path \')" "valid_chain_list l" "valid_path (rec_join l) " "(\k' \'. (k', \') \ set ((k, \) # l) \ valid_path \')" "valid_chain_list ((k, \) # l)" "boundary_chain (set ((k,\) # l))" for k \ l proof (cases "l = []") case True with that show "valid_path (rec_join ((k, \) # l))" by auto next case False then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) then show "valid_path (rec_join ((k, \) # l))" proof (simp, intro conjI impI) assume k_eq_1: "k = 1" have 0:"valid_path \" using that by auto have 1:"pathfinish \ = pathstart (rec_join (h#l'))" using that(6) k_eq_1 l_nempty by auto show "valid_path (\ +++ rec_join (h#l'))" using 0 1 valid_path_join that(4) l_nempty by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using that(7) by (auto simp add: boundary_chain_def) have "valid_path \" using that by auto then have 0: "valid_path (reversepath \)" - using Cauchy_Integral_Theorem.valid_path_imp_reverse + using valid_path_imp_reverse by auto have 1: "pathfinish (reversepath \) = pathstart (rec_join (h#l'))" using that(6) k_eq_neg_1 l_nempty by auto show "valid_path ((reversepath \) +++ rec_join (h#l'))" - using 0 1 Cauchy_Integral_Theorem.valid_path_join that(4) l_nempty by blast + using 0 1 valid_path_join that(4) l_nempty by blast qed qed have "\ps. valid_chain_list ps \ (\i f p psa. ps = (i, f) # p # psa \ ((i = 1 \ pathfinish f \ pathstart (rec_join (p # psa)) \ i \ 1 \ pathfinish (reversepath f) \ pathstart (rec_join (p # psa))) \ \ valid_chain_list (p # psa)))" by (smt coeff_cube_to_path.elims valid_chain_list.elims(3)) moreover have "boundary_chain (set l)" by (meson Cons.prems(1) boundary_chain_def set_subset_Cons subset_eq) ultimately show ?case using * Cons by (metis (no_types) list.set_intros(2) prod.collapse valid_chain_list.simps(3)) qed lemma pathstart_rec_join_1: "pathstart (rec_join ((1, \) # l)) = pathstart \" proof (cases "l = []") case True then show "pathstart (rec_join ((1, \) # l)) = pathstart \" by simp next case False then obtain h l' where "l = h#l'" by (meson rec_join.elims) then show "pathstart (rec_join ((1, \) # l)) = pathstart \" by simp qed lemma pathstart_rec_join_2: "pathstart (rec_join ((-1, \) # l)) = pathstart (reversepath \)" proof cases assume "l = []" then show "pathstart (rec_join ((- 1, \) # l)) = pathstart (reversepath \)" by simp next assume "l \ [] " then obtain h l' where "l = h#l'" by (meson rec_join.elims) then show "pathstart (rec_join ((- 1, \) # l)) = pathstart (reversepath \)" by simp qed lemma pathstart_rec_join: "pathstart (rec_join ((1, \) # l)) = pathstart \" "pathstart (rec_join ((-1, \) # l)) = pathstart (reversepath \)" using pathstart_rec_join_1 pathstart_rec_join_2 by auto lemma line_integral_exists_on_rec_join: assumes boundary_chain: "boundary_chain (set l)" and valid_chain_list: "valid_chain_list l" and valid_path: "\k \. (k, \) \ set l \ valid_path \" and line_integral_exists: "\(k, \) \ set l. line_integral_exists F basis \" shows "line_integral_exists F basis (rec_join l)" using assms proof (induction l) case Nil then show ?case proof (simp add: line_integral_exists_def) have "\x. (vector_derivative (\x. 0) (at x)) = 0" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x)" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x within {0..1})" using Derivative.vector_derivative_const_at by auto then have 0: "\x\{0..1}. (vector_derivative (\x. 0) (at x within{0..1})) = 0" by (simp add: gamma_deriv_at_within) have "(\x\{0..1}. (\b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) = 0)" by (simp add: 0) then have " ((\x. \b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) has_integral 0) {0..1}" by (meson has_integral_is_0) then show "(\x. \b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) integrable_on {0..1}" by auto qed next case (Cons a l) obtain k \ where aeq: "a = (k,\)" by force show ?case unfolding aeq proof cases assume l_empty: "l = []" then show "line_integral_exists F basis (rec_join ((k, \) # l))" using Cons.prems aeq line_integral_on_reverse_path(2) by fastforce next assume "l \ []" then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) show "line_integral_exists F basis (rec_join ((k, \) # l))" proof (auto simp add: l_nempty) assume k_eq_1: "k = 1" have 0: "line_integral_exists F basis \" using Cons.prems(4) aeq by auto have 1: "line_integral_exists F basis (rec_join l)" by (metis (mono_tags) Cons boundary_chain_def list.set_intros(2) valid_chain_list.elims(3) valid_chain_list.simps(3)) have 2: "valid_path \" using Cons aeq by auto have 3:"valid_path (rec_join l)" by (metis (no_types) Cons.prems boundary_chain_def joined_is_valid l_nempty set_subset_Cons subsetCE valid_chain_list.simps(3)) show "line_integral_exists F basis (\ +++ rec_join (h#l'))" using line_integral_distrib(2)[OF 0 1 2 3] assms l_nempty by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using Cons aeq by (simp add: boundary_chain_def) have gamma_valid: "valid_path \" using Cons aeq by auto then have 2: "valid_path (reversepath \)" - using Cauchy_Integral_Theorem.valid_path_imp_reverse by auto + using valid_path_imp_reverse by auto have "line_integral_exists F basis \" using Cons aeq by auto then have 0: "line_integral_exists F basis (reversepath \)" using line_integral_on_reverse_path(2) gamma_valid by auto have 1: "line_integral_exists F basis (rec_join l)" using Cons aeq by (metis (mono_tags) boundary_chain_def insert_iff list.set(2) list.set_intros(2) valid_chain_list.elims(3) valid_chain_list.simps(3)) have 3:"valid_path (rec_join l)" by (metis (no_types) Cons.prems(1) Cons.prems(2) Cons.prems(3) boundary_chain_def joined_is_valid l_nempty set_subset_Cons subsetCE valid_chain_list.simps(3)) show "line_integral_exists F basis ((reversepath \) +++ rec_join (h#l'))" using line_integral_distrib(2)[OF 0 1 2 3] assms l_nempty by auto qed qed qed lemma line_integral_exists_rec_join_cons: assumes "line_integral_exists F basis (rec_join ((1,\) # l))" "(\k' \'. (k', \') \ set ((1,\) # l) \ valid_path \')" "finite basis" shows "line_integral_exists F basis (\ +++ (rec_join l))" proof cases assume l_empty: "l = []" show "line_integral_exists F basis (\ +++ rec_join l)" using assms(2) line_integral_distrib(2)[OF assms(1) line_integral_exists_on_degenerate_path[OF assms(3)], of "0"] using degenerate_path_is_valid_path by (fastforce simp add: l_empty) next assume "l \ []" then obtain h l' where "l = h#l'" by (meson rec_join.elims) then show "line_integral_exists F basis (\ +++ rec_join l)" using assms by auto qed lemma line_integral_exists_rec_join_cons_2: assumes "line_integral_exists F basis (rec_join ((-1,\) # l))" "(\k' \'. (k', \') \ set ((1,\) # l) \ valid_path \')" "finite basis" shows "line_integral_exists F basis ((reversepath \) +++ (rec_join l))" proof cases assume l_empty: "l = []" show "line_integral_exists F basis ((reversepath \) +++ rec_join l)" using assms(2) line_integral_distrib(2)[OF assms(1) line_integral_exists_on_degenerate_path[OF assms(3)], of "0"] using degenerate_path_is_valid_path by (auto simp add: l_empty) next assume "l \ []" then obtain h l' where "l = h#l'" by (meson rec_join.elims) with assms show "line_integral_exists F basis ((reversepath \) +++ rec_join l)" using assms by auto qed lemma line_integral_exists_on_rec_join': assumes boundary_chain: "boundary_chain (set l)" and valid_chain_list: "valid_chain_list l" and valid_path: "\k \. (k, \) \ set l \ valid_path \" and line_integral_exists: "line_integral_exists F basis (rec_join l)" and finite_basis: "finite basis" shows "\(k, \) \ set l. line_integral_exists F basis \" using assms proof (induction l) case Nil show ?case by (simp add: line_integral_exists_def) next case ass: (Cons a l) obtain k \ where k_gamma:"a = (k,\)" by fastforce show ?case apply (auto simp add: k_gamma) proof - show "line_integral_exists F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have 0: "line_integral_exists F basis (\ +++ (rec_join l))" using line_integral_exists_rec_join_cons k_eq_1 k_gamma ass(4) ass(5) ass(6) by auto have 2: "valid_path \" using ass k_gamma by auto show "line_integral_exists F basis \" using line_integral_exists_joinD1[OF 0 2] by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using ass k_gamma by (simp add: boundary_chain_def) have 0: "line_integral_exists F basis ((reversepath \) +++ (rec_join l))" using line_integral_exists_rec_join_cons_2[OF ] k_eq_neg_1 k_gamma ass(4) ass(5) ass(6) by fastforce have gamma_valid: "valid_path \" using ass k_gamma by auto then have 2: "valid_path (reversepath \)" - using Cauchy_Integral_Theorem.valid_path_imp_reverse by auto + using valid_path_imp_reverse by auto have "line_integral_exists F basis (reversepath \)" using line_integral_exists_joinD1[OF 0 2] by auto then show "line_integral_exists F basis (\)" using line_integral_on_reverse_path(2)[OF 2] reversepath_reversepath by auto qed next have 0:"boundary_chain (set l)" using ass(2) by (auto simp add: boundary_chain_def) have 1:"valid_chain_list l" using ass(3) apply (auto simp add: k_gamma) by (metis valid_chain_list.elims(3) valid_chain_list.simps(3)) have 2:"(\k \. (k, \) \ set (l) \ valid_path \)" using ass(4) by auto have 3: "valid_path (rec_join l)" using joined_is_valid[OF 0] 1 2 by auto have 4: "line_integral_exists F basis (rec_join l)" proof(cases "k = 1") assume k_eq_1: "k = 1" have 0: "line_integral_exists F basis (\ +++ (rec_join l))" using line_integral_exists_rec_join_cons k_eq_1 k_gamma ass(4) ass(5) ass(6) by auto show "line_integral_exists F basis (rec_join l)" using line_integral_exists_joinD2[OF 0 3] by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using ass k_gamma by (simp add: boundary_chain_def) have 0: "line_integral_exists F basis ((reversepath \) +++ (rec_join l))" using line_integral_exists_rec_join_cons_2[OF ] k_eq_neg_1 k_gamma ass(4) ass(5) ass(6) by fastforce show "line_integral_exists F basis (rec_join l)" using line_integral_exists_joinD2[OF 0 3] by auto qed show "\a b. (a, b) \ set l \ line_integral_exists F basis b" using 0 1 2 3 4 ass(1)[OF 0 1 2] ass(6) by fastforce qed qed inductive chain_subdiv_path where I: "chain_subdiv_path \ (set l)" if "distinct l" "rec_join l = \" "valid_chain_list l" lemma valid_path_equiv_valid_chain_list: assumes path_eq_chain: "chain_subdiv_path \ one_chain" and "boundary_chain one_chain" "\(k, \) \ one_chain. valid_path \" shows "valid_path \" proof - obtain l where l_props: "set l = one_chain" "distinct l" "rec_join l = \" "valid_chain_list l" using chain_subdiv_path.cases path_eq_chain by force show "valid_path \" using joined_is_valid assms l_props by blast qed lemma line_integral_rec_join_cons: assumes "line_integral_exists F basis \" "line_integral_exists F basis (rec_join ((l)))" "(\k' \'. (k', \') \ set ((1,\) # l) \ valid_path \')" "finite basis" shows "line_integral F basis (rec_join ((1,\) # l)) = line_integral F basis (\ +++ (rec_join l))" proof cases assume l_empty: "l = []" show "line_integral F basis (rec_join ((1,\) # l)) = line_integral F basis (\ +++ (rec_join l))" using assms line_integral_distrib(1)[OF assms(1) line_integral_exists_on_degenerate_path[OF assms(4)], of "0"] apply (auto simp add: l_empty) using degenerate_path_is_valid_path line_integral_degenerate_path by fastforce next assume "l \ []" then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) show "line_integral F basis (rec_join ((1,\) # l)) = line_integral F basis (\ +++ (rec_join l))" using assms by (auto simp add: l_nempty) qed lemma line_integral_rec_join_cons_2: assumes "line_integral_exists F basis \" "line_integral_exists F basis (rec_join ((l)))" "(\k' \'. (k', \') \ set ((-1,\) # l) \ valid_path \')" "finite basis" shows "line_integral F basis (rec_join ((-1,\) # l)) = line_integral F basis ((reversepath \) +++ (rec_join l))" proof cases assume l_empty: "l = []" have 0: "line_integral_exists F basis (reversepath \)" using assms line_integral_on_reverse_path(2) by fastforce have 1: "valid_path (reversepath \)" using assms by fastforce show "line_integral F basis (rec_join ((-1,\) # l)) = line_integral F basis ((reversepath \) +++ (rec_join l))" using assms line_integral_distrib(1)[OF 0 line_integral_exists_on_degenerate_path[OF assms(4)], of "0"] apply (auto simp add: l_empty) using degenerate_path_is_valid_path line_integral_degenerate_path by fastforce next assume "l \ []" then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) show "line_integral F basis (rec_join ((-1,\) # l)) = line_integral F basis ((reversepath \) +++ (rec_join l))" using assms by (auto simp add: l_nempty) qed lemma one_chain_line_integral_rec_join: assumes l_props: "set l = one_chain" "distinct l" "valid_chain_list l" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "\(k::int, \) \ one_chain. line_integral_exists F basis \" and valid_path: "\(k::int, \) \ one_chain. valid_path \" and finite_basis: "finite basis" shows "line_integral F basis (rec_join l) = one_chain_line_integral F basis one_chain" proof - have 0: "sum_list (map (\((k::int), \). (k::int) * (line_integral F basis \)) l) = one_chain_line_integral F basis one_chain" unfolding one_chain_line_integral_def using l_props Groups_List.comm_monoid_add_class.sum.distinct_set_conv_list[OF l_props(2), of "(\(k, \). (k::int) * (line_integral F basis \))"] by auto have "valid_chain_list l \ boundary_chain (set l) \ (\(k::int, \) \ set l. line_integral_exists F basis \) \ (\(k::int, \) \ set l. valid_path \) \ line_integral F basis (rec_join l) = sum_list (map (\(k::int, \). k * (line_integral F basis \)) l)" proof (induction l) case Nil show ?case unfolding line_integral_def boundary_chain_def apply (auto) proof have "\x. (vector_derivative (\x. 0) (at x)) = 0" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x)" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x within {0..1})" using Derivative.vector_derivative_const_at by auto then have 0: "\x\{0..1}. (vector_derivative (\x. 0) (at x within{0..1})) = 0" by (metis (no_types) box_real(2) vector_derivative_within_cbox zero_less_one) have "(\x\{0..1}. (\b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) = 0)" by (simp add: 0) then show "((\x. \b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) has_integral 0) {0..1}" by (meson has_integral_is_0) qed next case ass: (Cons a l) obtain k::"int" and \::"one_cube" where props: "a = (k,\)" proof let ?k2 = "fst a" let ?\2 = "snd a" show "a = (?k2, ?\2)" by auto qed have "line_integral_exists F basis (rec_join (a # l))" using line_integral_exists_on_rec_join[OF ass(3) ass(2)] ass(5) ass(4) by blast have "boundary_chain (set l)" by (meson ass.prems(2) boundary_chain_def list.set_intros(2)) have val_l: "\f i. (i,f) \ set l \ valid_path f" using ass.prems(4) by fastforce have vcl_l: "valid_chain_list l" by (metis (no_types) ass.prems(1) valid_chain_list.elims(3) valid_chain_list.simps(3)) have line_integral_exists_on_joined: "line_integral_exists F basis (rec_join l)" by (metis \boundary_chain (set l)\ \line_integral_exists F basis (rec_join (a # l))\ emptyE val_l vcl_l joined_is_valid line_integral_exists_joinD2 line_integral_exists_on_rec_join list.set(1) neq_Nil_conv rec_join.simps(3)) have "valid_path (rec_join (a # l))" using joined_is_valid ass(5) ass(3) ass(2) by blast then have joined_is_valid: "valid_path (rec_join l)" using \boundary_chain (set l)\ val_l vcl_l joined_is_valid by blast show ?case proof (clarsimp, cases) assume k_eq_1: "(k::int) = 1" have line_integral_exists_on_gamma: "line_integral_exists F basis \" using ass props by auto have gamma_is_valid: "valid_path \" using ass props by auto have line_int_rw: "line_integral F basis (rec_join ((k, \) # l)) = line_integral F basis (\ +++ rec_join l)" proof - have gam_int: "line_integral_exists F basis \" using ass props by auto have rec_join_int: "line_integral_exists F basis (rec_join l)" using line_integral_exists_on_rec_join using line_integral_exists_on_joined by blast show ?thesis using line_integral_rec_join_cons[OF gam_int rec_join_int] ass k_eq_1 finite_basis props by force qed show "line_integral F basis (rec_join (a # l)) = (case a of (x, \) \ real_of_int x * line_integral F basis \) + (\(x, \)\l. real_of_int x * line_integral F basis \)" apply (simp add: props line_int_rw) using line_integral_distrib[OF line_integral_exists_on_gamma line_integral_exists_on_joined gamma_is_valid joined_is_valid] ass k_eq_1 vcl_l by (auto simp: boundary_chain_def props) next assume "k \ 1" then have k_eq_neg_1: "k = -1" using ass props by (auto simp add: boundary_chain_def) have line_integral_exists_on_gamma: "line_integral_exists F basis (reversepath \)" using line_integral_on_reverse_path ass props by auto have gamma_is_valid: "valid_path (reversepath \)" - using Cauchy_Integral_Theorem.valid_path_imp_reverse ass props by auto + using valid_path_imp_reverse ass props by auto have line_int_rw: "line_integral F basis (rec_join ((k, \) # l)) = line_integral F basis ((reversepath \) +++ rec_join l)" proof - have gam_int: "line_integral_exists F basis \" using ass props by auto have rec_join_int: "line_integral_exists F basis (rec_join l)" using line_integral_exists_on_rec_join using line_integral_exists_on_joined by blast show ?thesis using line_integral_rec_join_cons_2[OF gam_int rec_join_int] using ass k_eq_neg_1 using finite_basis props by blast qed show "line_integral F basis (rec_join (a # l)) = (case a of (x, \) \ real_of_int x * line_integral F basis \) + (\(x, \)\l. real_of_int x * line_integral F basis \)" apply (simp add: props line_int_rw) using line_integral_distrib[OF line_integral_exists_on_gamma line_integral_exists_on_joined gamma_is_valid joined_is_valid] props ass line_integral_on_reverse_path(1)[of "\" "F" "basis"] k_eq_neg_1 using \boundary_chain (set l)\ vcl_l by auto qed qed then have 1:"line_integral F basis (rec_join l) = sum_list (map (\(k::int, \). k * (line_integral F basis \)) l)" using l_props assms by auto then show ?thesis using 0 1 by auto qed lemma line_integral_on_path_eq_line_integral_on_equiv_chain: assumes path_eq_chain: "chain_subdiv_path \ one_chain" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "\(k::int, \) \ one_chain. line_integral_exists F basis \" and valid_path: "\(k::int, \) \ one_chain. valid_path \" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain = line_integral F basis \" "line_integral_exists F basis \" "valid_path \" proof - obtain l where l_props: "set l = one_chain" "distinct l" "rec_join l = \" "valid_chain_list l" using chain_subdiv_path.cases path_eq_chain by force show "line_integral_exists F basis \" using line_integral_exists_on_rec_join assms l_props by blast show "valid_path \" using joined_is_valid assms l_props by blast have "line_integral F basis (rec_join l) = one_chain_line_integral F basis one_chain" using one_chain_line_integral_rec_join l_props assms by auto then show "one_chain_line_integral F basis one_chain = line_integral F basis \" using l_props by auto qed lemma line_integral_on_path_eq_line_integral_on_equiv_chain': assumes path_eq_chain: "chain_subdiv_path \ one_chain" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "line_integral_exists F basis \" and valid_path: "\(k, \) \ one_chain. valid_path \" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain = line_integral F basis \" "\(k, \) \ one_chain. line_integral_exists F basis \" proof - obtain l where l_props: "set l = one_chain" "distinct l" "rec_join l = \" "valid_chain_list l" using chain_subdiv_path.cases path_eq_chain by force show 0: "\(k, \) \ one_chain. line_integral_exists F basis \" using line_integral_exists_on_rec_join' assms l_props by blast show "one_chain_line_integral F basis one_chain = line_integral F basis \" using line_integral_on_path_eq_line_integral_on_equiv_chain(1)[OF assms(1) assms(2) 0 assms(4) assms(5)] by auto qed definition chain_subdiv_chain where "chain_subdiv_chain one_chain1 subdiv \ \f. (\(f ` one_chain1)) = subdiv \ (\c\one_chain1. chain_subdiv_path (coeff_cube_to_path c) (f c)) \ pairwise (\ p p'. f p \ f p' = {}) one_chain1 \ (\x \ one_chain1. finite (f x))" lemma chain_subdiv_chain_character: shows "chain_subdiv_chain one_chain1 subdiv \ (\f. \(f ` one_chain1) = subdiv \ (\(k, \)\one_chain1. if k = 1 then chain_subdiv_path \ (f (k, \)) else chain_subdiv_path (reversepath \) (f (k, \))) \ (\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {}) \ (\x\one_chain1. finite (f x)))" unfolding chain_subdiv_chain_def by (safe; intro exI conjI iffI; fastforce simp add: pairwise_def) lemma chain_subdiv_chain_imp_finite_subdiv: assumes "finite one_chain1" "chain_subdiv_chain one_chain1 subdiv" shows "finite subdiv" using assms by(auto simp add: chain_subdiv_chain_def) lemma valid_subdiv_imp_valid_one_chain: assumes chain1_eq_chain2: "chain_subdiv_chain one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and valid_path: "\(k, \) \ subdiv. valid_path \" shows "\(k, \) \ one_chain1. valid_path \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\(k,\)\one_chain1. if k = 1 then chain_subdiv_path \ (f(k,\)) else chain_subdiv_path (reversepath \) (f(k,\)))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" using chain1_eq_chain2 chain_subdiv_chain_character by auto have "\ k \. (k,\) \ one_chain1\ valid_path \" proof- fix k \ assume ass: "(k,\) \ one_chain1" show "valid_path \" proof (cases "k = 1") assume k_eq_1: "k = 1" then have i:"chain_subdiv_path \ (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms apply (simp add: boundary_chain_def) by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast show ?thesis using valid_path_equiv_valid_chain_list[OF i ii iv] by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) then have i:"chain_subdiv_path (reversepath \) (f(k,\))" using f_props(2) ass using \k \ 1\ by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms by (auto simp add: boundary_chain_def) have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have "valid_path (reversepath \)" using valid_path_equiv_valid_chain_list[OF i ii iv] by auto then show ?thesis - using reversepath_reversepath Cauchy_Integral_Theorem.valid_path_imp_reverse + using reversepath_reversepath valid_path_imp_reverse by force qed qed then show valid_path1: "\(k, \) \ one_chain1. valid_path \" by auto qed lemma one_chain_line_integral_eq_line_integral_on_sudivision: assumes chain1_eq_chain2: "chain_subdiv_chain one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and line_integral_exists_on_chain2: "\(k, \) \ subdiv. line_integral_exists F basis \" and valid_path: "\(k, \) \ subdiv. valid_path \" and finite_chain1: "finite one_chain1" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" "\(k, \) \ one_chain1. line_integral_exists F basis \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\(k,\)\one_chain1. if k = 1 then chain_subdiv_path \ (f(k,\)) else chain_subdiv_path (reversepath \) (f(k,\)))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" "(\x \ one_chain1. finite (f x))" using chain1_eq_chain2 chain_subdiv_chain_character by (auto simp add: pairwise_def chain_subdiv_chain_def) then have 0: "one_chain_line_integral F basis subdiv = one_chain_line_integral F basis (\(f ` one_chain1))" by auto have finite_chain2: "finite subdiv" using finite_chain1 f_props(1) f_props(4) apply (simp add: image_def) using f_props(1) by auto have "\ k \. (k,\) \ one_chain1\ line_integral_exists F basis \" proof- fix k \ assume ass: "(k,\) \ one_chain1" show "line_integral_exists F basis \" proof (cases "k = 1") assume k_eq_1: "k = 1" then have i:"chain_subdiv_path \ (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms by (auto simp add: boundary_chain_def) have iii:"\(k, \)\f (k, \). line_integral_exists F basis \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast show ?thesis using line_integral_on_path_eq_line_integral_on_equiv_chain(2)[OF i ii iii iv finite_basis] by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) then have i:"chain_subdiv_path (reversepath \) (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms by (auto simp add: boundary_chain_def) have iii:"\(k, \)\f (k, \). line_integral_exists F basis \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have x:"line_integral_exists F basis (reversepath \)" using line_integral_on_path_eq_line_integral_on_equiv_chain(2)[OF i ii iii iv finite_basis] by auto have "valid_path (reversepath \)" using line_integral_on_path_eq_line_integral_on_equiv_chain(3)[OF i ii iii iv finite_basis] by auto then show ?thesis using line_integral_on_reverse_path(2) reversepath_reversepath x by fastforce qed qed then show line_integral_exists_on_chain1: "\(k, \) \ one_chain1. line_integral_exists F basis \" by auto have 1: "one_chain_line_integral F basis (\(f ` one_chain1)) = one_chain_line_integral F basis one_chain1" proof - have 0:"one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" proof - have finite: "\chain \ (f ` one_chain1). finite chain" using f_props(1) finite_chain2 by (meson Sup_upper finite_subset) have disj: " \A\f ` one_chain1. \B\f ` one_chain1. A \ B \ A \ B = {}" by (metis (no_types, hide_lams) f_props(3) image_iff) show "one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" using Groups_Big.comm_monoid_add_class.sum.Union_disjoint[OF finite disj] one_chain_line_integral_def by auto qed have 1:"(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = one_chain_line_integral F basis one_chain1" proof - have "(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1. k*(line_integral F basis \))" proof - have i:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \))" proof - have "inj_on f (one_chain1 - {p. f p = {}})" unfolding inj_on_def using f_props(3) by blast then have 0: "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \)))" using Groups_Big.comm_monoid_add_class.sum.reindex by auto have "\ k \. (k, \) \ (one_chain1 - {p. f p = {}}) \ one_chain_line_integral F basis (f(k, \)) = k* (line_integral F basis \)" proof- fix k \ assume ass: "(k, \) \ (one_chain1 - {p. f p = {}})" have bchain: "boundary_chain (f(k,\))" using f_props(1) boundary_chain2 ass by (auto simp add: boundary_chain_def) have wexist: "\(k, \)\(f(k,\)). line_integral_exists F basis \" using f_props(1) line_integral_exists_on_chain2 ass by blast have vpath: " \(k, \)\(f(k, \)). valid_path \" using f_props(1) assms ass by blast show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using f_props(2) k_eq_1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using f_props(2) k_eq_neg1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then have "one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path(1) ass line_integral_exists_on_chain1 valid_subdiv_imp_valid_one_chain[OF chain1_eq_chain2 boundary_chain1 boundary_chain2 valid_path] by force then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed qed then have "(\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \))) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" by (auto intro!: Finite_Cartesian_Product.sum_cong_aux) then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" using 0 by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ (k, \) \ {(k',\'). k'* (line_integral F basis \') = 0}" proof- fix k::"int" fix \::"one_cube" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto have bchain: "boundary_chain (f(k,\))" using f_props(1) boundary_chain2 ass by (auto simp add: boundary_chain_def) have wexist: "\(k, \)\(f(k,\)). line_integral_exists F basis \" using f_props(1) line_integral_exists_on_chain2 ass by blast have vpath: " \(k, \)\(f(k, \)). valid_path \" using f_props(1) assms ass by blast have "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using f_props(2) k_eq_1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using f_props(2) k_eq_neg1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then have "one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path(1) ass line_integral_exists_on_chain1 valid_subdiv_imp_valid_one_chain[OF chain1_eq_chain2 boundary_chain1 boundary_chain2 valid_path] by force then show "one_chain_line_integral F basis (f (k::int, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed then show "(k, \) \ {(k'::int, \'). k' * line_integral F basis \' = 0}" using zero_line_integral by auto qed then have ii:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" proof - have "\one_chain. one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}})) \ one_chain_line_integral F basis one_chain = 0" proof - fix one_chain assume "one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}}))" then show "one_chain_line_integral F basis one_chain = 0" by (auto simp add: one_chain_line_integral_def) qed then have 0:"(\one_chain \ f ` (one_chain1) - (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" using comm_monoid_add_class.sum.neutral by auto then have "(\one_chain \ f ` (one_chain1). one_chain_line_integral F basis one_chain) - (\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" proof - have finte: "finite (f ` one_chain1)" using finite_chain1 by auto show ?thesis using Groups_Big.sum_diff[OF finte, of " (f ` (one_chain1 - {p. f p = {}}))" "one_chain_line_integral F basis"] 0 by auto qed then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ f (k, \) \ {chain. one_chain_line_integral F basis chain = 0}" proof- fix k::"int" fix \::"one_cube" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have "one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto then show "f (k, \) \ {p'. (one_chain_line_integral F basis p' = 0)}" by auto qed then have iii:"(\(k::int,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \)) = (\(k::int,\)\one_chain1. k*(line_integral F basis \))" proof- have "\ k \. (k,\)\one_chain1 - (one_chain1 - {p. f p = {}}) \ k* (line_integral F basis \) = 0" proof- fix k \ assume ass: "(k,\)\one_chain1 - (one_chain1 - {p. f p = {}})" then have "f(k, \) = {}" by auto then have "one_chain_line_integral F basis (f(k, \)) = 0" by (auto simp add: one_chain_line_integral_def) then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto have bchain: "boundary_chain (f(k,\))" using f_props(1) boundary_chain2 ass by (auto simp add: boundary_chain_def) have wexist: "\(k, \)\(f(k,\)). line_integral_exists F basis \" using f_props(1) line_integral_exists_on_chain2 ass by blast have vpath: " \(k, \)\(f(k, \)). valid_path \" using f_props(1) assms ass by blast have "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using f_props(2) k_eq_1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using f_props(2) k_eq_neg1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then have "one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path(1) ass line_integral_exists_on_chain1 valid_subdiv_imp_valid_one_chain[OF chain1_eq_chain2 boundary_chain1 boundary_chain2 valid_path] by force then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed then show "k* (line_integral F basis \) = 0" using zero_line_integral by auto qed then have "\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k* (line_integral F basis \) = 0" by auto then have "(\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.comm_monoid_add_class.sum.neutral [of "one_chain1 - (one_chain1 - {p. f p = {}})" "(\(k::int,\). k* (line_integral F basis \))"] by (simp add: split_beta) then have "(\(k::int,\)\one_chain1. k*(line_integral F basis \)) - (\(k::int,\)\ (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.sum_diff[OF finite_chain1] by (metis (no_types) Diff_subset \(\(k, \)\one_chain1 - (one_chain1 - {p. f p = {}}). k * line_integral F basis \) = 0\ \\f B. B \ one_chain1 \ sum f (one_chain1 - B) = sum f one_chain1 - sum f B\) then show ?thesis by auto qed show ?thesis using i ii iii by auto qed then show ?thesis using one_chain_line_integral_def by auto qed show ?thesis using 0 1 by auto qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" using 0 1 by auto qed lemma one_chain_line_integral_eq_line_integral_on_sudivision': assumes chain1_eq_chain2: "chain_subdiv_chain one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and line_integral_exists_on_chain1: "\(k, \) \ one_chain1. line_integral_exists F basis \" and valid_path: "\(k, \) \ subdiv. valid_path \" and finite_chain1: "finite one_chain1" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" "\(k, \) \ subdiv. line_integral_exists F basis \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\(k,\)\one_chain1. if k = 1 then chain_subdiv_path \ (f(k,\)) else chain_subdiv_path (reversepath \) (f(k,\)))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" "(\x \ one_chain1. finite (f x))" using chain1_eq_chain2 chain_subdiv_chain_character by (auto simp add: pairwise_def chain_subdiv_chain_def) have finite_chain2: "finite subdiv" using finite_chain1 f_props(1) f_props(4) by blast have "\k \. (k, \) \ subdiv \ line_integral_exists F basis \" proof - fix k \ assume ass: "(k, \) \ subdiv" then obtain k' \' where kp_gammap: "(k',\') \ one_chain1" "(k,\) \ f(k',\')" using f_props(1) by fastforce show "line_integral_exists F basis \" proof (cases "k' = 1") assume k_eq_1: "k' = 1" then have i:"chain_subdiv_path \' (f(k',\'))" using f_props(2) kp_gammap ass by auto have ii:"boundary_chain (f(k',\'))" using f_props(1) ass assms kp_gammap by (meson UN_I boundary_chain_def) have iii:"line_integral_exists F basis \'" using assms kp_gammap by blast have "iv": " \(k, \)\f (k', \'). valid_path \" using f_props(1) ass assms kp_gammap by blast show ?thesis using line_integral_on_path_eq_line_integral_on_equiv_chain'(2)[OF i ii iii iv finite_basis] kp_gammap by auto next assume "k' \ 1" then have k_eq_neg1: "k' = -1" using boundary_chain1 kp_gammap by (auto simp add: boundary_chain_def) then have i:"chain_subdiv_path (reversepath \') (f(k',\'))" using f_props(2) kp_gammap by auto have ii:"boundary_chain (f(k',\'))" using f_props(1) assms kp_gammap by (meson UN_I boundary_chain_def) have iii: " \(k, \)\f (k', \'). valid_path \" using f_props(1) ass assms kp_gammap by blast have iv: "valid_path (reversepath \')" using valid_path_equiv_valid_chain_list[OF i ii iii] by force have "line_integral_exists F basis \'" using assms kp_gammap by blast then have x: "line_integral_exists F basis (reversepath \')" using iv line_integral_on_reverse_path(2) valid_path_reversepath by fastforce show ?thesis using line_integral_on_path_eq_line_integral_on_equiv_chain'(2)[OF i ii x iii finite_basis] kp_gammap by auto qed qed then show "\(k, \)\subdiv. line_integral_exists F basis \" by auto then show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" using one_chain_line_integral_eq_line_integral_on_sudivision(1) assms by auto qed lemma line_integral_sum_gen: assumes finite_basis: "finite basis" and line_integral_exists: "line_integral_exists F basis1 \" "line_integral_exists F basis2 \" and basis_partition: "basis1 \ basis2 = basis" "basis1 \ basis2 = {}" shows "line_integral F basis \ = (line_integral F basis1 \) + (line_integral F basis2 \)" "line_integral_exists F basis \" apply (simp add: line_integral_def) proof - have 0: "integral {0..1} (\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))) = integral {0..1} (\x. \b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + integral {0..1} (\x. \b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))" using Henstock_Kurzweil_Integration.integral_add line_integral_exists by (auto simp add: line_integral_exists_def) have 1: "integral {0..1} (\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) = integral {0..1} (\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)))" by (metis (mono_tags, lifting) basis_partition finite_Un finite_basis sum.union_disjoint) show "integral {0..1} (\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) = integral {0..1} (\x. \b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + integral {0..1} (\x. \b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))" using 0 1 by auto have 2: "((\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))) has_integral integral {0..1} (\x. \b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + integral {0..1} (\x. \b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))) {0..1}" using Henstock_Kurzweil_Integration.has_integral_add line_integral_exists has_integral_integral apply (auto simp add: line_integral_exists_def) by blast have 3: "(\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) = (\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)))" by (metis (mono_tags, lifting) basis_partition finite_Un finite_basis sum.union_disjoint) show "line_integral_exists F basis \" apply(auto simp add: line_integral_exists_def has_integral_integral) using 2 3 using has_integral_integrable_integral by fastforce qed definition common_boundary_sudivision_exists where "common_boundary_sudivision_exists one_chain1 one_chain2 \ \subdiv. chain_subdiv_chain one_chain1 subdiv \ chain_subdiv_chain one_chain2 subdiv \ (\(k, \) \ subdiv. valid_path \)\ boundary_chain subdiv" lemma common_boundary_sudivision_commutative: "(common_boundary_sudivision_exists one_chain1 one_chain2) = (common_boundary_sudivision_exists one_chain2 one_chain1)" apply (simp add: common_boundary_sudivision_exists_def) by blast lemma common_subdivision_imp_eq_line_integral: assumes "(common_boundary_sudivision_exists one_chain1 one_chain2)" "boundary_chain one_chain1" "boundary_chain one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" "finite one_chain1" "finite one_chain2" "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain2. line_integral_exists F basis \" proof - obtain subdiv where subdiv_props: "chain_subdiv_chain one_chain1 subdiv" "chain_subdiv_chain one_chain2 subdiv" "(\(k, \) \ subdiv. valid_path \)" "(boundary_chain subdiv)" using assms by (auto simp add: common_boundary_sudivision_exists_def) have i: "\(k, \)\subdiv. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision'(2)[OF subdiv_props(1) assms(2) subdiv_props(4) assms(4) subdiv_props(3) assms(5) assms(7)] by auto show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using one_chain_line_integral_eq_line_integral_on_sudivision'(1)[OF subdiv_props(1) assms(2) subdiv_props(4) assms(4) subdiv_props(3) assms(5) assms(7)] one_chain_line_integral_eq_line_integral_on_sudivision(1)[OF subdiv_props(2) assms(3) subdiv_props(4) i subdiv_props(3) assms(6) assms(7)] by auto show "\(k, \)\one_chain2. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision(2)[OF subdiv_props(2) assms(3) subdiv_props(4) i subdiv_props(3) assms(6) assms(7)] by auto qed definition common_sudiv_exists where "common_sudiv_exists one_chain1 one_chain2 \ \subdiv ps1 ps2. chain_subdiv_chain (one_chain1 - ps1) subdiv \ chain_subdiv_chain (one_chain2 - ps2) subdiv \ (\(k, \) \ subdiv. valid_path \) \ (boundary_chain subdiv) \ (\(k, \) \ ps1. point_path \) \ (\(k, \) \ ps2. point_path \)" lemma common_sudiv_exists_comm: shows "common_sudiv_exists C1 C2 = common_sudiv_exists C2 C1" by (auto simp add: common_sudiv_exists_def) lemma line_integral_degenerate_chain: assumes "(\(k, \) \ chain. point_path \)" assumes "finite basis" shows "one_chain_line_integral F basis chain = 0" proof (simp add: one_chain_line_integral_def) have "\(k,g)\chain. line_integral F basis g = 0" using assms line_integral_point_path by blast then have "\(k,g)\chain. real_of_int k * line_integral F basis g = 0" by auto then have "\p. p \ chain \ (case p of (i, f) \ real_of_int i * line_integral F basis f) = 0" by fastforce then show "(\x\chain. case x of (k, g) \ real_of_int k * line_integral F basis g) = 0" by simp qed lemma gen_common_subdiv_imp_common_subdiv: shows "(common_sudiv_exists one_chain1 one_chain2) = (\ps1 ps2. (common_boundary_sudivision_exists (one_chain1 - ps1) (one_chain2 - ps2)) \ (\(k, \)\ps1. point_path \) \ (\(k, \)\ps2. point_path \))" by (auto simp add: common_sudiv_exists_def common_boundary_sudivision_exists_def) lemma common_subdiv_imp_gen_common_subdiv: assumes "(common_boundary_sudivision_exists one_chain1 one_chain2)" shows "(common_sudiv_exists one_chain1 one_chain2)" using assms apply (auto simp add: common_sudiv_exists_def common_boundary_sudivision_exists_def) by (metis Diff_empty all_not_in_conv) lemma one_chain_line_integral_point_paths: assumes "finite one_chain" assumes "finite basis" assumes "(\(k, \)\ps. point_path \)" shows "one_chain_line_integral F basis (one_chain - ps) = one_chain_line_integral F basis (one_chain)" proof- have 0:"(\x\ps. case x of (k, g) \ (real_of_int k * line_integral F basis g) = 0)" using line_integral_point_path assms by force show "one_chain_line_integral F basis (one_chain - ps) = one_chain_line_integral F basis one_chain" unfolding one_chain_line_integral_def using 0 \finite one_chain\ by (force simp add: intro: comm_monoid_add_class.sum.mono_neutral_left) qed lemma boundary_chain_diff: assumes "boundary_chain one_chain" shows "boundary_chain (one_chain - s)" using assms by (auto simp add: boundary_chain_def) lemma gen_common_subdivision_imp_eq_line_integral: assumes "(common_sudiv_exists one_chain1 one_chain2)" "boundary_chain one_chain1" "boundary_chain one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" "finite one_chain1" "finite one_chain2" "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain2. line_integral_exists F basis \" proof - obtain ps1 ps2 where gen_subdiv: "(common_boundary_sudivision_exists (one_chain1 - ps1) (one_chain2 - ps2))" "(\(k, \)\ps1. point_path \)" " (\(k, \)\ps2. point_path \)" using assms(1) gen_common_subdiv_imp_common_subdiv by blast show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using one_chain_line_integral_point_paths gen_common_subdiv_imp_common_subdiv assms(2-7) gen_subdiv common_subdivision_imp_eq_line_integral(1)[OF gen_subdiv(1) boundary_chain_diff[OF assms(2)] boundary_chain_diff[OF assms(3)]] by auto show "\(k, \)\one_chain2. line_integral_exists F basis \" proof- obtain subdiv where subdiv_props: "chain_subdiv_chain (one_chain1-ps1) subdiv" "chain_subdiv_chain (one_chain2-ps2) subdiv" "(\(k, \) \ subdiv. valid_path \)" "(boundary_chain subdiv)" using gen_subdiv(1) by (auto simp add: common_boundary_sudivision_exists_def) have "\(k, \)\subdiv. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision'(2)[OF subdiv_props(1) boundary_chain_diff[OF assms(2)] subdiv_props(4)] assms(4) subdiv_props(3) assms(5) assms(7) by blast then have i: "\(k, \)\one_chain2-ps2. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision(2)[OF subdiv_props(2) boundary_chain_diff[OF assms(3)] subdiv_props(4)] subdiv_props(3) assms(6) assms(7) by blast then show ?thesis using gen_subdiv(3) line_integral_exists_point_path[OF assms(7)] by blast qed qed lemma common_sudiv_exists_refl: assumes "common_sudiv_exists C1 C2" shows "common_sudiv_exists C2 C1" using assms apply(simp add: common_sudiv_exists_def) by auto lemma chain_subdiv_path_singleton: shows "chain_subdiv_path \ {(1,\)}" proof - have "rec_join [(1,\)] = \" by (simp add: joinpaths_def) then have "set [(1,\)] = {(1, \)}" "distinct [(1,\)]" "rec_join [(1,\)] = \" "valid_chain_list [(1,\)]" by auto then show ?thesis by (metis (no_types) chain_subdiv_path.intros) qed lemma chain_subdiv_path_singleton_reverse: shows "chain_subdiv_path (reversepath \) {(-1,\)}" proof - have "rec_join [(-1,\)] = reversepath \" by (simp add: joinpaths_def) then have "set [(-1,\)] = {(- 1, \)}" "distinct [(-1,\)]" "rec_join [(-1,\)] = reversepath \" "valid_chain_list [(-1,\)]" by auto then have "chain_subdiv_path (reversepath \) (set [(- 1, \)])" using chain_subdiv_path.intros by blast then show ?thesis by simp qed lemma chain_subdiv_chain_refl: assumes "boundary_chain C" shows "chain_subdiv_chain C C" using chain_subdiv_path_singleton chain_subdiv_path_singleton_reverse assms unfolding chain_subdiv_chain_def boundary_chain_def pairwise_def using case_prodI2 coeff_cube_to_path.simps by (rule_tac x="\x. {x}" in exI) auto (*path reparam_weaketrization*) definition reparam_weak where "reparam_weak \1 \2 \ \\. (\x\{0..1}. \1 x = (\2 o \) x) \ \ piecewise_C1_differentiable_on {0..1} \ \(0) = 0 \ \(1) = 1 \ \ ` {0..1} = {0..1}" definition reparam where "reparam \1 \2 \ \\. (\x\{0..1}. \1 x = (\2 o \) x) \ \ piecewise_C1_differentiable_on {0..1} \ \(0) = 0 \ \(1) = 1 \ bij_betw \ {0..1} {0..1} \ \ -` {0..1} \ {0..1} \ (\x\{0..1}. finite (\ -` {x}))" lemma reparam_weak_eq_refl: shows "reparam_weak \1 \1" unfolding reparam_weak_def apply (rule_tac x="id" in exI) by (auto simp add: id_def piecewise_C1_differentiable_on_def C1_differentiable_on_def continuous_on_id) lemma line_integral_exists_smooth_one_base: assumes "\ C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities of f*) "continuous_on (path_image \) (\x. F x \ b)" shows "line_integral_exists F {b} \" proof- have gamma2_differentiable: "(\x \ {0 .. 1}. \ differentiable at x)" using assms(1) by (auto simp add: valid_path_def C1_differentiable_on_eq) then have gamma2_b_component_differentiable: "(\x \ {0 .. 1}. (\x. (\ x) \ b) differentiable at x)" by auto then have "(\x. (\ x) \ b) differentiable_on {0..1}" using differentiable_at_withinI by (auto simp add: differentiable_on_def) then have gama2_cont_comp: "continuous_on {0..1} (\x. (\ x) \ b)" using differentiable_imp_continuous_on by auto have gamma2_cont:"continuous_on {0..1} \" using assms(1) C1_differentiable_imp_continuous_on by (auto simp add: valid_path_def) have iii: "continuous_on {0..1} (\x. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))" proof- have 0: "continuous_on {0..1} (\x. F (\ x) \ b)" using assms(2) continuous_on_compose[OF gamma2_cont] by (auto simp add: path_image_def) obtain D where D: "(\x\{0..1}. (\ has_vector_derivative D x) (at x)) \ continuous_on {0..1} D" using assms(1) by (auto simp add: C1_differentiable_on_def) then have *:"\x\{0..1}. vector_derivative \ (at x within{0..1}) = D x" using vector_derivative_at vector_derivative_at_within_ivl by fastforce then have "continuous_on {0..1} (\x. vector_derivative \ (at x within{0..1}))" using continuous_on_eq D by force then have 1: "continuous_on {0..1} (\x. (vector_derivative \ (at x within{0..1})) \ b)" by(auto intro!: continuous_intros) show ?thesis using continuous_on_mult[OF 0 1] by auto qed then have "(\x. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) integrable_on {0..1}" using integrable_continuous_real by auto then show "line_integral_exists F {b} \" by(auto simp add: line_integral_exists_def) qed lemma contour_integral_primitive_lemma: fixes f :: "complex \ complex" and g :: "real \ complex" assumes "a \ b" and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) has_integral (f(g b) - f(g a))) {a..b}" proof - obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" using assms by (auto simp: piecewise_differentiable_on_def) have cfg: "continuous_on {a..b} (\x. f (g x))" apply (rule continuous_on_compose [OF cg, unfolded o_def]) using assms apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) done { fix x::real assume a: "a < x" and b: "x < b" and xk: "x \ k" then have "g differentiable at x within {a..b}" using k by (simp add: differentiable_at_withinI) then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" by (simp add: has_vector_derivative_def scaleR_conv_of_real) have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" by (simp add: has_field_derivative_def) have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" using diff_chain_within [OF gdiff fdiff] by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) } note * = this show ?thesis apply (rule fundamental_theorem_of_calculus_interior_strong) using k assms cfg * apply (auto simp: at_within_Icc_at) done qed lemma line_integral_primitive_lemma: (*Only for real normed fields, i.e. vectors with products*) fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a::{euclidean_space,real_normed_field}" and g :: "real \ 'a" assumes "\(a::'a). a \ s \ (f has_field_derivative (f' a)) (at a within s)" and "g piecewise_differentiable_on {0::real..1}" "\x. x \ {0..1} \ g x \ s" and "base_vec \ Basis" shows "((\x. ((f'(g x)) * (vector_derivative g (at x within {0..1}))) \ base_vec) has_integral (((f(g 1)) \ base_vec - (f(g 0)) \ base_vec))) {0..1}" proof - obtain k where k: "finite k" "\x\{0..1} - k. g differentiable (at x within {0..1})" and cg: "continuous_on {0..1} g" using assms by (auto simp: piecewise_differentiable_on_def) have cfg: "continuous_on {0..1} (\x. f (g x))" apply (rule continuous_on_compose [OF cg, unfolded o_def]) using assms apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) done { fix x::real assume a: "0 < x" and b: "x < 1" and xk: "x \ k" then have "g differentiable at x within {0..1}" using k by (simp add: differentiable_at_withinI) then have "(g has_vector_derivative vector_derivative g (at x within {0..1})) (at x within {0..1})" by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) then have gdiff: "(g has_derivative (\u. of_real u * vector_derivative g (at x within {0..1}))) (at x within {0..1})" by (simp add: has_vector_derivative_def scaleR_conv_of_real) have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {0..1})" using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {0..1})" by (simp add: has_field_derivative_def) have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {0..1})) (at x within {0..1})" using diff_chain_within [OF gdiff fdiff] by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) } then have *: "\x. x\{0<..<1} - k \ ((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {0..1})) (at x within {0..1})" by auto have "((\x. ((f'(g x))) * ((vector_derivative g (at x within {0..1})))) has_integral (((f(g 1)) - (f(g 0))))) {0..1}" using fundamental_theorem_of_calculus_interior_strong[OF k(1) zero_le_one _ cfg] using k assms cfg * by (auto simp: at_within_Icc_at) then have "((\x. (((f'(g x))) * ((vector_derivative g (at x within {0..1})))) \ base_vec) has_integral (((f(g 1)) - (f(g 0)))) \ base_vec) {0..1}" using has_integral_componentwise_iff assms(4) by fastforce then show ?thesis using inner_mult_left' by (simp add: inner_mult_left' inner_diff_left) qed lemma reparam_eq_line_integrals: assumes reparam: "reparam \1 \2" and pw_smooth: "\2 piecewise_C1_differentiable_on {0..1}" and (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities of f*) cont: "continuous_on (path_image \2) (\x. F x \ b)" and line_integral_ex: "line_integral_exists F {b} \2" (*We need to remove this and work on special cases like conservative fields and field/line combinations that satisfy the improper integrals theorem assumptions*) shows "line_integral F {b} \1 = line_integral F {b} \2" "line_integral_exists F {b} \1" proof- obtain \ where phi: "(\x\{0..1}. \1 x = (\2 o \) x)"" \ piecewise_C1_differentiable_on {0..1}"" \(0) = 0"" \(1) = 1""bij_betw \ {0..1} {0..1}" "\ -` {0..1} \ {0..1}" "\x\{0..1}. finite (\ -` {x})" using reparam by (auto simp add: reparam_def) obtain s where s: "finite s" "\ C1_differentiable_on {0..1} - s" using phi by(auto simp add: reparam_def piecewise_C1_differentiable_on_def) let ?s = "s \ {0..1}" have s_inter: "finite ?s" "\ C1_differentiable_on {0..1} - ?s" using s apply blast by (metis Diff_Compl Diff_Diff_Int Diff_eq inf.commute s(2)) have cont_phi: "continuous_on {0..1} \" using phi by(auto simp add: reparam_def piecewise_C1_differentiable_on_imp_continuous_on) obtain s' D where s'_D: "finite s'" "(\x \ {0 .. 1} - s'. \2 differentiable at x)" "(\x\{0..1} - s'. (\2 has_vector_derivative D x) (at x)) \ continuous_on ({0..1} - s') D" using pw_smooth apply (auto simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) by (simp add: vector_derivative_works) let ?s' = "s' \ {0..1}" have gamma2_differentiable: "finite ?s'" "(\x \ {0 .. 1} - ?s'. \2 differentiable at x)" "(\x\{0..1} - ?s'. (\2 has_vector_derivative D x) (at x)) \ continuous_on ({0..1} - ?s') D" using s'_D apply blast using s'_D(2) apply auto[1] by (metis Diff_Int2 inf_top.left_neutral s'_D(3)) then have gamma2_b_component_differentiable: "(\x \ {0 .. 1} - ?s'. (\x. (\2 x) \ b) differentiable at x)" using differentiable_inner by force then have "(\x. (\2 x) \ b) differentiable_on {0..1} - ?s'" using differentiable_at_withinI by (auto simp add: differentiable_on_def) then have gama2_cont_comp: "continuous_on ({0..1} - ?s') (\x. (\2 x) \ b)" using differentiable_imp_continuous_on by auto (**********Additional needed assumptions ************) have s_in01: "?s \ {0..1}" by auto have s'_in01: "?s' \ {0..1}" by auto have phi_backimg_s': "\ -` {0..1} \ {0..1}" using phi by auto have "inj_on \ {0..1}" using phi(5) by (auto simp add: bij_betw_def) have bij_phi: "bij_betw \ {0..1} {0..1}" using phi(5) by auto have finite_bck_img_single: "\x\{0..1}. finite (\ -` {x})" using phi by auto then have finite_bck_img_single_s': " \x\?s'. finite (\ -` {x})" by auto have gamma2_line_integrable: "(\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b)) integrable_on {0..1}" using line_integral_ex by (simp add: line_integral_exists_def) (****************************************************************) have finite_neg_img: "finite (\ -` ?s')" using finite_bck_img_single by (metis Int_iff finite_Int gamma2_differentiable(1) image_vimage_eq inf_img_fin_dom') have gamma2_cont:"continuous_on ({0..1} - ?s') \2" using gamma2_differentiable by (meson continuous_at_imp_continuous_on differentiable_imp_continuous_within) have iii: "continuous_on ({0..1} - ?s') (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" proof- have 0: "continuous_on ({0..1} - ?s') (\x. F (\2 x) \ b)" using cont continuous_on_compose[OF gamma2_cont] continuous_on_compose2 gamma2_cont unfolding path_image_def by fastforce have D: "(\x\{0..1} - ?s'. (\2 has_vector_derivative D x) (at x)) \ continuous_on ({0..1} - ?s') D" using gamma2_differentiable by auto then have *:"\x\{0..1} - ?s'. vector_derivative \2 (at x within{0..1}) = D x" using vector_derivative_at vector_derivative_at_within_ivl by fastforce then have "continuous_on ({0..1} - ?s') (\x. vector_derivative \2 (at x within{0..1}))" using continuous_on_eq D by metis then have 1: "continuous_on ({0..1} - ?s') (\x. (vector_derivative \2 (at x within{0..1})) \ b)" by (auto intro!: continuous_intros) show ?thesis using continuous_on_mult[OF 0 1] by auto qed have iv: "\(0) \ \(1)" using phi(3) phi(4) by (simp add: reparam_def) have v: "\`{0..1} \ {0..1}" using phi by (auto simp add: reparam_def bij_betw_def) obtain D where D_props: "(\x\{0..1} - ?s. (\ has_vector_derivative D x) (at x))" using s by (auto simp add: C1_differentiable_on_def) then have "(\x. x \ ({0..1} - ?s) \ (\ has_vector_derivative D x) (at x within {0..1}))" using has_vector_derivative_at_within by blast then have vi: "(\x. x \ ({0..1} - ?s) \ (\ has_real_derivative D x) (at x within {0..1}))" using has_field_derivative_iff_has_vector_derivative by blast have a:"((\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) ({0..1})" proof- have a: "integral {\ 1..\ 0} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b)) = 0" using integral_singleton integral_empty iv by (simp add: phi(3) phi(4)) have b: "((\x. D x *\<^sub>R (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b)) - integral {\ 1..\ 0} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" apply(rule has_integral_substitution_general_'[OF s_inter(1) zero_le_one gamma2_differentiable(1) v gamma2_line_integrable iii cont_phi finite_bck_img_single_s']) proof- have "surj_on {0..1} \" using bij_phi by (metis (full_types) bij_betw_def image_subsetI rangeI) then show "surj_on ?s' \" using bij_phi s'_in01 by blast show "inj_on \ ({0..1} \ (?s \ \ -` ?s'))" proof- have i: "inj_on \ {0..1}" using bij_phi using bij_betw_def by blast have ii: "({0..1} \ (?s \ \ -` ?s')) = {0..1}" using phi_backimg_s' s_in01 s'_in01 by blast show ?thesis using i ii by auto qed show "\x. x \ {0..1} - ?s \ (\ has_real_derivative D x) (at x within {0..1})" using vi by blast qed show ?thesis using a b by auto qed then have b: "integral {0..1} (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" by auto have gamma2_vec_diffable: "\x::real. x \ {0..1} - ((\ -` ?s') \ ?s) \ ((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- fix x::real assume ass: "x \ {0..1} - ((\ -` ?s') \ ?s)" have zer_le_x_le_1:"0\ x \ x \ 1" using ass by simp show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- have **: "\2 differentiable at (\ x)" using gamma2_differentiable(2) ass v by blast have ***: " \ differentiable at x" using s ass by (auto simp add: C1_differentiable_on_eq) then show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" using differentiable_chain_at[OF *** **] by (auto simp add: vector_derivative_works) qed qed then have gamma2_vec_deriv_within: "\x::real. x \ {0..1} - ((\ -` ?s') \ ?s) \ vector_derivative (\2 \ \) (at x) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF gamma2_vec_diffable] by auto have "\x\{0..1} - ((\ -` ?s') \ ?s). D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" proof fix x::real assume ass: "x \ {0..1} -((\ -` ?s') \ ?s)" then have 0: "\ differentiable (at x)" using s by (auto simp add: C1_differentiable_on_def differentiable_def has_vector_derivative_def) obtain D2 where "(\ has_vector_derivative D2) (at x)" using D_props ass by blast have "\ x \ {0..1} - ?s'" using phi(5) ass by (metis Diff_Un Diff_iff Int_iff bij_betw_def image_eqI vimageI) then have 1: "\2 differentiable (at (\ x))" using gamma2_differentiable by auto have 3:" vector_derivative \2 (at (\ x)) = vector_derivative \2 (at (\ x) within {0..1})" proof- have *:"0\ \ x \ \ x \ 1" using phi(5) ass using \\ x \ {0..1} - s' \ {0..1}\ by auto then have **:"(\2 has_vector_derivative (vector_derivative \2 (at (\ x)))) (at (\ x))" using 1 vector_derivative_works by auto show ?thesis using * vector_derivative_at_within_ivl[OF **] by auto qed show "D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = vector_derivative (\2 \ \) (at x within {0..1}) \ b" using vector_derivative_chain_at[OF 0 1] apply (auto simp add: gamma2_vec_deriv_within[OF ass, symmetric] 3[symmetric]) using D_props ass vector_derivative_at by fastforce qed then have c:"\x. x\({0..1} -((\ -` ?s') \ ?s)) \ D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" by auto then have d: "integral ({0..1}) (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b))" proof- have "negligible ((\ -` ?s') \ ?s)" using finite_neg_img s(1) by auto then show ?thesis using c integral_spike by (metis(no_types,lifting)) qed have phi_in_int: "(\x. x \ {0..1} \ \ x \ {0..1})" using phi using v by blast then have e: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))){0..1}" proof- have "negligible ?s" using s_inter(1) by auto have 0: "negligible ((\ -` ?s') \ ?s)" using finite_neg_img s(1) by auto have c':"\ x\ {0..1} - ((\ -` ?s') \ ?s). D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" using c by blast have has_integral_spike_eq': "\s t f g y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using a has_integral_spike_eq'[OF 0 c'] by blast qed then have f: "((\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" proof- assume ass: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" have *:"\x\{0..1} - ( ((\ -` ?s') \ ?s) \ {0,1}). (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) x = (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) x" proof- have "\x\{0<..<1} - (\ -` ?s' \ ?s). (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" proof have i: "open ({0<..<1} - ((\ -` ?s') \ ?s))" using open_diff s(1) open_greaterThanLessThan finite_neg_img by (simp add: open_diff) have ii: "\x\{0<..<1::real} - (\ -` ?s' \ ?s). (\2 \ \) x = \1 x" using phi(1) by auto fix x::real assume ass: " x \ {0<..<1::real} - ((\ -` ?s') \ ?s)" then have iii: "(\2 \ \ has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" by (metis (no_types) Diff_iff add.commute add_strict_mono ass atLeastAtMost_iff gamma2_vec_deriv_within gamma2_vec_diffable greaterThanLessThan_iff less_irrefl not_le) (*Most crucial but annoying step*) then have iv:"(\1 has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" using has_derivative_transform_within_open i ii ass apply(auto simp add: has_vector_derivative_def) apply (meson ass has_derivative_transform_within_open i ii) apply (meson ass has_derivative_transform_within_open i ii) by (meson ass has_derivative_transform_within_open i ii) have v: "0 \ x" "x \ 1" using ass by auto have 0: "vector_derivative \1 (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iv v(1) v(2) zero_less_one] by force have 1: "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iii v(1) v(2) zero_less_one] by force then have "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative \1 (at x within {0..1})" using 0 1 by auto then show "vector_derivative (\2 \ \) (at x within {0..1}) \ b = vector_derivative \1 (at x within {0..1}) \ b" by auto qed then have i: "\x\{0..1} - ( ((\ -` ?s') \ ?s)\{0,1}). (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" by auto have ii: "\x\{0..1} - (((\ -` ?s') \ ?s)\{0,1}). F (\1 x) \ b = F (\2 (\ x)) \ b" using phi(1) by auto show ?thesis using i ii by metis qed have **: "negligible ((\ -` ?s') \ ?s \ {0, 1})" using s(1) finite_neg_img by auto have has_integral_spike_eq': "\s t g f y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using has_integral_spike_eq'[OF ** *] ass by blast qed then show "line_integral_exists F {b} \1" using phi by(auto simp add: line_integral_exists_def) have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b))" using integral_unique[OF e] integral_unique[OF f] by metis moreover have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" using b d phi by (auto simp add:) ultimately show "line_integral F {b} \1 = line_integral F {b} \2" using phi by(auto simp add: line_integral_def) qed lemma reparam_weak_eq_line_integrals: assumes "reparam_weak \1 \2" "\2 C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities of f*) "continuous_on (path_image \2) (\x. F x \ b)" shows "line_integral F {b} \1 = line_integral F {b} \2" "line_integral_exists F {b} \1" proof- obtain \ where phi: "(\x\{0..1}. \1 x = (\2 o \) x)"" \ piecewise_C1_differentiable_on {0..1}"" \(0) = 0"" \(1) = 1"" \ ` {0..1} = {0..1}" using assms(1) by (auto simp add: reparam_weak_def) obtain s where s: "finite s" "\ C1_differentiable_on {0..1} - s" using phi by(auto simp add: reparam_weak_def piecewise_C1_differentiable_on_def) have cont_phi: "continuous_on {0..1} \" using phi by(auto simp add: reparam_weak_def piecewise_C1_differentiable_on_imp_continuous_on) have gamma2_differentiable: "(\x \ {0 .. 1}. \2 differentiable at x)" using assms(2) by (auto simp add: valid_path_def C1_differentiable_on_eq) then have gamma2_b_component_differentiable: "(\x \ {0 .. 1}. (\x. (\2 x) \ b) differentiable at x)" by auto then have "(\x. (\2 x) \ b) differentiable_on {0..1}" using differentiable_at_withinI by (auto simp add: differentiable_on_def) then have gama2_cont_comp: "continuous_on {0..1} (\x. (\2 x) \ b)" using differentiable_imp_continuous_on by auto have gamma2_cont:"continuous_on {0..1} \2" using assms(2) C1_differentiable_imp_continuous_on by (auto simp add: valid_path_def) have iii: "continuous_on {0..1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" proof- have 0: "continuous_on {0..1} (\x. F (\2 x) \ b)" using assms(3) continuous_on_compose[OF gamma2_cont] by (auto simp add: path_image_def) obtain D where D: "(\x\{0..1}. (\2 has_vector_derivative D x) (at x)) \ continuous_on {0..1} D" using assms(2) by (auto simp add: C1_differentiable_on_def) then have *:"\x\{0..1}. vector_derivative \2 (at x within{0..1}) = D x" using vector_derivative_at vector_derivative_at_within_ivl by fastforce then have "continuous_on {0..1} (\x. vector_derivative \2 (at x within{0..1}))" using continuous_on_eq D by force then have 1: "continuous_on {0..1} (\x. (vector_derivative \2 (at x within{0..1})) \ b)" by (auto intro!: continuous_intros) show ?thesis using continuous_on_mult[OF 0 1] by auto qed have iv: "\(0) \ \(1)" using phi(3) phi(4) by (simp add: reparam_weak_def) have v: "\`{0..1} \ {0..1}" using phi(5) by (simp add: reparam_weak_def) obtain D where D_props: "(\x\{0..1} - s. (\ has_vector_derivative D x) (at x))" using s by (auto simp add: C1_differentiable_on_def) then have "(\x. x \ ({0..1} -s) \ (\ has_vector_derivative D x) (at x within {0..1}))" using has_vector_derivative_at_within by blast then have vi: "(\x. x \ ({0..1} - s) \ (\ has_real_derivative D x) (at x within {0..1}))" using has_field_derivative_iff_has_vector_derivative by blast have a:"((\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" using has_integral_substitution_strong[OF s(1) zero_le_one iv v iii cont_phi vi] by simp then have b: "integral {0..1} (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" by auto have gamma2_vec_diffable: "\x::real. x \ {0..1} -s \ ((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- fix x::real assume ass: "x \ {0..1} -s" have zer_le_x_le_1:"0\ x \ x \ 1" using ass by auto show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- have **: "\2 differentiable at (\ x)" using phi gamma2_differentiable by (auto simp add: zer_le_x_le_1) have ***: " \ differentiable at x" using s ass by (auto simp add: C1_differentiable_on_eq) then show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" using differentiable_chain_at[OF *** **] by (auto simp add: vector_derivative_works) qed qed then have gamma2_vec_deriv_within: "\x::real. x \ {0..1} -s \ vector_derivative (\2 \ \) (at x) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF gamma2_vec_diffable] by auto have "\x\{0..1} - s. D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" proof fix x::real assume ass: "x \ {0..1} -s" then have 0: "\ differentiable (at x)" using s by (auto simp add: C1_differentiable_on_def differentiable_def has_vector_derivative_def) obtain D2 where "(\ has_vector_derivative D2) (at x)" using D_props ass by blast have "\ x \ {0..1}" using phi(5) ass by (auto simp add: reparam_weak_def) then have 1: "\2 differentiable (at (\ x))" using gamma2_differentiable by auto have 3:" vector_derivative \2 (at (\ x)) = vector_derivative \2 (at (\ x) within {0..1})" proof- have *:"0\ \ x \ \ x \ 1" using phi(5) ass by auto then have **:"(\2 has_vector_derivative (vector_derivative \2 (at (\ x)))) (at (\ x))" using 1 vector_derivative_works by auto show ?thesis using * vector_derivative_at_within_ivl[OF **] by auto qed show "D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = vector_derivative (\2 \ \) (at x within {0..1}) \ b" using vector_derivative_chain_at[OF 0 1] apply (auto simp add: gamma2_vec_deriv_within[OF ass, symmetric] 3[symmetric]) using D_props ass vector_derivative_at by fastforce qed then have c:"\x. x\({0..1} -s) \ D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" by auto then have d: "integral ({0..1}) (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b))" proof- have "negligible s" using s(1) by auto then show ?thesis using c integral_spike by (metis(no_types,lifting)) qed have phi_in_int: "(\x. x \ {0..1} \ \ x \ {0..1})" using phi by(auto simp add:) then have e: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))){0..1}" proof- have 0:"negligible s" using s(1) by auto have c':"\ x\ {0..1} -s. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" using c by auto have has_integral_spike_eq': "\s t f g y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using a has_integral_spike_eq'[OF 0 c'] by blast qed then have f: "((\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" proof- assume ass: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" have *:"\x\{0..1} - (s\{0,1}). (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) x = (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) x" proof- have "\x\{0<..<1} - s. (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" proof have i: "open ({0<..<1} - s)" using open_diff s open_greaterThanLessThan by blast have ii: "\x\{0<..<1::real} - s. (\2 \ \) x = \1 x" using phi(1) by auto fix x::real assume ass: " x \ {0<..<1::real} - s" then have iii: "(\2 \ \ has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" using has_vector_derivative_at_within gamma2_vec_deriv_within gamma2_vec_diffable by auto (*Most crucial but annoying step*) then have iv:"(\1 has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" using has_derivative_transform_within_open i ii ass apply(auto simp add: has_vector_derivative_def) by force have v: "0 \ x" "x \ 1" using ass by auto have 0: "vector_derivative \1 (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iv v(1) v(2) zero_less_one] by force have 1: "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iii v(1) v(2) zero_less_one] by force then have "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative \1 (at x within {0..1})" using 0 1 by auto then show "vector_derivative (\2 \ \) (at x within {0..1}) \ b = vector_derivative \1 (at x within {0..1}) \ b" by auto qed then have i: "\x\{0..1} - (s\{0,1}). (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" by auto have ii: "\x\{0..1} - (s\{0,1}). F (\1 x) \ b = F (\2 (\ x)) \ b" using phi(1) by auto show ?thesis using i ii by auto qed have **: "negligible (s\{0,1})" using s(1) by auto have has_integral_spike_eq': "\s t g f y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using has_integral_spike_eq'[OF ** *] ass by blast qed then show "line_integral_exists F {b} \1" using phi by(auto simp add: line_integral_exists_def) have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b))" using integral_unique[OF e] integral_unique[OF f] by metis moreover have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" using b d phi by (auto simp add:) ultimately show "line_integral F {b} \1 = line_integral F {b} \2" using phi by(auto simp add: line_integral_def) qed lemma line_integral_sum_basis: assumes "finite (basis::('a::euclidean_space) set)" "\b\basis. line_integral_exists F {b} \" shows "line_integral F basis \ = (\b\basis. line_integral F {b} \)" "line_integral_exists F basis \" using assms proof(induction basis) show "line_integral F {} \ = (\b\{}. line_integral F {b} \)" by(auto simp add: line_integral_def) show "\b\{}. line_integral_exists F {b} \ \ line_integral_exists F {} \" by(simp add: line_integral_exists_def integrable_0) next fix basis::"('a::euclidean_space) set" fix x::"'a::euclidean_space" fix \ assume ind_hyp: "(\b\basis. line_integral_exists F {b} \ \ line_integral_exists F basis \)" "(\b\basis. line_integral_exists F {b} \ \ line_integral F basis \ = (\b\basis. line_integral F {b} \))" assume step: "finite basis" "x \ basis" "\b\insert x basis. line_integral_exists F {b} \" then have 0: "line_integral_exists F {x} \" by auto have 1:"line_integral_exists F basis \" using ind_hyp step by auto then show "line_integral_exists F (insert x basis) \" using step(1) step(2) line_integral_sum_gen(2)[OF _ 0 1] by simp have 3: "finite (insert x basis)" using step(1) by auto have "line_integral F basis \ = (\b\basis. line_integral F {b} \)" using ind_hyp step by auto then show "line_integral F (insert x basis) \ = (\b\insert x basis. line_integral F {b} \)" using step(1) step(2) line_integral_sum_gen(1)[OF 3 0 1] by force qed lemma reparam_weak_eq_line_integrals_basis: assumes "reparam_weak \1 \2" "\2 C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities*) "\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" shows "line_integral F basis \1 = line_integral F basis \2" "line_integral_exists F basis \1" proof- show "line_integral_exists F basis \1" using reparam_weak_eq_line_integrals(2)[OF assms(1) assms(2)] assms(3-4) line_integral_sum_basis(2)[OF assms(4)] by(simp add: subset_iff) show "line_integral F basis \1 = line_integral F basis \2" using reparam_weak_eq_line_integrals[OF assms(1) assms(2)] assms(3-4) line_integral_sum_basis(1)[OF assms(4)] line_integral_exists_smooth_one_base[OF assms(2)] by(simp add: subset_iff) qed lemma reparam_eq_line_integrals_basis: assumes "reparam \1 \2" "\2 piecewise_C1_differentiable_on {0..1}" "\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" "\b\basis. line_integral_exists F {b} \2" (*We need to remove this and work on special cases like conservative fields and field/line combinations that satisfy the improper integrals theorem assumptions*) shows "line_integral F basis \1 = line_integral F basis \2" "line_integral_exists F basis \1" proof- show "line_integral_exists F basis \1" using reparam_eq_line_integrals(2)[OF assms(1) assms(2)] assms(3-5) line_integral_sum_basis(2)[OF assms(4)] by(simp add: subset_iff) show "line_integral F basis \1 = line_integral F basis \2" using reparam_eq_line_integrals[OF assms(1) assms(2)] assms(3-5) line_integral_sum_basis(1)[OF assms(4)] by(simp add: subset_iff) qed lemma line_integral_exists_smooth: assumes "\ C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities*) "\(b::'a::euclidean_space) \basis. continuous_on (path_image \) (\x. F x \ b)" "finite basis" shows "line_integral_exists F basis \" using reparam_weak_eq_line_integrals_basis(2)[OF reparam_weak_eq_refl[where ?\1.0 = \]] assms by fastforce lemma smooth_path_imp_reverse: assumes "g C1_differentiable_on {0..1}" shows "(reversepath g) C1_differentiable_on {0..1}" using assms continuous_on_const apply (auto simp: reversepath_def) apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) apply (auto simp: C1_differentiable_on_eq) apply (simp add: finite_vimageI inj_on_def) done lemma piecewise_smooth_path_imp_reverse: assumes "g piecewise_C1_differentiable_on {0..1}" shows "(reversepath g) piecewise_C1_differentiable_on {0..1}" using assms valid_path_reversepath using valid_path_def by blast definition chain_reparam_weak_chain where "chain_reparam_weak_chain one_chain1 one_chain2 \ \f. bij f \ f ` one_chain1 = one_chain2 \ (\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam_weak \ (snd (f(k,\))) else reparam_weak \ (reversepath (snd (f(k,\)))))" lemma chain_reparam_weak_chain_line_integral: assumes "chain_reparam_weak_chain one_chain1 one_chain2" "\(k2,\2)\one_chain2. \2 C1_differentiable_on {0..1}" "\(k2,\2)\one_chain2.\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" and bound1: "boundary_chain one_chain1" and bound2: "boundary_chain one_chain2" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- obtain f where f: "bij f" "(\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam_weak \ (snd (f(k,\))) else reparam_weak \ (reversepath (snd (f(k,\)))))" "f ` one_chain1 = one_chain2" using assms(1) by (auto simp add: chain_reparam_weak_chain_def) have 0:" \x\one_chain1. (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" proof {fix k1 \1 assume ass1: "(k1,\1) \one_chain1" have "real_of_int k1 * line_integral F basis \1 = (case (f (k1,\1)) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" proof(cases) assume ass2: "k1 = 1" let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2\ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam_weak \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam_weak \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) C1_differentiable_on {0..1}" using f(3) assms(2) ass1 smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) show ?thesis using line_integral_on_reverse_path(1)[OF i line_integral_exists_smooth[OF 1 2 ]] assms by auto qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 3 by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') next assume "k1 \ 1" then have ass2: "k1 = -1" using bound1 ass1 f(3) unfolding boundary_chain_def by force let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam_weak \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) C1_differentiable_on {0..1}" using f(3) assms(2) ass1 smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) show ?thesis using line_integral_on_reverse_path(1)[OF i line_integral_exists_smooth[OF 1 2 assms(4)]] by auto qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 3 by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam_weak \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') qed } then show "\x. x \ one_chain1 \ (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" by (auto simp add: case_prod_beta') qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using 0 by(simp add: one_chain_line_integral_def sum_bij[OF f(1) _ f(3)] split_beta) show "\(k, \)\one_chain1. line_integral_exists F basis \" using 0 by blast qed definition chain_reparam_chain where "chain_reparam_chain one_chain1 one_chain2 \ \f. bij f \ f ` one_chain1 = one_chain2 \ (\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam \ (snd (f(k,\))) else reparam \ (reversepath (snd (f(k,\)))))" definition chain_reparam_weak_path::"((real) \ (real * real)) \ ((int * ((real) \ (real * real))) set) \ bool" where "chain_reparam_weak_path \ one_chain \ \l. set l = one_chain \ distinct l \ reparam \ (rec_join l) \ valid_chain_list l \ l \ []" lemma chain_reparam_chain_line_integral: assumes "chain_reparam_chain one_chain1 one_chain2" "\(k2,\2)\one_chain2. \2 piecewise_C1_differentiable_on {0..1}" "\(k2,\2)\one_chain2.\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" and bound1: "boundary_chain one_chain1" and bound2: "boundary_chain one_chain2" and line: "\(k2,\2)\one_chain2. (\b\basis. line_integral_exists F {b} \2)" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- obtain f where f: "bij f" "(\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam \ (snd (f(k,\))) else reparam \ (reversepath (snd (f(k,\)))))" "f ` one_chain1 = one_chain2" using assms(1) by (auto simp add: chain_reparam_chain_def) have integ_exist_b: "\(k1,\1)\one_chain1. \b\basis. line_integral_exists F {b} (snd (f (k1, \1)))" using line f by fastforce have valid_cubes: "\(k1,\1)\one_chain1. valid_path (snd (f (k1, \1)))" using assms(2) f(3) valid_path_def by fastforce have integ_rev_exist_b: "\(k1,\1)\one_chain1. \b\basis. line_integral_exists F {b} (reversepath (snd (f (k1, \1))))" using line_integral_on_reverse_path(2) integ_exist_b valid_cubes by blast have 0:" \x\one_chain1. (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" proof {fix k1 \1 assume ass1: "(k1,\1) \one_chain1" have "real_of_int k1 * line_integral F basis \1 = (case (f (k1,\1)) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" proof(cases) assume ass2: "k1 = 1" let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2\ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] integ_exist_b ass1 ass2 ass3 by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 piecewise_smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) have ii: "line_integral_exists F basis (snd (f (k1, \1)))" using assms(4) line_integral_sum_basis(2) integ_exist_b ass1 by fastforce show ?thesis using i ii line_integral_on_reverse_path(1) valid_path_reversepath by blast qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] integ_rev_exist_b ass1 ass2 ass3 3 by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') next assume "k1 \ 1" then have ass2: "k1 = -1" using bound1 ass1 f(3) unfolding boundary_chain_def by force let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 piecewise_smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) show ?thesis using line_integral_on_reverse_path(1)[OF i] integ_rev_exist_b using ass1 assms(4) line_integral_sum_basis(2) by fastforce qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 3 using ass1 integ_rev_exist_b by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 using ass1 integ_exist_b by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') qed } then show "\x. x \ one_chain1 \ (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" by (auto simp add: case_prod_beta') qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using 0 by (simp add: one_chain_line_integral_def sum_bij[OF f(1) _ f(3)] prod.case_eq_if) show "\(k, \)\one_chain1. line_integral_exists F basis \" using 0 by blast qed lemma path_image_rec_join: fixes \::"real \ (real \ real)" fixes k::int fixes l shows "\k \. (k, \) \ set l \ valid_chain_list l \ path_image \ \ path_image (rec_join l)" proof(induction l) case Nil then show ?case by auto next case ass: (Cons a l) obtain k' \' where a: "a = (k',\')" by force have "path_image \ \ path_image (rec_join ((k',\') # l))" proof(cases) assume "l=[]" then show "path_image \ \ path_image (rec_join ((k',\') # l))" using ass(2-3) a by(auto simp add:) next assume "l\[]" then obtain b l' where b_l: "l = b # l'" by (meson rec_join.elims) obtain k'' \'' where b: "b = (k'',\'')" by force show ?thesis using ass path_image_reversepath b_l path_image_join by(fastforce simp add: a) qed then show ?case using a by auto qed lemma path_image_rec_join_2: fixes l shows "l \ [] \ valid_chain_list l \ path_image (rec_join l) \ (\(k, \) \ set l. path_image \)" proof(induction l) case Nil then show ?case by auto next case ass: (Cons a l) obtain k' \' where a: "a = (k',\')" by force have "path_image (rec_join (a # l)) \ (\(k, y)\set (a # l). path_image y)" proof(cases) assume "l=[]" then show "path_image (rec_join (a # l)) \ (\(k, y)\set (a # l). path_image y)" using step a by(auto simp add:) next assume "l\[]" then obtain b l' where b_l: "l = b # l'" by (meson rec_join.elims) obtain k'' \'' where b: "b = (k'',\'')" by force show ?thesis using ass path_image_reversepath b_l path_image_join apply(auto simp add: a) (*Excuse the ugliness of the next script*) apply blast by fastforce qed then show ?case using a by auto qed lemma continuous_on_closed_UN: assumes "finite S" shows "((\s. s \ S \ closed s) \ (\s. s \ S \ continuous_on s f) \ continuous_on (\S) f)" using assms proof(induction S) case empty then show ?case by auto next case (insert x F) then show ?case using continuous_on_closed_Un closed_Union by (simp add: closed_Union continuous_on_closed_Un) qed lemma chain_reparam_weak_path_line_integral: assumes path_eq_chain: "chain_reparam_weak_path \ one_chain" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "\b\basis. \(k::int, \) \ one_chain. line_integral_exists F {b} \" and valid_path: "\(k::int, \) \ one_chain. valid_path \" and finite_basis: "finite basis" and cont: "\b\basis. \(k,\2)\one_chain. continuous_on (path_image \2) (\x. F x \ b)" and finite_one_chain: "finite one_chain" shows "line_integral F basis \ = one_chain_line_integral F basis one_chain" "line_integral_exists F basis \" (*"valid_path \" This cannot be proven, see the crappy assumption of derivs/eq_pw_smooth :( *) proof - obtain l where l_props: "set l = one_chain" "distinct l" "reparam \ (rec_join l)" "valid_chain_list l" "l \ []" using chain_reparam_weak_path_def assms by auto have bchain_l: "boundary_chain (set l)" using l_props boundary_chain by (simp add: boundary_chain_def) have cont_forall: "\b\basis. continuous_on (\(k, \)\one_chain. path_image \) (\x. F x \ b)" proof fix b assume ass: "b \ basis" have "continuous_on (\((path_image \ snd) ` one_chain)) (\x. F x \ b)" apply(rule continuous_on_closed_UN[where ?S = "(path_image o snd) ` one_chain" ]) proof show "finite one_chain" using finite_one_chain by auto show "\s. s \ (path_image \ snd) ` one_chain \ closed s" using closed_valid_path_image[OF] valid_path by fastforce show "\s. s \ (path_image \ snd) ` one_chain \ continuous_on s (\x. F x \ b)" using cont ass by force qed then show "continuous_on (\(k, \)\one_chain. path_image \) (\x. F x \ b)" by (auto simp add: Union_eq case_prod_beta) qed show "line_integral_exists F basis \" proof (rule reparam_eq_line_integrals_basis[OF l_props(3) _ _ finite_basis]) let ?\2.0="rec_join l" show "?\2.0 piecewise_C1_differentiable_on {0..1}" apply(simp only: valid_path_def[symmetric]) apply(rule joined_is_valid) using assms l_props by auto show "\b\basis. continuous_on (path_image (rec_join l)) (\x. F x \ b)" using path_image_rec_join_2[OF l_props(5) l_props(4)] l_props(1) using cont_forall continuous_on_subset by blast show "\b\basis. line_integral_exists F {b} (rec_join l)" proof fix b assume ass: "b\basis" show "line_integral_exists F {b} (rec_join l)" proof (rule line_integral_exists_on_rec_join) show "boundary_chain (set l)" using l_props boundary_chain by auto show "valid_chain_list l" using l_props by auto show "\k \. (k, \) \ set l \ valid_path \" using l_props assms by auto show "\(k, \)\set l. line_integral_exists F {b} \" using l_props line_integral_exists ass by blast qed qed qed show "line_integral F basis \ = one_chain_line_integral F basis one_chain" proof- have i: "line_integral F basis (rec_join l) = one_chain_line_integral F basis one_chain" proof (rule one_chain_line_integral_rec_join) show "set l = one_chain" "distinct l" "valid_chain_list l" using l_props by auto show "boundary_chain one_chain" using boundary_chain by auto show "\(k, \)\one_chain. line_integral_exists F basis \" using line_integral_sum_basis(2)[OF finite_basis] line_integral_exists by blast show "\(k, \)\one_chain. valid_path \" using valid_path by auto show "finite basis" using finite_basis by auto qed have ii: "line_integral F basis \ = line_integral F basis (rec_join l)" proof (rule reparam_eq_line_integrals_basis[OF l_props(3) _ _ finite_basis]) let ?\2.0="rec_join l" show "?\2.0 piecewise_C1_differentiable_on {0..1}" apply(simp only: valid_path_def[symmetric]) apply(rule joined_is_valid) using assms l_props by auto show "\b\basis. continuous_on (path_image (rec_join l)) (\x. F x \ b)" using path_image_rec_join_2[OF l_props(5) l_props(4)] l_props(1) using cont_forall continuous_on_subset by blast show "\b\basis. line_integral_exists F {b} (rec_join l)" proof fix b assume ass: "b\basis" show "line_integral_exists F {b} (rec_join l)" proof (rule line_integral_exists_on_rec_join) show "boundary_chain (set l)" using l_props boundary_chain by auto show "valid_chain_list l" using l_props by auto show "\k \. (k, \) \ set l \ valid_path \" using l_props assms by auto show "\(k, \)\set l. line_integral_exists F {b} \" using l_props line_integral_exists ass by blast qed qed qed show "line_integral F basis \ = one_chain_line_integral F basis one_chain" using i ii by auto qed qed definition chain_reparam_chain' where "chain_reparam_chain' one_chain1 subdiv \ \f. ((\(f ` one_chain1)) = subdiv) \ (\cube \ one_chain1. chain_reparam_weak_path (rec_join [cube]) (f cube)) \ (\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {}) \ (\x \ one_chain1. finite (f x))" lemma chain_reparam_chain'_imp_finite_subdiv: assumes "finite one_chain1" "chain_reparam_chain' one_chain1 subdiv" shows "finite subdiv" using assms by(auto simp add: chain_reparam_chain'_def) lemma chain_reparam_chain'_line_integral: assumes chain1_eq_chain2: "chain_reparam_chain' one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and line_integral_exists_on_chain2: "\b\basis. \(k::int, \) \ subdiv. line_integral_exists F {b} \" and valid_path: "\(k, \) \ subdiv. valid_path \" and valid_path_2: "\(k, \) \ one_chain1. valid_path \" and finite_chain1: "finite one_chain1" and finite_basis: "finite basis" and cont_field: " \b\basis. \(k, \2)\subdiv. continuous_on (path_image \2) (\x. F x \ b)" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" "\(k, \) \ one_chain1. line_integral_exists F basis \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\cube \ one_chain1. chain_reparam_weak_path (rec_join [cube]) (f cube))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" "(\x \ one_chain1. finite (f x))" using chain1_eq_chain2 by (auto simp add: chain_reparam_chain'_def) then have 0: "one_chain_line_integral F basis subdiv = one_chain_line_integral F basis (\(f ` one_chain1))" by auto {fix k \ assume ass:"(k, \) \ one_chain1" have "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" then have i:"chain_reparam_weak_path \ (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms unfolding boundary_chain_def by blast have iii:"\b\basis. \(k, \)\f (k, \). line_integral_exists F {b} \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have v: "\b\basis. \(k, \2)\f (k, \). continuous_on (path_image \2) (\x. F x \ b)" using f_props(1) ass assms by blast have "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using chain_reparam_weak_path_line_integral[OF i ii iii iv finite_basis v] ass f_props(4) by (auto) then show "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) then have i:"chain_reparam_weak_path (reversepath \) (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms unfolding boundary_chain_def by blast have iii:"\b\basis. \(k, \)\f (k, \). line_integral_exists F {b} \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have v: "\b\basis. \(k, \2)\f (k, \). continuous_on (path_image \2) (\x. F x \ b)" using f_props(1) ass assms by blast have x:"line_integral_exists F basis (reversepath \) \ one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using chain_reparam_weak_path_line_integral[OF i ii iii iv finite_basis v] ass f_props(4) by auto have "valid_path (reversepath \)" using f_props(1) ass assms by auto then have "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path reversepath_reversepath x ass by metis then show "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k::int, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed} note cube_line_integ = this have finite_chain2: "finite subdiv" using finite_chain1 f_props(1) f_props(4) by auto show line_integral_exists_on_chain1: "\(k, \) \ one_chain1. line_integral_exists F basis \" using cube_line_integ by auto have 1: "one_chain_line_integral F basis (\(f ` one_chain1)) = one_chain_line_integral F basis one_chain1" proof - have 0:"one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" proof - have finite: "\chain \ (f ` one_chain1). finite chain" using f_props(1) finite_chain2 by (meson Sup_upper finite_subset) have disj: " \A\f ` one_chain1. \B\f ` one_chain1. A \ B \ A \ B = {}" apply(simp add:image_def) using f_props(3) by metis show "one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" using Groups_Big.comm_monoid_add_class.sum.Union_disjoint[OF finite disj] one_chain_line_integral_def by auto qed have 1:"(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = one_chain_line_integral F basis one_chain1" proof - have "(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1. k*(line_integral F basis \))" proof - have i:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \))" proof - have "inj_on f (one_chain1 - {p. f p = {}})" unfolding inj_on_def using f_props(3) by force then have 0: "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \)))" using Groups_Big.comm_monoid_add_class.sum.reindex by auto have "\ k \. (k, \) \ (one_chain1 - {p. f p = {}}) \ one_chain_line_integral F basis (f(k, \)) = k* (line_integral F basis \)" using cube_line_integ by auto then have "(\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \))) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" by (auto intro!: Finite_Cartesian_Product.sum_cong_aux) then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" using 0 by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ (k, \) \ {(k',\'). k'* (line_integral F basis \') = 0}" proof- fix k::"int" fix \::"real\real*real" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto show "(k, \) \ {(k'::int, \'). k' * line_integral F basis \' = 0}" using zero_line_integral cube_line_integ ass by force qed then have ii:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" proof - have "\one_chain. one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}})) \ one_chain_line_integral F basis one_chain = 0" proof - fix one_chain assume "one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}}))" then have "one_chain = {}" by auto then show "one_chain_line_integral F basis one_chain = 0" by (auto simp add: one_chain_line_integral_def) qed then have 0:"(\one_chain \ f ` (one_chain1) - (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" using Groups_Big.comm_monoid_add_class.sum.neutral by auto then have "(\one_chain \ f ` (one_chain1). one_chain_line_integral F basis one_chain) - (\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" proof - have finte: "finite (f ` one_chain1)" using finite_chain1 by auto show ?thesis using Groups_Big.sum_diff[OF finte, of " (f ` (one_chain1 - {p. f p = {}}))" "one_chain_line_integral F basis"] 0 by auto qed then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ f (k, \) \ {chain. one_chain_line_integral F basis chain = 0}" proof- fix k::"int" fix \::"real\real*real" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have "one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto then show "f (k, \) \ {p'. (one_chain_line_integral F basis p' = 0)}" by auto qed then have iii:"(\(k::int,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \)) = (\(k::int,\)\one_chain1. k*(line_integral F basis \))" proof- have "\ k \. (k,\)\one_chain1 - (one_chain1 - {p. f p = {}}) \ k* (line_integral F basis \) = 0" proof- fix k \ assume ass: "(k,\)\one_chain1 - (one_chain1 - {p. f p = {}})" then have "f(k, \) = {}" by auto then have "one_chain_line_integral F basis (f(k, \)) = 0" by (auto simp add: one_chain_line_integral_def) then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto then show "k* (line_integral F basis \) = 0" using zero_line_integral cube_line_integ ass by force qed then have "\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k* (line_integral F basis \) = 0" by auto then have "(\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.comm_monoid_add_class.sum.neutral [of "one_chain1 - (one_chain1 - {p. f p = {}})" "(\(k::int,\). k* (line_integral F basis \))"] by (simp add: split_beta) then have "(\(k::int,\)\one_chain1. k*(line_integral F basis \)) - (\(k::int,\)\ (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.sum_diff[OF finite_chain1] by (metis (no_types) Diff_subset \(\(k, \)\one_chain1 - (one_chain1 - {p. f p = {}}). k * line_integral F basis \) = 0\ \\f B. B \ one_chain1 \ sum f (one_chain1 - B) = sum f one_chain1 - sum f B\) then show ?thesis by auto qed show ?thesis using i ii iii by auto qed then show ?thesis using one_chain_line_integral_def by auto qed show ?thesis using 0 1 by auto qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" using 0 1 by auto qed lemma chain_reparam_chain'_line_integral_smooth_cubes: assumes "chain_reparam_chain' one_chain1 one_chain2" "\(k2,\2)\one_chain2. \2 C1_differentiable_on {0..1}" "\b\basis.\(k2,\2)\one_chain2. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" "finite one_chain1" "boundary_chain one_chain1" "boundary_chain one_chain2" "\(k,\)\one_chain1. valid_path \" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- {fix b assume "b \ basis" fix k \ assume "(k, \)\one_chain2" have "line_integral_exists F {b} \" apply(rule line_integral_exists_smooth) using \(k, \) \ one_chain2\ assms(2) apply blast using assms using \(k, \) \ one_chain2\ \b \ basis\ apply blast using \b \ basis\ by blast} then have a: "\b\basis. \(k, \)\one_chain2. line_integral_exists F {b} \" by auto have b: "\(k2,\2)\one_chain2. valid_path \2" using assms(2) by (simp add: C1_differentiable_imp_piecewise case_prod_beta valid_path_def) show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" by (rule chain_reparam_chain'_line_integral[OF assms(1) assms(6) assms(7) a b assms(8) assms(5) assms(4) assms(3)]) show "\(k, \)\one_chain1. line_integral_exists F basis \" by (rule chain_reparam_chain'_line_integral[OF assms(1) assms(6) assms(7) a b assms(8) assms(5) assms(4) assms(3)]) qed lemma chain_subdiv_path_pathimg_subset: assumes "chain_subdiv_path \ subdiv" shows "\(k',\')\subdiv. (path_image \') \ path_image \" using assms chain_subdiv_path.simps path_image_rec_join by(auto simp add: chain_subdiv_path.simps) lemma reparam_path_image: assumes "reparam \1 \2" shows "path_image \1 = path_image \2" using assms apply (auto simp add: reparam_def path_image_def image_def bij_betw_def) apply (force dest!: equalityD2) done lemma chain_reparam_weak_path_pathimg_subset: assumes "chain_reparam_weak_path \ subdiv" shows "\(k',\')\subdiv. (path_image \') \ path_image \" using assms apply(auto simp add: chain_reparam_weak_path_def) using path_image_rec_join reparam_path_image by blast lemma chain_subdiv_chain_pathimg_subset': assumes "chain_subdiv_chain one_chain subdiv" assumes "(k,\)\ subdiv" shows " \k' \'. (k',\')\ one_chain \ path_image \ \ path_image \'" using assms unfolding chain_subdiv_chain_def pairwise_def apply auto by (metis chain_subdiv_path.cases coeff_cube_to_path.simps path_image_rec_join path_image_reversepath) lemma chain_subdiv_chain_pathimg_subset: assumes "chain_subdiv_chain one_chain subdiv" shows "\ (path_image ` {\. \k. (k,\)\ subdiv}) \ \ (path_image ` {\. \k. (k,\)\ one_chain})" using assms unfolding chain_subdiv_chain_def pairwise_def apply auto by (metis UN_iff assms chain_subdiv_chain_pathimg_subset' subsetCE case_prodI2) lemma chain_reparam_chain'_pathimg_subset': assumes "chain_reparam_chain' one_chain subdiv" assumes "(k,\)\ subdiv" shows " \k' \'. (k',\')\ one_chain \ path_image \ \ path_image \'" using assms chain_reparam_weak_path_pathimg_subset apply(auto simp add: chain_reparam_chain'_def set_eq_iff) using path_image_reversepath case_prodE case_prodD old.prod.exhaust apply (simp add: list.distinct(1) list.inject rec_join.elims) by (smt case_prodD coeff_cube_to_path.simps rec_join.simps(2) reversepath_simps(2) surj_pair) definition common_reparam_exists:: "(int \ (real \ real \ real)) set \ (int \ (real \ real \ real)) set \ bool" where "common_reparam_exists one_chain1 one_chain2 \ (\subdiv ps1 ps2. chain_reparam_chain' (one_chain1 - ps1) subdiv \ chain_reparam_chain' (one_chain2 - ps2) subdiv \ (\(k, \)\subdiv. \ C1_differentiable_on {0..1}) \ boundary_chain subdiv \ (\(k, \)\ps1. point_path \) \ (\(k, \)\ps2. point_path \))" lemma common_reparam_exists_imp_eq_line_integral: assumes finite_basis: "finite basis" and "finite one_chain1" "finite one_chain2" "boundary_chain (one_chain1::(int \ (real \ real \ real)) set)" "boundary_chain (one_chain2::(int \ (real \ real \ real)) set)" " \(k2, \2)\one_chain2. \b\basis. continuous_on (path_image \2) (\x. F x \ b)" "(common_reparam_exists one_chain1 one_chain2)" "(\(k, \)\one_chain1. valid_path \)" "(\(k, \)\one_chain2. valid_path \)" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- obtain subdiv ps1 ps2 where props: "chain_reparam_chain' (one_chain1 - ps1) subdiv" "chain_reparam_chain' (one_chain2 - ps2) subdiv " "(\(k, \)\subdiv. \ C1_differentiable_on {0..1})" "boundary_chain subdiv" "(\(k, \)\ps1. point_path \)" "(\(k, \)\ps2. point_path \)" using assms by (auto simp add: common_reparam_exists_def) have subdiv_valid: "(\(k, \)\subdiv. valid_path \)" apply(simp add: valid_path_def) using props(3) using C1_differentiable_imp_piecewise by blast have onechain_boundary1: "boundary_chain (one_chain1 - ps1)" using assms(4) by (auto simp add: boundary_chain_def) have onechain_boundary2: "boundary_chain (one_chain2 - ps1)" using assms(5) by (auto simp add: boundary_chain_def) {fix k2 \2 b assume ass: "(k2, \2)\subdiv "" b\basis" have "\ k \. (k, \) \ subdiv \ \k' \'. (k', \') \ one_chain2 \ path_image \ \ path_image \'" by (meson chain_reparam_chain'_pathimg_subset' props Diff_subset subsetCE) then have "continuous_on (path_image \2) (\x. F x \ b)" using assms(6) continuous_on_subset[where ?f = " (\x. F x \ b)"] ass apply(auto simp add: subset_iff) by (metis (mono_tags, lifting) case_prodD)} then have cont_field: "\b\basis. \(k2, \2)\subdiv. continuous_on (path_image \2) (\x. F x \ b)" by auto have one_chain1_ps_valid: "(\(k, \)\one_chain1 - ps1. valid_path \)" using assms by auto have one_chain2_ps_valid: "(\(k, \)\one_chain2 - ps1. valid_path \)" using assms by auto have 0: "one_chain_line_integral F basis (one_chain1 - ps1) = one_chain_line_integral F basis subdiv" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(1) props(3) cont_field finite_basis]) using props assms apply blast using props assms using onechain_boundary1 apply blast using props assms apply blast using one_chain1_ps_valid by blast have 1:"\(k, \)\(one_chain1 - ps1). line_integral_exists F basis \" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(1) props(3) cont_field finite_basis]) using props assms apply blast using props assms using onechain_boundary1 apply blast using props assms apply blast using one_chain1_ps_valid by blast have 2: "one_chain_line_integral F basis (one_chain2 - ps2) = one_chain_line_integral F basis subdiv" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(2) props(3) cont_field finite_basis]) using props assms apply blast apply (simp add: assms(5) boundary_chain_diff) apply (simp add: props(4)) by (simp add: assms(9)) have 3:"\(k, \)\(one_chain2 - ps2). line_integral_exists F basis \" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(2) props(3) cont_field finite_basis]) using props assms apply blast apply (simp add: assms(5) boundary_chain_diff) apply (simp add: props(4)) by (simp add: assms(9)) show line_int_ex_chain1: "\(k, \)\one_chain1. line_integral_exists F basis \" using 0 using "1" finite_basis line_integral_exists_point_path props(5) by fastforce then show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using 0 1 2 3 using assms(2-3) finite_basis one_chain_line_integral_point_paths props(5) props(6) by auto qed definition subcube :: "real \ real \ (int \ (real \ real \ real)) \ (int \ (real \ real \ real))" where "subcube a b cube = (fst cube, subpath a b (snd cube))" lemma subcube_valid_path: assumes "valid_path (snd cube)" "a \ {0..1}" "b \ {0..1}" shows "valid_path (snd (subcube a b cube))" using valid_path_subpath[OF assms] by (auto simp add: subcube_def) end