# HG changeset patch # User immler # Date 1545940850 -3600 # Node ID 0f31dd2e540d1792e875a84252fd7ddc179389e7 # Parent f9bf65d90b698687c596ee35180e53670194ec45 generalized to big sum diff -r f9bf65d90b69 -r 0f31dd2e540d src/HOL/Analysis/Convex_Euclidean_Space.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 "(\v\y ` {1..k}. sum u {i \ {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 "(\v\y ` {1..k}. sum u {i \ {1..k}. y i = v} *\<^sub>R v) = x" - using sum_image_gen[OF fin, of "\i. u i *\<^sub>R y i" y, symmetric] + using sum.image_gen[OF fin, of "\i. u i *\<^sub>R y i" y, symmetric] unfolding scaleR_left.sum using obt(3) by auto ultimately have "\S u. finite S \ S \ p \ (\x\S. 0 \ u x) \ sum u S = 1 \ (\v\S. u v *\<^sub>R v) = x" @@ -2683,8 +2683,8 @@ by (auto simp: sum_constant_scaleR) } then have "(\x = 1..card S. u (f x)) = 1" "(\i = 1..card S. u (f i) *\<^sub>R f i) = y" - unfolding sum_image_gen[OF *(1), of "\x. u (f x) *\<^sub>R f x" f] - and sum_image_gen[OF *(1), of "\x. u (f x)" f] + unfolding sum.image_gen[OF *(1), of "\x. u (f x) *\<^sub>R f x" f] + and sum.image_gen[OF *(1), of "\x. u (f x)" f] unfolding f using sum.cong [of S S "\y. (\x\{x \ {1..card S}. f x = y}. u (f x) *\<^sub>R f x)" "\v. u v *\<^sub>R v"] using sum.cong [of S S "\y. (\x\{x \ {1..card S}. f x = y}. u (f x))" u] diff -r f9bf65d90b69 -r 0f31dd2e540d src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- 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 ((\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))) \ norm (\j = 0..s. \(x,K)\{xk \ \. 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 "\ < e/2" proof - diff -r f9bf65d90b69 -r 0f31dd2e540d src/HOL/Analysis/Linear_Algebra.thy --- 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 \ T" - shows "sum (\y. sum g {x. x \ S \ 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: "(\x. x \ y = x \ z) \ y = z" proof assume "\x. x \ y = x \ z" diff -r f9bf65d90b69 -r 0f31dd2e540d src/HOL/Groups_Big.thy --- 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 (\x. F (g x) {y. y \ B \ R x y}) A = F (\y. F (\x. g x y) {x. x \ A \ R x y}) B" by (simp add: inter_filter) (rule swap) +lemma image_gen: + assumes fin: "finite S" + shows "F h S = F (\y. F h {x. x \ S \ g x = y}) (g ` S)" +proof - + have "{y. y\ g`S \ g x = y} = {g x}" if "x \ S" for x + using that by auto + then have "F h S = F (\x. F (\y. h x) {y. y\ g`S \ g x = y}) S" + by simp + also have "\ = F (\y. F h {x. x \ S \ 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 \ T" + shows "F (\y. F h {x. x \ S \ 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 \ -(* TODO generalization candidates *) - -lemma (in comm_monoid_add) sum_image_gen: - assumes fin: "finite S" - shows "sum g S = sum (\y. sum g {x. x \ S \ f x = y}) (f ` S)" -proof - - have "{y. y\ f`S \ f x = y} = {f x}" if "x \ S" for x - using that by auto - then have "sum g S = sum (\x. sum (\y. g x) {y. y\ f`S \ f x = y}) S" - by simp - also have "\ = sum (\y. sum g {x. x \ S \ f x = y}) (f ` S)" - by (rule sum.swap_restrict [OF fin finite_imageI [OF fin]]) - finally show ?thesis . -qed - subsubsection \Properties in more restricted classes of structures\ @@ -756,7 +760,7 @@ also have "\ \ sum (\y. sum g {x. x\t \ i x = y}) (i ` t)" using assms(2-4) by (auto intro!: sum_mono2 sum_nonneg) also have "\ \ 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