# HG changeset patch # User wenzelm # Date 1434316951 -7200 # Node ID d1a9d098f870a61eb097852457cafb590ec06b3b # Parent f690cb540385cd1304335867809166a613cd7e6b# Parent 051b200f7578b576b61307c1960f0ff693ae2ab8 merged diff -r 051b200f7578 -r d1a9d098f870 src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Sun Jun 14 23:22:08 2015 +0200 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Sun Jun 14 23:22:31 2015 +0200 @@ -4266,54 +4266,50 @@ 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)] + obtain p where p: "p tagged_division_of cbox a b" "d1 fine p" "d2 fine p" + using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter + by metis 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)]]] + then have "\((\(x, k)\p. content k *\<^sub>R f x) - i) \ k\ < (i \ k - j \ k) / 3" + "\((\(x, k)\p. content k *\<^sub>R g x) - j) \ k\ < (i \ k - j \ k) / 3" + using k norm_bound_Basis_lt d1 d2 p + by blast+ 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 = cbox 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_cbox[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" - 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 - apply auto - done - ultimately show False - unfolding inner_simps by(rule *) + show ?thesis + proof (cases "\a b. s = cbox a b") + case True + with lem assms show ?thesis + by auto + next + case False + show ?thesis + proof (rule ccontr) + assume "\ i\k \ j\k" + then have ij: "(i\k - j\k) / 3 > 0" + by auto + note has_integral_altD[OF _ False 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_cbox[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" + 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" + by (rule lem[OF w1(1) w2(1)]) (simp add: assms(4)) + ultimately show False + unfolding inner_simps by(rule *) + qed + qed qed lemma integral_component_le: