--- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Mon Dec 11 17:28:32 2017 +0100
+++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Tue Dec 12 10:01:14 2017 +0100
@@ -379,6 +379,28 @@
"f holomorphic_on A \<Longrightarrow> (\<lambda>x. c *\<^sub>R f x) holomorphic_on A"
by (auto simp: scaleR_conv_of_real intro!: holomorphic_intros)
+lemma holomorphic_on_Un [holomorphic_intros]:
+ assumes "f holomorphic_on A" "f holomorphic_on B" "open A" "open B"
+ shows "f holomorphic_on (A \<union> B)"
+ using assms by (auto simp: holomorphic_on_def at_within_open[of _ A]
+ at_within_open[of _ B] at_within_open[of _ "A \<union> B"] open_Un)
+
+lemma holomorphic_on_If_Un [holomorphic_intros]:
+ assumes "f holomorphic_on A" "g holomorphic_on B" "open A" "open B"
+ assumes "\<And>z. z \<in> A \<Longrightarrow> z \<in> B \<Longrightarrow> f z = g z"
+ shows "(\<lambda>z. if z \<in> A then f z else g z) holomorphic_on (A \<union> B)" (is "?h holomorphic_on _")
+proof (intro holomorphic_on_Un)
+ note \<open>f holomorphic_on A\<close>
+ also have "f holomorphic_on A \<longleftrightarrow> ?h holomorphic_on A"
+ by (intro holomorphic_cong) auto
+ finally show \<dots> .
+next
+ note \<open>g holomorphic_on B\<close>
+ also have "g holomorphic_on B \<longleftrightarrow> ?h holomorphic_on B"
+ using assms by (intro holomorphic_cong) auto
+ finally show \<dots> .
+qed (insert assms, auto)
+
lemma DERIV_deriv_iff_field_differentiable:
"DERIV f x :> deriv f x \<longleftrightarrow> f field_differentiable at x"
unfolding field_differentiable_def by (metis DERIV_imp_deriv)
--- a/src/HOL/Analysis/Infinite_Set_Sum.thy Mon Dec 11 17:28:32 2017 +0100
+++ b/src/HOL/Analysis/Infinite_Set_Sum.thy Tue Dec 12 10:01:14 2017 +0100
@@ -256,6 +256,27 @@
show "f abs_summable_on insert x A" by simp
qed
+lemma abs_summable_sum:
+ assumes "\<And>x. x \<in> A \<Longrightarrow> f x abs_summable_on B"
+ shows "(\<lambda>y. \<Sum>x\<in>A. f x y) abs_summable_on B"
+ using assms unfolding abs_summable_on_def by (intro Bochner_Integration.integrable_sum)
+
+lemma abs_summable_Re: "f abs_summable_on A \<Longrightarrow> (\<lambda>x. Re (f x)) abs_summable_on A"
+ by (simp add: abs_summable_on_def)
+
+lemma abs_summable_Im: "f abs_summable_on A \<Longrightarrow> (\<lambda>x. Im (f x)) abs_summable_on A"
+ by (simp add: abs_summable_on_def)
+
+lemma abs_summable_on_finite_diff:
+ assumes "f abs_summable_on A" "A \<subseteq> B" "finite (B - A)"
+ shows "f abs_summable_on B"
+proof -
+ have "f abs_summable_on (A \<union> (B - A))"
+ by (intro abs_summable_on_union assms abs_summable_on_finite)
+ also from assms have "A \<union> (B - A) = B" by blast
+ finally show ?thesis .
+qed
+
lemma abs_summable_on_reindex_bij_betw:
assumes "bij_betw g A B"
shows "(\<lambda>x. f (g x)) abs_summable_on A \<longleftrightarrow> f abs_summable_on B"
@@ -407,6 +428,27 @@
lemma infsetsum_all_0: "(\<And>x. x \<in> A \<Longrightarrow> f x = 0) \<Longrightarrow> infsetsum f A = 0"
by simp
+lemma infsetsum_nonneg: "(\<And>x. x \<in> A \<Longrightarrow> f x \<ge> (0::real)) \<Longrightarrow> infsetsum f A \<ge> 0"
+ unfolding infsetsum_def by (rule Bochner_Integration.integral_nonneg) auto
+
+lemma sum_infsetsum:
+ assumes "\<And>x. x \<in> A \<Longrightarrow> f x abs_summable_on B"
+ shows "(\<Sum>x\<in>A. \<Sum>\<^sub>ay\<in>B. f x y) = (\<Sum>\<^sub>ay\<in>B. \<Sum>x\<in>A. f x y)"
+ using assms by (simp add: infsetsum_def abs_summable_on_def Bochner_Integration.integral_sum)
+
+lemma Re_infsetsum: "f abs_summable_on A \<Longrightarrow> Re (infsetsum f A) = (\<Sum>\<^sub>ax\<in>A. Re (f x))"
+ by (simp add: infsetsum_def abs_summable_on_def)
+
+lemma Im_infsetsum: "f abs_summable_on A \<Longrightarrow> Im (infsetsum f A) = (\<Sum>\<^sub>ax\<in>A. Im (f x))"
+ by (simp add: infsetsum_def abs_summable_on_def)
+
+lemma infsetsum_of_real:
+ shows "infsetsum (\<lambda>x. of_real (f x)
+ :: 'a :: {real_normed_algebra_1,banach,second_countable_topology,real_inner}) A =
+ of_real (infsetsum f A)"
+ unfolding infsetsum_def
+ by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_inner_left[of 1]]) auto
+
lemma infsetsum_finite [simp]: "finite A \<Longrightarrow> infsetsum f A = (\<Sum>x\<in>A. f x)"
by (simp add: infsetsum_def lebesgue_integral_count_space_finite)
--- a/src/HOL/Computational_Algebra/Euclidean_Algorithm.thy Mon Dec 11 17:28:32 2017 +0100
+++ b/src/HOL/Computational_Algebra/Euclidean_Algorithm.thy Tue Dec 12 10:01:14 2017 +0100
@@ -602,6 +602,9 @@
by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
qed
+lemma prime_factorization_Suc_0 [simp]: "prime_factorization (Suc 0) = {#}"
+ unfolding One_nat_def [symmetric] using prime_factorization_1 .
+
instance int :: normalization_euclidean_semiring ..
instance int :: euclidean_ring_gcd
--- a/src/HOL/Series.thy Mon Dec 11 17:28:32 2017 +0100
+++ b/src/HOL/Series.thy Tue Dec 12 10:01:14 2017 +0100
@@ -1223,4 +1223,32 @@
ultimately show ?thesis by simp
qed
+lemma summable_bounded_partials:
+ fixes f :: "nat \<Rightarrow> 'a :: {real_normed_vector,complete_space}"
+ assumes bound: "eventually (\<lambda>x0. \<forall>a\<ge>x0. \<forall>b>a. norm (sum f {a<..b}) \<le> g a) sequentially"
+ assumes g: "g \<longlonglongrightarrow> 0"
+ shows "summable f" unfolding summable_iff_convergent'
+proof (intro Cauchy_convergent CauchyI', goal_cases)
+ case (1 \<epsilon>)
+ with g have "eventually (\<lambda>x. \<bar>g x\<bar> < \<epsilon>) sequentially"
+ by (auto simp: tendsto_iff)
+ from eventually_conj[OF this bound] obtain x0 where x0:
+ "\<And>x. x \<ge> x0 \<Longrightarrow> \<bar>g x\<bar> < \<epsilon>" "\<And>a b. x0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> norm (sum f {a<..b}) \<le> g a"
+ unfolding eventually_at_top_linorder by auto
+
+ show ?case
+ proof (intro exI[of _ x0] allI impI)
+ fix m n assume mn: "x0 \<le> m" "m < n"
+ have "dist (sum f {..m}) (sum f {..n}) = norm (sum f {..n} - sum f {..m})"
+ by (simp add: dist_norm norm_minus_commute)
+ also have "sum f {..n} - sum f {..m} = sum f ({..n} - {..m})"
+ using mn by (intro Groups_Big.sum_diff [symmetric]) auto
+ also have "{..n} - {..m} = {m<..n}" using mn by auto
+ also have "norm (sum f {m<..n}) \<le> g m" using mn by (intro x0) auto
+ also have "\<dots> \<le> \<bar>g m\<bar>" by simp
+ also have "\<dots> < \<epsilon>" using mn by (intro x0) auto
+ finally show "dist (sum f {..m}) (sum f {..n}) < \<epsilon>" .
+ qed
+qed
+
end