# HG changeset patch # User wenzelm # Date 1378482901 -7200 # Node ID f41ab5a7df97cea92c044528f39bc569ff796c1e # Parent 63958e9e0073c5e984e5588be0e7cf9d59a5e69c tuned proofs; diff -r 63958e9e0073 -r f41ab5a7df97 src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Fri Sep 06 17:26:58 2013 +0200 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Fri Sep 06 17:55:01 2013 +0200 @@ -3125,102 +3125,210 @@ subsection {* Cauchy-type criterion for integrability. *} (* XXXXXXX *) -lemma integrable_cauchy: fixes f::"'n::ordered_euclidean_space \ 'a::{real_normed_vector,complete_space}" +lemma integrable_cauchy: + fixes f :: "'n::ordered_euclidean_space \ 'a::{real_normed_vector,complete_space}" shows "f integrable_on {a..b} \ - (\e>0.\d. gauge d \ (\p1 p2. p1 tagged_division_of {a..b} \ d fine p1 \ - p2 tagged_division_of {a..b} \ d fine p2 - \ norm(setsum (\(x,k). content k *\<^sub>R f x) p1 - - setsum (\(x,k). content k *\<^sub>R f x) p2) < e))" (is "?l = (\e>0. \d. ?P e d)") -proof assume ?l + (\e>0.\d. gauge d \ + (\p1 p2. p1 tagged_division_of {a..b} \ d fine p1 \ + p2 tagged_division_of {a..b} \ d fine p2 \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p1 - + setsum (\(x,k). content k *\<^sub>R f x) p2) < e))" + (is "?l = (\e>0. \d. ?P e d)") +proof + assume ?l then guess y unfolding integrable_on_def has_integral .. note y=this - show "\e>0. \d. ?P e d" proof(rule,rule) case goal1 hence "e/2 > 0" by auto + show "\e>0. \d. ?P e d" + 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] - show ?case apply(rule_tac x=d in exI,rule,rule d) apply(rule,rule,rule,(erule conjE)+) - proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b}" "d fine p1" "p2 tagged_division_of {a..b}" "d fine p2" + show ?case + apply (rule_tac x=d in exI) + apply rule + apply (rule d) + apply rule + apply rule + apply rule + apply (erule conjE)+ + proof - + fix p1 p2 + assume as: "p1 tagged_division_of {a..b}" "d fine p1" + "p2 tagged_division_of {a..b}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" - apply(rule dist_triangle_half_l[where y=y,unfolded dist_norm]) + apply (rule dist_triangle_half_l[where y=y,unfolded dist_norm]) using d(2)[OF conjI[OF as(1-2)]] d(2)[OF conjI[OF as(3-4)]] . - qed qed -next assume "\e>0. \d. ?P e d" hence "\n::nat. \d. ?P (inverse(real (n + 1))) d" by auto + qed + qed +next + assume "\e>0. \d. ?P e d" + then have "\n::nat. \d. ?P (inverse(real (n + 1))) d" + by auto from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format],rule_format] - have "\n. gauge (\x. \{d i x |i. i \ {0..n}})" apply(rule gauge_inters) using d(1) by auto - hence "\n. \p. p tagged_division_of {a..b} \ (\x. \{d i x |i. i \ {0..n}}) fine p" apply- - proof case goal1 from this[of n] show ?case apply(drule_tac fine_division_exists) by auto qed + have "\n. gauge (\x. \{d i x |i. i \ {0..n}})" + apply (rule gauge_inters) + using d(1) + apply auto + done + then have "\n. \p. p tagged_division_of {a..b} \ (\x. \{d i x |i. i \ {0..n}}) fine p" + apply - + proof + case goal1 + from this[of n] + show ?case + apply (drule_tac fine_division_exists) + apply auto + done + qed from choice[OF this] guess p .. note p = conjunctD2[OF this[rule_format]] - have dp:"\i n. i\n \ d i fine p n" using p(2) unfolding fine_inters by auto + have dp: "\i n. i\n \ d i fine p n" + using p(2) unfolding fine_inters by auto have "Cauchy (\n. setsum (\(x,k). content k *\<^sub>R (f x)) (p n))" - proof(rule CauchyI) case goal1 then guess N unfolding real_arch_inv[of e] .. note N=this - show ?case apply(rule_tac x=N in exI) - proof(rule,rule,rule,rule) fix m n assume mn:"N \ m" "N \ n" have *:"N = (N - 1) + 1" using N by auto + proof (rule CauchyI) + case goal1 + then guess N unfolding real_arch_inv[of e] .. note N=this + show ?case + apply (rule_tac x=N in exI) + proof (rule, rule, rule, rule) + fix m n + assume mn: "N \ m" "N \ n" + have *: "N = (N - 1) + 1" using N by auto show "norm ((\(x, k)\p m. content k *\<^sub>R f x) - (\(x, k)\p n. content k *\<^sub>R f x)) < e" - apply(rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]]) apply(subst *) apply(rule d(2)) - using dp p(1) using mn by auto - qed qed + apply (rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]]) + apply(subst *) + apply(rule d(2)) + using dp p(1) + using mn + apply auto + done + qed + qed then guess y unfolding convergent_eq_cauchy[symmetric] .. note y=this[THEN LIMSEQ_D] - show ?l unfolding integrable_on_def has_integral apply(rule_tac x=y in exI) - proof(rule,rule) fix e::real assume "e>0" hence *:"e/2 > 0" by auto - then guess N1 unfolding real_arch_inv[of "e/2"] .. note N1=this hence N1':"N1 = N1 - 1 + 1" by auto + show ?l + unfolding integrable_on_def has_integral + apply (rule_tac x=y in exI) + proof (rule, rule) + fix e :: real + assume "e>0" + then have *:"e/2 > 0" by auto + then guess N1 unfolding real_arch_inv[of "e/2"] .. note N1=this + then have N1': "N1 = N1 - 1 + 1" + 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)" - apply(rule_tac x="d (N1 + N2)" in exI) apply rule defer - proof(rule,rule,erule conjE) show "gauge (d (N1 + N2))" using d by auto - fix q assume as:"q tagged_division_of {a..b}" "d (N1 + N2) fine q" - have *:"inverse (real (N1 + N2 + 1)) < e / 2" apply(rule less_trans) using N1 by auto - show "norm ((\(x, k)\q. content k *\<^sub>R f x) - y) < e" apply(rule norm_triangle_half_r) - apply(rule less_trans[OF _ *]) apply(subst N1', rule d(2)[of "p (N1+N2)"]) defer + 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)" + apply (rule_tac x="d (N1 + N2)" in exI) + apply rule + defer + proof (rule, rule, erule conjE) + show "gauge (d (N1 + N2))" + using d by auto + fix q + assume as: "q tagged_division_of {a..b}" "d (N1 + N2) fine q" + have *: "inverse (real (N1 + N2 + 1)) < e / 2" + apply (rule less_trans) + using N1 + apply auto + done + show "norm ((\(x, k)\q. content k *\<^sub>R f x) - y) < e" + apply (rule norm_triangle_half_r) + apply (rule less_trans[OF _ *]) + apply (subst N1', rule d(2)[of "p (N1+N2)"]) + defer using N2[rule_format,of "N1+N2"] - using as dp[of "N1 - 1 + 1 + N2" "N1 + N2"] using p(1)[of "N1 + N2"] using N1 by auto qed qed qed + using as dp[of "N1 - 1 + 1 + N2" "N1 + N2"] + using p(1)[of "N1 + N2"] + using N1 + apply auto + done + qed + qed +qed + subsection {* Additivity of integral on abutting intervals. *} lemma interval_split: - fixes a::"'a::ordered_euclidean_space" assumes "k \ Basis" + fixes a :: "'a::ordered_euclidean_space" + assumes "k \ Basis" shows "{a..b} \ {x. x\k \ c} = {a .. (\i\Basis. (if i = k then min (b\k) c else b\i) *\<^sub>R i)}" "{a..b} \ {x. x\k \ c} = {(\i\Basis. (if i = k then max (a\k) c else a\i) *\<^sub>R i) .. b}" - apply(rule_tac[!] set_eqI) unfolding Int_iff mem_interval mem_Collect_eq using assms - by auto - -lemma content_split: fixes a::"'a::ordered_euclidean_space" assumes "k\Basis" shows - "content {a..b} = content({a..b} \ {x. x\k \ c}) + content({a..b} \ {x. x\k >= c})" + apply (rule_tac[!] set_eqI) + unfolding Int_iff mem_interval mem_Collect_eq + using assms + apply auto + done + +lemma content_split: + fixes a :: "'a::ordered_euclidean_space" + assumes "k \ Basis" + shows "content {a..b} = content({a..b} \ {x. x\k \ c}) + content({a..b} \ {x. x\k \ c})" proof cases note simps = interval_split[OF assms] content_closed_interval_cases eucl_le[where 'a='a] - have *:"Basis = insert k (Basis - {k})" "\x. finite (Basis-{x})" "\x. x\Basis-{x}" + have *: "Basis = insert k (Basis - {k})" "\x. finite (Basis-{x})" "\x. x\Basis-{x}" using assms by auto - have *:"\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" + have *: "\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" "(\i\Basis. b\i - a\i) = (\i\Basis-{k}. b\i - a\i) * (b\k - a\k)" - apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto - assume as:"a\b" moreover have "\x. min (b \ k) c = max (a \ k) c - \ x* (b\k - a\k) = x*(max (a \ k) c - a \ k) + x*(b \ k - max (a \ k) c)" - by (auto simp add:field_simps) - moreover have **:"(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = + apply (subst *(1)) + defer + apply (subst *(1)) + unfolding setprod_insert[OF *(2-)] + apply auto + done + assume as: "a \ b" + moreover + have "\x. min (b \ k) c = max (a \ k) c \ + x * (b\k - a\k) = x * (max (a \ k) c - a \ k) + x * (b \ k - max (a \ k) c)" + by (auto simp add: field_simps) + moreover + have **: "(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = (\i\Basis. (if i = k then min (b \ k) c else b \ i) - a \ i)" "(\i\Basis. b \ i - ((\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i)) = (\i\Basis. b \ i - (if i = k then max (a \ k) c else a \ i))" by (auto intro!: setprod_cong) have "\ a \ k \ c \ \ c \ b \ k \ False" - unfolding not_le using as[unfolded eucl_le[where 'a='a],rule_format,of k] assms by auto - ultimately show ?thesis using assms unfolding simps ** - unfolding *(1)[of "\i x. b\i - x"] *(1)[of "\i x. x - a\i"] unfolding *(2) + unfolding not_le + using as[unfolded eucl_le[where 'a='a],rule_format,of k] assms + by auto + ultimately show ?thesis + using assms + unfolding simps ** + unfolding *(1)[of "\i x. b\i - x"] *(1)[of "\i x. x - a\i"] + unfolding *(2) by auto next - assume "\ a \ b" then have "{a .. b} = {}" + assume "\ a \ b" + then have "{a .. b} = {}" unfolding interval_eq_empty by (auto simp: eucl_le[where 'a='a] not_le) - then show ?thesis by auto + then show ?thesis + by auto qed -lemma division_split_left_inj: fixes type::"'a::ordered_euclidean_space" - assumes "d division_of i" "k1 \ d" "k2 \ d" "k1 \ k2" - "k1 \ {x::'a. x\k \ c} = k2 \ {x. x\k \ c}"and k:"k\Basis" +lemma division_split_left_inj: + fixes type :: "'a::ordered_euclidean_space" + assumes "d division_of i" + and "k1 \ d" + and "k2 \ d" + and "k1 \ k2" + and "k1 \ {x::'a. x\k \ c} = k2 \ {x. x\k \ c}" + and k: "k\Basis" 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}) = {})" +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}) = {})" 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 - have **:"\s t u. s \ t = {} \ u \ s \ u \ t \ u = {}" by auto - show ?thesis unfolding uv1 uv2 * apply(rule **[OF d(5)[OF assms(2-4)]]) - defer apply(subst assms(5)[unfolded uv1 uv2]) unfolding uv1 uv2 by auto qed + have **: "\s t u. s \ t = {} \ u \ s \ u \ t \ u = {}" + by auto + show ?thesis + unfolding uv1 uv2 * apply(rule **[OF d(5)[OF assms(2-4)]]) + defer + apply (subst assms(5)[unfolded uv1 uv2]) + unfolding uv1 uv2 + apply auto + done +qed lemma division_split_right_inj: fixes type::"'a::ordered_euclidean_space" assumes "d division_of i" "k1 \ d" "k2 \ d" "k1 \ k2"