generalized to big sum
authorimmler
Thu, 27 Dec 2018 21:00:50 +0100
changeset 69510 0f31dd2e540d
parent 69509 f9bf65d90b69
child 69511 4cc5e4a528f8
generalized to big sum
src/HOL/Analysis/Convex_Euclidean_Space.thy
src/HOL/Analysis/Henstock_Kurzweil_Integration.thy
src/HOL/Analysis/Linear_Algebra.thy
src/HOL/Groups_Big.thy
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy	Thu Dec 27 19:48:41 2018 +0100
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy	Thu Dec 27 21:00:50 2018 +0100
@@ -2637,9 +2637,9 @@
     }
     moreover
     have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v}) = 1"
-      unfolding sum_image_gen[OF fin, symmetric] using obt(2) by auto
+      unfolding sum.image_gen[OF fin, symmetric] using obt(2) by auto
     moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
-      using sum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
+      using sum.image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
       unfolding scaleR_left.sum using obt(3) by auto
     ultimately
     have "\<exists>S u. finite S \<and> S \<subseteq> p \<and> (\<forall>x\<in>S. 0 \<le> u x) \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = x"
@@ -2683,8 +2683,8 @@
         by (auto simp: sum_constant_scaleR)
     }
     then have "(\<Sum>x = 1..card S. u (f x)) = 1" "(\<Sum>i = 1..card S. u (f i) *\<^sub>R f i) = y"
-      unfolding sum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f]
-        and sum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
+      unfolding sum.image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f]
+        and sum.image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
       unfolding f
       using sum.cong [of S S "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
       using sum.cong [of S S "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card S}. f x = y}. u (f x))" u]
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Thu Dec 27 19:48:41 2018 +0100
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Thu Dec 27 21:00:50 2018 +0100
@@ -5924,7 +5924,7 @@
       next
         have "norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f (m x) x) - (\<Sum>(x,K)\<in>\<D>. integral K (f (m x))))
               \<le> norm (\<Sum>j = 0..s. \<Sum>(x,K)\<in>{xk \<in> \<D>. m (fst xk) = j}. content K *\<^sub>R f (m x) x - integral K (f (m x)))"
-          apply (subst sum_group)
+          apply (subst sum.group)
           using s by (auto simp: sum_subtractf split_def p'(1))
         also have "\<dots> < e/2"
         proof -
--- a/src/HOL/Analysis/Linear_Algebra.thy	Thu Dec 27 19:48:41 2018 +0100
+++ b/src/HOL/Analysis/Linear_Algebra.thy	Thu Dec 27 21:00:50 2018 +0100
@@ -106,12 +106,6 @@
   using sum_norm_le[OF K] sum_constant[symmetric]
   by simp
 
-lemma sum_group:
-  assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T"
-  shows "sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) T = sum g S"
-  unfolding sum_image_gen[OF fS, of g f]
-  by (auto intro: sum.neutral sum.mono_neutral_right[OF fT fST])
-
 lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = z"
 proof
   assume "\<forall>x. x \<bullet> y = x \<bullet> z"
--- a/src/HOL/Groups_Big.thy	Thu Dec 27 19:48:41 2018 +0100
+++ b/src/HOL/Groups_Big.thy	Thu Dec 27 21:00:50 2018 +0100
@@ -439,6 +439,25 @@
     F (\<lambda>x. F (g x) {y. y \<in> B \<and> R x y}) A = F (\<lambda>y. F (\<lambda>x. g x y) {x. x \<in> A \<and> R x y}) B"
   by (simp add: inter_filter) (rule swap)
 
+lemma image_gen:
+  assumes fin: "finite S"
+  shows "F h S = F (\<lambda>y. F h {x. x \<in> S \<and> g x = y}) (g ` S)"
+proof -
+  have "{y. y\<in> g`S \<and> g x = y} = {g x}" if "x \<in> S" for x
+    using that by auto
+  then have "F h S = F (\<lambda>x. F (\<lambda>y. h x) {y. y\<in> g`S \<and> g x = y}) S"
+    by simp
+  also have "\<dots> = F (\<lambda>y. F h {x. x \<in> S \<and> g x = y}) (g ` S)"
+    by (rule swap_restrict [OF fin finite_imageI [OF fin]])
+  finally show ?thesis .
+qed
+
+lemma group:
+  assumes fS: "finite S" and fT: "finite T" and fST: "g ` S \<subseteq> T"
+  shows "F (\<lambda>y. F h {x. x \<in> S \<and> g x = y}) T = F h S"
+  unfolding image_gen[OF fS, of h g]
+  by (auto intro: neutral mono_neutral_right[OF fT fST])
+
 lemma Plus:
   fixes A :: "'b set" and B :: "'c set"
   assumes fin: "finite A" "finite B"
@@ -534,21 +553,6 @@
 in [(@{const_syntax sum}, K sum_tr')] end
 \<close>
 
-(* TODO generalization candidates *)
-
-lemma (in comm_monoid_add) sum_image_gen:
-  assumes fin: "finite S"
-  shows "sum g S = sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) (f ` S)"
-proof -
-  have "{y. y\<in> f`S \<and> f x = y} = {f x}" if "x \<in> S" for x
-    using that by auto
-  then have "sum g S = sum (\<lambda>x. sum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
-    by simp
-  also have "\<dots> = sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) (f ` S)"
-    by (rule sum.swap_restrict [OF fin finite_imageI [OF fin]])
-  finally show ?thesis .
-qed
-
 
 subsubsection \<open>Properties in more restricted classes of structures\<close>
 
@@ -756,7 +760,7 @@
   also have "\<dots> \<le> sum (\<lambda>y. sum g {x. x\<in>t \<and> i x = y}) (i ` t)"
     using assms(2-4) by (auto intro!: sum_mono2 sum_nonneg)
   also have "\<dots> \<le> sum g t"
-    using assms by (auto simp: sum_image_gen[symmetric])
+    using assms by (auto simp: sum.image_gen[symmetric])
   finally show ?thesis .
 qed