diff -r 56a408fa2440 -r 28c1f4f5335f src/HOL/Complex_Analysis/Contour_Integration.thy --- a/src/HOL/Complex_Analysis/Contour_Integration.thy Sat Aug 12 10:09:29 2023 +0100 +++ b/src/HOL/Complex_Analysis/Contour_Integration.thy Mon Aug 21 18:38:25 2023 +0100 @@ -44,14 +44,14 @@ unfolding contour_integrable_on_def contour_integral_def by blast lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" - apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def) + unfolding contour_integral_def has_contour_integral_def contour_integrable_on_def using has_integral_unique by blast lemma has_contour_integral_eqpath: - "\(f has_contour_integral y) p; f contour_integrable_on \; + "\(f has_contour_integral y) p; f contour_integrable_on \; contour_integral p f = contour_integral \ f\ \ (f has_contour_integral y) \" -using contour_integrable_on_def contour_integral_unique by auto + using contour_integrable_on_def contour_integral_unique by auto lemma has_contour_integral_integral: "f contour_integrable_on i \ (f has_contour_integral (contour_integral i f)) i" @@ -329,12 +329,12 @@ qed lemma contour_integrable_join [simp]: - "\valid_path g1; valid_path g2\ + "\valid_path g1; valid_path g2\ \ f contour_integrable_on (g1 +++ g2) \ f contour_integrable_on g1 \ f contour_integrable_on g2" -using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast + using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast lemma contour_integral_join [simp]: - "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ + "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ \ contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) @@ -353,11 +353,7 @@ using assms by (auto simp: has_contour_integral) then have i: "i = integral {a..1} (\x. f (g x) * vector_derivative g (at x)) + integral {0..a} (\x. f (g x) * vector_derivative g (at x))" - apply (rule has_integral_unique) - apply (subst add.commute) - apply (subst Henstock_Kurzweil_Integration.integral_combine) - using assms * integral_unique by auto - + by (smt (verit, ccfv_threshold) Henstock_Kurzweil_Integration.integral_combine a add.commute atLeastAtMost_iff has_integral_iff) have vd1: "vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" if "0 \ x" "x + a < 1" "x \ (\x. x - a) ` S" for x unfolding shiftpath_def @@ -371,8 +367,7 @@ then show "(g has_vector_derivative vector_derivative g (at (a + x))) (at (x + a))" by (metis add.commute vector_derivative_works) qed - then - show "((\x. g (a + x)) has_vector_derivative vector_derivative g (at (x + a))) (at x)" + then show "((\x. g (a + x)) has_vector_derivative vector_derivative g (at (x + a))) (at x)" by (auto simp: field_simps) show "0 < dist (1 - a) x" using that by auto @@ -474,8 +469,8 @@ lemma contour_integrable_on_shiftpath_eq: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" -using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto + shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" + using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto lemma contour_integral_shiftpath: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" @@ -556,26 +551,17 @@ lemma contour_integrable_subpath: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" shows "f contour_integrable_on (subpath u v g)" -proof (cases u v rule: linorder_class.le_cases) - case le - then show ?thesis - by (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) -next - case ge - with assms show ?thesis - by (metis (no_types, lifting) contour_integrable_on_def contour_integrable_reversepath_eq has_contour_integral_subpath reversepath_subpath valid_path_subpath) -qed + by (smt (verit, ccfv_threshold) assms contour_integrable_on_def contour_integrable_reversepath_eq + has_contour_integral_subpath reversepath_subpath valid_path_subpath) lemma has_integral_contour_integral_subpath: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "(((\x. f(g x) * vector_derivative g (at x))) + shows "((\x. f(g x) * vector_derivative g (at x)) has_integral contour_integral (subpath u v g) f) {u..v}" - using assms + (is "(?fg has_integral _)_") proof - - have "(\r. f (g r) * vector_derivative g (at r)) integrable_on {u..v}" - by (metis (full_types) assms(1) assms(3) assms(4) atLeastAtMost_iff atLeastatMost_subset_iff contour_integrable_on integrable_on_subinterval) - then have "((\r. f (g r) * vector_derivative g (at r)) has_integral integral {u..v} (\r. f (g r) * vector_derivative g (at r))) {u..v}" - by blast + have "(?fg has_integral integral {u..v} ?fg) {u..v}" + using assms contour_integrable_on integrable_on_subinterval by fastforce then show ?thesis by (metis (full_types) assms contour_integral_unique has_contour_integral_subpath) qed @@ -653,9 +639,9 @@ next fix x assume "x \ {0..1} - ({0, 1} \ g -` A \ {0<..<1})" hence "g x \ path_image g - A" by (auto simp: path_image_def) - from assms(4)[OF this] and assms(3) - show "f' (g' x) * vector_derivative g' (at x) = f (g x) * vector_derivative g (at x)" by simp - qed + with assms show "f' (g' x) * vector_derivative g' (at x) = f (g x) * vector_derivative g (at x)" + by simp +qed text \Contour integral along a segment on the real axis\ @@ -664,7 +650,7 @@ fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ - ((\x. f (of_real x)) has_integral I) {Re a..Re b}" + ((\x. f (of_real x)) has_integral I) {Re a..Re b}" proof - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" by (simp_all add: complex_eq_iff) @@ -738,9 +724,8 @@ 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 - using assms cfg * + } then show ?thesis + using assms cfg by (force simp: at_within_Icc_at intro: fundamental_theorem_of_calculus_interior_strong [OF \finite K\]) qed @@ -759,7 +744,7 @@ shows "(f' has_contour_integral 0) g" using assms by (metis diff_self contour_integral_primitive) -text\Existence of path integral for continuous function\ + lemma contour_integrable_continuous_linepath: assumes "continuous_on (closed_segment a b) f" shows "f contour_integrable_on (linepath a b)" @@ -951,10 +936,9 @@ by fastforce lemma contour_integrable_sum: - "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ + "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ \ (\x. sum (\a. f a x) s) contour_integrable_on p" - unfolding contour_integrable_on_def - by (metis has_contour_integral_sum) + unfolding contour_integrable_on_def by (metis has_contour_integral_sum) lemma contour_integrable_neg_iff: "(\x. -f x) contour_integrable_on g \ f contour_integrable_on g" @@ -1027,10 +1011,10 @@ apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) done - } note fj = this - show ?thesis + } + then show ?thesis using f k unfolding has_contour_integral_linepath - by (simp add: linepath_def has_integral_combine [OF _ _ fi fj]) + by (simp add: linepath_def has_integral_combine [OF _ _ fi]) qed lemma continuous_on_closed_segment_transform: @@ -1107,10 +1091,10 @@ then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\z. vector_derivative g (at (fst z)))" and hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" by auto - have "continuous_on (cbox (0, 0) (1, 1)) ((\(y1, y2). f y1 y2) \ (\w. ((g \ fst) w, (h \ snd) w)))" - apply (intro gcon hcon continuous_intros | simp)+ - apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) - done + have "continuous_on ((\x. (g (fst x), h (snd x))) ` cbox (0,0) (1,1)) (\(y1, y2). f y1 y2)" + by (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) + then have "continuous_on (cbox (0, 0) (1, 1)) ((\(y1, y2). f y1 y2) \ (\w. ((g \ fst) w, (h \ snd) w)))" + by (intro gcon hcon continuous_intros | simp)+ then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" by auto have "integral {0..1} (\x. contour_integral h (f (g x)) * vector_derivative g (at x)) = @@ -1186,7 +1170,7 @@ lemma valid_path_polynomial_function: fixes p :: "real \ 'a::euclidean_space" shows "polynomial_function p \ valid_path p" -by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) + by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) lemma valid_path_subpath_trivial [simp]: fixes g :: "real \ 'a::euclidean_space" @@ -1199,15 +1183,15 @@ where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" lemma pathstart_part_circlepath [simp]: - "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" -by (metis part_circlepath_def pathstart_def pathstart_linepath) + "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" + by (metis part_circlepath_def pathstart_def pathstart_linepath) lemma pathfinish_part_circlepath [simp]: - "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" -by (metis part_circlepath_def pathfinish_def pathfinish_linepath) + "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" + by (metis part_circlepath_def pathfinish_def pathfinish_linepath) lemma reversepath_part_circlepath[simp]: - "reversepath (part_circlepath z r s t) = part_circlepath z r t s" + "reversepath (part_circlepath z r s t) = part_circlepath z r t s" unfolding part_circlepath_def reversepath_def linepath_def by (auto simp:algebra_simps) @@ -1284,24 +1268,12 @@ lemma in_path_image_part_circlepath: assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" shows "norm(w - z) = r" -proof - - have "w \ {c. dist z c = r}" - by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) - thus ?thesis - by (simp add: dist_norm norm_minus_commute) -qed + by (smt (verit) assms dist_norm mem_Collect_eq norm_minus_commute path_image_part_circlepath_subset sphere_def subsetD) lemma path_image_part_circlepath_subset': assumes "r \ 0" shows "path_image (part_circlepath z r s t) \ sphere z r" -proof (cases "s \ t") - case True - thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp -next - case False - thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms - by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all -qed + by (smt (verit) assms path_image_part_circlepath_subset reversepath_part_circlepath reversepath_simps(2)) lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) @@ -1458,7 +1430,7 @@ have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" by (auto intro!: B [unfolded path_image_def image_def]) show ?thesis - apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) + using has_integral_bound [where 'a=real, simplified, OF _ **] using assms le * "2" \r > 0\ by (auto simp add: norm_mult vector_derivative_part_circlepath) qed qed @@ -1520,13 +1492,16 @@ qed have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" by force + have "\x n. \s \ t; \s - t\ \ 2 * pi; 0 \ x; x < 1; + x * (t - s) = 2 * (real_of_int n * pi)\ + \ x = 0" + by (rule ccontr) (auto simp: 2 field_split_simps abs_mult dest: of_int_leD) + then show ?thesis using False apply (simp add: simple_path_def loop_free_def) apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) apply (subst abs_away) apply (auto simp: 1) - apply (rule ccontr) - apply (auto simp: 2 field_split_simps abs_mult dest: of_int_leD) done qed @@ -1719,16 +1694,18 @@ and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" using eventually_happens [OF eventually_conj] by (fastforce simp: contour_integrable_on path_image_def) - have Ble: "B * e / (\B\ + 1) \ e" - using \0 \ B\ \0 < e\ by (simp add: field_split_simps) have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" proof (intro exI conjI ballI) show "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ e" if "x \ {0..1}" for x - apply (rule order_trans [OF _ Ble]) - using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ - apply (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le] simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) - done + proof - + have "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ B * e / (\B\ + 1)" + using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ + by (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le] simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) + also have "\ \ e" + using \0 \ B\ \0 < e\ by (simp add: field_split_simps) + finally show ?thesis . + qed qed (rule inta) } then show lintg: "l contour_integrable_on \"