diff -r 04aeda922be2 -r e9b2782c4f99 src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Sun Jun 03 15:49:55 2012 +0200 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Mon Jun 04 09:07:23 2012 +0200 @@ -2319,7 +2319,7 @@ 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 smt + have *:"\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" using [[z3_with_extensions]] by smt note le_less_trans[OF component_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 euclidean_simps by(rule *) qed @@ -3050,7 +3050,7 @@ hence "d \ {a..b}" using s(2)[OF k(1)] unfolding k by auto note di = this[unfolded mem_interval,THEN spec[where x=i]] hence xyi:"y$$i \ x$$i" unfolding y_def unfolding i xi euclidean_lambda_beta'[OF i(2)] if_P[OF refl] apply(cases) apply(subst if_P,assumption) unfolding if_not_P not_le using as(2) - using assms(2)[unfolded content_eq_0] using i(2) by smt+ + using assms(2)[unfolded content_eq_0] using i(2) using [[z3_with_extensions]] by smt+ thus "y \ x" unfolding euclidean_eq[where 'a='a] using i by auto have *:"{..i. 0) {.. {a..b}" "ball 0 (max B1 B2) \ {c..d}" have **:"ball 0 B1 \ ball (0::'n::ordered_euclidean_space) (max B1 B2)" "ball 0 B2 \ ball (0::'n::ordered_euclidean_space) (max B1 B2)" by auto have *:"\ga gc ha hc fa fc::real. abs(ga - i) < e / 3 \ abs(gc - i) < e / 3 \ abs(ha - j) < e / 3 \ - abs(hc - j) < e / 3 \ abs(i - j) < e / 3 \ ga \ fa \ fa \ ha \ gc \ fc \ fc \ hc\ abs(fa - fc) < e" by smt + abs(hc - j) < e / 3 \ abs(i - j) < e / 3 \ ga \ fa \ fa \ ha \ gc \ fc \ fc \ hc\ abs(fa - fc) < e" using [[z3_with_extensions]] by smt show "norm (integral {a..b} (\x. if x \ s then f x else 0) - integral {c..d} (\x. if x \ s then f x else 0)) < e" unfolding real_norm_def apply(rule *, safe) unfolding real_norm_def[THEN sym] apply(rule B1(2),rule order_trans,rule **,rule as(1))