# HG changeset patch # User wenzelm # Date 1378761062 -7200 # Node ID c24892032eea9013293499c1797b3ca50a178c73 # Parent c413adedef469f88996675a5f5493da27a75f355 tuned proofs; diff -r c413adedef46 -r c24892032eea src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Mon Sep 09 20:24:15 2013 +0200 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Mon Sep 09 23:11:02 2013 +0200 @@ -2999,7 +2999,7 @@ done lemma integrable_setsum: - "finite t \ \a \ t.(f a) integrable_on s \ (\x. setsum (\a. f a x) t) integrable_on s" + "finite t \ \a \ t. (f a) integrable_on s \ (\x. setsum (\a. f a x) t) integrable_on s" unfolding integrable_on_def apply (drule bchoice) using has_integral_setsum[of t] @@ -3141,7 +3141,12 @@ proof (rule, rule) case goal1 then have "e/2 > 0" by auto - then guess d apply- apply(drule y[rule_format]) by(erule exE,erule conjE) note d=this[rule_format] + then guess d + apply - + apply (drule y[rule_format]) + apply (elim exE conjE) + done + note d=this[rule_format] show ?case apply (rule_tac x=d in exI) apply rule @@ -3215,7 +3220,8 @@ by auto guess N2 using y[OF *] .. note N2=this show "\d. gauge d \ - (\p. p tagged_division_of {a..b} \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - y) < e)" + (\p. p tagged_division_of {a..b} \ d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R f x) - y) < e)" apply (rule_tac x="d (N1 + N2)" in exI) apply rule defer @@ -3315,7 +3321,8 @@ shows "content(k1 \ {x. x\k \ c}) = 0" proof - note d=division_ofD[OF assms(1)] - have *: "\a b::'a. \ c. (content({a..b} \ {x. x\k \ c}) = 0 \ interior({a..b} \ {x. x\k \ c}) = {})" + have *: "\(a::'a) b c. content ({a..b} \ {x. x\k \ c}) = 0 \ + interior({a..b} \ {x. x\k \ c}) = {}" unfolding interval_split[OF k] content_eq_0_interior by auto guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this guess u2 v2 using d(4)[OF assms(3)] apply-by(erule exE)+ note uv2=this @@ -3344,8 +3351,8 @@ have *: "\a b::'a. \c. content({a..b} \ {x. x\k \ c}) = 0 \ interior({a..b} \ {x. x\k \ c}) = {}" unfolding interval_split[OF k] content_eq_0_interior by auto - guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this - guess u2 v2 using d(4)[OF assms(3)] apply-by(erule exE)+ note uv2=this + guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this + guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this have **: "\s t u. s \ t = {} \ u \ s \ u \ t \ u = {}" by auto show ?thesis @@ -3361,7 +3368,7 @@ lemma tagged_division_split_left_inj: fixes x1 :: "'a::ordered_euclidean_space" assumes "d tagged_division_of i" - and "(x1,k1) \ d" + and "(x1, k1) \ d" and "(x2, k2) \ d" and "k1 \ k2" and "k1 \ {x. x\k \ c} = k2 \ {x. x\k \ c}" @@ -3389,7 +3396,7 @@ and "k1 \ k2" and "k1 \ {x. x\k \ c} = k2 \ {x. x\k \ c}" and k: "k \ Basis" - shows "content(k1 \ {x. x\k \ c}) = 0" + shows "content (k1 \ {x. x\k \ c}) = 0" proof - have *: "\a b c. (a,b) \ c \ b \ snd ` c" unfolding image_iff @@ -3472,8 +3479,10 @@ case goal1 then have e: "e/2 > 0" by auto - guess d1 using has_integralD[OF assms(1)[unfolded interval_split[OF k]] e] . note d1=this[unfolded interval_split[symmetric,OF k]] - guess d2 using has_integralD[OF assms(2)[unfolded interval_split[OF k]] e] . note d2=this[unfolded interval_split[symmetric,OF k]] + guess d1 using has_integralD[OF assms(1)[unfolded interval_split[OF k]] e] . + note d1=this[unfolded interval_split[symmetric,OF k]] + guess d2 using has_integralD[OF assms(2)[unfolded interval_split[OF k]] e] . + note d2=this[unfolded interval_split[symmetric,OF k]] let ?d = "\x. if x\k = c then (d1 x \ d2 x) else ball x (abs(x\k - c)) \ d1 x \ d2 x" show ?case apply (rule_tac x="?d" in exI) @@ -3486,7 +3495,8 @@ show "gauge ?d" using d1(1) d2(1) unfolding gauge_def by auto fix p - assume "p tagged_division_of {a..b}" "?d fine p" note p = this tagged_division_ofD[OF this(1)] + assume "p tagged_division_of {a..b}" "?d fine p" + note p = this tagged_division_ofD[OF this(1)] have lem0: "\x kk. (x, kk) \ p \ kk \ {x. x\k \ c} \ {} \ x\k \ c" "\x kk. (x, kk) \ p \ kk \ {x. x\k \ c} \ {} \ x\k \ c" @@ -3701,7 +3711,8 @@ qed also note setsum_addf[symmetric] also have *: "\x. x \ p \ - (\(x,ka). content (ka \ {x. x \ k \ c}) *\<^sub>R f x) x + (\(x,ka). content (ka \ {x. c \ x \ k}) *\<^sub>R f x) x = + (\(x,ka). content (ka \ {x. x \ k \ c}) *\<^sub>R f x) x + + (\(x,ka). content (ka \ {x. c \ x \ k}) *\<^sub>R f x) x = (\(x,ka). content ka *\<^sub>R f x) x" unfolding split_paired_all split_conv proof - @@ -3728,170 +3739,296 @@ subsection {* A sort of converse, integrability on subintervals. *} -lemma tagged_division_union_interval: fixes a::"'a::ordered_euclidean_space" - assumes "p1 tagged_division_of ({a..b} \ {x. x\k \ (c::real)})" "p2 tagged_division_of ({a..b} \ {x. x\k \ c})" - and k:"k\Basis" +lemma tagged_division_union_interval: + fixes a :: "'a::ordered_euclidean_space" + assumes "p1 tagged_division_of ({a..b} \ {x. x\k \ (c::real)})" + and "p2 tagged_division_of ({a..b} \ {x. x\k \ c})" + and k: "k \ Basis" shows "(p1 \ p2) tagged_division_of ({a..b})" -proof- have *:"{a..b} = ({a..b} \ {x. x\k \ c}) \ ({a..b} \ {x. x\k \ c})" by auto - show ?thesis apply(subst *) apply(rule tagged_division_union[OF assms(1-2)]) - unfolding interval_split[OF k] interior_closed_interval using k - by(auto simp add: eucl_less[where 'a='a] elim!: ballE[where x=k]) qed - -lemma has_integral_separate_sides: fixes f::"'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "(f has_integral i) ({a..b})" "e>0" and k:"k\Basis" - obtains d where "gauge d" "(\p1 p2. p1 tagged_division_of ({a..b} \ {x. x\k \ c}) \ d fine p1 \ - p2 tagged_division_of ({a..b} \ {x. x\k \ c}) \ d fine p2 - \ norm((setsum (\(x,k). content k *\<^sub>R f x) p1 + - setsum (\(x,k). content k *\<^sub>R f x) p2) - i) < e)" -proof- guess d using has_integralD[OF assms(1-2)] . note d=this - show ?thesis apply(rule that[of d]) apply(rule d) apply(rule,rule,rule,(erule conjE)+) - proof- fix p1 p2 assume "p1 tagged_division_of {a..b} \ {x. x \ k \ c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this - assume "p2 tagged_division_of {a..b} \ {x. c \ x \ k}" "d fine p2" note p2=tagged_division_ofD[OF this(1)] this +proof - + have *: "{a..b} = ({a..b} \ {x. x\k \ c}) \ ({a..b} \ {x. x\k \ c})" + by auto + show ?thesis + apply (subst *) + apply (rule tagged_division_union[OF assms(1-2)]) + unfolding interval_split[OF k] interior_closed_interval + using k + apply (auto simp add: eucl_less[where 'a='a] elim!: ballE[where x=k]) + done +qed + +lemma has_integral_separate_sides: + fixes f :: "'a::ordered_euclidean_space \ 'b::real_normed_vector" + assumes "(f has_integral i) ({a..b})" + and "e > 0" + and k: "k \ Basis" + obtains d where "gauge d" + "\p1 p2. p1 tagged_division_of ({a..b} \ {x. x\k \ c}) \ d fine p1 \ + p2 tagged_division_of ({a..b} \ {x. x\k \ c}) \ d fine p2 \ + norm ((setsum (\(x,k). content k *\<^sub>R f x) p1 + setsum (\(x,k). content k *\<^sub>R f x) p2) - i) < e" +proof - + guess d using has_integralD[OF assms(1-2)] . note d=this + show ?thesis + apply (rule that[of d]) + apply (rule d) + apply rule + apply rule + apply rule + apply (elim conjE) + proof - + fix p1 p2 + assume "p1 tagged_division_of {a..b} \ {x. x \ k \ c}" "d fine p1" + note p1=tagged_division_ofD[OF this(1)] this + assume "p2 tagged_division_of {a..b} \ {x. c \ x \ k}" "d fine p2" + note p2=tagged_division_ofD[OF this(1)] this note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this - have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" - apply(subst setsum_Un_zero) apply(rule p1 p2)+ apply(rule) unfolding split_paired_all split_conv - proof- fix a b assume ab:"(a,b) \ p1 \ p2" - have "(a,b) \ p1" using ab by auto from p1(4)[OF this] guess u v apply-by(erule exE)+ note uv =this - have "b \ {x. x\k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce - moreover have "interior {x::'a. x \ k = c} = {}" - proof(rule ccontr) case goal1 then obtain x where x:"x\interior {x::'a. x\k = c}" by auto + have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = + norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" + apply (subst setsum_Un_zero) + apply (rule p1 p2)+ + apply rule + unfolding split_paired_all split_conv + proof - + fix a b + assume ab: "(a, b) \ p1 \ p2" + have "(a, b) \ p1" + using ab by auto + from p1(4)[OF this] guess u v by (elim exE) note uv = this + have "b \ {x. x\k = c}" + using ab p1(3)[of a b] p2(3)[of a b] by fastforce + moreover + have "interior {x::'a. x \ k = c} = {}" + proof (rule ccontr) + assume "\ ?thesis" + then obtain x where x: "x \ interior {x::'a. x\k = c}" + by auto then guess e unfolding mem_interior .. note e=this - have x:"x\k = c" using x interior_subset by fastforce - have *:"\i. i\Basis \ \(x - (x + (e / 2) *\<^sub>R k)) \ i\ - = (if i = k then e/2 else 0)" using e k by (auto simp: inner_simps inner_not_same_Basis) + have x: "x\k = c" + using x interior_subset by fastforce + have *: "\i. i \ Basis \ \(x - (x + (e / 2) *\<^sub>R k)) \ i\ = (if i = k then e/2 else 0)" + using e k by (auto simp: inner_simps inner_not_same_Basis) have "(\i\Basis. \(x - (x + (e / 2 ) *\<^sub>R k)) \ i\) = - (\i\Basis. (if i = k then e / 2 else 0))" apply(rule setsum_cong2) apply(subst *) by auto - also have "... < e" apply(subst setsum_delta) using e by auto + (\i\Basis. (if i = k then e / 2 else 0))" + apply (rule setsum_cong2) + apply (subst *) + apply auto + done + also have "\ < e" + apply (subst setsum_delta) + using e + apply auto + done finally have "x + (e/2) *\<^sub>R k \ ball x e" unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) - hence "x + (e/2) *\<^sub>R k \ {x. x\k = c}" using e by auto - thus False unfolding mem_Collect_eq using e x k by (auto simp: inner_simps) - qed ultimately have "content b = 0" unfolding uv content_eq_0_interior apply-apply(drule interior_mono) by auto - thus "content b *\<^sub>R f a = 0" by auto + then have "x + (e/2) *\<^sub>R k \ {x. x\k = c}" + using e by auto + then show False + unfolding mem_Collect_eq using e x k by (auto simp: inner_simps) + qed + ultimately have "content b = 0" + unfolding uv content_eq_0_interior + apply - + apply (drule interior_mono) + apply auto + done + then show "content b *\<^sub>R f a = 0" + by auto qed auto - also have "\ < e" by(rule k d(2) p12 fine_union p1 p2)+ - finally show "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . qed qed + also have "\ < e" + by (rule k d(2) p12 fine_union p1 p2)+ + finally show "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . + qed +qed lemma integrable_split[intro]: - fixes f::"'a::ordered_euclidean_space \ 'b::{real_normed_vector,complete_space}" - assumes "f integrable_on {a..b}" and k:"k\Basis" - shows "f integrable_on ({a..b} \ {x. x\k \ c})" (is ?t1) and "f integrable_on ({a..b} \ {x. x\k \ c})" (is ?t2) -proof- guess y using assms(1) unfolding integrable_on_def .. note y=this + fixes f :: "'a::ordered_euclidean_space \ 'b::{real_normed_vector,complete_space}" + assumes "f integrable_on {a..b}" + and k: "k \ Basis" + shows "f integrable_on ({a..b} \ {x. x\k \ c})" (is ?t1) + and "f integrable_on ({a..b} \ {x. x\k \ c})" (is ?t2) +proof - + guess y using assms(1) unfolding integrable_on_def .. note y=this def b' \ "\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i::'a" def a' \ "\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i::'a" - show ?t1 ?t2 unfolding interval_split[OF k] integrable_cauchy unfolding interval_split[symmetric,OF k] - proof(rule_tac[!] allI impI)+ fix e::real assume "e>0" hence "e/2>0" by auto + show ?t1 ?t2 + unfolding interval_split[OF k] integrable_cauchy + unfolding interval_split[symmetric,OF k] + proof (rule_tac[!] allI impI)+ + fix e :: real + assume "e > 0" + then have "e/2>0" + by auto from has_integral_separate_sides[OF y this k,of c] guess d . note d=this[rule_format] - let ?P = "\A. \d. gauge d \ (\p1 p2. p1 tagged_division_of {a..b} \ A \ d fine p1 - \ p2 tagged_division_of {a..b} \ A \ d fine p2 \ + let ?P = "\A. \d. gauge d \ (\p1 p2. p1 tagged_division_of {a..b} \ A \ d fine p1 \ + p2 tagged_division_of {a..b} \ A \ d fine p2 \ norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e)" - show "?P {x. x \ k \ c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule) - proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p1 - \ p2 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p2" + show "?P {x. x \ k \ c}" + apply (rule_tac x=d in exI) + apply rule + apply (rule d) + apply rule + apply rule + apply rule + proof - + fix p1 p2 + assume as: "p1 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p1 \ + p2 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" - proof- guess p using fine_division_exists[OF d(1), of a' b] . note p=this - show ?thesis using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]] + proof - + guess p using fine_division_exists[OF d(1), of a' b] . note p=this + show ?thesis + using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]] using as unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] - using p using assms by(auto simp add:algebra_simps) - qed qed - show "?P {x. x \ k \ c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule) - proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p1 - \ p2 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p2" + using p using assms + by (auto simp add: algebra_simps) + qed + qed + show "?P {x. x \ k \ c}" + apply (rule_tac x=d in exI) + apply rule + apply (rule d) + apply rule + apply rule + apply rule + proof - + fix p1 p2 + assume as: "p1 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p1 \ + p2 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" - proof- guess p using fine_division_exists[OF d(1), of a b'] . note p=this - show ?thesis using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]] - using as unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] - using p using assms by(auto simp add:algebra_simps) qed qed qed qed + proof - + guess p using fine_division_exists[OF d(1), of a b'] . note p=this + show ?thesis + using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]] + using as + unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] + using p + using assms + by (auto simp add:algebra_simps) + qed + qed + qed +qed + subsection {* Generalized notion of additivity. *} definition "neutral opp = (SOME x. \y. opp x y = y \ opp y x = y)" -definition operative :: "('a \ 'a \ 'a) \ (('b::ordered_euclidean_space) set \ 'a) \ bool" where - "operative opp f \ - (\a b. content {a..b} = 0 \ f {a..b} = neutral(opp)) \ - (\a b c. \k\Basis. f({a..b}) = - opp (f({a..b} \ {x. x\k \ c})) - (f({a..b} \ {x. x\k \ c})))" - -lemma operativeD[dest]: fixes type::"'a::ordered_euclidean_space" assumes "operative opp f" - shows "\a b. content {a..b} = 0 \ f {a..b::'a} = neutral(opp)" - "\a b c k. k\Basis \ f({a..b}) = opp (f({a..b} \ {x. x\k \ c})) (f({a..b} \ {x. x\k \ c}))" +definition operative :: "('a \ 'a \ 'a) \ (('b::ordered_euclidean_space) set \ 'a) \ bool" + where "operative opp f \ + (\a b. content {a..b} = 0 \ f {a..b} = neutral opp) \ + (\a b c. \k\Basis. f {a..b} = opp (f({a..b} \ {x. x\k \ c})) (f ({a..b} \ {x. x\k \ c})))" + +lemma operativeD[dest]: + fixes type :: "'a::ordered_euclidean_space" + assumes "operative opp f" + shows "\a b::'a. content {a..b} = 0 \ f {a..b} = neutral opp" + and "\a b c k. k \ Basis \ f {a..b} = + opp (f ({a..b} \ {x. x\k \ c})) (f ({a..b} \ {x. x\k \ c}))" using assms unfolding operative_def by auto -lemma operative_trivial: - "operative opp f \ content({a..b}) = 0 \ f({a..b}) = neutral opp" +lemma operative_trivial: "operative opp f \ content({a..b}) = 0 \ f({a..b}) = neutral opp" unfolding operative_def by auto -lemma property_empty_interval: - "(\a b. content({a..b}) = 0 \ P({a..b})) \ P {}" +lemma property_empty_interval: "\a b. content {a..b} = 0 \ P {a..b} \ P {}" using content_empty unfolding empty_as_interval by auto lemma operative_empty: "operative opp f \ f {} = neutral opp" - unfolding operative_def apply(rule property_empty_interval) by auto + unfolding operative_def by (rule property_empty_interval) auto + subsection {* Using additivity of lifted function to encode definedness. *} -lemma forall_option: "(\x. P x) \ P None \ (\x. P(Some x))" - by (metis option.nchotomy) - -lemma exists_option: "(\x. P x) \ P None \ (\x. P(Some x))" +lemma forall_option: "(\x. P x) \ P None \ (\x. P (Some x))" by (metis option.nchotomy) -fun lifted -where - "lifted (opp::'a\'a\'b) (Some x) (Some y) = Some (opp x y)" +lemma exists_option: "(\x. P x) \ P None \ (\x. P (Some x))" + by (metis option.nchotomy) + +fun lifted where + "lifted (opp :: 'a \ 'a \ 'b) (Some x) (Some y) = Some (opp x y)" | "lifted opp None _ = (None::'b option)" | "lifted opp _ None = None" lemma lifted_simp_1[simp]: "lifted opp v None = None" by (induct v) auto -definition "monoidal opp \ (\x y. opp x y = opp y x) \ - (\x y z. opp x (opp y z) = opp (opp x y) z) \ - (\x. opp (neutral opp) x = x)" +definition "monoidal opp \ + (\x y. opp x y = opp y x) \ + (\x y z. opp x (opp y z) = opp (opp x y) z) \ + (\x. opp (neutral opp) x = x)" lemma monoidalI: assumes "\x y. opp x y = opp y x" - "\x y z. opp x (opp y z) = opp (opp x y) z" - "\x. opp (neutral opp) x = x" shows "monoidal opp" + and "\x y z. opp x (opp y z) = opp (opp x y) z" + and "\x. opp (neutral opp) x = x" + shows "monoidal opp" unfolding monoidal_def using assms by fastforce lemma monoidal_ac: assumes "monoidal opp" - shows "opp (neutral opp) a = a" "opp a (neutral opp) = a" "opp a b = opp b a" - "opp (opp a b) c = opp a (opp b c)" "opp a (opp b c) = opp b (opp a c)" + shows "opp (neutral opp) a = a" + and "opp a (neutral opp) = a" + and "opp a b = opp b a" + and "opp (opp a b) c = opp a (opp b c)" + and "opp a (opp b c) = opp b (opp a c)" using assms unfolding monoidal_def by metis+ -lemma monoidal_simps[simp]: assumes "monoidal opp" - shows "opp (neutral opp) a = a" "opp a (neutral opp) = a" +lemma monoidal_simps[simp]: + assumes "monoidal opp" + shows "opp (neutral opp) a = a" + and "opp a (neutral opp) = a" using monoidal_ac[OF assms] by auto -lemma neutral_lifted[cong]: assumes "monoidal opp" - shows "neutral (lifted opp) = Some(neutral opp)" - apply(subst neutral_def) apply(rule some_equality) apply(rule,induct_tac y) prefer 3 +lemma neutral_lifted[cong]: + assumes "monoidal opp" + shows "neutral (lifted opp) = Some (neutral opp)" + apply (subst neutral_def) + apply (rule some_equality) + apply rule + apply (induct_tac y) + prefer 3 proof - - fix x assume "\y. lifted opp x y = y \ lifted opp y x = y" - thus "x = Some (neutral opp)" - apply(induct x) defer - apply rule apply(subst neutral_def) apply(subst eq_commute,rule some_equality) - apply(rule,erule_tac x="Some y" in allE) defer apply(erule_tac x="Some x" in allE) + fix x + assume "\y. lifted opp x y = y \ lifted opp y x = y" + then show "x = Some (neutral opp)" + apply (induct x) + defer + apply rule + apply (subst neutral_def) + apply (subst eq_commute) + apply(rule some_equality) + apply rule + apply (erule_tac x="Some y" in allE) + defer + apply (erule_tac x="Some x" in allE) apply auto done -qed(auto simp add:monoidal_ac[OF assms]) - -lemma monoidal_lifted[intro]: assumes "monoidal opp" shows "monoidal(lifted opp)" - unfolding monoidal_def forall_option neutral_lifted[OF assms] using monoidal_ac[OF assms] by auto +qed (auto simp add:monoidal_ac[OF assms]) + +lemma monoidal_lifted[intro]: + assumes "monoidal opp" + shows "monoidal (lifted opp)" + unfolding monoidal_def forall_option neutral_lifted[OF assms] + using monoidal_ac[OF assms] + by auto definition "support opp f s = {x. x\s \ f x \ neutral opp}" -definition "fold' opp e s \ (if finite s then Finite_Set.fold opp e s else e)" -definition "iterate opp s f \ fold' (\x a. opp (f x) a) (neutral opp) (support opp f s)" - -lemma support_subset[intro]:"support opp f s \ s" unfolding support_def by auto -lemma support_empty[simp]:"support opp f {} = {}" using support_subset[of opp f "{}"] by auto - -lemma comp_fun_commute_monoidal[intro]: assumes "monoidal opp" shows "comp_fun_commute opp" - unfolding comp_fun_commute_def using monoidal_ac[OF assms] by auto +definition "fold' opp e s = (if finite s then Finite_Set.fold opp e s else e)" +definition "iterate opp s f = fold' (\x a. opp (f x) a) (neutral opp) (support opp f s)" + +lemma support_subset[intro]: "support opp f s \ s" + unfolding support_def by auto + +lemma support_empty[simp]: "support opp f {} = {}" + using support_subset[of opp f "{}"] by auto + +lemma comp_fun_commute_monoidal[intro]: + assumes "monoidal opp" + shows "comp_fun_commute opp" + unfolding comp_fun_commute_def + using monoidal_ac[OF assms] + by auto lemma support_clauses: "\f g s. support opp f {} = {}" @@ -3902,98 +4039,190 @@ "\f g s. support opp f (s \ t) = (support opp f s) \ (support opp f t)" "\f g s. support opp f (s - t) = (support opp f s) - (support opp f t)" "\f g s. support opp g (f ` s) = f ` (support opp (g o f) s)" -unfolding support_def by auto - -lemma finite_support[intro]:"finite s \ finite (support opp f s)" unfolding support_def by auto -lemma iterate_empty[simp]:"iterate opp {} f = neutral opp" +lemma finite_support[intro]: "finite s \ finite (support opp f s)" + unfolding support_def by auto + +lemma iterate_empty[simp]: "iterate opp {} f = neutral opp" unfolding iterate_def fold'_def by auto -lemma iterate_insert[simp]: assumes "monoidal opp" "finite s" - shows "iterate opp (insert x s) f = (if x \ s then iterate opp s f else opp (f x) (iterate opp s f))" -proof(cases "x\s") case True hence *:"insert x s = s" by auto +lemma iterate_insert[simp]: + assumes "monoidal opp" + and "finite s" + shows "iterate opp (insert x s) f = + (if x \ s then iterate opp s f else opp (f x) (iterate opp s f))" +proof (cases "x \ s") + case True + then have *: "insert x s = s" + by auto show ?thesis unfolding iterate_def if_P[OF True] * by auto -next case False note x=this +next + case False + note x = this note * = comp_fun_commute.comp_comp_fun_commute [OF comp_fun_commute_monoidal[OF assms(1)]] - show ?thesis proof(cases "f x = neutral opp") - case True show ?thesis unfolding iterate_def if_not_P[OF x] support_clauses if_P[OF True] - unfolding True monoidal_simps[OF assms(1)] by auto - next case False show ?thesis unfolding iterate_def fold'_def if_not_P[OF x] support_clauses if_not_P[OF False] - apply(subst comp_fun_commute.fold_insert[OF * finite_support, simplified comp_def]) - using `finite s` unfolding support_def using False x by auto qed qed + show ?thesis + proof (cases "f x = neutral opp") + case True + show ?thesis + unfolding iterate_def if_not_P[OF x] support_clauses if_P[OF True] + unfolding True monoidal_simps[OF assms(1)] + by auto + next + case False + show ?thesis + unfolding iterate_def fold'_def if_not_P[OF x] support_clauses if_not_P[OF False] + apply (subst comp_fun_commute.fold_insert[OF * finite_support, simplified comp_def]) + using `finite s` + unfolding support_def + using False x + apply auto + done + qed +qed lemma iterate_some: - assumes "monoidal opp" "finite s" - shows "iterate (lifted opp) s (\x. Some(f x)) = Some (iterate opp s f)" using assms(2) -proof(induct s) case empty thus ?case using assms by auto -next case (insert x F) show ?case apply(subst iterate_insert) prefer 3 apply(subst if_not_P) - defer unfolding insert(3) lifted.simps apply rule using assms insert by auto qed + assumes "monoidal opp" + and "finite s" + shows "iterate (lifted opp) s (\x. Some(f x)) = Some (iterate opp s f)" + using assms(2) +proof (induct s) + case empty + then show ?case + using assms by auto +next + case (insert x F) + show ?case + apply (subst iterate_insert) + prefer 3 + apply (subst if_not_P) + defer + unfolding insert(3) lifted.simps + apply rule + using assms insert + apply auto + done +qed + + subsection {* Two key instances of additivity. *} -lemma neutral_add[simp]: - "neutral op + = (0::_::comm_monoid_add)" unfolding neutral_def - apply(rule some_equality) defer apply(erule_tac x=0 in allE) by auto +lemma neutral_add[simp]: "neutral op + = (0::'a::comm_monoid_add)" + unfolding neutral_def + apply (rule some_equality) + defer + apply (erule_tac x=0 in allE) + apply auto + done lemma operative_content[intro]: "operative (op +) content" - unfolding operative_def neutral_add apply safe - unfolding content_split[symmetric] .. + unfolding operative_def neutral_add + apply safe + unfolding content_split[symmetric] + apply rule + done lemma neutral_monoid: "neutral ((op +)::('a::comm_monoid_add) \ 'a \ 'a) = 0" by (rule neutral_add) (* FIXME: duplicate *) -lemma monoidal_monoid[intro]: - shows "monoidal ((op +)::('a::comm_monoid_add) \ 'a \ 'a)" - unfolding monoidal_def neutral_monoid by(auto simp add: algebra_simps) - -lemma operative_integral: fixes f::"'a::ordered_euclidean_space \ 'b::banach" +lemma monoidal_monoid[intro]: "monoidal ((op +)::('a::comm_monoid_add) \ 'a \ 'a)" + unfolding monoidal_def neutral_monoid + by (auto simp add: algebra_simps) + +lemma operative_integral: + fixes f :: "'a::ordered_euclidean_space \ 'b::banach" shows "operative (lifted(op +)) (\i. if f integrable_on i then Some(integral i f) else None)" - unfolding operative_def unfolding neutral_lifted[OF monoidal_monoid] neutral_add - apply(rule,rule,rule,rule) defer apply(rule allI ballI)+ -proof- - fix a b c and k :: 'a assume k:"k\Basis" + unfolding operative_def + unfolding neutral_lifted[OF monoidal_monoid] neutral_add + apply rule + apply rule + apply rule + apply rule + defer + apply (rule allI ballI)+ +proof - + fix a b c + fix k :: 'a + assume k: "k \ Basis" show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) = lifted op + (if f integrable_on {a..b} \ {x. x \ k \ c} then Some (integral ({a..b} \ {x. x \ k \ c}) f) else None) (if f integrable_on {a..b} \ {x. c \ x \ k} then Some (integral ({a..b} \ {x. c \ x \ k}) f) else None)" - proof(cases "f integrable_on {a..b}") - case True show ?thesis unfolding if_P[OF True] using k apply- - unfolding if_P[OF integrable_split(1)[OF True]] unfolding if_P[OF integrable_split(2)[OF True]] - unfolding lifted.simps option.inject apply(rule integral_unique) apply(rule has_integral_split[OF _ _ k]) - apply(rule_tac[!] integrable_integral integrable_split)+ using True k by auto - next case False have "(\ (f integrable_on {a..b} \ {x. x \ k \ c})) \ (\ ( f integrable_on {a..b} \ {x. c \ x \ k}))" - proof(rule ccontr) case goal1 hence "f integrable_on {a..b}" apply- unfolding integrable_on_def - apply(rule_tac x="integral ({a..b} \ {x. x \ k \ c}) f + integral ({a..b} \ {x. x \ k \ c}) f" in exI) - apply(rule has_integral_split[OF _ _ k]) apply(rule_tac[!] integrable_integral) by auto - thus False using False by auto - qed thus ?thesis using False by auto - qed next - fix a b assume as:"content {a..b::'a} = 0" - thus "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) = Some 0" - unfolding if_P[OF integrable_on_null[OF as]] using has_integral_null_eq[OF as] by auto qed + proof (cases "f integrable_on {a..b}") + case True + show ?thesis + unfolding if_P[OF True] + using k + apply - + unfolding if_P[OF integrable_split(1)[OF True]] + unfolding if_P[OF integrable_split(2)[OF True]] + unfolding lifted.simps option.inject + apply (rule integral_unique) + apply (rule has_integral_split[OF _ _ k]) + apply (rule_tac[!] integrable_integral integrable_split)+ + using True k + apply auto + done + next + case False + have "\ (f integrable_on {a..b} \ {x. x \ k \ c}) \ \ ( f integrable_on {a..b} \ {x. c \ x \ k})" + proof (rule ccontr) + assume "\ ?thesis" + then have "f integrable_on {a..b}" + apply - + unfolding integrable_on_def + apply (rule_tac x="integral ({a..b} \ {x. x \ k \ c}) f + integral ({a..b} \ {x. x \ k \ c}) f" in exI) + apply (rule has_integral_split[OF _ _ k]) + apply (rule_tac[!] integrable_integral) + apply auto + done + then show False + using False by auto + qed + then show ?thesis + using False by auto + qed +next + fix a b :: 'a + assume as: "content {a..b} = 0" + then show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) = Some 0" + unfolding if_P[OF integrable_on_null[OF as]] + using has_integral_null_eq[OF as] + by auto +qed + subsection {* Points of division of a partition. *} definition "division_points (k::('a::ordered_euclidean_space) set) d = - {(j,x). j\Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ - (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" - -lemma division_points_finite: fixes i::"('a::ordered_euclidean_space) set" - assumes "d division_of i" shows "finite (division_points i d)" -proof- note assm = division_ofD[OF assms] + {(j,x). j \ Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ + (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" + +lemma division_points_finite: + fixes i :: "'a::ordered_euclidean_space set" + assumes "d division_of i" + shows "finite (division_points i d)" +proof - + note assm = division_ofD[OF assms] let ?M = "\j. {(j,x)|x. (interval_lowerbound i)\j < x \ x < (interval_upperbound i)\j \ - (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" - have *:"division_points i d = \(?M ` Basis)" + (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" + have *: "division_points i d = \(?M ` Basis)" unfolding division_points_def by auto - show ?thesis unfolding * using assm by auto qed - -lemma division_points_subset: fixes a::"'a::ordered_euclidean_space" - assumes "d division_of {a..b}" "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" and k:"k\Basis" - shows "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} - \ division_points ({a..b}) d" (is ?t1) and - "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} - \ division_points ({a..b}) d" (is ?t2) -proof- note assm = division_ofD[OF assms(1)] - have *:"\i\Basis. a\i \ b\i" + show ?thesis + unfolding * using assm by auto +qed + +lemma division_points_subset: + fixes a :: "'a::ordered_euclidean_space" + assumes "d division_of {a..b}" + and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" + and k: "k \ Basis" + shows "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ l \ {x. x\k \ c} \ {}} \ + division_points ({a..b}) d" (is ?t1) + and "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} \ + division_points ({a..b}) d" (is ?t2) +proof - + note assm = division_ofD[OF assms(1)] + have *: "\i\Basis. a\i \ b\i" "\i\Basis. a\i \ (\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i" "\i\Basis. (\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i \ b\i" "min (b \ k) c = c" "max (a \ k) c = c" @@ -4003,83 +4232,148 @@ unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding * unfolding subset_eq - apply(rule) + apply rule unfolding mem_Collect_eq split_beta - apply(erule bexE conjE)+ - apply(simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms) - apply(erule exE conjE)+ + apply (erule bexE conjE)+ + apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms) + apply (erule exE conjE)+ proof - fix i l x assume as:"a \ fst x < snd x" "snd x < (if fst x = k then c else b \ fst x)" + fix i l x + assume as: + "a \ fst x < snd x" "snd x < (if fst x = k then c else b \ fst x)" "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" - "i = l \ {x. x \ k \ c}" "l \ d" "l \ {x. x \ k \ c} \ {}" and fstx:"fst x \Basis" + "i = l \ {x. x \ k \ c}" "l \ d" "l \ {x. x \ k \ c} \ {}" + and fstx: "fst x \ Basis" from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this - have *:"\i\Basis. u \ i \ (\i\Basis. (if i = k then min (v \ k) c else v \ i) *\<^sub>R i) \ i" + have *: "\i\Basis. u \ i \ (\i\Basis. (if i = k then min (v \ k) c else v \ i) *\<^sub>R i) \ i" using as(6) unfolding l interval_split[OF k] interval_ne_empty as . - have **:"\i\Basis. u\i \ v\i" using l using as(6) unfolding interval_ne_empty[symmetric] by auto + have **: "\i\Basis. u\i \ v\i" + using l using as(6) unfolding interval_ne_empty[symmetric] by auto show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" apply (rule bexI[OF _ `l \ d`]) using as(1-3,5) fstx unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as - by (auto split: split_if_asm) + apply (auto split: split_if_asm) + done show "snd x < b \ fst x" using as(2) `c < b\k` by (auto split: split_if_asm) qed show ?t2 unfolding division_points_def interval_split[OF k, of a b] - unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding * - unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta - apply(erule bexE conjE)+ - apply(simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms) - apply(erule exE conjE)+ + unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] + unfolding * + unfolding subset_eq + apply rule + unfolding mem_Collect_eq split_beta + apply (erule bexE conjE)+ + apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms) + apply (erule exE conjE)+ proof - fix i l x assume as:"(if fst x = k then c else a \ fst x) < snd x" "snd x < b \ fst x" + fix i l x + assume as: + "(if fst x = k then c else a \ fst x) < snd x" "snd x < b \ fst x" "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" - "i = l \ {x. c \ x \ k}" "l \ d" "l \ {x. c \ x \ k} \ {}" and fstx:"fst x \ Basis" - from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this - have *:"\i\Basis. (\i\Basis. (if i = k then max (u \ k) c else u \ i) *\<^sub>R i) \ i \ v \ i" + "i = l \ {x. c \ x \ k}" "l \ d" "l \ {x. c \ x \ k} \ {}" + and fstx: "fst x \ Basis" + from assm(4)[OF this(5)] guess u v by (elim exE) note l=this + have *: "\i\Basis. (\i\Basis. (if i = k then max (u \ k) c else u \ i) *\<^sub>R i) \ i \ v \ i" using as(6) unfolding l interval_split[OF k] interval_ne_empty as . - have **:"\i\Basis. u\i \ v\i" using l using as(6) unfolding interval_ne_empty[symmetric] by auto + have **: "\i\Basis. u\i \ v\i" + using l using as(6) unfolding interval_ne_empty[symmetric] by auto show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" apply (rule bexI[OF _ `l \ d`]) using as(1-3,5) fstx unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as - by (auto split: split_if_asm) + apply (auto split: split_if_asm) + done show "a \ fst x < snd x" using as(1) `a\k < c` by (auto split: split_if_asm) qed qed -lemma division_points_psubset: fixes a::"'a::ordered_euclidean_space" - assumes "d division_of {a..b}" "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" - "l \ d" "interval_lowerbound l\k = c \ interval_upperbound l\k = c" and k:"k\Basis" - shows "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} - \ division_points ({a..b}) d" (is "?D1 \ ?D") - "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} - \ division_points ({a..b}) d" (is "?D2 \ ?D") -proof- have ab:"\i\Basis. a\i \ b\i" using assms(2) by(auto intro!:less_imp_le) - guess u v using division_ofD(4)[OF assms(1,5)] apply-by(erule exE)+ note l=this - have uv:"\i\Basis. u\i \ v\i" "\i\Basis. a\i \ u\i \ v\i \ b\i" +lemma division_points_psubset: + fixes a :: "'a::ordered_euclidean_space" + assumes "d division_of {a..b}" + and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" + and "l \ d" + and "interval_lowerbound l\k = c \ interval_upperbound l\k = c" + and k: "k \ Basis" + shows "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ + division_points ({a..b}) d" (is "?D1 \ ?D") + and "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ + division_points ({a..b}) d" (is "?D2 \ ?D") +proof - + have ab: "\i\Basis. a\i \ b\i" + using assms(2) by (auto intro!:less_imp_le) + guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this + have uv: "\i\Basis. u\i \ v\i" "\i\Basis. a\i \ u\i \ v\i \ b\i" using division_ofD(2,2,3)[OF assms(1,5)] unfolding l interval_ne_empty - unfolding subset_eq apply- defer apply(erule_tac x=u in ballE, erule_tac x=v in ballE) unfolding mem_interval by auto - have *:"interval_upperbound ({a..b} \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" - "interval_upperbound ({a..b} \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" - unfolding interval_split[OF k] apply(subst interval_bounds) prefer 3 apply(subst interval_bounds) - unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab k by auto - have "\x. x \ ?D - ?D1" using assms(2-) apply-apply(erule disjE) - apply(rule_tac x="(k,(interval_lowerbound l)\k)" in exI) defer - apply(rule_tac x="(k,(interval_upperbound l)\k)" in exI) - unfolding division_points_def unfolding interval_bounds[OF ab] by(auto simp add:*) - thus "?D1 \ ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) using k by auto - - have *:"interval_lowerbound ({a..b} \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" - "interval_lowerbound ({a..b} \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" - unfolding interval_split[OF k] apply(subst interval_bounds) prefer 3 apply(subst interval_bounds) - unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab k by auto - have "\x. x \ ?D - ?D2" using assms(2-) apply-apply(erule disjE) - apply(rule_tac x="(k,(interval_lowerbound l)\k)" in exI) defer - apply(rule_tac x="(k,(interval_upperbound l)\k)" in exI) - unfolding division_points_def unfolding interval_bounds[OF ab] by(auto simp add:*) - thus "?D2 \ ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4) k]) by auto qed + unfolding subset_eq + apply - + defer + apply (erule_tac x=u in ballE) + apply (erule_tac x=v in ballE) + unfolding mem_interval + apply auto + done + have *: "interval_upperbound ({a..b} \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" + "interval_upperbound ({a..b} \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" + unfolding interval_split[OF k] + apply (subst interval_bounds) + prefer 3 + apply (subst interval_bounds) + unfolding l interval_bounds[OF uv(1)] + using uv[rule_format,of k] ab k + apply auto + done + have "\x. x \ ?D - ?D1" + using assms(2-) + apply - + apply (erule disjE) + apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI) + defer + apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI) + unfolding division_points_def + unfolding interval_bounds[OF ab] + apply (auto simp add:*) + done + then show "?D1 \ ?D" + apply - + apply rule + apply (rule division_points_subset[OF assms(1-4)]) + using k + apply auto + done + + have *: "interval_lowerbound ({a..b} \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" + "interval_lowerbound ({a..b} \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" + unfolding interval_split[OF k] + apply (subst interval_bounds) + prefer 3 + apply (subst interval_bounds) + unfolding l interval_bounds[OF uv(1)] + using uv[rule_format,of k] ab k + apply auto + done + have "\x. x \ ?D - ?D2" + using assms(2-) + apply - + apply (erule disjE) + apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI) + defer + apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI) + unfolding division_points_def + unfolding interval_bounds[OF ab] + apply (auto simp add:*) + done + then show "?D2 \ ?D" + apply - + apply rule + apply (rule division_points_subset[OF assms(1-4) k]) + apply auto + done +qed + subsection {* Preservation by divisions and tagged divisions. *} @@ -4091,138 +4385,308 @@ lemma iterate_expand_cases: "iterate opp s f = (if finite(support opp f s) then iterate opp (support opp f s) f else neutral opp)" - apply(cases) apply(subst if_P,assumption) unfolding iterate_def support_support fold'_def by auto - -lemma iterate_image: assumes "monoidal opp" "inj_on f s" + apply cases + apply (subst if_P, assumption) + unfolding iterate_def support_support fold'_def + apply auto + done + +lemma iterate_image: + assumes "monoidal opp" + and "inj_on f s" shows "iterate opp (f ` s) g = iterate opp s (g \ f)" -proof- have *:"\s. finite s \ \x\s. \y\s. f x = f y \ x = y \ - iterate opp (f ` s) g = iterate opp s (g \ f)" - proof- case goal1 show ?case using goal1 - proof(induct s) case empty thus ?case using assms(1) by auto - next case (insert x s) show ?case unfolding iterate_insert[OF assms(1) insert(1)] - unfolding if_not_P[OF insert(2)] apply(subst insert(3)[symmetric]) - unfolding image_insert defer apply(subst iterate_insert[OF assms(1)]) - apply(rule finite_imageI insert)+ apply(subst if_not_P) - unfolding image_iff o_def using insert(2,4) by auto - qed qed +proof - + have *: "\s. finite s \ \x\s. \y\s. f x = f y \ x = y \ + iterate opp (f ` s) g = iterate opp s (g \ f)" + proof - + case goal1 + then show ?case + proof (induct s) + case empty + then show ?case + using assms(1) by auto + next + case (insert x s) + show ?case + unfolding iterate_insert[OF assms(1) insert(1)] + unfolding if_not_P[OF insert(2)] + apply (subst insert(3)[symmetric]) + unfolding image_insert + defer + apply (subst iterate_insert[OF assms(1)]) + apply (rule finite_imageI insert)+ + apply (subst if_not_P) + unfolding image_iff o_def + using insert(2,4) + apply auto + done + qed + qed show ?thesis - apply(cases "finite (support opp g (f ` s))") - apply(subst (1) iterate_support[symmetric],subst (2) iterate_support[symmetric]) - unfolding support_clauses apply(rule *)apply(rule finite_imageD,assumption) unfolding inj_on_def[symmetric] - apply(rule subset_inj_on[OF assms(2) support_subset])+ - apply(subst iterate_expand_cases) unfolding support_clauses apply(simp only: if_False) - apply(subst iterate_expand_cases) apply(subst if_not_P) by auto qed - - -(* This lemma about iterations comes up in a few places. *) + apply (cases "finite (support opp g (f ` s))") + apply (subst (1) iterate_support[symmetric],subst (2) iterate_support[symmetric]) + unfolding support_clauses + apply (rule *) + apply (rule finite_imageD,assumption) + unfolding inj_on_def[symmetric] + apply (rule subset_inj_on[OF assms(2) support_subset])+ + apply (subst iterate_expand_cases) + unfolding support_clauses + apply (simp only: if_False) + apply (subst iterate_expand_cases) + apply (subst if_not_P) + apply auto + done +qed + + +(* This lemma about iterations comes up in a few places. *) lemma iterate_nonzero_image_lemma: - assumes "monoidal opp" "finite s" "g(a) = neutral opp" - "\x\s. \y\s. f x = f y \ x \ y \ g(f x) = neutral opp" + assumes "monoidal opp" + and "finite s" "g(a) = neutral opp" + and "\x\s. \y\s. f x = f y \ x \ y \ g(f x) = neutral opp" shows "iterate opp {f x | x. x \ s \ f x \ a} g = iterate opp s (g \ f)" -proof- have *:"{f x |x. x \ s \ ~(f x = a)} = f ` {x. x \ s \ ~(f x = a)}" by auto - have **:"support opp (g \ f) {x \ s. f x \ a} = support opp (g \ f) s" +proof - + have *: "{f x |x. x \ s \ f x \ a} = f ` {x. x \ s \ f x \ a}" + by auto + have **: "support opp (g \ f) {x \ s. f x \ a} = support opp (g \ f) s" unfolding support_def using assms(3) by auto - show ?thesis unfolding * - apply(subst iterate_support[symmetric]) unfolding support_clauses - apply(subst iterate_image[OF assms(1)]) defer - apply(subst(2) iterate_support[symmetric]) apply(subst **) - unfolding inj_on_def using assms(3,4) unfolding support_def by auto qed + show ?thesis + unfolding * + apply (subst iterate_support[symmetric]) + unfolding support_clauses + apply (subst iterate_image[OF assms(1)]) + defer + apply (subst(2) iterate_support[symmetric]) + apply (subst **) + unfolding inj_on_def + using assms(3,4) + unfolding support_def + apply auto + done +qed lemma iterate_eq_neutral: - assumes "monoidal opp" "\x \ s. (f(x) = neutral opp)" - shows "(iterate opp s f = neutral opp)" -proof- have *:"support opp f s = {}" unfolding support_def using assms(2) by auto - show ?thesis apply(subst iterate_support[symmetric]) - unfolding * using assms(1) by auto qed - -lemma iterate_op: assumes "monoidal opp" "finite s" - shows "iterate opp s (\x. opp (f x) (g x)) = opp (iterate opp s f) (iterate opp s g)" using assms(2) -proof(induct s) case empty thus ?case unfolding iterate_insert[OF assms(1)] using assms(1) by auto -next case (insert x F) show ?case unfolding iterate_insert[OF assms(1) insert(1)] if_not_P[OF insert(2)] insert(3) - unfolding monoidal_ac[OF assms(1)] by(rule refl) qed - -lemma iterate_eq: assumes "monoidal opp" "\x. x \ s \ f x = g x" + assumes "monoidal opp" + and "\x \ s. f x = neutral opp" + shows "iterate opp s f = neutral opp" +proof - + have *: "support opp f s = {}" + unfolding support_def using assms(2) by auto + show ?thesis + apply (subst iterate_support[symmetric]) + unfolding * + using assms(1) + apply auto + done +qed + +lemma iterate_op: + assumes "monoidal opp" + and "finite s" + shows "iterate opp s (\x. opp (f x) (g x)) = opp (iterate opp s f) (iterate opp s g)" + using assms(2) +proof (induct s) + case empty + then show ?case + unfolding iterate_insert[OF assms(1)] using assms(1) by auto +next + case (insert x F) + show ?case + unfolding iterate_insert[OF assms(1) insert(1)] if_not_P[OF insert(2)] insert(3) + by (simp add: monoidal_ac[OF assms(1)]) +qed + +lemma iterate_eq: + assumes "monoidal opp" + and "\x. x \ s \ f x = g x" shows "iterate opp s f = iterate opp s g" -proof- have *:"support opp g s = support opp f s" +proof - + have *: "support opp g s = support opp f s" unfolding support_def using assms(2) by auto show ?thesis - proof(cases "finite (support opp f s)") - case False thus ?thesis apply(subst iterate_expand_cases,subst(2) iterate_expand_cases) - unfolding * by auto - next def su \ "support opp f s" + proof (cases "finite (support opp f s)") + case False + then show ?thesis + apply (subst iterate_expand_cases) + apply (subst(2) iterate_expand_cases) + unfolding * + apply auto + done + next + def su \ "support opp f s" case True note support_subset[of opp f s] - thus ?thesis apply- apply(subst iterate_support[symmetric],subst(2) iterate_support[symmetric]) unfolding * using True + then show ?thesis + apply - + apply (subst iterate_support[symmetric]) + apply (subst(2) iterate_support[symmetric]) + unfolding * + using True unfolding su_def[symmetric] - proof(induct su) case empty show ?case by auto - next case (insert x s) show ?case unfolding iterate_insert[OF assms(1) insert(1)] - unfolding if_not_P[OF insert(2)] apply(subst insert(3)) - defer apply(subst assms(2)[of x]) using insert by auto qed qed qed - -lemma nonempty_witness: assumes "s \ {}" obtains x where "x \ s" using assms by auto - -lemma operative_division: fixes f::"('a::ordered_euclidean_space) set \ 'b" - assumes "monoidal opp" "operative opp f" "d division_of {a..b}" + proof (induct su) + case empty + show ?case by auto + next + case (insert x s) + show ?case + unfolding iterate_insert[OF assms(1) insert(1)] + unfolding if_not_P[OF insert(2)] + apply (subst insert(3)) + defer + apply (subst assms(2)[of x]) + using insert + apply auto + done + qed + qed +qed + +lemma nonempty_witness: + assumes "s \ {}" + obtains x where "x \ s" + using assms by auto + +lemma operative_division: + fixes f :: "'a::ordered_euclidean_space set \ 'b" + assumes "monoidal opp" + and "operative opp f" + and "d division_of {a..b}" shows "iterate opp d f = f {a..b}" -proof- def C \ "card (division_points {a..b} d)" thus ?thesis using assms - proof(induct C arbitrary:a b d rule:full_nat_induct) +proof - + def C \ "card (division_points {a..b} d)" + then show ?thesis + using assms + proof (induct C arbitrary: a b d rule: full_nat_induct) case goal1 - { presume *:"content {a..b} \ 0 \ ?case" - thus ?case apply-apply(cases) defer apply assumption - proof- assume as:"content {a..b} = 0" - show ?case unfolding operativeD(1)[OF assms(2) as] apply(rule iterate_eq_neutral[OF goal1(2)]) - proof fix x assume x:"x\d" - then guess u v apply(drule_tac division_ofD(4)[OF goal1(4)]) by(erule exE)+ - thus "f x = neutral opp" using division_of_content_0[OF as goal1(4)] - using operativeD(1)[OF assms(2)] x by auto - qed qed } + { presume *: "content {a..b} \ 0 \ ?case" + then show ?case + apply - + apply cases + defer + apply assumption + proof - + assume as: "content {a..b} = 0" + show ?case + unfolding operativeD(1)[OF assms(2) as] + apply(rule iterate_eq_neutral[OF goal1(2)]) + proof + fix x + assume x: "x \ d" + then guess u v + apply (drule_tac division_ofD(4)[OF goal1(4)]) + apply (elim exE) + done + then show "f x = neutral opp" + using division_of_content_0[OF as goal1(4)] + using operativeD(1)[OF assms(2)] x + by auto + qed + qed + } assume "content {a..b} \ 0" note ab = this[unfolded content_lt_nz[symmetric] content_pos_lt_eq] - hence ab':"\i\Basis. a\i \ b\i" by (auto intro!: less_imp_le) show ?case - proof(cases "division_points {a..b} d = {}") - case True have d':"\i\d. \u v. i = {u..v} \ + then have ab': "\i\Basis. a\i \ b\i" + by (auto intro!: less_imp_le) + show ?case + proof (cases "division_points {a..b} d = {}") + case True + have d': "\i\d. \u v. i = {u..v} \ (\j\Basis. u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j)" - unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule) - apply(rule_tac x=a in exI,rule_tac x=b in exI) apply(rule,rule refl) + unfolding forall_in_division[OF goal1(4)] + apply rule + apply rule + apply rule + apply (rule_tac x=a in exI) + apply (rule_tac x=b in exI) + apply rule + apply (rule refl) proof - fix u v and j :: 'a assume j:"j\Basis" assume as:"{u..v} \ d" note division_ofD(3)[OF goal1(4) this] - hence uv:"\i\Basis. u\i \ v\i" "u\j \ v\j" using j unfolding interval_ne_empty by auto - have *:"\p r Q. \ j\Basis \ p \ r \ (\x\d. Q x) \ p \ r \ (Q {u..v})" using as j by auto + fix u v + fix j :: 'a + assume j: "j \ Basis" + assume as: "{u..v} \ d" + note division_ofD(3)[OF goal1(4) this] + then have uv: "\i\Basis. u\i \ v\i" "u\j \ v\j" + using j unfolding interval_ne_empty by auto + have *: "\p r Q. \ j\Basis \ p \ r \ (\x\d. Q x) \ p \ r \ Q {u..v}" + using as j by auto have "(j, u\j) \ division_points {a..b} d" "(j, v\j) \ division_points {a..b} d" using True by auto note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps] note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]] - moreover have "a\j \ u\j" "v\j \ b\j" using division_ofD(2,2,3)[OF goal1(4) as] - unfolding subset_eq apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE) - unfolding interval_ne_empty mem_interval using j by auto + moreover + have "a\j \ u\j" "v\j \ b\j" + using division_ofD(2,2,3)[OF goal1(4) as] + unfolding subset_eq + apply - + apply (erule_tac x=u in ballE,erule_tac[3] x=v in ballE) + unfolding interval_ne_empty mem_interval + using j + apply auto + done ultimately show "u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j" unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by auto qed have "(1/2) *\<^sub>R (a+b) \ {a..b}" unfolding mem_interval using ab by(auto intro!: less_imp_le simp: inner_simps) note this[unfolded division_ofD(6)[OF goal1(4),symmetric] Union_iff] - then guess i .. note i=this guess u v using d'[rule_format,OF i(1)] apply-by(erule exE conjE)+ note uv=this + then guess i .. note i=this + guess u v using d'[rule_format,OF i(1)] by (elim exE conjE) note uv=this have "{a..b} \ d" - proof- { presume "i = {a..b}" thus ?thesis using i by auto } - { presume "u = a" "v = b" thus "i = {a..b}" using uv by auto } - show "u = a" "v = b" unfolding euclidean_eq_iff[where 'a='a] - proof(safe) - fix j :: 'a assume j:"j\Basis" + proof - + { presume "i = {a..b}" then show ?thesis using i by auto } + { presume "u = a" "v = b" then show "i = {a..b}" using uv by auto } + show "u = a" "v = b" + unfolding euclidean_eq_iff[where 'a='a] + proof safe + fix j :: 'a + assume j: "j \ Basis" note i(2)[unfolded uv mem_interval,rule_format,of j] - thus "u \ j = a \ j" "v \ j = b \ j" using uv(2)[rule_format,of j] j by (auto simp: inner_simps) - qed qed - hence *:"d = insert {a..b} (d - {{a..b}})" by auto - have "iterate opp (d - {{a..b}}) f = neutral opp" apply(rule iterate_eq_neutral[OF goal1(2)]) - proof fix x assume x:"x \ d - {{a..b}}" hence "x\d" by auto note d'[rule_format,OF this] - then guess u v apply-by(erule exE conjE)+ note uv=this - have "u\a \ v\b" using x[unfolded uv] by auto - then obtain j where "u\j \ a\j \ v\j \ b\j" and j:"j\Basis" unfolding euclidean_eq_iff[where 'a='a] by auto - hence "u\j = v\j" using uv(2)[rule_format,OF j] by auto - hence "content {u..v} = 0" unfolding content_eq_0 apply(rule_tac x=j in bexI) using j by auto - thus "f x = neutral opp" unfolding uv(1) by(rule operativeD(1)[OF goal1(3)]) - qed thus "iterate opp d f = f {a..b}" apply-apply(subst *) - apply(subst iterate_insert[OF goal1(2)]) using goal1(2,4) by auto - next case False hence "\x. x\division_points {a..b} d" by auto - then guess k c unfolding split_paired_Ex apply- unfolding division_points_def mem_Collect_eq split_conv - by(erule exE conjE)+ note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']] + then show "u \ j = a \ j" and "v \ j = b \ j" + using uv(2)[rule_format,of j] j by (auto simp: inner_simps) + qed + qed + then have *: "d = insert {a..b} (d - {{a..b}})" + by auto + have "iterate opp (d - {{a..b}}) f = neutral opp" + apply (rule iterate_eq_neutral[OF goal1(2)]) + proof + fix x + assume x: "x \ d - {{a..b}}" + then have "x\d" + by auto note d'[rule_format,OF this] + then guess u v by (elim exE conjE) note uv=this + have "u \ a \ v \ b" + using x[unfolded uv] by auto + then obtain j where "u\j \ a\j \ v\j \ b\j" and j: "j \ Basis" + unfolding euclidean_eq_iff[where 'a='a] by auto + then have "u\j = v\j" + using uv(2)[rule_format,OF j] by auto + then have "content {u..v} = 0" + unfolding content_eq_0 + apply (rule_tac x=j in bexI) + using j + apply auto + done + then show "f x = neutral opp" + unfolding uv(1) by (rule operativeD(1)[OF goal1(3)]) + qed + then show "iterate opp d f = f {a..b}" + apply - + apply (subst *) + apply (subst iterate_insert[OF goal1(2)]) + using goal1(2,4) + apply auto + done + next + case False + then have "\x. x \ division_points {a..b} d" + by auto + then guess k c + unfolding split_paired_Ex + unfolding division_points_def mem_Collect_eq split_conv + apply (elim exE conjE) + done + note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']] from this(3) guess j .. note j=this def d1 \ "{l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" def d2 \ "{l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" @@ -4230,354 +4694,730 @@ def ca \ "(\i\Basis. (if i = k then c else a\i) *\<^sub>R i)::'a" note division_points_psubset[OF goal1(4) ab kc(1-2) j] note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)] - hence *:"(iterate opp d1 f) = f ({a..b} \ {x. x\k \ c})" "(iterate opp d2 f) = f ({a..b} \ {x. x\k \ c})" - apply- unfolding interval_split[OF kc(4)] apply(rule_tac[!] goal1(1)[rule_format]) + then have *: "(iterate opp d1 f) = f ({a..b} \ {x. x\k \ c})" + "(iterate opp d2 f) = f ({a..b} \ {x. x\k \ c})" + unfolding interval_split[OF kc(4)] + apply (rule_tac[!] goal1(1)[rule_format]) using division_split[OF goal1(4), where k=k and c=c] - unfolding interval_split[OF kc(4)] d1_def[symmetric] d2_def[symmetric] unfolding goal1(2) Suc_le_mono - using goal1(2-3) using division_points_finite[OF goal1(4)] using kc(4) by auto + unfolding interval_split[OF kc(4)] d1_def[symmetric] d2_def[symmetric] + unfolding goal1(2) Suc_le_mono + using goal1(2-3) + using division_points_finite[OF goal1(4)] + using kc(4) + apply auto + done have "f {a..b} = opp (iterate opp d1 f) (iterate opp d2 f)" (is "_ = ?prev") - unfolding * apply(rule operativeD(2)) using goal1(3) using kc(4) by auto + unfolding * + apply (rule operativeD(2)) + using goal1(3) + using kc(4) + apply auto + done also have "iterate opp d1 f = iterate opp d (\l. f(l \ {x. x\k \ c}))" - unfolding d1_def apply(rule iterate_nonzero_image_lemma[unfolded o_def]) - unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+ - unfolding empty_as_interval[symmetric] apply(rule content_empty) - proof(rule,rule,rule,erule conjE) fix l y assume as:"l \ d" "y \ d" "l \ {x. x \ k \ c} = y \ {x. x \ k \ c}" "l \ y" - from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this - show "f (l \ {x. x \ k \ c}) = neutral opp" unfolding l interval_split[OF kc(4)] - apply(rule operativeD(1) goal1)+ unfolding interval_split[symmetric,OF kc(4)] apply(rule division_split_left_inj) - apply(rule goal1) unfolding l[symmetric] apply(rule as(1),rule as(2)) by(rule kc(4) as)+ - qed also have "iterate opp d2 f = iterate opp d (\l. f(l \ {x. x\k \ c}))" - unfolding d2_def apply(rule iterate_nonzero_image_lemma[unfolded o_def]) - unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+ - unfolding empty_as_interval[symmetric] apply(rule content_empty) - proof(rule,rule,rule,erule conjE) fix l y assume as:"l \ d" "y \ d" "l \ {x. c \ x \ k} = y \ {x. c \ x \ k}" "l \ y" - from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this - show "f (l \ {x. x \ k \ c}) = neutral opp" unfolding l interval_split[OF kc(4)] - apply(rule operativeD(1) goal1)+ unfolding interval_split[symmetric,OF kc(4)] apply(rule division_split_right_inj) - apply(rule goal1) unfolding l[symmetric] apply(rule as(1),rule as(2)) by(rule as kc(4))+ - qed also have *:"\x\d. f x = opp (f (x \ {x. x \ k \ c})) (f (x \ {x. c \ x \ k}))" - unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule,rule operativeD(2)) using goal1(3) kc by auto - have "opp (iterate opp d (\l. f (l \ {x. x \ k \ c}))) (iterate opp d (\l. f (l \ {x. c \ x \ k}))) - = iterate opp d f" apply(subst(3) iterate_eq[OF _ *[rule_format]]) prefer 3 - apply(rule iterate_op[symmetric]) using goal1 by auto + unfolding d1_def + apply (rule iterate_nonzero_image_lemma[unfolded o_def]) + unfolding empty_as_interval + apply (rule goal1 division_of_finite operativeD[OF goal1(3)])+ + unfolding empty_as_interval[symmetric] + apply (rule content_empty) + proof (rule, rule, rule, erule conjE) + fix l y + assume as: "l \ d" "y \ d" "l \ {x. x \ k \ c} = y \ {x. x \ k \ c}" "l \ y" + from division_ofD(4)[OF goal1(4) this(1)] guess u v by (elim exE) note l=this + show "f (l \ {x. x \ k \ c}) = neutral opp" + unfolding l interval_split[OF kc(4)] + apply (rule operativeD(1) goal1)+ + unfolding interval_split[symmetric,OF kc(4)] + apply (rule division_split_left_inj) + apply (rule goal1) + unfolding l[symmetric] + apply (rule as(1), rule as(2)) + apply (rule kc(4) as)+ + done + qed + also have "iterate opp d2 f = iterate opp d (\l. f(l \ {x. x\k \ c}))" + unfolding d2_def + apply (rule iterate_nonzero_image_lemma[unfolded o_def]) + unfolding empty_as_interval + apply (rule goal1 division_of_finite operativeD[OF goal1(3)])+ + unfolding empty_as_interval[symmetric] + apply (rule content_empty) + proof (rule, rule, rule, erule conjE) + fix l y + assume as: "l \ d" "y \ d" "l \ {x. c \ x \ k} = y \ {x. c \ x \ k}" "l \ y" + from division_ofD(4)[OF goal1(4) this(1)] guess u v by (elim exE) note l=this + show "f (l \ {x. x \ k \ c}) = neutral opp" + unfolding l interval_split[OF kc(4)] + apply (rule operativeD(1) goal1)+ + unfolding interval_split[symmetric,OF kc(4)] + apply (rule division_split_right_inj) + apply (rule goal1) + unfolding l[symmetric] + apply (rule as(1)) + apply (rule as(2)) + apply (rule as kc(4))+ + done + qed also have *: "\x\d. f x = opp (f (x \ {x. x \ k \ c})) (f (x \ {x. c \ x \ k}))" + unfolding forall_in_division[OF goal1(4)] + apply (rule, rule, rule, rule operativeD(2)) + using goal1(3) kc + by auto + have "opp (iterate opp d (\l. f (l \ {x. x \ k \ c}))) (iterate opp d (\l. f (l \ {x. c \ x \ k}))) = + iterate opp d f" + apply (subst(3) iterate_eq[OF _ *[rule_format]]) + prefer 3 + apply (rule iterate_op[symmetric]) + using goal1 + apply auto + done finally show ?thesis by auto - qed qed qed - -lemma iterate_image_nonzero: assumes "monoidal opp" - "finite s" "\x\s. \y\s. ~(x = y) \ f x = f y \ g(f x) = neutral opp" - shows "iterate opp (f ` s) g = iterate opp s (g \ f)" using assms -proof(induct rule:finite_subset_induct[OF assms(2) subset_refl]) - case goal1 show ?case using assms(1) by auto -next case goal2 have *:"\x y. y = neutral opp \ x = opp y x" using assms(1) by auto - show ?case unfolding image_insert apply(subst iterate_insert[OF assms(1)]) - apply(rule finite_imageI goal2)+ - apply(cases "f a \ f ` F") unfolding if_P if_not_P apply(subst goal2(4)[OF assms(1) goal2(1)]) defer - apply(subst iterate_insert[OF assms(1) goal2(1)]) defer - apply(subst iterate_insert[OF assms(1) goal2(1)]) - unfolding if_not_P[OF goal2(3)] defer unfolding image_iff defer apply(erule bexE) - apply(rule *) unfolding o_def apply(rule_tac y=x in goal2(7)[rule_format]) - using goal2 unfolding o_def by auto qed - -lemma operative_tagged_division: assumes "monoidal opp" "operative opp f" "d tagged_division_of {a..b}" - shows "iterate(opp) d (\(x,l). f l) = f {a..b}" -proof- have *:"(\(x,l). f l) = (f o snd)" unfolding o_def by(rule,auto) note assm = tagged_division_ofD[OF assms(3)] - have "iterate(opp) d (\(x,l). f l) = iterate opp (snd ` d) f" unfolding * - apply(rule iterate_image_nonzero[symmetric,OF assms(1)]) apply(rule tagged_division_of_finite assms)+ - unfolding Ball_def split_paired_All snd_conv apply(rule,rule,rule,rule,rule,rule,rule,erule conjE) - proof- fix a b aa ba assume as:"(a, b) \ d" "(aa, ba) \ d" "(a, b) \ (aa, ba)" "b = ba" - guess u v using assm(4)[OF as(1)] apply-by(erule exE)+ note uv=this - show "f b = neutral opp" unfolding uv apply(rule operativeD(1)[OF assms(2)]) - unfolding content_eq_0_interior using tagged_division_ofD(5)[OF assms(3) as(1-3)] - unfolding as(4)[symmetric] uv by auto - qed also have "\ = f {a..b}" + qed + qed +qed + +lemma iterate_image_nonzero: + assumes "monoidal opp" + and "finite s" + and "\x\s. \y\s. x \ y \ f x = f y \ g (f x) = neutral opp" + shows "iterate opp (f ` s) g = iterate opp s (g \ f)" + using assms +proof (induct rule: finite_subset_induct[OF assms(2) subset_refl]) + case goal1 + show ?case + using assms(1) by auto +next + case goal2 + have *: "\x y. y = neutral opp \ x = opp y x" + using assms(1) by auto + show ?case + unfolding image_insert + apply (subst iterate_insert[OF assms(1)]) + apply (rule finite_imageI goal2)+ + apply (cases "f a \ f ` F") + unfolding if_P if_not_P + apply (subst goal2(4)[OF assms(1) goal2(1)]) + defer + apply (subst iterate_insert[OF assms(1) goal2(1)]) + defer + apply (subst iterate_insert[OF assms(1) goal2(1)]) + unfolding if_not_P[OF goal2(3)] + defer unfolding image_iff + defer + apply (erule bexE) + apply (rule *) + unfolding o_def + apply (rule_tac y=x in goal2(7)[rule_format]) + using goal2 + unfolding o_def + apply auto + done +qed + +lemma operative_tagged_division: + assumes "monoidal opp" + and "operative opp f" + and "d tagged_division_of {a..b}" + shows "iterate opp d (\(x,l). f l) = f {a..b}" +proof - + have *: "(\(x,l). f l) = f \ snd" + unfolding o_def by rule auto note assm = tagged_division_ofD[OF assms(3)] + have "iterate opp d (\(x,l). f l) = iterate opp (snd ` d) f" + unfolding * + apply (rule iterate_image_nonzero[symmetric,OF assms(1)]) + apply (rule tagged_division_of_finite assms)+ + unfolding Ball_def split_paired_All snd_conv + apply (rule, rule, rule, rule, rule, rule, rule, erule conjE) + proof - + fix a b aa ba + assume as: "(a, b) \ d" "(aa, ba) \ d" "(a, b) \ (aa, ba)" "b = ba" + guess u v using assm(4)[OF as(1)] by (elim exE) note uv=this + show "f b = neutral opp" + unfolding uv + apply (rule operativeD(1)[OF assms(2)]) + unfolding content_eq_0_interior + using tagged_division_ofD(5)[OF assms(3) as(1-3)] + unfolding as(4)[symmetric] uv + apply auto + done + qed + also have "\ = f {a..b}" using operative_division[OF assms(1-2) division_of_tagged_division[OF assms(3)]] . - finally show ?thesis . qed + finally show ?thesis . +qed + subsection {* Additivity of content. *} lemma setsum_iterate: - assumes "finite s" shows "setsum f s = iterate op + s f" + assumes "finite s" + shows "setsum f s = iterate op + s f" proof - have *: "setsum f s = setsum f (support op + f s)" apply (rule setsum_mono_zero_right) - unfolding support_def neutral_monoid using assms by auto + unfolding support_def neutral_monoid + using assms + apply auto + done then show ?thesis unfolding * iterate_def fold'_def setsum.eq_fold unfolding neutral_monoid by (simp add: comp_def) qed -lemma additive_content_division: assumes "d division_of {a..b}" - shows "setsum content d = content({a..b})" +lemma additive_content_division: + assumes "d division_of {a..b}" + shows "setsum content d = content {a..b}" unfolding operative_division[OF monoidal_monoid operative_content assms,symmetric] - apply(subst setsum_iterate) using assms by auto + apply (subst setsum_iterate) + using assms + apply auto + done lemma additive_content_tagged_division: assumes "d tagged_division_of {a..b}" - shows "setsum (\(x,l). content l) d = content({a..b})" + shows "setsum (\(x,l). content l) d = content {a..b}" unfolding operative_tagged_division[OF monoidal_monoid operative_content assms,symmetric] - apply(subst setsum_iterate) using assms by auto + apply (subst setsum_iterate) + using assms + apply auto + done + subsection {* Finally, the integral of a constant *} lemma has_integral_const[intro]: - "((\x. c) has_integral (content({a..b::'a::ordered_euclidean_space}) *\<^sub>R c)) ({a..b})" - unfolding has_integral apply(rule,rule,rule_tac x="\x. ball x 1" in exI) - apply(rule,rule gauge_trivial)apply(rule,rule,erule conjE) - unfolding split_def apply(subst scaleR_left.setsum[symmetric, unfolded o_def]) - defer apply(subst additive_content_tagged_division[unfolded split_def]) apply assumption by auto + fixes a b :: "'a::ordered_euclidean_space" + shows "((\x. c) has_integral (content {a..b} *\<^sub>R c)) {a..b}" + unfolding has_integral + apply rule + apply rule + apply (rule_tac x="\x. ball x 1" in exI) + apply rule + apply (rule gauge_trivial) + apply rule + apply rule + apply (erule conjE) + unfolding split_def + apply (subst scaleR_left.setsum[symmetric, unfolded o_def]) + defer + apply (subst additive_content_tagged_division[unfolded split_def]) + apply assumption + apply auto + done lemma integral_const[simp]: fixes a b :: "'a::ordered_euclidean_space" shows "integral {a .. b} (\x. c) = content {a .. b} *\<^sub>R c" by (rule integral_unique) (rule has_integral_const) + subsection {* Bounds on the norm of Riemann sums and the integral itself. *} -lemma dsum_bound: assumes "p division_of {a..b}" "norm(c) \ e" - shows "norm(setsum (\l. content l *\<^sub>R c) p) \ e * content({a..b})" (is "?l \ ?r") - apply(rule order_trans,rule norm_setsum) unfolding norm_scaleR setsum_left_distrib[symmetric] - apply(rule order_trans[OF mult_left_mono],rule assms,rule setsum_abs_ge_zero) - apply(subst mult_commute) apply(rule mult_left_mono) - apply(rule order_trans[of _ "setsum content p"]) apply(rule eq_refl,rule setsum_cong2) - apply(subst abs_of_nonneg) unfolding additive_content_division[OF assms(1)] -proof- from order_trans[OF norm_ge_zero[of c] assms(2)] show "0 \ e" . - fix x assume "x\p" from division_ofD(4)[OF assms(1) this] guess u v apply-by(erule exE)+ - thus "0 \ content x" using content_pos_le by auto -qed(insert assms,auto) - -lemma rsum_bound: assumes "p tagged_division_of {a..b}" "\x\{a..b}. norm(f x) \ e" - shows "norm(setsum (\(x,k). content k *\<^sub>R f x) p) \ e * content({a..b})" -proof(cases "{a..b} = {}") case True - show ?thesis using assms(1) unfolding True tagged_division_of_trivial by auto -next case False show ?thesis - apply(rule order_trans,rule norm_setsum) unfolding split_def norm_scaleR - apply(rule order_trans[OF setsum_mono]) apply(rule mult_left_mono[OF _ abs_ge_zero, of _ e]) defer - unfolding setsum_left_distrib[symmetric] apply(subst mult_commute) apply(rule mult_left_mono) - apply(rule order_trans[of _ "setsum (content \ snd) p"]) apply(rule eq_refl,rule setsum_cong2) - apply(subst o_def, rule abs_of_nonneg) - proof- show "setsum (content \ snd) p \ content {a..b}" apply(rule eq_refl) - unfolding additive_content_tagged_division[OF assms(1),symmetric] split_def by auto +lemma dsum_bound: + assumes "p division_of {a..b}" + and "norm c \ e" + shows "norm (setsum (\l. content l *\<^sub>R c) p) \ e * content({a..b})" + apply (rule order_trans) + apply (rule norm_setsum) + unfolding norm_scaleR setsum_left_distrib[symmetric] + apply (rule order_trans[OF mult_left_mono]) + apply (rule assms) + apply (rule setsum_abs_ge_zero) + apply (subst mult_commute) + apply (rule mult_left_mono) + apply (rule order_trans[of _ "setsum content p"]) + apply (rule eq_refl) + apply (rule setsum_cong2) + apply (subst abs_of_nonneg) + unfolding additive_content_division[OF assms(1)] +proof - + from order_trans[OF norm_ge_zero[of c] assms(2)] + show "0 \ e" . + fix x assume "x \ p" + from division_ofD(4)[OF assms(1) this] guess u v by (elim exE) + then show "0 \ content x" + using content_pos_le by auto +qed (insert assms, auto) + +lemma rsum_bound: + assumes "p tagged_division_of {a..b}" + and "\x\{a..b}. norm (f x) \ e" + shows "norm (setsum (\(x,k). content k *\<^sub>R f x) p) \ e * content {a..b}" +proof (cases "{a..b} = {}") + case True + show ?thesis + using assms(1) unfolding True tagged_division_of_trivial by auto +next + case False + show ?thesis + apply (rule order_trans) + apply (rule norm_setsum) + unfolding split_def norm_scaleR + apply (rule order_trans[OF setsum_mono]) + apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e]) + defer + unfolding setsum_left_distrib[symmetric] + apply (subst mult_commute) + apply (rule mult_left_mono) + apply (rule order_trans[of _ "setsum (content \ snd) p"]) + apply (rule eq_refl) + apply (rule setsum_cong2) + apply (subst o_def) + apply (rule abs_of_nonneg) + proof - + show "setsum (content \ snd) p \ content {a..b}" + apply (rule eq_refl) + unfolding additive_content_tagged_division[OF assms(1),symmetric] split_def + apply auto + done guess w using nonempty_witness[OF False] . - thus "e\0" apply-apply(rule order_trans) defer apply(rule assms(2)[rule_format],assumption) by auto - fix xk assume *:"xk\p" guess x k using surj_pair[of xk] apply-by(erule exE)+ note xk = this *[unfolded this] - from tagged_division_ofD(4)[OF assms(1) xk(2)] guess u v apply-by(erule exE)+ note uv=this - show "0\ content (snd xk)" unfolding xk snd_conv uv by(rule content_pos_le) - show "norm (f (fst xk)) \ e" unfolding xk fst_conv using tagged_division_ofD(2,3)[OF assms(1) xk(2)] assms(2) by auto - qed qed + then show "e \ 0" + apply - + apply (rule order_trans) + defer + apply (rule assms(2)[rule_format]) + apply assumption + apply auto + done + fix xk + assume *: "xk \ p" + guess x k using surj_pair[of xk] by (elim exE) note xk = this *[unfolded this] + from tagged_division_ofD(4)[OF assms(1) xk(2)] guess u v by (elim exE) note uv=this + show "0 \ content (snd xk)" + unfolding xk snd_conv uv by(rule content_pos_le) + show "norm (f (fst xk)) \ e" + unfolding xk fst_conv using tagged_division_ofD(2,3)[OF assms(1) xk(2)] assms(2) by auto + qed +qed lemma rsum_diff_bound: - assumes "p tagged_division_of {a..b}" "\x\{a..b}. norm(f x - g x) \ e" - shows "norm(setsum (\(x,k). content k *\<^sub>R f x) p - setsum (\(x,k). content k *\<^sub>R g x) p) \ e * content({a..b})" - apply(rule order_trans[OF _ rsum_bound[OF assms]]) apply(rule eq_refl) apply(rule arg_cong[where f=norm]) - unfolding setsum_subtractf[symmetric] apply(rule setsum_cong2) unfolding scaleR_diff_right by auto - -lemma has_integral_bound: fixes f::"'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "0 \ B" "(f has_integral i) ({a..b})" "\x\{a..b}. norm(f x) \ B" + assumes "p tagged_division_of {a..b}" + and "\x\{a..b}. norm (f x - g x) \ e" + shows "norm (setsum (\(x,k). content k *\<^sub>R f x) p - setsum (\(x,k). content k *\<^sub>R g x) p) \ + e * content {a..b}" + apply (rule order_trans[OF _ rsum_bound[OF assms]]) + apply (rule eq_refl) + apply (rule arg_cong[where f=norm]) + unfolding setsum_subtractf[symmetric] + apply (rule setsum_cong2) + unfolding scaleR_diff_right + apply auto + done + +lemma has_integral_bound: + fixes f :: "'a::ordered_euclidean_space \ 'b::real_normed_vector" + assumes "0 \ B" + and "(f has_integral i) {a..b}" + and "\x\{a..b}. norm (f x) \ B" shows "norm i \ B * content {a..b}" -proof- let ?P = "content {a..b} > 0" { presume "?P \ ?thesis" - thus ?thesis proof(cases ?P) case False - hence *:"content {a..b} = 0" using content_lt_nz by auto - hence **:"i = 0" using assms(2) apply(subst has_integral_null_eq[symmetric]) by auto - show ?thesis unfolding * ** using assms(1) by auto - qed auto } assume ab:?P - { presume "\ ?thesis \ False" thus ?thesis by auto } - assume "\ ?thesis" hence *:"norm i - B * content {a..b} > 0" by auto - from assms(2)[unfolded has_integral,rule_format,OF *] guess d apply-by(erule exE conjE)+ note d=this[rule_format] +proof - + let ?P = "content {a..b} > 0" + { + presume "?P \ ?thesis" + then show ?thesis + proof (cases ?P) + case False + then have *: "content {a..b} = 0" + using content_lt_nz by auto + hence **: "i = 0" + using assms(2) + apply (subst has_integral_null_eq[symmetric]) + apply auto + done + show ?thesis + unfolding * ** using assms(1) by auto + qed auto + } + assume ab: ?P + { presume "\ ?thesis \ False" then show ?thesis by auto } + assume "\ ?thesis" + then have *: "norm i - B * content {a..b} > 0" + by auto + from assms(2)[unfolded has_integral,rule_format,OF *] + guess d by (elim exE conjE) note d=this[rule_format] from fine_division_exists[OF this(1), of a b] guess p . note p=this - have *:"\s B. norm s \ B \ \ (norm (s - i) < norm i - B)" - proof- case goal1 thus ?case unfolding not_less - using norm_triangle_sub[of i s] unfolding norm_minus_commute by auto - qed show False using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto qed + have *: "\s B. norm s \ B \ \ norm (s - i) < norm i - B" + proof - + case goal1 + then show ?case + unfolding not_less + using norm_triangle_sub[of i s] + unfolding norm_minus_commute + by auto + qed + show False + using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto +qed + subsection {* Similar theorems about relationship among components. *} -lemma rsum_component_le: fixes f::"'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "p tagged_division_of {a..b}" "\x\{a..b}. (f x)\i \ (g x)\i" +lemma rsum_component_le: + fixes f :: "'a::ordered_euclidean_space \ 'b::euclidean_space" + assumes "p tagged_division_of {a..b}" + and "\x\{a..b}. (f x)\i \ (g x)\i" shows "(setsum (\(x,k). content k *\<^sub>R f x) p)\i \ (setsum (\(x,k). content k *\<^sub>R g x) p)\i" - unfolding inner_setsum_left apply(rule setsum_mono) apply safe -proof- fix a b assume ab:"(a,b) \ p" note assm = tagged_division_ofD(2-4)[OF assms(1) ab] - from this(3) guess u v apply-by(erule exE)+ note b=this - show "(content b *\<^sub>R f a) \ i \ (content b *\<^sub>R g a) \ i" unfolding b - unfolding inner_simps real_scaleR_def apply(rule mult_left_mono) - defer apply(rule content_pos_le,rule assms(2)[rule_format]) using assm by auto qed + unfolding inner_setsum_left + apply (rule setsum_mono) + apply safe +proof - + fix a b + assume ab: "(a, b) \ p" + note assm = tagged_division_ofD(2-4)[OF assms(1) ab] + from this(3) guess u v by (elim exE) note b=this + show "(content b *\<^sub>R f a) \ i \ (content b *\<^sub>R g a) \ i" + unfolding b + unfolding inner_simps real_scaleR_def + apply (rule mult_left_mono) + defer + apply (rule content_pos_le,rule assms(2)[rule_format]) + using assm + apply auto + done +qed lemma has_integral_component_le: - fixes f g::"'a::ordered_euclidean_space \ 'b::euclidean_space" + fixes f g :: "'a::ordered_euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" - assumes "(f has_integral i) s" "(g has_integral j) s" "\x\s. (f x)\k \ (g x)\k" + assumes "(f has_integral i) s" "(g has_integral j) s" + and "\x\s. (f x)\k \ (g x)\k" shows "i\k \ j\k" proof - have lem:"\a b i (j::'b). \g f::'a \ 'b. (f has_integral i) ({a..b}) \ (g has_integral j) ({a..b}) \ \x\{a..b}. (f x)\k \ (g x)\k \ i\k \ j\k" proof (rule ccontr) case goal1 - then have *: "0 < (i\k - j\k) / 3" by auto - guess d1 using goal1(1)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d1=this[rule_format] - guess d2 using goal1(2)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d2=this[rule_format] + then have *: "0 < (i\k - j\k) / 3" + by auto + guess d1 using goal1(1)[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d1=this[rule_format] + guess d2 using goal1(2)[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d2=this[rule_format] guess p using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter . note p = this(1) conjunctD2[OF this(2)] note le_less_trans[OF Basis_le_norm[OF k]] note this[OF d1(2)[OF conjI[OF p(1,2)]]] this[OF d2(2)[OF conjI[OF p(1,3)]]] - thus False + then show False unfolding inner_simps using rsum_component_le[OF p(1) goal1(3)] by (simp add: abs_real_def split: split_if_asm) - qed let ?P = "\a b. s = {a..b}" - { presume "\ ?P \ ?thesis" thus ?thesis proof(cases ?P) - case True then guess a b apply-by(erule exE)+ note s=this - show ?thesis apply(rule lem) using assms[unfolded s] by auto - qed auto } assume as:"\ ?P" - { presume "\ ?thesis \ False" thus ?thesis by auto } - assume "\ i\k \ j\k" hence ij:"(i\k - j\k) / 3 > 0" by auto + qed + let ?P = "\a b. s = {a..b}" + { + presume "\ ?P \ ?thesis" + then show ?thesis + proof (cases ?P) + case True + then guess a b by (elim exE) note s=this + show ?thesis + apply (rule lem) + using assms[unfolded s] + apply auto + done + qed auto + } + assume as: "\ ?P" + { presume "\ ?thesis \ False" then show ?thesis by auto } + assume "\ i\k \ j\k" + then have ij: "(i\k - j\k) / 3 > 0" + by auto note has_integral_altD[OF _ as this] from this[OF assms(2)] this[OF assms(3)] guess B1 B2 . note B=this[rule_format] - have "bounded (ball 0 B1 \ ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+ - from bounded_subset_closed_interval[OF this] guess a b apply- by(erule exE)+ + have "bounded (ball 0 B1 \ ball (0::'a) B2)" + unfolding bounded_Un by(rule conjI bounded_ball)+ + from bounded_subset_closed_interval[OF this] guess a b by (elim exE) note ab = conjunctD2[OF this[unfolded Un_subset_iff]] guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this] guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this] - have *:"\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" + have *: "\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" by (simp add: abs_real_def split: split_if_asm) - note le_less_trans[OF Basis_le_norm[OF k]] note this[OF w1(2)] this[OF w2(2)] moreover - have "w1\k \ w2\k" apply(rule lem[OF w1(1) w2(1)]) using assms by auto ultimately - show False unfolding inner_simps by(rule *) + note le_less_trans[OF Basis_le_norm[OF k]] + note this[OF w1(2)] this[OF w2(2)] + moreover + have "w1\k \ w2\k" + apply (rule lem[OF w1(1) w2(1)]) + using assms + apply auto + done + ultimately show False + unfolding inner_simps by(rule *) qed -lemma integral_component_le: fixes g f::"'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "k\Basis" "f integrable_on s" "g integrable_on s" "\x\s. (f x)\k \ (g x)\k" +lemma integral_component_le: + fixes g f :: "'a::ordered_euclidean_space \ 'b::euclidean_space" + assumes "k \ Basis" + and "f integrable_on s" "g integrable_on s" + and "\x\s. (f x)\k \ (g x)\k" shows "(integral s f)\k \ (integral s g)\k" - apply(rule has_integral_component_le) using integrable_integral assms by auto - -lemma has_integral_component_nonneg: fixes f::"'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "k\Basis" "(f has_integral i) s" "\x\s. 0 \ (f x)\k" shows "0 \ i\k" - using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] using assms(3-) by auto - -lemma integral_component_nonneg: fixes f::"'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "k\Basis" "f integrable_on s" "\x\s. 0 \ (f x)\k" shows "0 \ (integral s f)\k" - apply(rule has_integral_component_nonneg) using assms by auto - -lemma has_integral_component_neg: fixes f::"'a::ordered_euclidean_space \ 'b::ordered_euclidean_space" - assumes "k\Basis" "(f has_integral i) s" "\x\s. (f x)\k \ 0"shows "i\k \ 0" - using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) by auto + apply (rule has_integral_component_le) + using integrable_integral assms + apply auto + done + +lemma has_integral_component_nonneg: + fixes f :: "'a::ordered_euclidean_space \ 'b::euclidean_space" + assumes "k \ Basis" + and "(f has_integral i) s" + and "\x\s. 0 \ (f x)\k" + shows "0 \ i\k" + using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] + using assms(3-) + by auto + +lemma integral_component_nonneg: + fixes f :: "'a::ordered_euclidean_space \ 'b::euclidean_space" + assumes "k \ Basis" + and "f integrable_on s" "\x\s. 0 \ (f x)\k" + shows "0 \ (integral s f)\k" + apply (rule has_integral_component_nonneg) + using assms + apply auto + done + +lemma has_integral_component_neg: + fixes f :: "'a::ordered_euclidean_space \ 'b::ordered_euclidean_space" + assumes "k \ Basis" + and "(f has_integral i) s" + and "\x\s. (f x)\k \ 0" + shows "i\k \ 0" + using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) + by auto lemma has_integral_component_lbound: - fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "(f has_integral i) {a..b}" "\x\{a..b}. B \ f(x)\k" "k\Basis" + fixes f :: "'a::ordered_euclidean_space => 'b::ordered_euclidean_space" + assumes "(f has_integral i) {a..b}" + and "\x\{a..b}. B \ f(x)\k" + and "k \ Basis" shows "B * content {a..b} \ i\k" using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) - by (auto simp add:field_simps) + by (auto simp add: field_simps) lemma has_integral_component_ubound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "(f has_integral i) {a..b}" "\x\{a..b}. f x\k \ B" "k\Basis" - shows "i\k \ B * content({a..b})" - using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) - by(auto simp add:field_simps) - -lemma integral_component_lbound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "f integrable_on {a..b}" "\x\{a..b}. B \ f(x)\k" "k\Basis" - shows "B * content({a..b}) \ (integral({a..b}) f)\k" - apply(rule has_integral_component_lbound) using assms unfolding has_integral_integral by auto - -lemma integral_component_ubound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "f integrable_on {a..b}" "\x\{a..b}. f(x)\k \ B" "k\Basis" - shows "(integral({a..b}) f)\k \ B * content({a..b})" - apply(rule has_integral_component_ubound) using assms unfolding has_integral_integral by auto + assumes "(f has_integral i) {a..b}" + and "\x\{a..b}. f x\k \ B" + and "k \ Basis" + shows "i\k \ B * content {a..b}" + using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) + by (auto simp add: field_simps) + +lemma integral_component_lbound: + fixes f :: "'a::ordered_euclidean_space \ 'b::ordered_euclidean_space" + assumes "f integrable_on {a..b}" + and "\x\{a..b}. B \ f(x)\k" + and "k \ Basis" + shows "B * content {a..b} \ (integral({a..b}) f)\k" + apply (rule has_integral_component_lbound) + using assms + unfolding has_integral_integral + apply auto + done + +lemma integral_component_ubound: + fixes f :: "'a::ordered_euclidean_space \ 'b::ordered_euclidean_space" + assumes "f integrable_on {a..b}" + and "\x\{a..b}. f x\k \ B" + and "k \ Basis" + shows "(integral {a..b} f)\k \ B * content {a..b}" + apply (rule has_integral_component_ubound) + using assms + unfolding has_integral_integral + apply auto + done + subsection {* Uniform limit of integrable functions is integrable. *} -lemma integrable_uniform_limit: fixes f::"'a::ordered_euclidean_space \ 'b::banach" - assumes "\e>0. \g. (\x\{a..b}. norm(f x - g x) \ e) \ g integrable_on {a..b}" +lemma integrable_uniform_limit: + fixes f :: "'a::ordered_euclidean_space \ 'b::banach" + assumes "\e>0. \g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" shows "f integrable_on {a..b}" -proof- { presume *:"content {a..b} > 0 \ ?thesis" - show ?thesis apply cases apply(rule *,assumption) - unfolding content_lt_nz integrable_on_def using has_integral_null by auto } - assume as:"content {a..b} > 0" - have *:"\P. \e>(0::real). P e \ \n::nat. P (inverse (real n+1))" by auto +proof - + { + presume *: "content {a..b} > 0 \ ?thesis" + show ?thesis + apply cases + apply (rule *) + apply assumption + unfolding content_lt_nz integrable_on_def + using has_integral_null + apply auto + done + } + assume as: "content {a..b} > 0" + have *: "\P. \e>(0::real). P e \ \n::nat. P (inverse (real n + 1))" + by auto from choice[OF *[OF assms]] guess g .. note g=conjunctD2[OF this[rule_format],rule_format] from choice[OF allI[OF g(2)[unfolded integrable_on_def], of "\x. x"]] guess i .. note i=this[rule_format] - have "Cauchy i" unfolding Cauchy_def - proof(rule,rule) fix e::real assume "e>0" - hence "e / 4 / content {a..b} > 0" using as by(auto simp add:field_simps) - then guess M apply-apply(subst(asm) real_arch_inv) by(erule exE conjE)+ note M=this - show "\M. \m\M. \n\M. dist (i m) (i n) < e" apply(rule_tac x=M in exI,rule,rule,rule,rule) - proof- case goal1 have "e/4>0" using `e>0` by auto note * = i[unfolded has_integral,rule_format,OF this] - from *[of m] guess gm apply-by(erule conjE exE)+ note gm=this[rule_format] - from *[of n] guess gn apply-by(erule conjE exE)+ note gn=this[rule_format] + have "Cauchy i" + unfolding Cauchy_def + proof (rule, rule) + fix e :: real + assume "e>0" + then have "e / 4 / content {a..b} > 0" + using as by (auto simp add: field_simps) + then guess M + apply - + apply (subst(asm) real_arch_inv) + apply (elim exE conjE) + done + note M=this + show "\M. \m\M. \n\M. dist (i m) (i n) < e" + apply (rule_tac x=M in exI,rule,rule,rule,rule) + proof - + case goal1 + have "e/4>0" using `e>0` by auto + note * = i[unfolded has_integral,rule_format,OF this] + from *[of m] guess gm by (elim conjE exE) note gm=this[rule_format] + from *[of n] guess gn by (elim conjE exE) note gn=this[rule_format] from fine_division_exists[OF gauge_inter[OF gm(1) gn(1)], of a b] guess p . note p=this - have lem2:"\s1 s2 i1 i2. norm(s2 - s1) \ e/2 \ norm(s1 - i1) < e / 4 \ norm(s2 - i2) < e / 4 \norm(i1 - i2) < e" - proof- case goal1 have "norm (i1 - i2) \ norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" + have lem2: "\s1 s2 i1 i2. norm(s2 - s1) \ e/2 \ norm (s1 - i1) < e / 4 \ + norm (s2 - i2) < e / 4 \ norm (i1 - i2) < e" + proof - + case goal1 + have "norm (i1 - i2) \ norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] - using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:algebra_simps) - also have "\ < e" using goal1 unfolding norm_minus_commute by(auto simp add:algebra_simps) + using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] + by (auto simp add: algebra_simps) + also have "\ < e" + using goal1 + unfolding norm_minus_commute + by (auto simp add: algebra_simps) finally show ?case . qed - show ?case unfolding dist_norm apply(rule lem2) defer - apply(rule gm(2)[OF conjI[OF p(1)]],rule_tac[2] gn(2)[OF conjI[OF p(1)]]) - using conjunctD2[OF p(2)[unfolded fine_inter]] apply- apply assumption+ apply(rule order_trans) - apply(rule rsum_diff_bound[OF p(1), where e="2 / real M"]) - proof show "2 / real M * content {a..b} \ e / 2" unfolding divide_inverse - using M as by(auto simp add:field_simps) - fix x assume x:"x \ {a..b}" + show ?case + unfolding dist_norm + apply (rule lem2) + defer + apply (rule gm(2)[OF conjI[OF p(1)]],rule_tac[2] gn(2)[OF conjI[OF p(1)]]) + using conjunctD2[OF p(2)[unfolded fine_inter]] + apply - + apply assumption+ + apply (rule order_trans) + apply (rule rsum_diff_bound[OF p(1), where e="2 / real M"]) + proof + show "2 / real M * content {a..b} \ e / 2" + unfolding divide_inverse + using M as + by (auto simp add: field_simps) + fix x + assume x: "x \ {a..b}" have "norm (f x - g n x) + norm (f x - g m x) \ inverse (real n + 1) + inverse (real m + 1)" - using g(1)[OF x, of n] g(1)[OF x, of m] by auto - also have "\ \ inverse (real M) + inverse (real M)" apply(rule add_mono) - apply(rule_tac[!] le_imp_inverse_le) using goal1 M by auto - also have "\ = 2 / real M" unfolding divide_inverse by auto + using g(1)[OF x, of n] g(1)[OF x, of m] by auto + also have "\ \ inverse (real M) + inverse (real M)" + apply (rule add_mono) + apply (rule_tac[!] le_imp_inverse_le) + using goal1 M + apply auto + done + also have "\ = 2 / real M" + unfolding divide_inverse by auto finally show "norm (g n x - g m x) \ 2 / real M" using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] - by(auto simp add:algebra_simps simp add:norm_minus_commute) - qed qed qed + by (auto simp add: algebra_simps simp add: norm_minus_commute) + qed + qed + qed from this[unfolded convergent_eq_cauchy[symmetric]] guess s .. note s=this - show ?thesis unfolding integrable_on_def apply(rule_tac x=s in exI) unfolding has_integral - proof(rule,rule) - case goal1 hence *:"e/3 > 0" by auto + show ?thesis + unfolding integrable_on_def + apply (rule_tac x=s in exI) + unfolding has_integral + proof (rule, rule) + case goal1 + then have *: "e/3 > 0" by auto from LIMSEQ_D [OF s this] guess N1 .. note N1=this - from goal1 as have "e / 3 / content {a..b} > 0" by(auto simp add:field_simps) - from real_arch_invD[OF this] guess N2 apply-by(erule exE conjE)+ note N2=this + from goal1 as have "e / 3 / content {a..b} > 0" + by (auto simp add: field_simps) + from real_arch_invD[OF this] guess N2 by (elim exE conjE) note N2=this from i[of "N1 + N2",unfolded has_integral,rule_format,OF *] guess g' .. note g'=conjunctD2[OF this,rule_format] - have lem:"\sf sg i. norm(sf - sg) \ e / 3 \ norm(i - s) < e / 3 \ norm(sg - i) < e / 3 \ norm(sf - s) < e" - proof- case goal1 have "norm (sf - s) \ norm (sf - sg) + norm (sg - i) + norm (i - s)" + have lem: "\sf sg i. norm (sf - sg) \ e / 3 \ + norm(i - s) < e / 3 \ norm (sg - i) < e / 3 \ norm (sf - s) < e" + proof - + case goal1 + have "norm (sf - s) \ norm (sf - sg) + norm (sg - i) + norm (i - s)" using norm_triangle_ineq[of "sf - sg" "sg - s"] - using norm_triangle_ineq[of "sg - i" " i - s"] by(auto simp add:algebra_simps) - also have "\ < e" using goal1 unfolding norm_minus_commute by(auto simp add:algebra_simps) + using norm_triangle_ineq[of "sg - i" " i - s"] + by (auto simp add: algebra_simps) + also have "\ < e" + using goal1 + unfolding norm_minus_commute + by (auto simp add: algebra_simps) finally show ?case . qed - show ?case apply(rule_tac x=g' in exI) apply(rule,rule g') - proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \ g' fine p" note * = g'(2)[OF this] - show "norm ((\(x, k)\p. content k *\<^sub>R f x) - s) < e" apply-apply(rule lem[OF _ _ *]) - apply(rule order_trans,rule rsum_diff_bound[OF p[THEN conjunct1]]) apply(rule,rule g,assumption) - proof- have "content {a..b} < e / 3 * (real N2)" - using N2 unfolding inverse_eq_divide using as by(auto simp add:field_simps) - hence "content {a..b} < e / 3 * (real (N1 + N2) + 1)" - apply-apply(rule less_le_trans,assumption) using `e>0` by auto - thus "inverse (real (N1 + N2) + 1) * content {a..b} \ e / 3" - unfolding inverse_eq_divide by(auto simp add:field_simps) - show "norm (i (N1 + N2) - s) < e / 3" by(rule N1[rule_format],auto) - qed qed qed qed + show ?case + apply (rule_tac x=g' in exI) + apply rule + apply (rule g') + proof (rule, rule) + fix p + assume p: "p tagged_division_of {a..b} \ g' fine p" + note * = g'(2)[OF this] + show "norm ((\(x, k)\p. content k *\<^sub>R f x) - s) < e" + apply - + apply (rule lem[OF _ _ *]) + apply (rule order_trans) + apply (rule rsum_diff_bound[OF p[THEN conjunct1]]) + apply rule + apply (rule g) + apply assumption + proof - + have "content {a..b} < e / 3 * (real N2)" + using N2 unfolding inverse_eq_divide using as by (auto simp add: field_simps) + then have "content {a..b} < e / 3 * (real (N1 + N2) + 1)" + apply - + apply (rule less_le_trans,assumption) + using `e>0` + apply auto + done + then show "inverse (real (N1 + N2) + 1) * content {a..b} \ e / 3" + unfolding inverse_eq_divide + by (auto simp add: field_simps) + show "norm (i (N1 + N2) - s) < e / 3" + by (rule N1[rule_format]) auto + qed + qed + qed +qed + subsection {* Negligible sets. *} -definition "negligible (s::('a::ordered_euclidean_space) set) \ (\a b. ((indicator s :: 'a\real) has_integral 0) {a..b})" +definition "negligible (s:: 'a::ordered_euclidean_space set) \ + (\a b. ((indicator s :: 'a\real) has_integral 0) {a..b})" + subsection {* Negligibility of hyperplane. *} lemma vsum_nonzero_image_lemma: - assumes "finite s" "g(a) = 0" - "\x\s. \y\s. f x = f y \ x \ y \ g(f x) = 0" + assumes "finite s" + and "g a = 0" + and "\x\s. \y\s. f x = f y \ x \ y \ g (f x) = 0" shows "setsum g {f x |x. x \ s \ f x \ a} = setsum (g o f) s" - unfolding setsum_iterate[OF assms(1)] apply(subst setsum_iterate) defer - apply(rule iterate_nonzero_image_lemma) apply(rule assms monoidal_monoid)+ - unfolding assms using neutral_add unfolding neutral_add using assms by auto - -lemma interval_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "k\Basis" + unfolding setsum_iterate[OF assms(1)] + apply (subst setsum_iterate) + defer + apply (rule iterate_nonzero_image_lemma) + apply (rule assms monoidal_monoid)+ + unfolding assms + using neutral_add + unfolding neutral_add + using assms + apply auto + done + +lemma interval_doublesplit: + fixes a :: "'a::ordered_euclidean_space" + assumes "k \ Basis" shows "{a..b} \ {x . abs(x\k - c) \ (e::real)} = - {(\i\Basis. (if i = k then max (a\k) (c - e) else a\i) *\<^sub>R i) .. - (\i\Basis. (if i = k then min (b\k) (c + e) else b\i) *\<^sub>R i)}" -proof- have *:"\x c e::real. abs(x - c) \ e \ x \ c - e \ x \ c + e" by auto - have **:"\s P Q. s \ {x. P x \ Q x} = (s \ {x. Q x}) \ {x. P x}" by blast - show ?thesis unfolding * ** interval_split[OF assms] by(rule refl) qed + {(\i\Basis. (if i = k then max (a\k) (c - e) else a\i) *\<^sub>R i) .. + (\i\Basis. (if i = k then min (b\k) (c + e) else b\i) *\<^sub>R i)}" +proof - + have *: "\x c e::real. abs(x - c) \ e \ x \ c - e \ x \ c + e" + by auto + have **: "\s P Q. s \ {x. P x \ Q x} = (s \ {x. Q x}) \ {x. P x}" + by blast + show ?thesis + unfolding * ** interval_split[OF assms] by (rule refl) +qed lemma division_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "p division_of {a..b}" and k:"k\Basis" shows "{l \ {x. abs(x\k - c) \ e} |l. l \ p \ l \ {x. abs(x\k - c) \ e} \ {}} division_of ({a..b} \ {x. abs(x\k - c) \ e})"