# HG changeset patch # User blanchet # Date 1266437620 -3600 # Node ID a815c3f4eef2274249c2dad29af4ce987b0c8f89 # Parent 69fa4c39dab28528c319c988928aef012aca7ddb# Parent 3b9762ad372d5ac906c7e50106e231f3a459a17b merged diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Finite_Set.thy Wed Feb 17 21:13:40 2010 +0100 @@ -2034,6 +2034,31 @@ apply auto done +lemma setprod_mono: + fixes f :: "'a \ 'b\linordered_semidom" + assumes "\i\A. 0 \ f i \ f i \ g i" + shows "setprod f A \ setprod g A" +proof (cases "finite A") + case True + hence ?thesis "setprod f A \ 0" using subset_refl[of A] + proof (induct A rule: finite_subset_induct) + case (insert a F) + thus "setprod f (insert a F) \ setprod g (insert a F)" "0 \ setprod f (insert a F)" + unfolding setprod_insert[OF insert(1,3)] + using assms[rule_format,OF insert(2)] insert + by (auto intro: mult_mono mult_nonneg_nonneg) + qed auto + thus ?thesis by simp +qed auto + +lemma abs_setprod: + fixes f :: "'a \ 'b\{linordered_field,abs}" + shows "abs (setprod f A) = setprod (\x. abs (f x)) A" +proof (cases "finite A") + case True thus ?thesis + by induct (auto simp add: field_simps setprod_insert abs_mult) +qed auto + subsection {* Finite cardinality *} diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Library/Formal_Power_Series.thy Wed Feb 17 21:13:40 2010 +0100 @@ -2864,8 +2864,8 @@ then have nz: "pochhammer (1 + b - of_nat n) n \ 0" by (auto simp add: algebra_simps) - from nz kn have nz': "pochhammer (1 + b - of_nat n) k \ 0" - by (simp add: pochhammer_neq_0_mono) + from nz kn [simplified] have nz': "pochhammer (1 + b - of_nat n) k \ 0" + by (rule pochhammer_neq_0_mono) {assume k0: "k = 0 \ n =0" then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" using kn diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Library/Ramsey.thy Wed Feb 17 21:13:40 2010 +0100 @@ -111,7 +111,7 @@ have infYx': "infinite (Yx-{yx'})" using fields px by auto with fields px yx' Suc.prems have partfx': "part r s (Yx - {yx'}) (f \ insert yx')" - by (simp add: o_def part_Suc_imp_part part_subset [where ?YY=YY]) + by (simp add: o_def part_Suc_imp_part part_subset [where YY=YY and Y=Yx]) from Suc.hyps [OF infYx' partfx'] obtain Y' and t' where Y': "Y' \ Yx - {yx'}" "infinite Y'" "t' < s" diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Library/Word.thy --- a/src/HOL/Library/Word.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Library/Word.thy Wed Feb 17 21:13:40 2010 +0100 @@ -980,7 +980,8 @@ fix xs assume "length (norm_signed (\#xs)) = Suc (length xs)" thus "norm_signed (\#xs) = \#xs" - by (simp add: norm_signed_Cons norm_unsigned_equal split: split_if_asm) + by (simp add: norm_signed_Cons norm_unsigned_equal [THEN eqTrueI] + split: split_if_asm) next fix xs assume "length (norm_signed (\#xs)) = Suc (length xs)" diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Library/Zorn.thy --- a/src/HOL/Library/Zorn.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Library/Zorn.thy Wed Feb 17 21:13:40 2010 +0100 @@ -333,7 +333,7 @@ lemma antisym_init_seg_of: "r initial_segment_of s \ s initial_segment_of r \ r=s" -by(auto simp:init_seg_of_def) +unfolding init_seg_of_def by safe lemma Chain_init_seg_of_Union: "R \ Chain init_seg_of \ r\R \ r initial_segment_of \R" diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Wed Feb 17 21:13:40 2010 +0100 @@ -387,7 +387,7 @@ apply(rule_tac le_less_trans[of _ "e/2"]) by(auto intro!:mult_imp_div_pos_le simp add:group_simps) qed auto qed lemma has_derivative_at_alt: - "(f has_derivative f') (at (x::real^'n)) \ bounded_linear f' \ + "(f has_derivative f') (at x) \ bounded_linear f' \ (\e>0. \d>0. \y. norm(y - x) < d \ norm(f y - f x - f'(y - x)) \ e * norm(y - x))" using has_derivative_within_alt[where s=UNIV] unfolding within_UNIV by auto diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Multivariate_Analysis/Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1042,11 +1042,6 @@ shows "norm x \ norm y + norm (x - y)" using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps) -lemma norm_triangle_le: "norm(x::real ^ 'n) + norm y <= e ==> norm(x + y) <= e" - by (metis order_trans norm_triangle_ineq) -lemma norm_triangle_lt: "norm(x::real ^ 'n) + norm(y) < e ==> norm(x + y) < e" - by (metis basic_trans_rules(21) norm_triangle_ineq) - lemma component_le_norm: "\x$i\ <= norm x" apply (simp add: norm_vector_def) apply (rule member_le_setL2, simp_all) @@ -1275,6 +1270,22 @@ shows "dist y x1 < e / 2 \ dist y x2 < e / 2 \ dist x1 x2 < e" by (rule dist_triangle_half_l, simp_all add: dist_commute) + +lemma norm_triangle_half_r: + shows "norm (y - x1) < e / 2 \ norm (y - x2) < e / 2 \ norm (x1 - x2) < e" + using dist_triangle_half_r unfolding vector_dist_norm[THEN sym] by auto + +lemma norm_triangle_half_l: assumes "norm (x - y) < e / 2" "norm (x' - (y)) < e / 2" + shows "norm (x - x') < e" + using dist_triangle_half_l[OF assms[unfolded vector_dist_norm[THEN sym]]] + unfolding vector_dist_norm[THEN sym] . + +lemma norm_triangle_le: "norm(x) + norm y <= e ==> norm(x + y) <= e" + by (metis order_trans norm_triangle_ineq) + +lemma norm_triangle_lt: "norm(x) + norm(y) < e ==> norm(x + y) < e" + by (metis basic_trans_rules(21) norm_triangle_ineq) + lemma dist_triangle_add: fixes x y x' y' :: "'a::real_normed_vector" shows "dist (x + y) (x' + y') <= dist x x' + dist y y'" diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Multivariate_Analysis/Integration_MV.cert --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Integration_MV.cert Wed Feb 17 21:13:40 2010 +0100 @@ -0,0 +1,3270 @@ +tB2Atlor9W4pSnrAz5nHpw 907 0 +#2 := false +#299 := 0::real +decl uf_1 :: (-> T3 T2 real) +decl uf_10 :: (-> T4 T2) +decl uf_7 :: T4 +#15 := uf_7 +#22 := (uf_10 uf_7) +decl uf_2 :: (-> T1 T3) +decl uf_4 :: T1 +#11 := uf_4 +#91 := (uf_2 uf_4) +#902 := (uf_1 #91 #22) +#297 := -1::real +#1084 := (* -1::real #902) +decl uf_16 :: T1 +#50 := uf_16 +#78 := (uf_2 uf_16) +#799 := (uf_1 #78 #22) +#1267 := (+ #799 #1084) +#1272 := (>= #1267 0::real) +#1266 := (= #799 #902) +decl uf_9 :: T3 +#21 := uf_9 +#23 := (uf_1 uf_9 #22) +#905 := (= #23 #902) +decl uf_11 :: T3 +#24 := uf_11 +#850 := (uf_1 uf_11 #22) +#904 := (= #850 #902) +decl uf_6 :: (-> T2 T4) +#74 := (uf_6 #22) +#281 := (= uf_7 #74) +#922 := (ite #281 #905 #904) +decl uf_8 :: T3 +#18 := uf_8 +#848 := (uf_1 uf_8 #22) +#903 := (= #848 #902) +#60 := 0::int +decl uf_5 :: (-> T4 int) +#803 := (uf_5 #74) +#117 := -1::int +#813 := (* -1::int #803) +#16 := (uf_5 uf_7) +#916 := (+ #16 #813) +#917 := (<= #916 0::int) +#925 := (ite #917 #922 #903) +#6 := (:var 0 T2) +#19 := (uf_1 uf_8 #6) +#544 := (pattern #19) +#25 := (uf_1 uf_11 #6) +#543 := (pattern #25) +#92 := (uf_1 #91 #6) +#542 := (pattern #92) +#13 := (uf_6 #6) +#541 := (pattern #13) +#447 := (= #19 #92) +#445 := (= #25 #92) +#444 := (= #23 #92) +#20 := (= #13 uf_7) +#446 := (ite #20 #444 #445) +#120 := (* -1::int #16) +#14 := (uf_5 #13) +#121 := (+ #14 #120) +#119 := (>= #121 0::int) +#448 := (ite #119 #446 #447) +#545 := (forall (vars (?x3 T2)) (:pat #541 #542 #543 #544) #448) +#451 := (forall (vars (?x3 T2)) #448) +#548 := (iff #451 #545) +#546 := (iff #448 #448) +#547 := [refl]: #546 +#549 := [quant-intro #547]: #548 +#26 := (ite #20 #23 #25) +#127 := (ite #119 #26 #19) +#368 := (= #92 #127) +#369 := (forall (vars (?x3 T2)) #368) +#452 := (iff #369 #451) +#449 := (iff #368 #448) +#450 := [rewrite]: #449 +#453 := [quant-intro #450]: #452 +#392 := (~ #369 #369) +#390 := (~ #368 #368) +#391 := [refl]: #390 +#366 := [nnf-pos #391]: #392 +decl uf_3 :: (-> T1 T2 real) +#12 := (uf_3 uf_4 #6) +#132 := (= #12 #127) +#135 := (forall (vars (?x3 T2)) #132) +#370 := (iff #135 #369) +#4 := (:var 1 T1) +#8 := (uf_3 #4 #6) +#5 := (uf_2 #4) +#7 := (uf_1 #5 #6) +#9 := (= #7 #8) +#10 := (forall (vars (?x1 T1) (?x2 T2)) #9) +#113 := [asserted]: #10 +#371 := [rewrite* #113]: #370 +#17 := (< #14 #16) +#27 := (ite #17 #19 #26) +#28 := (= #12 #27) +#29 := (forall (vars (?x3 T2)) #28) +#136 := (iff #29 #135) +#133 := (iff #28 #132) +#130 := (= #27 #127) +#118 := (not #119) +#124 := (ite #118 #19 #26) +#128 := (= #124 #127) +#129 := [rewrite]: #128 +#125 := (= #27 #124) +#122 := (iff #17 #118) +#123 := [rewrite]: #122 +#126 := [monotonicity #123]: #125 +#131 := [trans #126 #129]: #130 +#134 := [monotonicity #131]: #133 +#137 := [quant-intro #134]: #136 +#114 := [asserted]: #29 +#138 := [mp #114 #137]: #135 +#372 := [mp #138 #371]: #369 +#367 := [mp~ #372 #366]: #369 +#454 := [mp #367 #453]: #451 +#550 := [mp #454 #549]: #545 +#738 := (not #545) +#928 := (or #738 #925) +#75 := (= #74 uf_7) +#906 := (ite #75 #905 #904) +#907 := (+ #803 #120) +#908 := (>= #907 0::int) +#909 := (ite #908 #906 #903) +#929 := (or #738 #909) +#931 := (iff #929 #928) +#933 := (iff #928 #928) +#934 := [rewrite]: #933 +#926 := (iff #909 #925) +#923 := (iff #906 #922) +#283 := (iff #75 #281) +#284 := [rewrite]: #283 +#924 := [monotonicity #284]: #923 +#920 := (iff #908 #917) +#910 := (+ #120 #803) +#913 := (>= #910 0::int) +#918 := (iff #913 #917) +#919 := [rewrite]: #918 +#914 := (iff #908 #913) +#911 := (= #907 #910) +#912 := [rewrite]: #911 +#915 := [monotonicity #912]: #914 +#921 := [trans #915 #919]: #920 +#927 := [monotonicity #921 #924]: #926 +#932 := [monotonicity #927]: #931 +#935 := [trans #932 #934]: #931 +#930 := [quant-inst]: #929 +#936 := [mp #930 #935]: #928 +#1300 := [unit-resolution #936 #550]: #925 +#989 := (= #16 #803) +#1277 := (= #803 #16) +#280 := [asserted]: #75 +#287 := [mp #280 #284]: #281 +#1276 := [symm #287]: #75 +#1278 := [monotonicity #1276]: #1277 +#1301 := [symm #1278]: #989 +#1302 := (not #989) +#1303 := (or #1302 #917) +#1304 := [th-lemma]: #1303 +#1305 := [unit-resolution #1304 #1301]: #917 +#950 := (not #917) +#949 := (not #925) +#951 := (or #949 #950 #922) +#952 := [def-axiom]: #951 +#1306 := [unit-resolution #952 #1305 #1300]: #922 +#937 := (not #922) +#1307 := (or #937 #905) +#938 := (not #281) +#939 := (or #937 #938 #905) +#940 := [def-axiom]: #939 +#1308 := [unit-resolution #940 #287]: #1307 +#1309 := [unit-resolution #1308 #1306]: #905 +#1356 := (= #799 #23) +#800 := (= #23 #799) +decl uf_15 :: T4 +#40 := uf_15 +#41 := (uf_5 uf_15) +#814 := (+ #41 #813) +#815 := (<= #814 0::int) +#836 := (not #815) +#158 := (* -1::int #41) +#1270 := (+ #16 #158) +#1265 := (>= #1270 0::int) +#1339 := (not #1265) +#1269 := (= #16 #41) +#1298 := (not #1269) +#286 := (= uf_7 uf_15) +#44 := (uf_10 uf_15) +#72 := (uf_6 #44) +#73 := (= #72 uf_15) +#277 := (= uf_15 #72) +#278 := (iff #73 #277) +#279 := [rewrite]: #278 +#276 := [asserted]: #73 +#282 := [mp #276 #279]: #277 +#1274 := [symm #282]: #73 +#729 := (= uf_7 #72) +decl uf_17 :: (-> int T4) +#611 := (uf_5 #72) +#991 := (uf_17 #611) +#1289 := (= #991 #72) +#992 := (= #72 #991) +#55 := (:var 0 T4) +#56 := (uf_5 #55) +#574 := (pattern #56) +#57 := (uf_17 #56) +#177 := (= #55 #57) +#575 := (forall (vars (?x7 T4)) (:pat #574) #177) +#195 := (forall (vars (?x7 T4)) #177) +#578 := (iff #195 #575) +#576 := (iff #177 #177) +#577 := [refl]: #576 +#579 := [quant-intro #577]: #578 +#405 := (~ #195 #195) +#403 := (~ #177 #177) +#404 := [refl]: #403 +#406 := [nnf-pos #404]: #405 +#58 := (= #57 #55) +#59 := (forall (vars (?x7 T4)) #58) +#196 := (iff #59 #195) +#193 := (iff #58 #177) +#194 := [rewrite]: #193 +#197 := [quant-intro #194]: #196 +#155 := [asserted]: #59 +#200 := [mp #155 #197]: #195 +#407 := [mp~ #200 #406]: #195 +#580 := [mp #407 #579]: #575 +#995 := (not #575) +#996 := (or #995 #992) +#997 := [quant-inst]: #996 +#1273 := [unit-resolution #997 #580]: #992 +#1290 := [symm #1273]: #1289 +#1293 := (= uf_7 #991) +#993 := (uf_17 #803) +#1287 := (= #993 #991) +#1284 := (= #803 #611) +#987 := (= #41 #611) +#1279 := (= #611 #41) +#1280 := [monotonicity #1274]: #1279 +#1281 := [symm #1280]: #987 +#1282 := (= #803 #41) +#1275 := [hypothesis]: #1269 +#1283 := [trans #1278 #1275]: #1282 +#1285 := [trans #1283 #1281]: #1284 +#1288 := [monotonicity #1285]: #1287 +#1291 := (= uf_7 #993) +#994 := (= #74 #993) +#1000 := (or #995 #994) +#1001 := [quant-inst]: #1000 +#1286 := [unit-resolution #1001 #580]: #994 +#1292 := [trans #287 #1286]: #1291 +#1294 := [trans #1292 #1288]: #1293 +#1295 := [trans #1294 #1290]: #729 +#1296 := [trans #1295 #1274]: #286 +#290 := (not #286) +#76 := (= uf_15 uf_7) +#77 := (not #76) +#291 := (iff #77 #290) +#288 := (iff #76 #286) +#289 := [rewrite]: #288 +#292 := [monotonicity #289]: #291 +#285 := [asserted]: #77 +#295 := [mp #285 #292]: #290 +#1297 := [unit-resolution #295 #1296]: false +#1299 := [lemma #1297]: #1298 +#1342 := (or #1269 #1339) +#1271 := (<= #1270 0::int) +#621 := (* -1::int #611) +#723 := (+ #16 #621) +#724 := (<= #723 0::int) +decl uf_12 :: T1 +#30 := uf_12 +#88 := (uf_2 uf_12) +#771 := (uf_1 #88 #44) +#45 := (uf_1 uf_9 #44) +#772 := (= #45 #771) +#796 := (not #772) +decl uf_14 :: T1 +#38 := uf_14 +#83 := (uf_2 uf_14) +#656 := (uf_1 #83 #44) +#1239 := (= #656 #771) +#1252 := (not #1239) +#1324 := (iff #1252 #796) +#1322 := (iff #1239 #772) +#1320 := (= #656 #45) +#661 := (= #45 #656) +#659 := (uf_1 uf_11 #44) +#664 := (= #656 #659) +#667 := (ite #277 #661 #664) +#657 := (uf_1 uf_8 #44) +#670 := (= #656 #657) +#622 := (+ #41 #621) +#623 := (<= #622 0::int) +#673 := (ite #623 #667 #670) +#84 := (uf_1 #83 #6) +#560 := (pattern #84) +#467 := (= #19 #84) +#465 := (= #25 #84) +#464 := (= #45 #84) +#43 := (= #13 uf_15) +#466 := (ite #43 #464 #465) +#159 := (+ #14 #158) +#157 := (>= #159 0::int) +#468 := (ite #157 #466 #467) +#561 := (forall (vars (?x5 T2)) (:pat #541 #560 #543 #544) #468) +#471 := (forall (vars (?x5 T2)) #468) +#564 := (iff #471 #561) +#562 := (iff #468 #468) +#563 := [refl]: #562 +#565 := [quant-intro #563]: #564 +#46 := (ite #43 #45 #25) +#165 := (ite #157 #46 #19) +#378 := (= #84 #165) +#379 := (forall (vars (?x5 T2)) #378) +#472 := (iff #379 #471) +#469 := (iff #378 #468) +#470 := [rewrite]: #469 +#473 := [quant-intro #470]: #472 +#359 := (~ #379 #379) +#361 := (~ #378 #378) +#358 := [refl]: #361 +#356 := [nnf-pos #358]: #359 +#39 := (uf_3 uf_14 #6) +#170 := (= #39 #165) +#173 := (forall (vars (?x5 T2)) #170) +#380 := (iff #173 #379) +#381 := [rewrite* #113]: #380 +#42 := (< #14 #41) +#47 := (ite #42 #19 #46) +#48 := (= #39 #47) +#49 := (forall (vars (?x5 T2)) #48) +#174 := (iff #49 #173) +#171 := (iff #48 #170) +#168 := (= #47 #165) +#156 := (not #157) +#162 := (ite #156 #19 #46) +#166 := (= #162 #165) +#167 := [rewrite]: #166 +#163 := (= #47 #162) +#160 := (iff #42 #156) +#161 := [rewrite]: #160 +#164 := [monotonicity #161]: #163 +#169 := [trans #164 #167]: #168 +#172 := [monotonicity #169]: #171 +#175 := [quant-intro #172]: #174 +#116 := [asserted]: #49 +#176 := [mp #116 #175]: #173 +#382 := [mp #176 #381]: #379 +#357 := [mp~ #382 #356]: #379 +#474 := [mp #357 #473]: #471 +#566 := [mp #474 #565]: #561 +#676 := (not #561) +#677 := (or #676 #673) +#658 := (= #657 #656) +#660 := (= #659 #656) +#662 := (ite #73 #661 #660) +#612 := (+ #611 #158) +#613 := (>= #612 0::int) +#663 := (ite #613 #662 #658) +#678 := (or #676 #663) +#680 := (iff #678 #677) +#682 := (iff #677 #677) +#683 := [rewrite]: #682 +#674 := (iff #663 #673) +#671 := (iff #658 #670) +#672 := [rewrite]: #671 +#668 := (iff #662 #667) +#665 := (iff #660 #664) +#666 := [rewrite]: #665 +#669 := [monotonicity #279 #666]: #668 +#626 := (iff #613 #623) +#615 := (+ #158 #611) +#618 := (>= #615 0::int) +#624 := (iff #618 #623) +#625 := [rewrite]: #624 +#619 := (iff #613 #618) +#616 := (= #612 #615) +#617 := [rewrite]: #616 +#620 := [monotonicity #617]: #619 +#627 := [trans #620 #625]: #626 +#675 := [monotonicity #627 #669 #672]: #674 +#681 := [monotonicity #675]: #680 +#684 := [trans #681 #683]: #680 +#679 := [quant-inst]: #678 +#685 := [mp #679 #684]: #677 +#1311 := [unit-resolution #685 #566]: #673 +#1312 := (not #987) +#1313 := (or #1312 #623) +#1314 := [th-lemma]: #1313 +#1315 := [unit-resolution #1314 #1281]: #623 +#645 := (not #623) +#698 := (not #673) +#699 := (or #698 #645 #667) +#700 := [def-axiom]: #699 +#1316 := [unit-resolution #700 #1315 #1311]: #667 +#686 := (not #667) +#1317 := (or #686 #661) +#687 := (not #277) +#688 := (or #686 #687 #661) +#689 := [def-axiom]: #688 +#1318 := [unit-resolution #689 #282]: #1317 +#1319 := [unit-resolution #1318 #1316]: #661 +#1321 := [symm #1319]: #1320 +#1323 := [monotonicity #1321]: #1322 +#1325 := [monotonicity #1323]: #1324 +#1145 := (* -1::real #771) +#1240 := (+ #656 #1145) +#1241 := (<= #1240 0::real) +#1249 := (not #1241) +#1243 := [hypothesis]: #1241 +decl uf_18 :: T3 +#80 := uf_18 +#1040 := (uf_1 uf_18 #44) +#1043 := (* -1::real #1040) +#1156 := (+ #771 #1043) +#1157 := (>= #1156 0::real) +#1189 := (not #1157) +#708 := (uf_1 #91 #44) +#1168 := (+ #708 #1043) +#1169 := (<= #1168 0::real) +#1174 := (or #1157 #1169) +#1177 := (not #1174) +#89 := (uf_1 #88 #6) +#552 := (pattern #89) +#81 := (uf_1 uf_18 #6) +#594 := (pattern #81) +#324 := (* -1::real #92) +#325 := (+ #81 #324) +#323 := (>= #325 0::real) +#317 := (* -1::real #89) +#318 := (+ #81 #317) +#319 := (<= #318 0::real) +#436 := (or #319 #323) +#437 := (not #436) +#601 := (forall (vars (?x11 T2)) (:pat #594 #552 #542) #437) +#440 := (forall (vars (?x11 T2)) #437) +#604 := (iff #440 #601) +#602 := (iff #437 #437) +#603 := [refl]: #602 +#605 := [quant-intro #603]: #604 +#326 := (not #323) +#320 := (not #319) +#329 := (and #320 #326) +#332 := (forall (vars (?x11 T2)) #329) +#441 := (iff #332 #440) +#438 := (iff #329 #437) +#439 := [rewrite]: #438 +#442 := [quant-intro #439]: #441 +#425 := (~ #332 #332) +#423 := (~ #329 #329) +#424 := [refl]: #423 +#426 := [nnf-pos #424]: #425 +#306 := (* -1::real #84) +#307 := (+ #81 #306) +#305 := (>= #307 0::real) +#308 := (not #305) +#301 := (* -1::real #81) +#79 := (uf_1 #78 #6) +#302 := (+ #79 #301) +#300 := (>= #302 0::real) +#298 := (not #300) +#311 := (and #298 #308) +#314 := (forall (vars (?x10 T2)) #311) +#335 := (and #314 #332) +#93 := (< #81 #92) +#90 := (< #89 #81) +#94 := (and #90 #93) +#95 := (forall (vars (?x11 T2)) #94) +#85 := (< #81 #84) +#82 := (< #79 #81) +#86 := (and #82 #85) +#87 := (forall (vars (?x10 T2)) #86) +#96 := (and #87 #95) +#336 := (iff #96 #335) +#333 := (iff #95 #332) +#330 := (iff #94 #329) +#327 := (iff #93 #326) +#328 := [rewrite]: #327 +#321 := (iff #90 #320) +#322 := [rewrite]: #321 +#331 := [monotonicity #322 #328]: #330 +#334 := [quant-intro #331]: #333 +#315 := (iff #87 #314) +#312 := (iff #86 #311) +#309 := (iff #85 #308) +#310 := [rewrite]: #309 +#303 := (iff #82 #298) +#304 := [rewrite]: #303 +#313 := [monotonicity #304 #310]: #312 +#316 := [quant-intro #313]: #315 +#337 := [monotonicity #316 #334]: #336 +#293 := [asserted]: #96 +#338 := [mp #293 #337]: #335 +#340 := [and-elim #338]: #332 +#427 := [mp~ #340 #426]: #332 +#443 := [mp #427 #442]: #440 +#606 := [mp #443 #605]: #601 +#1124 := (not #601) +#1180 := (or #1124 #1177) +#1142 := (* -1::real #708) +#1143 := (+ #1040 #1142) +#1144 := (>= #1143 0::real) +#1146 := (+ #1040 #1145) +#1147 := (<= #1146 0::real) +#1148 := (or #1147 #1144) +#1149 := (not #1148) +#1181 := (or #1124 #1149) +#1183 := (iff #1181 #1180) +#1185 := (iff #1180 #1180) +#1186 := [rewrite]: #1185 +#1178 := (iff #1149 #1177) +#1175 := (iff #1148 #1174) +#1172 := (iff #1144 #1169) +#1162 := (+ #1142 #1040) +#1165 := (>= #1162 0::real) +#1170 := (iff #1165 #1169) +#1171 := [rewrite]: #1170 +#1166 := (iff #1144 #1165) +#1163 := (= #1143 #1162) +#1164 := [rewrite]: #1163 +#1167 := [monotonicity #1164]: #1166 +#1173 := [trans #1167 #1171]: #1172 +#1160 := (iff #1147 #1157) +#1150 := (+ #1145 #1040) +#1153 := (<= #1150 0::real) +#1158 := (iff #1153 #1157) +#1159 := [rewrite]: #1158 +#1154 := (iff #1147 #1153) +#1151 := (= #1146 #1150) +#1152 := [rewrite]: #1151 +#1155 := [monotonicity #1152]: #1154 +#1161 := [trans #1155 #1159]: #1160 +#1176 := [monotonicity #1161 #1173]: #1175 +#1179 := [monotonicity #1176]: #1178 +#1184 := [monotonicity #1179]: #1183 +#1187 := [trans #1184 #1186]: #1183 +#1182 := [quant-inst]: #1181 +#1188 := [mp #1182 #1187]: #1180 +#1244 := [unit-resolution #1188 #606]: #1177 +#1190 := (or #1174 #1189) +#1191 := [def-axiom]: #1190 +#1245 := [unit-resolution #1191 #1244]: #1189 +#1054 := (+ #656 #1043) +#1055 := (<= #1054 0::real) +#1079 := (not #1055) +#607 := (uf_1 #78 #44) +#1044 := (+ #607 #1043) +#1045 := (>= #1044 0::real) +#1060 := (or #1045 #1055) +#1063 := (not #1060) +#567 := (pattern #79) +#428 := (or #300 #305) +#429 := (not #428) +#595 := (forall (vars (?x10 T2)) (:pat #567 #594 #560) #429) +#432 := (forall (vars (?x10 T2)) #429) +#598 := (iff #432 #595) +#596 := (iff #429 #429) +#597 := [refl]: #596 +#599 := [quant-intro #597]: #598 +#433 := (iff #314 #432) +#430 := (iff #311 #429) +#431 := [rewrite]: #430 +#434 := [quant-intro #431]: #433 +#420 := (~ #314 #314) +#418 := (~ #311 #311) +#419 := [refl]: #418 +#421 := [nnf-pos #419]: #420 +#339 := [and-elim #338]: #314 +#422 := [mp~ #339 #421]: #314 +#435 := [mp #422 #434]: #432 +#600 := [mp #435 #599]: #595 +#1066 := (not #595) +#1067 := (or #1066 #1063) +#1039 := (* -1::real #656) +#1041 := (+ #1040 #1039) +#1042 := (>= #1041 0::real) +#1046 := (or #1045 #1042) +#1047 := (not #1046) +#1068 := (or #1066 #1047) +#1070 := (iff #1068 #1067) +#1072 := (iff #1067 #1067) +#1073 := [rewrite]: #1072 +#1064 := (iff #1047 #1063) +#1061 := (iff #1046 #1060) +#1058 := (iff #1042 #1055) +#1048 := (+ #1039 #1040) +#1051 := (>= #1048 0::real) +#1056 := (iff #1051 #1055) +#1057 := [rewrite]: #1056 +#1052 := (iff #1042 #1051) +#1049 := (= #1041 #1048) +#1050 := [rewrite]: #1049 +#1053 := [monotonicity #1050]: #1052 +#1059 := [trans #1053 #1057]: #1058 +#1062 := [monotonicity #1059]: #1061 +#1065 := [monotonicity #1062]: #1064 +#1071 := [monotonicity #1065]: #1070 +#1074 := [trans #1071 #1073]: #1070 +#1069 := [quant-inst]: #1068 +#1075 := [mp #1069 #1074]: #1067 +#1246 := [unit-resolution #1075 #600]: #1063 +#1080 := (or #1060 #1079) +#1081 := [def-axiom]: #1080 +#1247 := [unit-resolution #1081 #1246]: #1079 +#1248 := [th-lemma #1247 #1245 #1243]: false +#1250 := [lemma #1248]: #1249 +#1253 := (or #1252 #1241) +#1254 := [th-lemma]: #1253 +#1310 := [unit-resolution #1254 #1250]: #1252 +#1326 := [mp #1310 #1325]: #796 +#1328 := (or #724 #772) +decl uf_13 :: T3 +#33 := uf_13 +#609 := (uf_1 uf_13 #44) +#773 := (= #609 #771) +#775 := (ite #724 #773 #772) +#32 := (uf_1 uf_9 #6) +#553 := (pattern #32) +#34 := (uf_1 uf_13 #6) +#551 := (pattern #34) +#456 := (= #32 #89) +#455 := (= #34 #89) +#457 := (ite #119 #455 #456) +#554 := (forall (vars (?x4 T2)) (:pat #541 #551 #552 #553) #457) +#460 := (forall (vars (?x4 T2)) #457) +#557 := (iff #460 #554) +#555 := (iff #457 #457) +#556 := [refl]: #555 +#558 := [quant-intro #556]: #557 +#143 := (ite #119 #34 #32) +#373 := (= #89 #143) +#374 := (forall (vars (?x4 T2)) #373) +#461 := (iff #374 #460) +#458 := (iff #373 #457) +#459 := [rewrite]: #458 +#462 := [quant-intro #459]: #461 +#362 := (~ #374 #374) +#364 := (~ #373 #373) +#365 := [refl]: #364 +#363 := [nnf-pos #365]: #362 +#31 := (uf_3 uf_12 #6) +#148 := (= #31 #143) +#151 := (forall (vars (?x4 T2)) #148) +#375 := (iff #151 #374) +#376 := [rewrite* #113]: #375 +#35 := (ite #17 #32 #34) +#36 := (= #31 #35) +#37 := (forall (vars (?x4 T2)) #36) +#152 := (iff #37 #151) +#149 := (iff #36 #148) +#146 := (= #35 #143) +#140 := (ite #118 #32 #34) +#144 := (= #140 #143) +#145 := [rewrite]: #144 +#141 := (= #35 #140) +#142 := [monotonicity #123]: #141 +#147 := [trans #142 #145]: #146 +#150 := [monotonicity #147]: #149 +#153 := [quant-intro #150]: #152 +#115 := [asserted]: #37 +#154 := [mp #115 #153]: #151 +#377 := [mp #154 #376]: #374 +#360 := [mp~ #377 #363]: #374 +#463 := [mp #360 #462]: #460 +#559 := [mp #463 #558]: #554 +#778 := (not #554) +#779 := (or #778 #775) +#714 := (+ #611 #120) +#715 := (>= #714 0::int) +#774 := (ite #715 #773 #772) +#780 := (or #778 #774) +#782 := (iff #780 #779) +#784 := (iff #779 #779) +#785 := [rewrite]: #784 +#776 := (iff #774 #775) +#727 := (iff #715 #724) +#717 := (+ #120 #611) +#720 := (>= #717 0::int) +#725 := (iff #720 #724) +#726 := [rewrite]: #725 +#721 := (iff #715 #720) +#718 := (= #714 #717) +#719 := [rewrite]: #718 +#722 := [monotonicity #719]: #721 +#728 := [trans #722 #726]: #727 +#777 := [monotonicity #728]: #776 +#783 := [monotonicity #777]: #782 +#786 := [trans #783 #785]: #782 +#781 := [quant-inst]: #780 +#787 := [mp #781 #786]: #779 +#1327 := [unit-resolution #787 #559]: #775 +#788 := (not #775) +#791 := (or #788 #724 #772) +#792 := [def-axiom]: #791 +#1329 := [unit-resolution #792 #1327]: #1328 +#1330 := [unit-resolution #1329 #1326]: #724 +#988 := (>= #622 0::int) +#1331 := (or #1312 #988) +#1332 := [th-lemma]: #1331 +#1333 := [unit-resolution #1332 #1281]: #988 +#761 := (not #724) +#1334 := (not #988) +#1335 := (or #1271 #1334 #761) +#1336 := [th-lemma]: #1335 +#1337 := [unit-resolution #1336 #1333 #1330]: #1271 +#1338 := (not #1271) +#1340 := (or #1269 #1338 #1339) +#1341 := [th-lemma]: #1340 +#1343 := [unit-resolution #1341 #1337]: #1342 +#1344 := [unit-resolution #1343 #1299]: #1339 +#990 := (>= #916 0::int) +#1345 := (or #1302 #990) +#1346 := [th-lemma]: #1345 +#1347 := [unit-resolution #1346 #1301]: #990 +#1348 := (not #990) +#1349 := (or #836 #1348 #1265) +#1350 := [th-lemma]: #1349 +#1351 := [unit-resolution #1350 #1347 #1344]: #836 +#1353 := (or #815 #800) +#801 := (uf_1 uf_13 #22) +#820 := (= #799 #801) +#823 := (ite #815 #820 #800) +#476 := (= #32 #79) +#475 := (= #34 #79) +#477 := (ite #157 #475 #476) +#568 := (forall (vars (?x6 T2)) (:pat #541 #551 #567 #553) #477) +#480 := (forall (vars (?x6 T2)) #477) +#571 := (iff #480 #568) +#569 := (iff #477 #477) +#570 := [refl]: #569 +#572 := [quant-intro #570]: #571 +#181 := (ite #157 #34 #32) +#383 := (= #79 #181) +#384 := (forall (vars (?x6 T2)) #383) +#481 := (iff #384 #480) +#478 := (iff #383 #477) +#479 := [rewrite]: #478 +#482 := [quant-intro #479]: #481 +#352 := (~ #384 #384) +#354 := (~ #383 #383) +#355 := [refl]: #354 +#353 := [nnf-pos #355]: #352 +#51 := (uf_3 uf_16 #6) +#186 := (= #51 #181) +#189 := (forall (vars (?x6 T2)) #186) +#385 := (iff #189 #384) +#386 := [rewrite* #113]: #385 +#52 := (ite #42 #32 #34) +#53 := (= #51 #52) +#54 := (forall (vars (?x6 T2)) #53) +#190 := (iff #54 #189) +#187 := (iff #53 #186) +#184 := (= #52 #181) +#178 := (ite #156 #32 #34) +#182 := (= #178 #181) +#183 := [rewrite]: #182 +#179 := (= #52 #178) +#180 := [monotonicity #161]: #179 +#185 := [trans #180 #183]: #184 +#188 := [monotonicity #185]: #187 +#191 := [quant-intro #188]: #190 +#139 := [asserted]: #54 +#192 := [mp #139 #191]: #189 +#387 := [mp #192 #386]: #384 +#402 := [mp~ #387 #353]: #384 +#483 := [mp #402 #482]: #480 +#573 := [mp #483 #572]: #568 +#634 := (not #568) +#826 := (or #634 #823) +#802 := (= #801 #799) +#804 := (+ #803 #158) +#805 := (>= #804 0::int) +#806 := (ite #805 #802 #800) +#827 := (or #634 #806) +#829 := (iff #827 #826) +#831 := (iff #826 #826) +#832 := [rewrite]: #831 +#824 := (iff #806 #823) +#821 := (iff #802 #820) +#822 := [rewrite]: #821 +#818 := (iff #805 #815) +#807 := (+ #158 #803) +#810 := (>= #807 0::int) +#816 := (iff #810 #815) +#817 := [rewrite]: #816 +#811 := (iff #805 #810) +#808 := (= #804 #807) +#809 := [rewrite]: #808 +#812 := [monotonicity #809]: #811 +#819 := [trans #812 #817]: #818 +#825 := [monotonicity #819 #822]: #824 +#830 := [monotonicity #825]: #829 +#833 := [trans #830 #832]: #829 +#828 := [quant-inst]: #827 +#834 := [mp #828 #833]: #826 +#1352 := [unit-resolution #834 #573]: #823 +#835 := (not #823) +#839 := (or #835 #815 #800) +#840 := [def-axiom]: #839 +#1354 := [unit-resolution #840 #1352]: #1353 +#1355 := [unit-resolution #1354 #1351]: #800 +#1357 := [symm #1355]: #1356 +#1358 := [trans #1357 #1309]: #1266 +#1359 := (not #1266) +#1360 := (or #1359 #1272) +#1361 := [th-lemma]: #1360 +#1362 := [unit-resolution #1361 #1358]: #1272 +#1085 := (uf_1 uf_18 #22) +#1099 := (* -1::real #1085) +#1112 := (+ #902 #1099) +#1113 := (<= #1112 0::real) +#1137 := (not #1113) +#960 := (uf_1 #88 #22) +#1100 := (+ #960 #1099) +#1101 := (>= #1100 0::real) +#1118 := (or #1101 #1113) +#1121 := (not #1118) +#1125 := (or #1124 #1121) +#1086 := (+ #1085 #1084) +#1087 := (>= #1086 0::real) +#1088 := (* -1::real #960) +#1089 := (+ #1085 #1088) +#1090 := (<= #1089 0::real) +#1091 := (or #1090 #1087) +#1092 := (not #1091) +#1126 := (or #1124 #1092) +#1128 := (iff #1126 #1125) +#1130 := (iff #1125 #1125) +#1131 := [rewrite]: #1130 +#1122 := (iff #1092 #1121) +#1119 := (iff #1091 #1118) +#1116 := (iff #1087 #1113) +#1106 := (+ #1084 #1085) +#1109 := (>= #1106 0::real) +#1114 := (iff #1109 #1113) +#1115 := [rewrite]: #1114 +#1110 := (iff #1087 #1109) +#1107 := (= #1086 #1106) +#1108 := [rewrite]: #1107 +#1111 := [monotonicity #1108]: #1110 +#1117 := [trans #1111 #1115]: #1116 +#1104 := (iff #1090 #1101) +#1093 := (+ #1088 #1085) +#1096 := (<= #1093 0::real) +#1102 := (iff #1096 #1101) +#1103 := [rewrite]: #1102 +#1097 := (iff #1090 #1096) +#1094 := (= #1089 #1093) +#1095 := [rewrite]: #1094 +#1098 := [monotonicity #1095]: #1097 +#1105 := [trans #1098 #1103]: #1104 +#1120 := [monotonicity #1105 #1117]: #1119 +#1123 := [monotonicity #1120]: #1122 +#1129 := [monotonicity #1123]: #1128 +#1132 := [trans #1129 #1131]: #1128 +#1127 := [quant-inst]: #1126 +#1133 := [mp #1127 #1132]: #1125 +#1363 := [unit-resolution #1133 #606]: #1121 +#1138 := (or #1118 #1137) +#1139 := [def-axiom]: #1138 +#1364 := [unit-resolution #1139 #1363]: #1137 +#1200 := (+ #799 #1099) +#1201 := (>= #1200 0::real) +#1231 := (not #1201) +#847 := (uf_1 #83 #22) +#1210 := (+ #847 #1099) +#1211 := (<= #1210 0::real) +#1216 := (or #1201 #1211) +#1219 := (not #1216) +#1222 := (or #1066 #1219) +#1197 := (* -1::real #847) +#1198 := (+ #1085 #1197) +#1199 := (>= #1198 0::real) +#1202 := (or #1201 #1199) +#1203 := (not #1202) +#1223 := (or #1066 #1203) +#1225 := (iff #1223 #1222) +#1227 := (iff #1222 #1222) +#1228 := [rewrite]: #1227 +#1220 := (iff #1203 #1219) +#1217 := (iff #1202 #1216) +#1214 := (iff #1199 #1211) +#1204 := (+ #1197 #1085) +#1207 := (>= #1204 0::real) +#1212 := (iff #1207 #1211) +#1213 := [rewrite]: #1212 +#1208 := (iff #1199 #1207) +#1205 := (= #1198 #1204) +#1206 := [rewrite]: #1205 +#1209 := [monotonicity #1206]: #1208 +#1215 := [trans #1209 #1213]: #1214 +#1218 := [monotonicity #1215]: #1217 +#1221 := [monotonicity #1218]: #1220 +#1226 := [monotonicity #1221]: #1225 +#1229 := [trans #1226 #1228]: #1225 +#1224 := [quant-inst]: #1223 +#1230 := [mp #1224 #1229]: #1222 +#1365 := [unit-resolution #1230 #600]: #1219 +#1232 := (or #1216 #1231) +#1233 := [def-axiom]: #1232 +#1366 := [unit-resolution #1233 #1365]: #1231 +[th-lemma #1366 #1364 #1362]: false +unsat +NQHwTeL311Tq3wf2s5BReA 419 0 +#2 := false +#194 := 0::real +decl uf_4 :: (-> T2 T3 real) +decl uf_6 :: (-> T1 T3) +decl uf_3 :: T1 +#21 := uf_3 +#25 := (uf_6 uf_3) +decl uf_5 :: T2 +#24 := uf_5 +#26 := (uf_4 uf_5 #25) +decl uf_7 :: T2 +#27 := uf_7 +#28 := (uf_4 uf_7 #25) +decl uf_10 :: T1 +#38 := uf_10 +#42 := (uf_6 uf_10) +decl uf_9 :: T2 +#33 := uf_9 +#43 := (uf_4 uf_9 #42) +#41 := (= uf_3 uf_10) +#44 := (ite #41 #43 #28) +#9 := 0::int +decl uf_2 :: (-> T1 int) +#39 := (uf_2 uf_10) +#226 := -1::int +#229 := (* -1::int #39) +#22 := (uf_2 uf_3) +#230 := (+ #22 #229) +#228 := (>= #230 0::int) +#236 := (ite #228 #44 #26) +#192 := -1::real +#244 := (* -1::real #236) +#642 := (+ #26 #244) +#643 := (<= #642 0::real) +#567 := (= #26 #236) +#227 := (not #228) +decl uf_1 :: (-> int T1) +#593 := (uf_1 #39) +#660 := (= #593 uf_10) +#594 := (= uf_10 #593) +#4 := (:var 0 T1) +#5 := (uf_2 #4) +#546 := (pattern #5) +#6 := (uf_1 #5) +#93 := (= #4 #6) +#547 := (forall (vars (?x1 T1)) (:pat #546) #93) +#96 := (forall (vars (?x1 T1)) #93) +#550 := (iff #96 #547) +#548 := (iff #93 #93) +#549 := [refl]: #548 +#551 := [quant-intro #549]: #550 +#448 := (~ #96 #96) +#450 := (~ #93 #93) +#451 := [refl]: #450 +#449 := [nnf-pos #451]: #448 +#7 := (= #6 #4) +#8 := (forall (vars (?x1 T1)) #7) +#97 := (iff #8 #96) +#94 := (iff #7 #93) +#95 := [rewrite]: #94 +#98 := [quant-intro #95]: #97 +#92 := [asserted]: #8 +#101 := [mp #92 #98]: #96 +#446 := [mp~ #101 #449]: #96 +#552 := [mp #446 #551]: #547 +#595 := (not #547) +#600 := (or #595 #594) +#601 := [quant-inst]: #600 +#654 := [unit-resolution #601 #552]: #594 +#680 := [symm #654]: #660 +#681 := (= uf_3 #593) +#591 := (uf_1 #22) +#658 := (= #591 #593) +#656 := (= #593 #591) +#652 := (= #39 #22) +#647 := (= #22 #39) +#290 := (<= #230 0::int) +#70 := (<= #22 #39) +#388 := (iff #70 #290) +#389 := [rewrite]: #388 +#341 := [asserted]: #70 +#390 := [mp #341 #389]: #290 +#646 := [hypothesis]: #228 +#648 := [th-lemma #646 #390]: #647 +#653 := [symm #648]: #652 +#657 := [monotonicity #653]: #656 +#659 := [symm #657]: #658 +#592 := (= uf_3 #591) +#596 := (or #595 #592) +#597 := [quant-inst]: #596 +#655 := [unit-resolution #597 #552]: #592 +#682 := [trans #655 #659]: #681 +#683 := [trans #682 #680]: #41 +#570 := (not #41) +decl uf_11 :: T2 +#47 := uf_11 +#59 := (uf_4 uf_11 #42) +#278 := (ite #41 #26 #59) +#459 := (* -1::real #278) +#637 := (+ #26 #459) +#639 := (>= #637 0::real) +#585 := (= #26 #278) +#661 := [hypothesis]: #41 +#587 := (or #570 #585) +#588 := [def-axiom]: #587 +#662 := [unit-resolution #588 #661]: #585 +#663 := (not #585) +#664 := (or #663 #639) +#665 := [th-lemma]: #664 +#666 := [unit-resolution #665 #662]: #639 +decl uf_8 :: T2 +#30 := uf_8 +#56 := (uf_4 uf_8 #42) +#357 := (* -1::real #56) +#358 := (+ #43 #357) +#356 := (>= #358 0::real) +#355 := (not #356) +#374 := (* -1::real #59) +#375 := (+ #56 #374) +#373 := (>= #375 0::real) +#376 := (not #373) +#381 := (and #355 #376) +#64 := (< #39 #39) +#67 := (ite #64 #43 #59) +#68 := (< #56 #67) +#53 := (uf_4 uf_5 #42) +#65 := (ite #64 #53 #43) +#66 := (< #65 #56) +#69 := (and #66 #68) +#382 := (iff #69 #381) +#379 := (iff #68 #376) +#370 := (< #56 #59) +#377 := (iff #370 #376) +#378 := [rewrite]: #377 +#371 := (iff #68 #370) +#368 := (= #67 #59) +#363 := (ite false #43 #59) +#366 := (= #363 #59) +#367 := [rewrite]: #366 +#364 := (= #67 #363) +#343 := (iff #64 false) +#344 := [rewrite]: #343 +#365 := [monotonicity #344]: #364 +#369 := [trans #365 #367]: #368 +#372 := [monotonicity #369]: #371 +#380 := [trans #372 #378]: #379 +#361 := (iff #66 #355) +#352 := (< #43 #56) +#359 := (iff #352 #355) +#360 := [rewrite]: #359 +#353 := (iff #66 #352) +#350 := (= #65 #43) +#345 := (ite false #53 #43) +#348 := (= #345 #43) +#349 := [rewrite]: #348 +#346 := (= #65 #345) +#347 := [monotonicity #344]: #346 +#351 := [trans #347 #349]: #350 +#354 := [monotonicity #351]: #353 +#362 := [trans #354 #360]: #361 +#383 := [monotonicity #362 #380]: #382 +#340 := [asserted]: #69 +#384 := [mp #340 #383]: #381 +#385 := [and-elim #384]: #355 +#394 := (* -1::real #53) +#395 := (+ #43 #394) +#393 := (>= #395 0::real) +#54 := (uf_4 uf_7 #42) +#402 := (* -1::real #54) +#403 := (+ #53 #402) +#401 := (>= #403 0::real) +#397 := (+ #43 #374) +#398 := (<= #397 0::real) +#412 := (and #393 #398 #401) +#73 := (<= #43 #59) +#72 := (<= #53 #43) +#74 := (and #72 #73) +#71 := (<= #54 #53) +#75 := (and #71 #74) +#415 := (iff #75 #412) +#406 := (and #393 #398) +#409 := (and #401 #406) +#413 := (iff #409 #412) +#414 := [rewrite]: #413 +#410 := (iff #75 #409) +#407 := (iff #74 #406) +#399 := (iff #73 #398) +#400 := [rewrite]: #399 +#392 := (iff #72 #393) +#396 := [rewrite]: #392 +#408 := [monotonicity #396 #400]: #407 +#404 := (iff #71 #401) +#405 := [rewrite]: #404 +#411 := [monotonicity #405 #408]: #410 +#416 := [trans #411 #414]: #415 +#342 := [asserted]: #75 +#417 := [mp #342 #416]: #412 +#418 := [and-elim #417]: #393 +#650 := (+ #26 #394) +#651 := (<= #650 0::real) +#649 := (= #26 #53) +#671 := (= #53 #26) +#669 := (= #42 #25) +#667 := (= #25 #42) +#668 := [monotonicity #661]: #667 +#670 := [symm #668]: #669 +#672 := [monotonicity #670]: #671 +#673 := [symm #672]: #649 +#674 := (not #649) +#675 := (or #674 #651) +#676 := [th-lemma]: #675 +#677 := [unit-resolution #676 #673]: #651 +#462 := (+ #56 #459) +#465 := (>= #462 0::real) +#438 := (not #465) +#316 := (ite #290 #278 #43) +#326 := (* -1::real #316) +#327 := (+ #56 #326) +#325 := (>= #327 0::real) +#324 := (not #325) +#439 := (iff #324 #438) +#466 := (iff #325 #465) +#463 := (= #327 #462) +#460 := (= #326 #459) +#457 := (= #316 #278) +#1 := true +#452 := (ite true #278 #43) +#455 := (= #452 #278) +#456 := [rewrite]: #455 +#453 := (= #316 #452) +#444 := (iff #290 true) +#445 := [iff-true #390]: #444 +#454 := [monotonicity #445]: #453 +#458 := [trans #454 #456]: #457 +#461 := [monotonicity #458]: #460 +#464 := [monotonicity #461]: #463 +#467 := [monotonicity #464]: #466 +#468 := [monotonicity #467]: #439 +#297 := (ite #290 #54 #53) +#305 := (* -1::real #297) +#306 := (+ #56 #305) +#307 := (<= #306 0::real) +#308 := (not #307) +#332 := (and #308 #324) +#58 := (= uf_10 uf_3) +#60 := (ite #58 #26 #59) +#52 := (< #39 #22) +#61 := (ite #52 #43 #60) +#62 := (< #56 #61) +#55 := (ite #52 #53 #54) +#57 := (< #55 #56) +#63 := (and #57 #62) +#335 := (iff #63 #332) +#281 := (ite #52 #43 #278) +#284 := (< #56 #281) +#287 := (and #57 #284) +#333 := (iff #287 #332) +#330 := (iff #284 #324) +#321 := (< #56 #316) +#328 := (iff #321 #324) +#329 := [rewrite]: #328 +#322 := (iff #284 #321) +#319 := (= #281 #316) +#291 := (not #290) +#313 := (ite #291 #43 #278) +#317 := (= #313 #316) +#318 := [rewrite]: #317 +#314 := (= #281 #313) +#292 := (iff #52 #291) +#293 := [rewrite]: #292 +#315 := [monotonicity #293]: #314 +#320 := [trans #315 #318]: #319 +#323 := [monotonicity #320]: #322 +#331 := [trans #323 #329]: #330 +#311 := (iff #57 #308) +#302 := (< #297 #56) +#309 := (iff #302 #308) +#310 := [rewrite]: #309 +#303 := (iff #57 #302) +#300 := (= #55 #297) +#294 := (ite #291 #53 #54) +#298 := (= #294 #297) +#299 := [rewrite]: #298 +#295 := (= #55 #294) +#296 := [monotonicity #293]: #295 +#301 := [trans #296 #299]: #300 +#304 := [monotonicity #301]: #303 +#312 := [trans #304 #310]: #311 +#334 := [monotonicity #312 #331]: #333 +#288 := (iff #63 #287) +#285 := (iff #62 #284) +#282 := (= #61 #281) +#279 := (= #60 #278) +#225 := (iff #58 #41) +#277 := [rewrite]: #225 +#280 := [monotonicity #277]: #279 +#283 := [monotonicity #280]: #282 +#286 := [monotonicity #283]: #285 +#289 := [monotonicity #286]: #288 +#336 := [trans #289 #334]: #335 +#179 := [asserted]: #63 +#337 := [mp #179 #336]: #332 +#339 := [and-elim #337]: #324 +#469 := [mp #339 #468]: #438 +#678 := [th-lemma #469 #677 #418 #385 #666]: false +#679 := [lemma #678]: #570 +#684 := [unit-resolution #679 #683]: false +#685 := [lemma #684]: #227 +#577 := (or #228 #567) +#578 := [def-axiom]: #577 +#645 := [unit-resolution #578 #685]: #567 +#686 := (not #567) +#687 := (or #686 #643) +#688 := [th-lemma]: #687 +#689 := [unit-resolution #688 #645]: #643 +#31 := (uf_4 uf_8 #25) +#245 := (+ #31 #244) +#246 := (<= #245 0::real) +#247 := (not #246) +#34 := (uf_4 uf_9 #25) +#48 := (uf_4 uf_11 #25) +#255 := (ite #228 #48 #34) +#264 := (* -1::real #255) +#265 := (+ #31 #264) +#263 := (>= #265 0::real) +#266 := (not #263) +#271 := (and #247 #266) +#40 := (< #22 #39) +#49 := (ite #40 #34 #48) +#50 := (< #31 #49) +#45 := (ite #40 #26 #44) +#46 := (< #45 #31) +#51 := (and #46 #50) +#272 := (iff #51 #271) +#269 := (iff #50 #266) +#260 := (< #31 #255) +#267 := (iff #260 #266) +#268 := [rewrite]: #267 +#261 := (iff #50 #260) +#258 := (= #49 #255) +#252 := (ite #227 #34 #48) +#256 := (= #252 #255) +#257 := [rewrite]: #256 +#253 := (= #49 #252) +#231 := (iff #40 #227) +#232 := [rewrite]: #231 +#254 := [monotonicity #232]: #253 +#259 := [trans #254 #257]: #258 +#262 := [monotonicity #259]: #261 +#270 := [trans #262 #268]: #269 +#250 := (iff #46 #247) +#241 := (< #236 #31) +#248 := (iff #241 #247) +#249 := [rewrite]: #248 +#242 := (iff #46 #241) +#239 := (= #45 #236) +#233 := (ite #227 #26 #44) +#237 := (= #233 #236) +#238 := [rewrite]: #237 +#234 := (= #45 #233) +#235 := [monotonicity #232]: #234 +#240 := [trans #235 #238]: #239 +#243 := [monotonicity #240]: #242 +#251 := [trans #243 #249]: #250 +#273 := [monotonicity #251 #270]: #272 +#178 := [asserted]: #51 +#274 := [mp #178 #273]: #271 +#275 := [and-elim #274]: #247 +#196 := (* -1::real #31) +#212 := (+ #26 #196) +#213 := (<= #212 0::real) +#214 := (not #213) +#197 := (+ #28 #196) +#195 := (>= #197 0::real) +#193 := (not #195) +#219 := (and #193 #214) +#23 := (< #22 #22) +#35 := (ite #23 #34 #26) +#36 := (< #31 #35) +#29 := (ite #23 #26 #28) +#32 := (< #29 #31) +#37 := (and #32 #36) +#220 := (iff #37 #219) +#217 := (iff #36 #214) +#209 := (< #31 #26) +#215 := (iff #209 #214) +#216 := [rewrite]: #215 +#210 := (iff #36 #209) +#207 := (= #35 #26) +#202 := (ite false #34 #26) +#205 := (= #202 #26) +#206 := [rewrite]: #205 +#203 := (= #35 #202) +#180 := (iff #23 false) +#181 := [rewrite]: #180 +#204 := [monotonicity #181]: #203 +#208 := [trans #204 #206]: #207 +#211 := [monotonicity #208]: #210 +#218 := [trans #211 #216]: #217 +#200 := (iff #32 #193) +#189 := (< #28 #31) +#198 := (iff #189 #193) +#199 := [rewrite]: #198 +#190 := (iff #32 #189) +#187 := (= #29 #28) +#182 := (ite false #26 #28) +#185 := (= #182 #28) +#186 := [rewrite]: #185 +#183 := (= #29 #182) +#184 := [monotonicity #181]: #183 +#188 := [trans #184 #186]: #187 +#191 := [monotonicity #188]: #190 +#201 := [trans #191 #199]: #200 +#221 := [monotonicity #201 #218]: #220 +#177 := [asserted]: #37 +#222 := [mp #177 #221]: #219 +#224 := [and-elim #222]: #214 +[th-lemma #224 #275 #689]: false +unsat +NX/HT1QOfbspC2LtZNKpBA 428 0 +#2 := false +decl uf_10 :: T1 +#38 := uf_10 +decl uf_3 :: T1 +#21 := uf_3 +#45 := (= uf_3 uf_10) +decl uf_1 :: (-> int T1) +decl uf_2 :: (-> T1 int) +#39 := (uf_2 uf_10) +#588 := (uf_1 #39) +#686 := (= #588 uf_10) +#589 := (= uf_10 #588) +#4 := (:var 0 T1) +#5 := (uf_2 #4) +#541 := (pattern #5) +#6 := (uf_1 #5) +#93 := (= #4 #6) +#542 := (forall (vars (?x1 T1)) (:pat #541) #93) +#96 := (forall (vars (?x1 T1)) #93) +#545 := (iff #96 #542) +#543 := (iff #93 #93) +#544 := [refl]: #543 +#546 := [quant-intro #544]: #545 +#454 := (~ #96 #96) +#456 := (~ #93 #93) +#457 := [refl]: #456 +#455 := [nnf-pos #457]: #454 +#7 := (= #6 #4) +#8 := (forall (vars (?x1 T1)) #7) +#97 := (iff #8 #96) +#94 := (iff #7 #93) +#95 := [rewrite]: #94 +#98 := [quant-intro #95]: #97 +#92 := [asserted]: #8 +#101 := [mp #92 #98]: #96 +#452 := [mp~ #101 #455]: #96 +#547 := [mp #452 #546]: #542 +#590 := (not #542) +#595 := (or #590 #589) +#596 := [quant-inst]: #595 +#680 := [unit-resolution #596 #547]: #589 +#687 := [symm #680]: #686 +#688 := (= uf_3 #588) +#22 := (uf_2 uf_3) +#586 := (uf_1 #22) +#684 := (= #586 #588) +#682 := (= #588 #586) +#678 := (= #39 #22) +#676 := (= #22 #39) +#9 := 0::int +#227 := -1::int +#230 := (* -1::int #39) +#231 := (+ #22 #230) +#296 := (<= #231 0::int) +#70 := (<= #22 #39) +#393 := (iff #70 #296) +#394 := [rewrite]: #393 +#347 := [asserted]: #70 +#395 := [mp #347 #394]: #296 +#229 := (>= #231 0::int) +decl uf_4 :: (-> T2 T3 real) +decl uf_6 :: (-> T1 T3) +#25 := (uf_6 uf_3) +decl uf_7 :: T2 +#27 := uf_7 +#28 := (uf_4 uf_7 #25) +decl uf_9 :: T2 +#33 := uf_9 +#34 := (uf_4 uf_9 #25) +#46 := (uf_6 uf_10) +decl uf_5 :: T2 +#24 := uf_5 +#47 := (uf_4 uf_5 #46) +#48 := (ite #45 #47 #34) +#256 := (ite #229 #48 #28) +#568 := (= #28 #256) +#648 := (not #568) +#194 := 0::real +#192 := -1::real +#265 := (* -1::real #256) +#640 := (+ #28 #265) +#642 := (>= #640 0::real) +#645 := (not #642) +#643 := [hypothesis]: #642 +decl uf_8 :: T2 +#30 := uf_8 +#31 := (uf_4 uf_8 #25) +#266 := (+ #31 #265) +#264 := (>= #266 0::real) +#267 := (not #264) +#26 := (uf_4 uf_5 #25) +decl uf_11 :: T2 +#41 := uf_11 +#42 := (uf_4 uf_11 #25) +#237 := (ite #229 #42 #26) +#245 := (* -1::real #237) +#246 := (+ #31 #245) +#247 := (<= #246 0::real) +#248 := (not #247) +#272 := (and #248 #267) +#40 := (< #22 #39) +#49 := (ite #40 #28 #48) +#50 := (< #31 #49) +#43 := (ite #40 #26 #42) +#44 := (< #43 #31) +#51 := (and #44 #50) +#273 := (iff #51 #272) +#270 := (iff #50 #267) +#261 := (< #31 #256) +#268 := (iff #261 #267) +#269 := [rewrite]: #268 +#262 := (iff #50 #261) +#259 := (= #49 #256) +#228 := (not #229) +#253 := (ite #228 #28 #48) +#257 := (= #253 #256) +#258 := [rewrite]: #257 +#254 := (= #49 #253) +#232 := (iff #40 #228) +#233 := [rewrite]: #232 +#255 := [monotonicity #233]: #254 +#260 := [trans #255 #258]: #259 +#263 := [monotonicity #260]: #262 +#271 := [trans #263 #269]: #270 +#251 := (iff #44 #248) +#242 := (< #237 #31) +#249 := (iff #242 #248) +#250 := [rewrite]: #249 +#243 := (iff #44 #242) +#240 := (= #43 #237) +#234 := (ite #228 #26 #42) +#238 := (= #234 #237) +#239 := [rewrite]: #238 +#235 := (= #43 #234) +#236 := [monotonicity #233]: #235 +#241 := [trans #236 #239]: #240 +#244 := [monotonicity #241]: #243 +#252 := [trans #244 #250]: #251 +#274 := [monotonicity #252 #271]: #273 +#178 := [asserted]: #51 +#275 := [mp #178 #274]: #272 +#277 := [and-elim #275]: #267 +#196 := (* -1::real #31) +#197 := (+ #28 #196) +#195 := (>= #197 0::real) +#193 := (not #195) +#213 := (* -1::real #34) +#214 := (+ #31 #213) +#212 := (>= #214 0::real) +#215 := (not #212) +#220 := (and #193 #215) +#23 := (< #22 #22) +#35 := (ite #23 #28 #34) +#36 := (< #31 #35) +#29 := (ite #23 #26 #28) +#32 := (< #29 #31) +#37 := (and #32 #36) +#221 := (iff #37 #220) +#218 := (iff #36 #215) +#209 := (< #31 #34) +#216 := (iff #209 #215) +#217 := [rewrite]: #216 +#210 := (iff #36 #209) +#207 := (= #35 #34) +#202 := (ite false #28 #34) +#205 := (= #202 #34) +#206 := [rewrite]: #205 +#203 := (= #35 #202) +#180 := (iff #23 false) +#181 := [rewrite]: #180 +#204 := [monotonicity #181]: #203 +#208 := [trans #204 #206]: #207 +#211 := [monotonicity #208]: #210 +#219 := [trans #211 #217]: #218 +#200 := (iff #32 #193) +#189 := (< #28 #31) +#198 := (iff #189 #193) +#199 := [rewrite]: #198 +#190 := (iff #32 #189) +#187 := (= #29 #28) +#182 := (ite false #26 #28) +#185 := (= #182 #28) +#186 := [rewrite]: #185 +#183 := (= #29 #182) +#184 := [monotonicity #181]: #183 +#188 := [trans #184 #186]: #187 +#191 := [monotonicity #188]: #190 +#201 := [trans #191 #199]: #200 +#222 := [monotonicity #201 #219]: #221 +#177 := [asserted]: #37 +#223 := [mp #177 #222]: #220 +#224 := [and-elim #223]: #193 +#644 := [th-lemma #224 #277 #643]: false +#646 := [lemma #644]: #645 +#647 := [hypothesis]: #568 +#649 := (or #648 #642) +#650 := [th-lemma]: #649 +#651 := [unit-resolution #650 #647 #646]: false +#652 := [lemma #651]: #648 +#578 := (or #229 #568) +#579 := [def-axiom]: #578 +#675 := [unit-resolution #579 #652]: #229 +#677 := [th-lemma #675 #395]: #676 +#679 := [symm #677]: #678 +#683 := [monotonicity #679]: #682 +#685 := [symm #683]: #684 +#587 := (= uf_3 #586) +#591 := (or #590 #587) +#592 := [quant-inst]: #591 +#681 := [unit-resolution #592 #547]: #587 +#689 := [trans #681 #685]: #688 +#690 := [trans #689 #687]: #45 +#571 := (not #45) +#54 := (uf_4 uf_11 #46) +#279 := (ite #45 #28 #54) +#465 := (* -1::real #279) +#632 := (+ #28 #465) +#633 := (<= #632 0::real) +#580 := (= #28 #279) +#656 := [hypothesis]: #45 +#582 := (or #571 #580) +#583 := [def-axiom]: #582 +#657 := [unit-resolution #583 #656]: #580 +#658 := (not #580) +#659 := (or #658 #633) +#660 := [th-lemma]: #659 +#661 := [unit-resolution #660 #657]: #633 +#57 := (uf_4 uf_8 #46) +#363 := (* -1::real #57) +#379 := (+ #47 #363) +#380 := (<= #379 0::real) +#381 := (not #380) +#364 := (+ #54 #363) +#362 := (>= #364 0::real) +#361 := (not #362) +#386 := (and #361 #381) +#59 := (uf_4 uf_7 #46) +#64 := (< #39 #39) +#67 := (ite #64 #59 #47) +#68 := (< #57 #67) +#65 := (ite #64 #47 #54) +#66 := (< #65 #57) +#69 := (and #66 #68) +#387 := (iff #69 #386) +#384 := (iff #68 #381) +#376 := (< #57 #47) +#382 := (iff #376 #381) +#383 := [rewrite]: #382 +#377 := (iff #68 #376) +#374 := (= #67 #47) +#369 := (ite false #59 #47) +#372 := (= #369 #47) +#373 := [rewrite]: #372 +#370 := (= #67 #369) +#349 := (iff #64 false) +#350 := [rewrite]: #349 +#371 := [monotonicity #350]: #370 +#375 := [trans #371 #373]: #374 +#378 := [monotonicity #375]: #377 +#385 := [trans #378 #383]: #384 +#367 := (iff #66 #361) +#358 := (< #54 #57) +#365 := (iff #358 #361) +#366 := [rewrite]: #365 +#359 := (iff #66 #358) +#356 := (= #65 #54) +#351 := (ite false #47 #54) +#354 := (= #351 #54) +#355 := [rewrite]: #354 +#352 := (= #65 #351) +#353 := [monotonicity #350]: #352 +#357 := [trans #353 #355]: #356 +#360 := [monotonicity #357]: #359 +#368 := [trans #360 #366]: #367 +#388 := [monotonicity #368 #385]: #387 +#346 := [asserted]: #69 +#389 := [mp #346 #388]: #386 +#391 := [and-elim #389]: #381 +#397 := (* -1::real #59) +#398 := (+ #47 #397) +#399 := (<= #398 0::real) +#409 := (* -1::real #54) +#410 := (+ #47 #409) +#408 := (>= #410 0::real) +#60 := (uf_4 uf_9 #46) +#402 := (* -1::real #60) +#403 := (+ #59 #402) +#404 := (<= #403 0::real) +#418 := (and #399 #404 #408) +#73 := (<= #59 #60) +#72 := (<= #47 #59) +#74 := (and #72 #73) +#71 := (<= #54 #47) +#75 := (and #71 #74) +#421 := (iff #75 #418) +#412 := (and #399 #404) +#415 := (and #408 #412) +#419 := (iff #415 #418) +#420 := [rewrite]: #419 +#416 := (iff #75 #415) +#413 := (iff #74 #412) +#405 := (iff #73 #404) +#406 := [rewrite]: #405 +#400 := (iff #72 #399) +#401 := [rewrite]: #400 +#414 := [monotonicity #401 #406]: #413 +#407 := (iff #71 #408) +#411 := [rewrite]: #407 +#417 := [monotonicity #411 #414]: #416 +#422 := [trans #417 #420]: #421 +#348 := [asserted]: #75 +#423 := [mp #348 #422]: #418 +#424 := [and-elim #423]: #399 +#637 := (+ #28 #397) +#639 := (>= #637 0::real) +#636 := (= #28 #59) +#666 := (= #59 #28) +#664 := (= #46 #25) +#662 := (= #25 #46) +#663 := [monotonicity #656]: #662 +#665 := [symm #663]: #664 +#667 := [monotonicity #665]: #666 +#668 := [symm #667]: #636 +#669 := (not #636) +#670 := (or #669 #639) +#671 := [th-lemma]: #670 +#672 := [unit-resolution #671 #668]: #639 +#468 := (+ #57 #465) +#471 := (<= #468 0::real) +#444 := (not #471) +#322 := (ite #296 #279 #47) +#330 := (* -1::real #322) +#331 := (+ #57 #330) +#332 := (<= #331 0::real) +#333 := (not #332) +#445 := (iff #333 #444) +#472 := (iff #332 #471) +#469 := (= #331 #468) +#466 := (= #330 #465) +#463 := (= #322 #279) +#1 := true +#458 := (ite true #279 #47) +#461 := (= #458 #279) +#462 := [rewrite]: #461 +#459 := (= #322 #458) +#450 := (iff #296 true) +#451 := [iff-true #395]: #450 +#460 := [monotonicity #451]: #459 +#464 := [trans #460 #462]: #463 +#467 := [monotonicity #464]: #466 +#470 := [monotonicity #467]: #469 +#473 := [monotonicity #470]: #472 +#474 := [monotonicity #473]: #445 +#303 := (ite #296 #60 #59) +#313 := (* -1::real #303) +#314 := (+ #57 #313) +#312 := (>= #314 0::real) +#311 := (not #312) +#338 := (and #311 #333) +#52 := (< #39 #22) +#61 := (ite #52 #59 #60) +#62 := (< #57 #61) +#53 := (= uf_10 uf_3) +#55 := (ite #53 #28 #54) +#56 := (ite #52 #47 #55) +#58 := (< #56 #57) +#63 := (and #58 #62) +#341 := (iff #63 #338) +#282 := (ite #52 #47 #279) +#285 := (< #282 #57) +#291 := (and #62 #285) +#339 := (iff #291 #338) +#336 := (iff #285 #333) +#327 := (< #322 #57) +#334 := (iff #327 #333) +#335 := [rewrite]: #334 +#328 := (iff #285 #327) +#325 := (= #282 #322) +#297 := (not #296) +#319 := (ite #297 #47 #279) +#323 := (= #319 #322) +#324 := [rewrite]: #323 +#320 := (= #282 #319) +#298 := (iff #52 #297) +#299 := [rewrite]: #298 +#321 := [monotonicity #299]: #320 +#326 := [trans #321 #324]: #325 +#329 := [monotonicity #326]: #328 +#337 := [trans #329 #335]: #336 +#317 := (iff #62 #311) +#308 := (< #57 #303) +#315 := (iff #308 #311) +#316 := [rewrite]: #315 +#309 := (iff #62 #308) +#306 := (= #61 #303) +#300 := (ite #297 #59 #60) +#304 := (= #300 #303) +#305 := [rewrite]: #304 +#301 := (= #61 #300) +#302 := [monotonicity #299]: #301 +#307 := [trans #302 #305]: #306 +#310 := [monotonicity #307]: #309 +#318 := [trans #310 #316]: #317 +#340 := [monotonicity #318 #337]: #339 +#294 := (iff #63 #291) +#288 := (and #285 #62) +#292 := (iff #288 #291) +#293 := [rewrite]: #292 +#289 := (iff #63 #288) +#286 := (iff #58 #285) +#283 := (= #56 #282) +#280 := (= #55 #279) +#226 := (iff #53 #45) +#278 := [rewrite]: #226 +#281 := [monotonicity #278]: #280 +#284 := [monotonicity #281]: #283 +#287 := [monotonicity #284]: #286 +#290 := [monotonicity #287]: #289 +#295 := [trans #290 #293]: #294 +#342 := [trans #295 #340]: #341 +#179 := [asserted]: #63 +#343 := [mp #179 #342]: #338 +#345 := [and-elim #343]: #333 +#475 := [mp #345 #474]: #444 +#673 := [th-lemma #475 #672 #424 #391 #661]: false +#674 := [lemma #673]: #571 +[unit-resolution #674 #690]: false +unsat +IL2powemHjRpCJYwmXFxyw 211 0 +#2 := false +#33 := 0::real +decl uf_11 :: (-> T5 T6 real) +decl uf_15 :: T6 +#28 := uf_15 +decl uf_16 :: T5 +#30 := uf_16 +#31 := (uf_11 uf_16 uf_15) +decl uf_12 :: (-> T7 T8 T5) +decl uf_14 :: T8 +#26 := uf_14 +decl uf_13 :: (-> T1 T7) +decl uf_8 :: T1 +#16 := uf_8 +#25 := (uf_13 uf_8) +#27 := (uf_12 #25 uf_14) +#29 := (uf_11 #27 uf_15) +#73 := -1::real +#84 := (* -1::real #29) +#85 := (+ #84 #31) +#74 := (* -1::real #31) +#75 := (+ #29 #74) +#112 := (>= #75 0::real) +#119 := (ite #112 #75 #85) +#127 := (* -1::real #119) +decl uf_17 :: T5 +#37 := uf_17 +#38 := (uf_11 uf_17 uf_15) +#102 := -1/3::real +#103 := (* -1/3::real #38) +#128 := (+ #103 #127) +#100 := 1/3::real +#101 := (* 1/3::real #31) +#129 := (+ #101 #128) +#130 := (<= #129 0::real) +#131 := (not #130) +#40 := 3::real +#39 := (- #31 #38) +#41 := (/ #39 3::real) +#32 := (- #29 #31) +#35 := (- #32) +#34 := (< #32 0::real) +#36 := (ite #34 #35 #32) +#42 := (< #36 #41) +#136 := (iff #42 #131) +#104 := (+ #101 #103) +#78 := (< #75 0::real) +#90 := (ite #78 #85 #75) +#109 := (< #90 #104) +#134 := (iff #109 #131) +#124 := (< #119 #104) +#132 := (iff #124 #131) +#133 := [rewrite]: #132 +#125 := (iff #109 #124) +#122 := (= #90 #119) +#113 := (not #112) +#116 := (ite #113 #85 #75) +#120 := (= #116 #119) +#121 := [rewrite]: #120 +#117 := (= #90 #116) +#114 := (iff #78 #113) +#115 := [rewrite]: #114 +#118 := [monotonicity #115]: #117 +#123 := [trans #118 #121]: #122 +#126 := [monotonicity #123]: #125 +#135 := [trans #126 #133]: #134 +#110 := (iff #42 #109) +#107 := (= #41 #104) +#93 := (* -1::real #38) +#94 := (+ #31 #93) +#97 := (/ #94 3::real) +#105 := (= #97 #104) +#106 := [rewrite]: #105 +#98 := (= #41 #97) +#95 := (= #39 #94) +#96 := [rewrite]: #95 +#99 := [monotonicity #96]: #98 +#108 := [trans #99 #106]: #107 +#91 := (= #36 #90) +#76 := (= #32 #75) +#77 := [rewrite]: #76 +#88 := (= #35 #85) +#81 := (- #75) +#86 := (= #81 #85) +#87 := [rewrite]: #86 +#82 := (= #35 #81) +#83 := [monotonicity #77]: #82 +#89 := [trans #83 #87]: #88 +#79 := (iff #34 #78) +#80 := [monotonicity #77]: #79 +#92 := [monotonicity #80 #89 #77]: #91 +#111 := [monotonicity #92 #108]: #110 +#137 := [trans #111 #135]: #136 +#72 := [asserted]: #42 +#138 := [mp #72 #137]: #131 +decl uf_1 :: T1 +#4 := uf_1 +#43 := (uf_13 uf_1) +#44 := (uf_12 #43 uf_14) +#45 := (uf_11 #44 uf_15) +#149 := (* -1::real #45) +#150 := (+ #38 #149) +#140 := (+ #93 #45) +#161 := (<= #150 0::real) +#168 := (ite #161 #140 #150) +#176 := (* -1::real #168) +#177 := (+ #103 #176) +#178 := (+ #101 #177) +#179 := (<= #178 0::real) +#180 := (not #179) +#46 := (- #45 #38) +#48 := (- #46) +#47 := (< #46 0::real) +#49 := (ite #47 #48 #46) +#50 := (< #49 #41) +#185 := (iff #50 #180) +#143 := (< #140 0::real) +#155 := (ite #143 #150 #140) +#158 := (< #155 #104) +#183 := (iff #158 #180) +#173 := (< #168 #104) +#181 := (iff #173 #180) +#182 := [rewrite]: #181 +#174 := (iff #158 #173) +#171 := (= #155 #168) +#162 := (not #161) +#165 := (ite #162 #150 #140) +#169 := (= #165 #168) +#170 := [rewrite]: #169 +#166 := (= #155 #165) +#163 := (iff #143 #162) +#164 := [rewrite]: #163 +#167 := [monotonicity #164]: #166 +#172 := [trans #167 #170]: #171 +#175 := [monotonicity #172]: #174 +#184 := [trans #175 #182]: #183 +#159 := (iff #50 #158) +#156 := (= #49 #155) +#141 := (= #46 #140) +#142 := [rewrite]: #141 +#153 := (= #48 #150) +#146 := (- #140) +#151 := (= #146 #150) +#152 := [rewrite]: #151 +#147 := (= #48 #146) +#148 := [monotonicity #142]: #147 +#154 := [trans #148 #152]: #153 +#144 := (iff #47 #143) +#145 := [monotonicity #142]: #144 +#157 := [monotonicity #145 #154 #142]: #156 +#160 := [monotonicity #157 #108]: #159 +#186 := [trans #160 #184]: #185 +#139 := [asserted]: #50 +#187 := [mp #139 #186]: #180 +#299 := (+ #140 #176) +#300 := (<= #299 0::real) +#290 := (= #140 #168) +#329 := [hypothesis]: #162 +#191 := (+ #29 #149) +#192 := (<= #191 0::real) +#51 := (<= #29 #45) +#193 := (iff #51 #192) +#194 := [rewrite]: #193 +#188 := [asserted]: #51 +#195 := [mp #188 #194]: #192 +#298 := (+ #75 #127) +#301 := (<= #298 0::real) +#284 := (= #75 #119) +#302 := [hypothesis]: #113 +#296 := (+ #85 #127) +#297 := (<= #296 0::real) +#285 := (= #85 #119) +#288 := (or #112 #285) +#289 := [def-axiom]: #288 +#303 := [unit-resolution #289 #302]: #285 +#304 := (not #285) +#305 := (or #304 #297) +#306 := [th-lemma]: #305 +#307 := [unit-resolution #306 #303]: #297 +#315 := (not #290) +#310 := (not #300) +#311 := (or #310 #112) +#308 := [hypothesis]: #300 +#309 := [th-lemma #308 #307 #138 #302 #187 #195]: false +#312 := [lemma #309]: #311 +#322 := [unit-resolution #312 #302]: #310 +#316 := (or #315 #300) +#313 := [hypothesis]: #310 +#314 := [hypothesis]: #290 +#317 := [th-lemma]: #316 +#318 := [unit-resolution #317 #314 #313]: false +#319 := [lemma #318]: #316 +#323 := [unit-resolution #319 #322]: #315 +#292 := (or #162 #290) +#293 := [def-axiom]: #292 +#324 := [unit-resolution #293 #323]: #162 +#325 := [th-lemma #324 #307 #138 #302 #195]: false +#326 := [lemma #325]: #112 +#286 := (or #113 #284) +#287 := [def-axiom]: #286 +#330 := [unit-resolution #287 #326]: #284 +#331 := (not #284) +#332 := (or #331 #301) +#333 := [th-lemma]: #332 +#334 := [unit-resolution #333 #330]: #301 +#335 := [th-lemma #326 #334 #195 #329 #138]: false +#336 := [lemma #335]: #161 +#327 := [unit-resolution #293 #336]: #290 +#328 := [unit-resolution #319 #327]: #300 +[th-lemma #326 #334 #195 #328 #187 #138]: false +unsat +GX51o3DUO/UBS3eNP2P9kA 285 0 +#2 := false +#7 := 0::real +decl uf_4 :: real +#16 := uf_4 +#40 := -1::real +#116 := (* -1::real uf_4) +decl uf_3 :: real +#11 := uf_3 +#117 := (+ uf_3 #116) +#128 := (<= #117 0::real) +#129 := (not #128) +#220 := 2/3::real +#221 := (* 2/3::real uf_3) +#222 := (+ #221 #116) +decl uf_2 :: real +#5 := uf_2 +#67 := 1/3::real +#68 := (* 1/3::real uf_2) +#233 := (+ #68 #222) +#243 := (<= #233 0::real) +#268 := (not #243) +#287 := [hypothesis]: #268 +#41 := (* -1::real uf_2) +decl uf_1 :: real +#4 := uf_1 +#42 := (+ uf_1 #41) +#79 := (>= #42 0::real) +#80 := (not #79) +#297 := (or #80 #243) +#158 := (+ uf_1 #116) +#159 := (<= #158 0::real) +#22 := (<= uf_1 uf_4) +#160 := (iff #22 #159) +#161 := [rewrite]: #160 +#155 := [asserted]: #22 +#162 := [mp #155 #161]: #159 +#200 := (* 1/3::real uf_3) +#198 := -4/3::real +#199 := (* -4/3::real uf_2) +#201 := (+ #199 #200) +#202 := (+ uf_1 #201) +#203 := (>= #202 0::real) +#258 := (not #203) +#292 := [hypothesis]: #79 +#293 := (or #80 #258) +#69 := -1/3::real +#70 := (* -1/3::real uf_3) +#186 := -2/3::real +#187 := (* -2/3::real uf_2) +#188 := (+ #187 #70) +#189 := (+ uf_1 #188) +#204 := (<= #189 0::real) +#205 := (ite #79 #203 #204) +#210 := (not #205) +#51 := (* -1::real uf_1) +#52 := (+ #51 uf_2) +#86 := (ite #79 #42 #52) +#94 := (* -1::real #86) +#95 := (+ #70 #94) +#96 := (+ #68 #95) +#97 := (<= #96 0::real) +#98 := (not #97) +#211 := (iff #98 #210) +#208 := (iff #97 #205) +#182 := 4/3::real +#183 := (* 4/3::real uf_2) +#184 := (+ #183 #70) +#185 := (+ #51 #184) +#190 := (ite #79 #185 #189) +#195 := (<= #190 0::real) +#206 := (iff #195 #205) +#207 := [rewrite]: #206 +#196 := (iff #97 #195) +#193 := (= #96 #190) +#172 := (+ #41 #70) +#173 := (+ uf_1 #172) +#170 := (+ uf_2 #70) +#171 := (+ #51 #170) +#174 := (ite #79 #171 #173) +#179 := (+ #68 #174) +#191 := (= #179 #190) +#192 := [rewrite]: #191 +#180 := (= #96 #179) +#177 := (= #95 #174) +#164 := (ite #79 #52 #42) +#167 := (+ #70 #164) +#175 := (= #167 #174) +#176 := [rewrite]: #175 +#168 := (= #95 #167) +#156 := (= #94 #164) +#165 := [rewrite]: #156 +#169 := [monotonicity #165]: #168 +#178 := [trans #169 #176]: #177 +#181 := [monotonicity #178]: #180 +#194 := [trans #181 #192]: #193 +#197 := [monotonicity #194]: #196 +#209 := [trans #197 #207]: #208 +#212 := [monotonicity #209]: #211 +#13 := 3::real +#12 := (- uf_2 uf_3) +#14 := (/ #12 3::real) +#6 := (- uf_1 uf_2) +#9 := (- #6) +#8 := (< #6 0::real) +#10 := (ite #8 #9 #6) +#15 := (< #10 #14) +#103 := (iff #15 #98) +#71 := (+ #68 #70) +#45 := (< #42 0::real) +#57 := (ite #45 #52 #42) +#76 := (< #57 #71) +#101 := (iff #76 #98) +#91 := (< #86 #71) +#99 := (iff #91 #98) +#100 := [rewrite]: #99 +#92 := (iff #76 #91) +#89 := (= #57 #86) +#83 := (ite #80 #52 #42) +#87 := (= #83 #86) +#88 := [rewrite]: #87 +#84 := (= #57 #83) +#81 := (iff #45 #80) +#82 := [rewrite]: #81 +#85 := [monotonicity #82]: #84 +#90 := [trans #85 #88]: #89 +#93 := [monotonicity #90]: #92 +#102 := [trans #93 #100]: #101 +#77 := (iff #15 #76) +#74 := (= #14 #71) +#60 := (* -1::real uf_3) +#61 := (+ uf_2 #60) +#64 := (/ #61 3::real) +#72 := (= #64 #71) +#73 := [rewrite]: #72 +#65 := (= #14 #64) +#62 := (= #12 #61) +#63 := [rewrite]: #62 +#66 := [monotonicity #63]: #65 +#75 := [trans #66 #73]: #74 +#58 := (= #10 #57) +#43 := (= #6 #42) +#44 := [rewrite]: #43 +#55 := (= #9 #52) +#48 := (- #42) +#53 := (= #48 #52) +#54 := [rewrite]: #53 +#49 := (= #9 #48) +#50 := [monotonicity #44]: #49 +#56 := [trans #50 #54]: #55 +#46 := (iff #8 #45) +#47 := [monotonicity #44]: #46 +#59 := [monotonicity #47 #56 #44]: #58 +#78 := [monotonicity #59 #75]: #77 +#104 := [trans #78 #102]: #103 +#39 := [asserted]: #15 +#105 := [mp #39 #104]: #98 +#213 := [mp #105 #212]: #210 +#259 := (or #205 #80 #258) +#260 := [def-axiom]: #259 +#294 := [unit-resolution #260 #213]: #293 +#295 := [unit-resolution #294 #292]: #258 +#296 := [th-lemma #287 #292 #295 #162]: false +#298 := [lemma #296]: #297 +#299 := [unit-resolution #298 #287]: #80 +#261 := (not #204) +#281 := (or #79 #261) +#262 := (or #205 #79 #261) +#263 := [def-axiom]: #262 +#282 := [unit-resolution #263 #213]: #281 +#300 := [unit-resolution #282 #299]: #261 +#290 := (or #79 #204 #243) +#276 := [hypothesis]: #261 +#288 := [hypothesis]: #80 +#289 := [th-lemma #288 #276 #162 #287]: false +#291 := [lemma #289]: #290 +#301 := [unit-resolution #291 #300 #299 #287]: false +#302 := [lemma #301]: #243 +#303 := (or #129 #268) +#223 := (* -4/3::real uf_3) +#224 := (+ #223 uf_4) +#234 := (+ #68 #224) +#244 := (<= #234 0::real) +#245 := (ite #128 #243 #244) +#250 := (not #245) +#107 := (+ #60 uf_4) +#135 := (ite #128 #107 #117) +#143 := (* -1::real #135) +#144 := (+ #70 #143) +#145 := (+ #68 #144) +#146 := (<= #145 0::real) +#147 := (not #146) +#251 := (iff #147 #250) +#248 := (iff #146 #245) +#235 := (ite #128 #233 #234) +#240 := (<= #235 0::real) +#246 := (iff #240 #245) +#247 := [rewrite]: #246 +#241 := (iff #146 #240) +#238 := (= #145 #235) +#225 := (ite #128 #222 #224) +#230 := (+ #68 #225) +#236 := (= #230 #235) +#237 := [rewrite]: #236 +#231 := (= #145 #230) +#228 := (= #144 #225) +#214 := (ite #128 #117 #107) +#217 := (+ #70 #214) +#226 := (= #217 #225) +#227 := [rewrite]: #226 +#218 := (= #144 #217) +#215 := (= #143 #214) +#216 := [rewrite]: #215 +#219 := [monotonicity #216]: #218 +#229 := [trans #219 #227]: #228 +#232 := [monotonicity #229]: #231 +#239 := [trans #232 #237]: #238 +#242 := [monotonicity #239]: #241 +#249 := [trans #242 #247]: #248 +#252 := [monotonicity #249]: #251 +#17 := (- uf_4 uf_3) +#19 := (- #17) +#18 := (< #17 0::real) +#20 := (ite #18 #19 #17) +#21 := (< #20 #14) +#152 := (iff #21 #147) +#110 := (< #107 0::real) +#122 := (ite #110 #117 #107) +#125 := (< #122 #71) +#150 := (iff #125 #147) +#140 := (< #135 #71) +#148 := (iff #140 #147) +#149 := [rewrite]: #148 +#141 := (iff #125 #140) +#138 := (= #122 #135) +#132 := (ite #129 #117 #107) +#136 := (= #132 #135) +#137 := [rewrite]: #136 +#133 := (= #122 #132) +#130 := (iff #110 #129) +#131 := [rewrite]: #130 +#134 := [monotonicity #131]: #133 +#139 := [trans #134 #137]: #138 +#142 := [monotonicity #139]: #141 +#151 := [trans #142 #149]: #150 +#126 := (iff #21 #125) +#123 := (= #20 #122) +#108 := (= #17 #107) +#109 := [rewrite]: #108 +#120 := (= #19 #117) +#113 := (- #107) +#118 := (= #113 #117) +#119 := [rewrite]: #118 +#114 := (= #19 #113) +#115 := [monotonicity #109]: #114 +#121 := [trans #115 #119]: #120 +#111 := (iff #18 #110) +#112 := [monotonicity #109]: #111 +#124 := [monotonicity #112 #121 #109]: #123 +#127 := [monotonicity #124 #75]: #126 +#153 := [trans #127 #151]: #152 +#106 := [asserted]: #21 +#154 := [mp #106 #153]: #147 +#253 := [mp #154 #252]: #250 +#269 := (or #245 #129 #268) +#270 := [def-axiom]: #269 +#304 := [unit-resolution #270 #253]: #303 +#305 := [unit-resolution #304 #302]: #129 +#271 := (not #244) +#306 := (or #128 #271) +#272 := (or #245 #128 #271) +#273 := [def-axiom]: #272 +#307 := [unit-resolution #273 #253]: #306 +#308 := [unit-resolution #307 #305]: #271 +#285 := (or #128 #244) +#274 := [hypothesis]: #271 +#275 := [hypothesis]: #129 +#278 := (or #204 #128 #244) +#277 := [th-lemma #276 #275 #274 #162]: false +#279 := [lemma #277]: #278 +#280 := [unit-resolution #279 #275 #274]: #204 +#283 := [unit-resolution #282 #280]: #79 +#284 := [th-lemma #275 #274 #283 #162]: false +#286 := [lemma #284]: #285 +[unit-resolution #286 #308 #305]: false +unsat +cebG074uorSr8ODzgTmcKg 97 0 +#2 := false +#18 := 0::real +decl uf_1 :: (-> T2 T1 real) +decl uf_5 :: T1 +#11 := uf_5 +decl uf_2 :: T2 +#4 := uf_2 +#20 := (uf_1 uf_2 uf_5) +#42 := -1::real +#53 := (* -1::real #20) +decl uf_3 :: T2 +#7 := uf_3 +#19 := (uf_1 uf_3 uf_5) +#54 := (+ #19 #53) +#63 := (<= #54 0::real) +#21 := (- #19 #20) +#22 := (< 0::real #21) +#23 := (not #22) +#74 := (iff #23 #63) +#57 := (< 0::real #54) +#60 := (not #57) +#72 := (iff #60 #63) +#64 := (not #63) +#67 := (not #64) +#70 := (iff #67 #63) +#71 := [rewrite]: #70 +#68 := (iff #60 #67) +#65 := (iff #57 #64) +#66 := [rewrite]: #65 +#69 := [monotonicity #66]: #68 +#73 := [trans #69 #71]: #72 +#61 := (iff #23 #60) +#58 := (iff #22 #57) +#55 := (= #21 #54) +#56 := [rewrite]: #55 +#59 := [monotonicity #56]: #58 +#62 := [monotonicity #59]: #61 +#75 := [trans #62 #73]: #74 +#41 := [asserted]: #23 +#76 := [mp #41 #75]: #63 +#5 := (:var 0 T1) +#8 := (uf_1 uf_3 #5) +#141 := (pattern #8) +#6 := (uf_1 uf_2 #5) +#140 := (pattern #6) +#45 := (* -1::real #8) +#46 := (+ #6 #45) +#44 := (>= #46 0::real) +#43 := (not #44) +#142 := (forall (vars (?x1 T1)) (:pat #140 #141) #43) +#49 := (forall (vars (?x1 T1)) #43) +#145 := (iff #49 #142) +#143 := (iff #43 #43) +#144 := [refl]: #143 +#146 := [quant-intro #144]: #145 +#80 := (~ #49 #49) +#82 := (~ #43 #43) +#83 := [refl]: #82 +#81 := [nnf-pos #83]: #80 +#9 := (< #6 #8) +#10 := (forall (vars (?x1 T1)) #9) +#50 := (iff #10 #49) +#47 := (iff #9 #43) +#48 := [rewrite]: #47 +#51 := [quant-intro #48]: #50 +#39 := [asserted]: #10 +#52 := [mp #39 #51]: #49 +#79 := [mp~ #52 #81]: #49 +#147 := [mp #79 #146]: #142 +#164 := (not #142) +#165 := (or #164 #64) +#148 := (* -1::real #19) +#149 := (+ #20 #148) +#150 := (>= #149 0::real) +#151 := (not #150) +#166 := (or #164 #151) +#168 := (iff #166 #165) +#170 := (iff #165 #165) +#171 := [rewrite]: #170 +#162 := (iff #151 #64) +#160 := (iff #150 #63) +#152 := (+ #148 #20) +#155 := (>= #152 0::real) +#158 := (iff #155 #63) +#159 := [rewrite]: #158 +#156 := (iff #150 #155) +#153 := (= #149 #152) +#154 := [rewrite]: #153 +#157 := [monotonicity #154]: #156 +#161 := [trans #157 #159]: #160 +#163 := [monotonicity #161]: #162 +#169 := [monotonicity #163]: #168 +#172 := [trans #169 #171]: #168 +#167 := [quant-inst]: #166 +#173 := [mp #167 #172]: #165 +[unit-resolution #173 #147 #76]: false +unsat +DKRtrJ2XceCkITuNwNViRw 57 0 +#2 := false +#4 := 0::real +decl uf_1 :: (-> T2 real) +decl uf_2 :: (-> T1 T1 T2) +decl uf_12 :: (-> T4 T1) +decl uf_4 :: T4 +#11 := uf_4 +#39 := (uf_12 uf_4) +decl uf_10 :: T4 +#27 := uf_10 +#38 := (uf_12 uf_10) +#40 := (uf_2 #38 #39) +#41 := (uf_1 #40) +#264 := (>= #41 0::real) +#266 := (not #264) +#43 := (= #41 0::real) +#44 := (not #43) +#131 := [asserted]: #44 +#272 := (or #43 #266) +#42 := (<= #41 0::real) +#130 := [asserted]: #42 +#265 := (not #42) +#270 := (or #43 #265 #266) +#271 := [th-lemma]: #270 +#273 := [unit-resolution #271 #130]: #272 +#274 := [unit-resolution #273 #131]: #266 +#6 := (:var 0 T1) +#5 := (:var 1 T1) +#7 := (uf_2 #5 #6) +#241 := (pattern #7) +#8 := (uf_1 #7) +#65 := (>= #8 0::real) +#242 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #241) #65) +#66 := (forall (vars (?x1 T1) (?x2 T1)) #65) +#245 := (iff #66 #242) +#243 := (iff #65 #65) +#244 := [refl]: #243 +#246 := [quant-intro #244]: #245 +#149 := (~ #66 #66) +#151 := (~ #65 #65) +#152 := [refl]: #151 +#150 := [nnf-pos #152]: #149 +#9 := (<= 0::real #8) +#10 := (forall (vars (?x1 T1) (?x2 T1)) #9) +#67 := (iff #10 #66) +#63 := (iff #9 #65) +#64 := [rewrite]: #63 +#68 := [quant-intro #64]: #67 +#60 := [asserted]: #10 +#69 := [mp #60 #68]: #66 +#147 := [mp~ #69 #150]: #66 +#247 := [mp #147 #246]: #242 +#267 := (not #242) +#268 := (or #267 #264) +#269 := [quant-inst]: #268 +[unit-resolution #269 #247 #274]: false +unsat +97KJAJfUio+nGchEHWvgAw 91 0 +#2 := false +#38 := 0::real +decl uf_1 :: (-> T1 T2 real) +decl uf_3 :: T2 +#5 := uf_3 +decl uf_4 :: T1 +#7 := uf_4 +#8 := (uf_1 uf_4 uf_3) +#35 := -1::real +#36 := (* -1::real #8) +decl uf_2 :: T1 +#4 := uf_2 +#6 := (uf_1 uf_2 uf_3) +#37 := (+ #6 #36) +#130 := (>= #37 0::real) +#155 := (not #130) +#43 := (= #6 #8) +#55 := (not #43) +#15 := (= #8 #6) +#16 := (not #15) +#56 := (iff #16 #55) +#53 := (iff #15 #43) +#54 := [rewrite]: #53 +#57 := [monotonicity #54]: #56 +#34 := [asserted]: #16 +#60 := [mp #34 #57]: #55 +#158 := (or #43 #155) +#39 := (<= #37 0::real) +#9 := (<= #6 #8) +#40 := (iff #9 #39) +#41 := [rewrite]: #40 +#32 := [asserted]: #9 +#42 := [mp #32 #41]: #39 +#154 := (not #39) +#156 := (or #43 #154 #155) +#157 := [th-lemma]: #156 +#159 := [unit-resolution #157 #42]: #158 +#160 := [unit-resolution #159 #60]: #155 +#10 := (:var 0 T2) +#12 := (uf_1 uf_2 #10) +#123 := (pattern #12) +#11 := (uf_1 uf_4 #10) +#122 := (pattern #11) +#44 := (* -1::real #12) +#45 := (+ #11 #44) +#46 := (<= #45 0::real) +#124 := (forall (vars (?x1 T2)) (:pat #122 #123) #46) +#49 := (forall (vars (?x1 T2)) #46) +#127 := (iff #49 #124) +#125 := (iff #46 #46) +#126 := [refl]: #125 +#128 := [quant-intro #126]: #127 +#62 := (~ #49 #49) +#64 := (~ #46 #46) +#65 := [refl]: #64 +#63 := [nnf-pos #65]: #62 +#13 := (<= #11 #12) +#14 := (forall (vars (?x1 T2)) #13) +#50 := (iff #14 #49) +#47 := (iff #13 #46) +#48 := [rewrite]: #47 +#51 := [quant-intro #48]: #50 +#33 := [asserted]: #14 +#52 := [mp #33 #51]: #49 +#61 := [mp~ #52 #63]: #49 +#129 := [mp #61 #128]: #124 +#144 := (not #124) +#145 := (or #144 #130) +#131 := (* -1::real #6) +#132 := (+ #8 #131) +#133 := (<= #132 0::real) +#146 := (or #144 #133) +#148 := (iff #146 #145) +#150 := (iff #145 #145) +#151 := [rewrite]: #150 +#142 := (iff #133 #130) +#134 := (+ #131 #8) +#137 := (<= #134 0::real) +#140 := (iff #137 #130) +#141 := [rewrite]: #140 +#138 := (iff #133 #137) +#135 := (= #132 #134) +#136 := [rewrite]: #135 +#139 := [monotonicity #136]: #138 +#143 := [trans #139 #141]: #142 +#149 := [monotonicity #143]: #148 +#152 := [trans #149 #151]: #148 +#147 := [quant-inst]: #146 +#153 := [mp #147 #152]: #145 +[unit-resolution #153 #129 #160]: false +unsat +flJYbeWfe+t2l/zsRqdujA 149 0 +#2 := false +#19 := 0::real +decl uf_1 :: (-> T1 T2 real) +decl uf_3 :: T2 +#5 := uf_3 +decl uf_4 :: T1 +#7 := uf_4 +#8 := (uf_1 uf_4 uf_3) +#44 := -1::real +#156 := (* -1::real #8) +decl uf_2 :: T1 +#4 := uf_2 +#6 := (uf_1 uf_2 uf_3) +#203 := (+ #6 #156) +#205 := (>= #203 0::real) +#9 := (= #6 #8) +#40 := [asserted]: #9 +#208 := (not #9) +#209 := (or #208 #205) +#210 := [th-lemma]: #209 +#211 := [unit-resolution #210 #40]: #205 +decl uf_5 :: T1 +#12 := uf_5 +#22 := (uf_1 uf_5 uf_3) +#160 := (* -1::real #22) +#161 := (+ #6 #160) +#207 := (>= #161 0::real) +#222 := (not #207) +#206 := (= #6 #22) +#216 := (not #206) +#62 := (= #8 #22) +#70 := (not #62) +#217 := (iff #70 #216) +#214 := (iff #62 #206) +#212 := (iff #206 #62) +#213 := [monotonicity #40]: #212 +#215 := [symm #213]: #214 +#218 := [monotonicity #215]: #217 +#23 := (= #22 #8) +#24 := (not #23) +#71 := (iff #24 #70) +#68 := (iff #23 #62) +#69 := [rewrite]: #68 +#72 := [monotonicity #69]: #71 +#43 := [asserted]: #24 +#75 := [mp #43 #72]: #70 +#219 := [mp #75 #218]: #216 +#225 := (or #206 #222) +#162 := (<= #161 0::real) +#172 := (+ #8 #160) +#173 := (>= #172 0::real) +#178 := (not #173) +#163 := (not #162) +#181 := (or #163 #178) +#184 := (not #181) +#10 := (:var 0 T2) +#15 := (uf_1 uf_4 #10) +#149 := (pattern #15) +#13 := (uf_1 uf_5 #10) +#148 := (pattern #13) +#11 := (uf_1 uf_2 #10) +#147 := (pattern #11) +#50 := (* -1::real #15) +#51 := (+ #13 #50) +#52 := (<= #51 0::real) +#76 := (not #52) +#45 := (* -1::real #13) +#46 := (+ #11 #45) +#47 := (<= #46 0::real) +#78 := (not #47) +#73 := (or #78 #76) +#83 := (not #73) +#150 := (forall (vars (?x1 T2)) (:pat #147 #148 #149) #83) +#86 := (forall (vars (?x1 T2)) #83) +#153 := (iff #86 #150) +#151 := (iff #83 #83) +#152 := [refl]: #151 +#154 := [quant-intro #152]: #153 +#55 := (and #47 #52) +#58 := (forall (vars (?x1 T2)) #55) +#87 := (iff #58 #86) +#84 := (iff #55 #83) +#85 := [rewrite]: #84 +#88 := [quant-intro #85]: #87 +#79 := (~ #58 #58) +#81 := (~ #55 #55) +#82 := [refl]: #81 +#80 := [nnf-pos #82]: #79 +#16 := (<= #13 #15) +#14 := (<= #11 #13) +#17 := (and #14 #16) +#18 := (forall (vars (?x1 T2)) #17) +#59 := (iff #18 #58) +#56 := (iff #17 #55) +#53 := (iff #16 #52) +#54 := [rewrite]: #53 +#48 := (iff #14 #47) +#49 := [rewrite]: #48 +#57 := [monotonicity #49 #54]: #56 +#60 := [quant-intro #57]: #59 +#41 := [asserted]: #18 +#61 := [mp #41 #60]: #58 +#77 := [mp~ #61 #80]: #58 +#89 := [mp #77 #88]: #86 +#155 := [mp #89 #154]: #150 +#187 := (not #150) +#188 := (or #187 #184) +#157 := (+ #22 #156) +#158 := (<= #157 0::real) +#159 := (not #158) +#164 := (or #163 #159) +#165 := (not #164) +#189 := (or #187 #165) +#191 := (iff #189 #188) +#193 := (iff #188 #188) +#194 := [rewrite]: #193 +#185 := (iff #165 #184) +#182 := (iff #164 #181) +#179 := (iff #159 #178) +#176 := (iff #158 #173) +#166 := (+ #156 #22) +#169 := (<= #166 0::real) +#174 := (iff #169 #173) +#175 := [rewrite]: #174 +#170 := (iff #158 #169) +#167 := (= #157 #166) +#168 := [rewrite]: #167 +#171 := [monotonicity #168]: #170 +#177 := [trans #171 #175]: #176 +#180 := [monotonicity #177]: #179 +#183 := [monotonicity #180]: #182 +#186 := [monotonicity #183]: #185 +#192 := [monotonicity #186]: #191 +#195 := [trans #192 #194]: #191 +#190 := [quant-inst]: #189 +#196 := [mp #190 #195]: #188 +#220 := [unit-resolution #196 #155]: #184 +#197 := (or #181 #162) +#198 := [def-axiom]: #197 +#221 := [unit-resolution #198 #220]: #162 +#223 := (or #206 #163 #222) +#224 := [th-lemma]: #223 +#226 := [unit-resolution #224 #221]: #225 +#227 := [unit-resolution #226 #219]: #222 +#199 := (or #181 #173) +#200 := [def-axiom]: #199 +#228 := [unit-resolution #200 #220]: #173 +[th-lemma #228 #227 #211]: false +unsat +rbrrQuQfaijtLkQizgEXnQ 222 0 +#2 := false +#4 := 0::real +decl uf_2 :: (-> T2 T1 real) +decl uf_5 :: T1 +#15 := uf_5 +decl uf_3 :: T2 +#7 := uf_3 +#20 := (uf_2 uf_3 uf_5) +decl uf_6 :: T2 +#17 := uf_6 +#18 := (uf_2 uf_6 uf_5) +#59 := -1::real +#73 := (* -1::real #18) +#106 := (+ #73 #20) +decl uf_1 :: real +#5 := uf_1 +#78 := (* -1::real #20) +#79 := (+ #18 #78) +#144 := (+ uf_1 #79) +#145 := (<= #144 0::real) +#148 := (ite #145 uf_1 #106) +#279 := (* -1::real #148) +#280 := (+ uf_1 #279) +#281 := (<= #280 0::real) +#289 := (not #281) +#72 := 1/2::real +#151 := (* 1/2::real #148) +#248 := (<= #151 0::real) +#162 := (= #151 0::real) +#24 := 2::real +#27 := (- #20 #18) +#28 := (<= uf_1 #27) +#29 := (ite #28 uf_1 #27) +#30 := (/ #29 2::real) +#31 := (+ #18 #30) +#32 := (= #31 #18) +#33 := (not #32) +#34 := (not #33) +#165 := (iff #34 #162) +#109 := (<= uf_1 #106) +#112 := (ite #109 uf_1 #106) +#118 := (* 1/2::real #112) +#123 := (+ #18 #118) +#129 := (= #18 #123) +#163 := (iff #129 #162) +#154 := (+ #18 #151) +#157 := (= #18 #154) +#160 := (iff #157 #162) +#161 := [rewrite]: #160 +#158 := (iff #129 #157) +#155 := (= #123 #154) +#152 := (= #118 #151) +#149 := (= #112 #148) +#146 := (iff #109 #145) +#147 := [rewrite]: #146 +#150 := [monotonicity #147]: #149 +#153 := [monotonicity #150]: #152 +#156 := [monotonicity #153]: #155 +#159 := [monotonicity #156]: #158 +#164 := [trans #159 #161]: #163 +#142 := (iff #34 #129) +#134 := (not #129) +#137 := (not #134) +#140 := (iff #137 #129) +#141 := [rewrite]: #140 +#138 := (iff #34 #137) +#135 := (iff #33 #134) +#132 := (iff #32 #129) +#126 := (= #123 #18) +#130 := (iff #126 #129) +#131 := [rewrite]: #130 +#127 := (iff #32 #126) +#124 := (= #31 #123) +#121 := (= #30 #118) +#115 := (/ #112 2::real) +#119 := (= #115 #118) +#120 := [rewrite]: #119 +#116 := (= #30 #115) +#113 := (= #29 #112) +#107 := (= #27 #106) +#108 := [rewrite]: #107 +#110 := (iff #28 #109) +#111 := [monotonicity #108]: #110 +#114 := [monotonicity #111 #108]: #113 +#117 := [monotonicity #114]: #116 +#122 := [trans #117 #120]: #121 +#125 := [monotonicity #122]: #124 +#128 := [monotonicity #125]: #127 +#133 := [trans #128 #131]: #132 +#136 := [monotonicity #133]: #135 +#139 := [monotonicity #136]: #138 +#143 := [trans #139 #141]: #142 +#166 := [trans #143 #164]: #165 +#105 := [asserted]: #34 +#167 := [mp #105 #166]: #162 +#283 := (not #162) +#284 := (or #283 #248) +#285 := [th-lemma]: #284 +#286 := [unit-resolution #285 #167]: #248 +#287 := [hypothesis]: #281 +#53 := (<= uf_1 0::real) +#54 := (not #53) +#6 := (< 0::real uf_1) +#55 := (iff #6 #54) +#56 := [rewrite]: #55 +#50 := [asserted]: #6 +#57 := [mp #50 #56]: #54 +#288 := [th-lemma #57 #287 #286]: false +#290 := [lemma #288]: #289 +#241 := (= uf_1 #148) +#242 := (= #106 #148) +#299 := (not #242) +#282 := (+ #106 #279) +#291 := (<= #282 0::real) +#296 := (not #291) +decl uf_4 :: T2 +#10 := uf_4 +#16 := (uf_2 uf_4 uf_5) +#260 := (+ #16 #78) +#261 := (>= #260 0::real) +#266 := (not #261) +#8 := (:var 0 T1) +#11 := (uf_2 uf_4 #8) +#234 := (pattern #11) +#9 := (uf_2 uf_3 #8) +#233 := (pattern #9) +#60 := (* -1::real #11) +#61 := (+ #9 #60) +#62 := (<= #61 0::real) +#179 := (not #62) +#235 := (forall (vars (?x1 T1)) (:pat #233 #234) #179) +#178 := (forall (vars (?x1 T1)) #179) +#238 := (iff #178 #235) +#236 := (iff #179 #179) +#237 := [refl]: #236 +#239 := [quant-intro #237]: #238 +#65 := (exists (vars (?x1 T1)) #62) +#68 := (not #65) +#175 := (~ #68 #178) +#180 := (~ #179 #179) +#177 := [refl]: #180 +#176 := [nnf-neg #177]: #175 +#12 := (<= #9 #11) +#13 := (exists (vars (?x1 T1)) #12) +#14 := (not #13) +#69 := (iff #14 #68) +#66 := (iff #13 #65) +#63 := (iff #12 #62) +#64 := [rewrite]: #63 +#67 := [quant-intro #64]: #66 +#70 := [monotonicity #67]: #69 +#51 := [asserted]: #14 +#71 := [mp #51 #70]: #68 +#173 := [mp~ #71 #176]: #178 +#240 := [mp #173 #239]: #235 +#269 := (not #235) +#270 := (or #269 #266) +#250 := (* -1::real #16) +#251 := (+ #20 #250) +#252 := (<= #251 0::real) +#253 := (not #252) +#271 := (or #269 #253) +#273 := (iff #271 #270) +#275 := (iff #270 #270) +#276 := [rewrite]: #275 +#267 := (iff #253 #266) +#264 := (iff #252 #261) +#254 := (+ #250 #20) +#257 := (<= #254 0::real) +#262 := (iff #257 #261) +#263 := [rewrite]: #262 +#258 := (iff #252 #257) +#255 := (= #251 #254) +#256 := [rewrite]: #255 +#259 := [monotonicity #256]: #258 +#265 := [trans #259 #263]: #264 +#268 := [monotonicity #265]: #267 +#274 := [monotonicity #268]: #273 +#277 := [trans #274 #276]: #273 +#272 := [quant-inst]: #271 +#278 := [mp #272 #277]: #270 +#293 := [unit-resolution #278 #240]: #266 +#90 := (* 1/2::real #20) +#102 := (+ #73 #90) +#89 := (* 1/2::real #16) +#103 := (+ #89 #102) +#100 := (>= #103 0::real) +#23 := (+ #16 #20) +#25 := (/ #23 2::real) +#26 := (<= #18 #25) +#98 := (iff #26 #100) +#91 := (+ #89 #90) +#94 := (<= #18 #91) +#97 := (iff #94 #100) +#99 := [rewrite]: #97 +#95 := (iff #26 #94) +#92 := (= #25 #91) +#93 := [rewrite]: #92 +#96 := [monotonicity #93]: #95 +#101 := [trans #96 #99]: #98 +#58 := [asserted]: #26 +#104 := [mp #58 #101]: #100 +#294 := [hypothesis]: #291 +#295 := [th-lemma #294 #104 #293 #286]: false +#297 := [lemma #295]: #296 +#298 := [hypothesis]: #242 +#300 := (or #299 #291) +#301 := [th-lemma]: #300 +#302 := [unit-resolution #301 #298 #297]: false +#303 := [lemma #302]: #299 +#246 := (or #145 #242) +#247 := [def-axiom]: #246 +#304 := [unit-resolution #247 #303]: #145 +#243 := (not #145) +#244 := (or #243 #241) +#245 := [def-axiom]: #244 +#305 := [unit-resolution #245 #304]: #241 +#306 := (not #241) +#307 := (or #306 #281) +#308 := [th-lemma]: #307 +[unit-resolution #308 #305 #290]: false +unsat +hwh3oeLAWt56hnKIa8Wuow 248 0 +#2 := false +#4 := 0::real +decl uf_2 :: (-> T2 T1 real) +decl uf_5 :: T1 +#15 := uf_5 +decl uf_6 :: T2 +#17 := uf_6 +#18 := (uf_2 uf_6 uf_5) +decl uf_4 :: T2 +#10 := uf_4 +#16 := (uf_2 uf_4 uf_5) +#66 := -1::real +#137 := (* -1::real #16) +#138 := (+ #137 #18) +decl uf_1 :: real +#5 := uf_1 +#80 := (* -1::real #18) +#81 := (+ #16 #80) +#201 := (+ uf_1 #81) +#202 := (<= #201 0::real) +#205 := (ite #202 uf_1 #138) +#352 := (* -1::real #205) +#353 := (+ uf_1 #352) +#354 := (<= #353 0::real) +#362 := (not #354) +#79 := 1/2::real +#244 := (* 1/2::real #205) +#322 := (<= #244 0::real) +#245 := (= #244 0::real) +#158 := -1/2::real +#208 := (* -1/2::real #205) +#211 := (+ #18 #208) +decl uf_3 :: T2 +#7 := uf_3 +#20 := (uf_2 uf_3 uf_5) +#117 := (+ #80 #20) +#85 := (* -1::real #20) +#86 := (+ #18 #85) +#188 := (+ uf_1 #86) +#189 := (<= #188 0::real) +#192 := (ite #189 uf_1 #117) +#195 := (* 1/2::real #192) +#198 := (+ #18 #195) +#97 := (* 1/2::real #20) +#109 := (+ #80 #97) +#96 := (* 1/2::real #16) +#110 := (+ #96 #109) +#107 := (>= #110 0::real) +#214 := (ite #107 #198 #211) +#217 := (= #18 #214) +#248 := (iff #217 #245) +#241 := (= #18 #211) +#246 := (iff #241 #245) +#247 := [rewrite]: #246 +#242 := (iff #217 #241) +#239 := (= #214 #211) +#234 := (ite false #198 #211) +#237 := (= #234 #211) +#238 := [rewrite]: #237 +#235 := (= #214 #234) +#232 := (iff #107 false) +#104 := (not #107) +#24 := 2::real +#23 := (+ #16 #20) +#25 := (/ #23 2::real) +#26 := (< #25 #18) +#108 := (iff #26 #104) +#98 := (+ #96 #97) +#101 := (< #98 #18) +#106 := (iff #101 #104) +#105 := [rewrite]: #106 +#102 := (iff #26 #101) +#99 := (= #25 #98) +#100 := [rewrite]: #99 +#103 := [monotonicity #100]: #102 +#111 := [trans #103 #105]: #108 +#65 := [asserted]: #26 +#112 := [mp #65 #111]: #104 +#233 := [iff-false #112]: #232 +#236 := [monotonicity #233]: #235 +#240 := [trans #236 #238]: #239 +#243 := [monotonicity #240]: #242 +#249 := [trans #243 #247]: #248 +#33 := (- #18 #16) +#34 := (<= uf_1 #33) +#35 := (ite #34 uf_1 #33) +#36 := (/ #35 2::real) +#37 := (- #18 #36) +#28 := (- #20 #18) +#29 := (<= uf_1 #28) +#30 := (ite #29 uf_1 #28) +#31 := (/ #30 2::real) +#32 := (+ #18 #31) +#27 := (<= #18 #25) +#38 := (ite #27 #32 #37) +#39 := (= #38 #18) +#40 := (not #39) +#41 := (not #40) +#220 := (iff #41 #217) +#141 := (<= uf_1 #138) +#144 := (ite #141 uf_1 #138) +#159 := (* -1/2::real #144) +#160 := (+ #18 #159) +#120 := (<= uf_1 #117) +#123 := (ite #120 uf_1 #117) +#129 := (* 1/2::real #123) +#134 := (+ #18 #129) +#114 := (<= #18 #98) +#165 := (ite #114 #134 #160) +#171 := (= #18 #165) +#218 := (iff #171 #217) +#215 := (= #165 #214) +#212 := (= #160 #211) +#209 := (= #159 #208) +#206 := (= #144 #205) +#203 := (iff #141 #202) +#204 := [rewrite]: #203 +#207 := [monotonicity #204]: #206 +#210 := [monotonicity #207]: #209 +#213 := [monotonicity #210]: #212 +#199 := (= #134 #198) +#196 := (= #129 #195) +#193 := (= #123 #192) +#190 := (iff #120 #189) +#191 := [rewrite]: #190 +#194 := [monotonicity #191]: #193 +#197 := [monotonicity #194]: #196 +#200 := [monotonicity #197]: #199 +#187 := (iff #114 #107) +#186 := [rewrite]: #187 +#216 := [monotonicity #186 #200 #213]: #215 +#219 := [monotonicity #216]: #218 +#184 := (iff #41 #171) +#176 := (not #171) +#179 := (not #176) +#182 := (iff #179 #171) +#183 := [rewrite]: #182 +#180 := (iff #41 #179) +#177 := (iff #40 #176) +#174 := (iff #39 #171) +#168 := (= #165 #18) +#172 := (iff #168 #171) +#173 := [rewrite]: #172 +#169 := (iff #39 #168) +#166 := (= #38 #165) +#163 := (= #37 #160) +#150 := (* 1/2::real #144) +#155 := (- #18 #150) +#161 := (= #155 #160) +#162 := [rewrite]: #161 +#156 := (= #37 #155) +#153 := (= #36 #150) +#147 := (/ #144 2::real) +#151 := (= #147 #150) +#152 := [rewrite]: #151 +#148 := (= #36 #147) +#145 := (= #35 #144) +#139 := (= #33 #138) +#140 := [rewrite]: #139 +#142 := (iff #34 #141) +#143 := [monotonicity #140]: #142 +#146 := [monotonicity #143 #140]: #145 +#149 := [monotonicity #146]: #148 +#154 := [trans #149 #152]: #153 +#157 := [monotonicity #154]: #156 +#164 := [trans #157 #162]: #163 +#135 := (= #32 #134) +#132 := (= #31 #129) +#126 := (/ #123 2::real) +#130 := (= #126 #129) +#131 := [rewrite]: #130 +#127 := (= #31 #126) +#124 := (= #30 #123) +#118 := (= #28 #117) +#119 := [rewrite]: #118 +#121 := (iff #29 #120) +#122 := [monotonicity #119]: #121 +#125 := [monotonicity #122 #119]: #124 +#128 := [monotonicity #125]: #127 +#133 := [trans #128 #131]: #132 +#136 := [monotonicity #133]: #135 +#115 := (iff #27 #114) +#116 := [monotonicity #100]: #115 +#167 := [monotonicity #116 #136 #164]: #166 +#170 := [monotonicity #167]: #169 +#175 := [trans #170 #173]: #174 +#178 := [monotonicity #175]: #177 +#181 := [monotonicity #178]: #180 +#185 := [trans #181 #183]: #184 +#221 := [trans #185 #219]: #220 +#113 := [asserted]: #41 +#222 := [mp #113 #221]: #217 +#250 := [mp #222 #249]: #245 +#356 := (not #245) +#357 := (or #356 #322) +#358 := [th-lemma]: #357 +#359 := [unit-resolution #358 #250]: #322 +#360 := [hypothesis]: #354 +#60 := (<= uf_1 0::real) +#61 := (not #60) +#6 := (< 0::real uf_1) +#62 := (iff #6 #61) +#63 := [rewrite]: #62 +#57 := [asserted]: #6 +#64 := [mp #57 #63]: #61 +#361 := [th-lemma #64 #360 #359]: false +#363 := [lemma #361]: #362 +#315 := (= uf_1 #205) +#316 := (= #138 #205) +#371 := (not #316) +#355 := (+ #138 #352) +#364 := (<= #355 0::real) +#368 := (not #364) +#87 := (<= #86 0::real) +#82 := (<= #81 0::real) +#90 := (and #82 #87) +#21 := (<= #18 #20) +#19 := (<= #16 #18) +#22 := (and #19 #21) +#91 := (iff #22 #90) +#88 := (iff #21 #87) +#89 := [rewrite]: #88 +#83 := (iff #19 #82) +#84 := [rewrite]: #83 +#92 := [monotonicity #84 #89]: #91 +#59 := [asserted]: #22 +#93 := [mp #59 #92]: #90 +#95 := [and-elim #93]: #87 +#366 := [hypothesis]: #364 +#367 := [th-lemma #366 #95 #112 #359]: false +#369 := [lemma #367]: #368 +#370 := [hypothesis]: #316 +#372 := (or #371 #364) +#373 := [th-lemma]: #372 +#374 := [unit-resolution #373 #370 #369]: false +#375 := [lemma #374]: #371 +#320 := (or #202 #316) +#321 := [def-axiom]: #320 +#376 := [unit-resolution #321 #375]: #202 +#317 := (not #202) +#318 := (or #317 #315) +#319 := [def-axiom]: #318 +#377 := [unit-resolution #319 #376]: #315 +#378 := (not #315) +#379 := (or #378 #354) +#380 := [th-lemma]: #379 +[unit-resolution #380 #377 #363]: false +unsat +WdMJH3tkMv/rps8y9Ukq5Q 86 0 +#2 := false +#37 := 0::real +decl uf_2 :: (-> T2 T1 real) +decl uf_4 :: T1 +#12 := uf_4 +decl uf_3 :: T2 +#5 := uf_3 +#13 := (uf_2 uf_3 uf_4) +#34 := -1::real +#140 := (* -1::real #13) +decl uf_1 :: real +#4 := uf_1 +#141 := (+ uf_1 #140) +#143 := (>= #141 0::real) +#6 := (:var 0 T1) +#7 := (uf_2 uf_3 #6) +#127 := (pattern #7) +#35 := (* -1::real #7) +#36 := (+ uf_1 #35) +#47 := (>= #36 0::real) +#134 := (forall (vars (?x2 T1)) (:pat #127) #47) +#49 := (forall (vars (?x2 T1)) #47) +#137 := (iff #49 #134) +#135 := (iff #47 #47) +#136 := [refl]: #135 +#138 := [quant-intro #136]: #137 +#67 := (~ #49 #49) +#58 := (~ #47 #47) +#66 := [refl]: #58 +#68 := [nnf-pos #66]: #67 +#10 := (<= #7 uf_1) +#11 := (forall (vars (?x2 T1)) #10) +#50 := (iff #11 #49) +#46 := (iff #10 #47) +#48 := [rewrite]: #46 +#51 := [quant-intro #48]: #50 +#32 := [asserted]: #11 +#52 := [mp #32 #51]: #49 +#69 := [mp~ #52 #68]: #49 +#139 := [mp #69 #138]: #134 +#149 := (not #134) +#150 := (or #149 #143) +#151 := [quant-inst]: #150 +#144 := [unit-resolution #151 #139]: #143 +#142 := (<= #141 0::real) +#38 := (<= #36 0::real) +#128 := (forall (vars (?x1 T1)) (:pat #127) #38) +#41 := (forall (vars (?x1 T1)) #38) +#131 := (iff #41 #128) +#129 := (iff #38 #38) +#130 := [refl]: #129 +#132 := [quant-intro #130]: #131 +#62 := (~ #41 #41) +#64 := (~ #38 #38) +#65 := [refl]: #64 +#63 := [nnf-pos #65]: #62 +#8 := (<= uf_1 #7) +#9 := (forall (vars (?x1 T1)) #8) +#42 := (iff #9 #41) +#39 := (iff #8 #38) +#40 := [rewrite]: #39 +#43 := [quant-intro #40]: #42 +#31 := [asserted]: #9 +#44 := [mp #31 #43]: #41 +#61 := [mp~ #44 #63]: #41 +#133 := [mp #61 #132]: #128 +#145 := (not #128) +#146 := (or #145 #142) +#147 := [quant-inst]: #146 +#148 := [unit-resolution #147 #133]: #142 +#45 := (= uf_1 #13) +#55 := (not #45) +#14 := (= #13 uf_1) +#15 := (not #14) +#56 := (iff #15 #55) +#53 := (iff #14 #45) +#54 := [rewrite]: #53 +#57 := [monotonicity #54]: #56 +#33 := [asserted]: #15 +#60 := [mp #33 #57]: #55 +#153 := (not #143) +#152 := (not #142) +#154 := (or #45 #152 #153) +#155 := [th-lemma]: #154 +[unit-resolution #155 #60 #148 #144]: false +unsat +V+IAyBZU/6QjYs6JkXx8LQ 57 0 +#2 := false +#4 := 0::real +decl uf_1 :: (-> T2 real) +decl uf_2 :: (-> T1 T1 T2) +decl uf_12 :: (-> T4 T1) +decl uf_4 :: T4 +#11 := uf_4 +#39 := (uf_12 uf_4) +decl uf_10 :: T4 +#27 := uf_10 +#38 := (uf_12 uf_10) +#40 := (uf_2 #38 #39) +#41 := (uf_1 #40) +#264 := (>= #41 0::real) +#266 := (not #264) +#43 := (= #41 0::real) +#44 := (not #43) +#131 := [asserted]: #44 +#272 := (or #43 #266) +#42 := (<= #41 0::real) +#130 := [asserted]: #42 +#265 := (not #42) +#270 := (or #43 #265 #266) +#271 := [th-lemma]: #270 +#273 := [unit-resolution #271 #130]: #272 +#274 := [unit-resolution #273 #131]: #266 +#6 := (:var 0 T1) +#5 := (:var 1 T1) +#7 := (uf_2 #5 #6) +#241 := (pattern #7) +#8 := (uf_1 #7) +#65 := (>= #8 0::real) +#242 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #241) #65) +#66 := (forall (vars (?x1 T1) (?x2 T1)) #65) +#245 := (iff #66 #242) +#243 := (iff #65 #65) +#244 := [refl]: #243 +#246 := [quant-intro #244]: #245 +#149 := (~ #66 #66) +#151 := (~ #65 #65) +#152 := [refl]: #151 +#150 := [nnf-pos #152]: #149 +#9 := (<= 0::real #8) +#10 := (forall (vars (?x1 T1) (?x2 T1)) #9) +#67 := (iff #10 #66) +#63 := (iff #9 #65) +#64 := [rewrite]: #63 +#68 := [quant-intro #64]: #67 +#60 := [asserted]: #10 +#69 := [mp #60 #68]: #66 +#147 := [mp~ #69 #150]: #66 +#247 := [mp #147 #246]: #242 +#267 := (not #242) +#268 := (or #267 #264) +#269 := [quant-inst]: #268 +[unit-resolution #269 #247 #274]: false +unsat diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Multivariate_Analysis/Integration_MV.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Integration_MV.thy Wed Feb 17 21:13:40 2010 +0100 @@ -0,0 +1,3465 @@ + +header {* Kurzweil-Henstock gauge integration in many dimensions. *} +(* Author: John Harrison + Translation from HOL light: Robert Himmelmann, TU Muenchen *) + +theory Integration_MV + imports Derivative SMT +begin + +declare [[smt_certificates="~~/src/HOL/Multivariate_Analysis/Integration_MV.cert"]] +declare [[smt_record=true]] +declare [[z3_proofs=true]] + +lemma conjunctD2: assumes "a \ b" shows a b using assms by auto +lemma conjunctD3: assumes "a \ b \ c" shows a b c using assms by auto +lemma conjunctD4: assumes "a \ b \ c \ d" shows a b c d using assms by auto +lemma conjunctD5: assumes "a \ b \ c \ d \ e" shows a b c d e using assms by auto + +declare smult_conv_scaleR[simp] + +subsection {* Some useful lemmas about intervals. *} + +lemma empty_as_interval: "{} = {1..0::real^'n}" + apply(rule set_ext,rule) defer unfolding vector_le_def mem_interval + using UNIV_witness[where 'a='n] apply(erule_tac exE,rule_tac x=x in allE) by auto + +lemma interior_subset_union_intervals: + assumes "i = {a..b::real^'n}" "j = {c..d}" "interior j \ {}" "i \ j \ s" "interior(i) \ interior(j) = {}" + shows "interior i \ interior s" proof- + have "{a<.. {c..d} = {}" using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5) + unfolding assms(1,2) interior_closed_interval by auto + moreover have "{a<.. {c..d} \ s" apply(rule order_trans,rule interval_open_subset_closed) + using assms(4) unfolding assms(1,2) by auto + ultimately show ?thesis apply-apply(rule interior_maximal) defer apply(rule open_interior) + unfolding assms(1,2) interior_closed_interval by auto qed + +lemma inter_interior_unions_intervals: fixes f::"(real^'n) set set" + assumes "finite f" "open s" "\t\f. \a b. t = {a..b}" "\t\f. s \ (interior t) = {}" + shows "s \ interior(\f) = {}" proof(rule ccontr,unfold ex_in_conv[THEN sym]) case goal1 + have lem1:"\x e s U. ball x e \ s \ interior U \ ball x e \ s \ U" apply rule defer apply(rule_tac Int_greatest) + unfolding open_subset_interior[OF open_ball] using interior_subset by auto + have lem2:"\x s P. \x\s. P x \ \x\insert x s. P x" by auto + have "\f. finite f \ (\t\f. \a b. t = {a..b}) \ (\x. x \ s \ interior (\f)) \ (\t\f. \x. \e>0. ball x e \ s \ t)" proof- case goal1 + thus ?case proof(induct rule:finite_induct) + case empty from this(2) guess x .. hence False unfolding Union_empty interior_empty by auto thus ?case by auto next + case (insert i f) guess x using insert(5) .. note x = this + then guess e unfolding open_contains_ball_eq[OF open_Int[OF assms(2) open_interior],rule_format] .. note e=this + guess a using insert(4)[rule_format,OF insertI1] .. then guess b .. note ab = this + show ?case proof(cases "x\i") case False hence "x \ UNIV - {a..b}" unfolding ab by auto + then guess d unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_interval],rule_format] .. + hence "0 < d" "ball x (min d e) \ UNIV - i" using e unfolding ab by auto + hence "ball x (min d e) \ s \ interior (\f)" using e unfolding lem1 by auto hence "x \ s \ interior (\f)" using `d>0` e by auto + hence "\t\f. \x e. 0 < e \ ball x e \ s \ t" apply-apply(rule insert(3)) using insert(4) by auto thus ?thesis by auto next + case True show ?thesis proof(cases "x\{a<.. a$k \ x$k \ b$k" unfolding mem_interval by(auto simp add:not_less) + hence "x$k = a$k \ x$k = b$k" using True unfolding ab and mem_interval apply(erule_tac x=k in allE) by auto + hence "\x. ball x (e/2) \ s \ (\f)" proof(erule_tac disjE) + let ?z = "x - (e/2) *\<^sub>R basis k" assume as:"x$k = a$k" have "ball ?z (e / 2) \ i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE) + fix y assume "y \ ball ?z (e / 2) \ i" hence "dist ?z y < e/2" and yi:"y\i" by auto + hence "\(?z - y) $ k\ < e/2" using component_le_norm[of "?z - y" k] unfolding vector_dist_norm by auto + hence "y$k < a$k" unfolding vector_component_simps vector_scaleR_component as using e[THEN conjunct1] by(auto simp add:field_simps) + hence "y \ i" unfolding ab mem_interval not_all by(rule_tac x=k in exI,auto) thus False using yi by auto qed + moreover have "ball ?z (e/2) \ s \ (\insert i f)" apply(rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) proof + fix y assume as:"y\ ball ?z (e/2)" have "norm (x - y) \ \e\ / 2 + norm (x - y - (e / 2) *\<^sub>R basis k)" + apply-apply(rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R basis k"]) + unfolding norm_scaleR norm_basis by auto + also have "\ < \e\ / 2 + \e\ / 2" apply(rule add_strict_left_mono) using as unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps) + finally show "y\ball x e" unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps) qed + ultimately show ?thesis apply(rule_tac x="?z" in exI) unfolding Union_insert by auto + next let ?z = "x + (e/2) *\<^sub>R basis k" assume as:"x$k = b$k" have "ball ?z (e / 2) \ i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE) + fix y assume "y \ ball ?z (e / 2) \ i" hence "dist ?z y < e/2" and yi:"y\i" by auto + hence "\(?z - y) $ k\ < e/2" using component_le_norm[of "?z - y" k] unfolding vector_dist_norm by auto + hence "y$k > b$k" unfolding vector_component_simps vector_scaleR_component as using e[THEN conjunct1] by(auto simp add:field_simps) + hence "y \ i" unfolding ab mem_interval not_all by(rule_tac x=k in exI,auto) thus False using yi by auto qed + moreover have "ball ?z (e/2) \ s \ (\insert i f)" apply(rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) proof + fix y assume as:"y\ ball ?z (e/2)" have "norm (x - y) \ \e\ / 2 + norm (x - y + (e / 2) *\<^sub>R basis k)" + apply-apply(rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R basis k"]) + unfolding norm_scaleR norm_basis by auto + also have "\ < \e\ / 2 + \e\ / 2" apply(rule add_strict_left_mono) using as unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps) + finally show "y\ball x e" unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps) qed + ultimately show ?thesis apply(rule_tac x="?z" in exI) unfolding Union_insert by auto qed + then guess x .. hence "x \ s \ interior (\f)" unfolding lem1[where U="\f",THEN sym] using centre_in_ball e[THEN conjunct1] by auto + thus ?thesis apply-apply(rule lem2,rule insert(3)) using insert(4) by auto qed qed qed qed note * = this + guess t using *[OF assms(1,3) goal1] .. from this(2) guess x .. then guess e .. + hence "x \ s" "x\interior t" defer using open_subset_interior[OF open_ball, of x e t] by auto + thus False using `t\f` assms(4) by auto qed +subsection {* Bounds on intervals where they exist. *} + +definition "interval_upperbound (s::(real^'n) set) = (\ i. Sup {a. \x\s. x$i = a})" + +definition "interval_lowerbound (s::(real^'n) set) = (\ i. Inf {a. \x\s. x$i = a})" + +lemma interval_upperbound[simp]: assumes "\i. a$i \ b$i" shows "interval_upperbound {a..b} = b" + using assms unfolding interval_upperbound_def Cart_eq Cart_lambda_beta apply-apply(rule,erule_tac x=i in allE) + apply(rule Sup_unique) unfolding setle_def apply rule unfolding mem_Collect_eq apply(erule bexE) unfolding mem_interval defer + apply(rule,rule) apply(rule_tac x="b$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=b in bexI) + unfolding mem_interval using assms by auto + +lemma interval_lowerbound[simp]: assumes "\i. a$i \ b$i" shows "interval_lowerbound {a..b} = a" + using assms unfolding interval_lowerbound_def Cart_eq Cart_lambda_beta apply-apply(rule,erule_tac x=i in allE) + apply(rule Inf_unique) unfolding setge_def apply rule unfolding mem_Collect_eq apply(erule bexE) unfolding mem_interval defer + apply(rule,rule) apply(rule_tac x="a$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=a in bexI) + unfolding mem_interval using assms by auto + +lemmas interval_bounds = interval_upperbound interval_lowerbound + +lemma interval_bounds'[simp]: assumes "{a..b}\{}" shows "interval_upperbound {a..b} = b" "interval_lowerbound {a..b} = a" + using assms unfolding interval_ne_empty by auto + +lemma interval_upperbound_1[simp]: "dest_vec1 a \ dest_vec1 b \ interval_upperbound {a..b} = (b::real^1)" + apply(rule interval_upperbound) by auto + +lemma interval_lowerbound_1[simp]: "dest_vec1 a \ dest_vec1 b \ interval_lowerbound {a..b} = (a::real^1)" + apply(rule interval_lowerbound) by auto + +lemmas interval_bound_1 = interval_upperbound_1 interval_lowerbound_1 + +subsection {* Content (length, area, volume...) of an interval. *} + +definition "content (s::(real^'n) set) = + (if s = {} then 0 else (\i\UNIV. (interval_upperbound s)$i - (interval_lowerbound s)$i))" + +lemma interval_not_empty:"\i. a$i \ b$i \ {a..b::real^'n} \ {}" + unfolding interval_eq_empty unfolding not_ex not_less by assumption + +lemma content_closed_interval: assumes "\i. a$i \ b$i" + shows "content {a..b} = (\i\UNIV. b$i - a$i)" + using interval_not_empty[OF assms] unfolding content_def interval_upperbound[OF assms] interval_lowerbound[OF assms] by auto + +lemma content_closed_interval': assumes "{a..b}\{}" shows "content {a..b} = (\i\UNIV. b$i - a$i)" + apply(rule content_closed_interval) using assms unfolding interval_ne_empty . + +lemma content_1:"dest_vec1 a \ dest_vec1 b \ content {a..b} = dest_vec1 b - dest_vec1 a" + using content_closed_interval[of a b] by auto + +lemma content_1':"a \ b \ content {vec1 a..vec1 b} = b - a" using content_1[of "vec a" "vec b"] by auto + +lemma content_unit[intro]: "content{0..1::real^'n} = 1" proof- + have *:"\i. 0$i \ (1::real^'n::finite)$i" by auto + have "0 \ {0..1::real^'n::finite}" unfolding mem_interval by auto + thus ?thesis unfolding content_def interval_bounds[OF *] using setprod_1 by auto qed + +lemma content_pos_le[intro]: "0 \ content {a..b}" proof(cases "{a..b}={}") + case False hence *:"\i. a $ i \ b $ i" unfolding interval_ne_empty by assumption + have "(\i\UNIV. interval_upperbound {a..b} $ i - interval_lowerbound {a..b} $ i) \ 0" + apply(rule setprod_nonneg) unfolding interval_bounds[OF *] using * apply(erule_tac x=x in allE) by auto + thus ?thesis unfolding content_def by(auto simp del:interval_bounds') qed(unfold content_def, auto) + +lemma content_pos_lt: assumes "\i. a$i < b$i" shows "0 < content {a..b}" +proof- have help_lemma1: "\i. a$i < b$i \ \i. a$i \ ((b$i)::real)" apply(rule,erule_tac x=i in allE) by auto + show ?thesis unfolding content_closed_interval[OF help_lemma1[OF assms]] apply(rule setprod_pos) + using assms apply(erule_tac x=x in allE) by auto qed + +lemma content_pos_lt_1: "dest_vec1 a < dest_vec1 b \ 0 < content({a..b})" + apply(rule content_pos_lt) by auto + +lemma content_eq_0: "content({a..b::real^'n}) = 0 \ (\i. b$i \ a$i)" proof(cases "{a..b} = {}") + case True thus ?thesis unfolding content_def if_P[OF True] unfolding interval_eq_empty apply- + apply(rule,erule exE) apply(rule_tac x=i in exI) by auto next + guess a using UNIV_witness[where 'a='n] .. case False note as=this[unfolded interval_eq_empty not_ex not_less] + show ?thesis unfolding content_def if_not_P[OF False] setprod_zero_iff[OF finite_UNIV] + apply(rule) apply(erule_tac[!] exE bexE) unfolding interval_bounds[OF as] apply(rule_tac x=x in exI) defer + apply(rule_tac x=i in bexI) using as apply(erule_tac x=i in allE) by auto qed + +lemma cond_cases:"(P \ Q x) \ (\ P \ Q y) \ Q (if P then x else y)" by auto + +lemma content_closed_interval_cases: + "content {a..b} = (if \i. a$i \ b$i then setprod (\i. b$i - a$i) UNIV else 0)" apply(rule cond_cases) + apply(rule content_closed_interval) unfolding content_eq_0 not_all not_le defer apply(erule exE,rule_tac x=x in exI) by auto + +lemma content_eq_0_interior: "content {a..b} = 0 \ interior({a..b}) = {}" + unfolding content_eq_0 interior_closed_interval interval_eq_empty by auto + +lemma content_eq_0_1: "content {a..b::real^1} = 0 \ dest_vec1 b \ dest_vec1 a" + unfolding content_eq_0 by auto + +lemma content_pos_lt_eq: "0 < content {a..b} \ (\i. a$i < b$i)" + apply(rule) defer apply(rule content_pos_lt,assumption) proof- assume "0 < content {a..b}" + hence "content {a..b} \ 0" by auto thus "\i. a$i < b$i" unfolding content_eq_0 not_ex not_le by auto qed + +lemma content_empty[simp]: "content {} = 0" unfolding content_def by auto + +lemma content_subset: assumes "{a..b} \ {c..d}" shows "content {a..b::real^'n} \ content {c..d}" proof(cases "{a..b}={}") + case True thus ?thesis using content_pos_le[of c d] by auto next + case False hence ab_ne:"\i. a $ i \ b $ i" unfolding interval_ne_empty by auto + hence ab_ab:"a\{a..b}" "b\{a..b}" unfolding mem_interval by auto + have "{c..d} \ {}" using assms False by auto + hence cd_ne:"\i. c $ i \ d $ i" using assms unfolding interval_ne_empty by auto + show ?thesis unfolding content_def unfolding interval_bounds[OF ab_ne] interval_bounds[OF cd_ne] + unfolding if_not_P[OF False] if_not_P[OF `{c..d} \ {}`] apply(rule setprod_mono,rule) proof fix i::'n + show "0 \ b $ i - a $ i" using ab_ne[THEN spec[where x=i]] by auto + show "b $ i - a $ i \ d $ i - c $ i" + using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(2),of i] + using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(1),of i] by auto qed qed + +lemma content_lt_nz: "0 < content {a..b} \ content {a..b} \ 0" + unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by auto + +subsection {* The notion of a gauge --- simply an open set containing the point. *} + +definition gauge where "gauge d \ (\x. x\(d x) \ open(d x))" + +lemma gaugeI:assumes "\x. x\g x" "\x. open (g x)" shows "gauge g" + using assms unfolding gauge_def by auto + +lemma gaugeD[dest]: assumes "gauge d" shows "x\d x" "open (d x)" using assms unfolding gauge_def by auto + +lemma gauge_ball_dependent: "\x. 0 < e x \ gauge (\x. ball x (e x))" + unfolding gauge_def by auto + +lemma gauge_ball[intro?]: "0 < e \ gauge (\x. ball x e)" unfolding gauge_def by auto + +lemma gauge_trivial[intro]: "gauge (\x. ball x 1)" apply(rule gauge_ball) by auto + +lemma gauge_inter: "gauge d1 \ gauge d2 \ gauge (\x. (d1 x) \ (d2 x))" + unfolding gauge_def by auto + +lemma gauge_inters: assumes "finite s" "\d\s. gauge (f d)" shows "gauge(\x. \ {f d x | d. d \ s})" proof- + have *:"\x. {f d x |d. d \ s} = (\d. f d x) ` s" by auto show ?thesis + unfolding gauge_def unfolding * + using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto qed + +lemma gauge_existence_lemma: "(\x. \d::real. p x \ 0 < d \ q d x) \ (\x. \d>0. p x \ q d x)" by(meson zero_less_one) + +subsection {* Divisions. *} + +definition division_of (infixl "division'_of" 40) where + "s division_of i \ + finite s \ + (\k\s. k \ i \ k \ {} \ (\a b. k = {a..b})) \ + (\k1\s. \k2\s. k1 \ k2 \ interior(k1) \ interior(k2) = {}) \ + (\s = i)" + +lemma division_ofD[dest]: assumes "s division_of i" + shows"finite s" "\k. k\s \ k \ i" "\k. k\s \ k \ {}" "\k. k\s \ (\a b. k = {a..b})" + "\k1 k2. \k1\s; k2\s; k1 \ k2\ \ interior(k1) \ interior(k2) = {}" "\s = i" using assms unfolding division_of_def by auto + +lemma division_ofI: + assumes "finite s" "\k. k\s \ k \ i" "\k. k\s \ k \ {}" "\k. k\s \ (\a b. k = {a..b})" + "\k1 k2. \k1\s; k2\s; k1 \ k2\ \ interior(k1) \ interior(k2) = {}" "\s = i" + shows "s division_of i" using assms unfolding division_of_def by auto + +lemma division_of_finite: "s division_of i \ finite s" + unfolding division_of_def by auto + +lemma division_of_self[intro]: "{a..b} \ {} \ {{a..b}} division_of {a..b}" + unfolding division_of_def by auto + +lemma division_of_trivial[simp]: "s division_of {} \ s = {}" unfolding division_of_def by auto + +lemma division_of_sing[simp]: "s division_of {a..a::real^'n} \ s = {{a..a}}" (is "?l = ?r") proof + assume ?r moreover { assume "s = {{a}}" moreover fix k assume "k\s" + ultimately have"\x y. k = {x..y}" apply(rule_tac x=a in exI)+ unfolding interval_sing[THEN conjunct1] by auto } + ultimately show ?l unfolding division_of_def interval_sing[THEN conjunct1] by auto next + assume ?l note as=conjunctD4[OF this[unfolded division_of_def interval_sing[THEN conjunct1]]] + { fix x assume x:"x\s" have "x={a}" using as(2)[rule_format,OF x] by auto } + moreover have "s \ {}" using as(4) by auto ultimately show ?r unfolding interval_sing[THEN conjunct1] by auto qed + +lemma elementary_empty: obtains p where "p division_of {}" + unfolding division_of_trivial by auto + +lemma elementary_interval: obtains p where "p division_of {a..b}" + by(metis division_of_trivial division_of_self) + +lemma division_contains: "s division_of i \ \x\i. \k\s. x \ k" + unfolding division_of_def by auto + +lemma forall_in_division: + "d division_of i \ ((\x\d. P x) \ (\a b. {a..b} \ d \ P {a..b}))" + unfolding division_of_def by fastsimp + +lemma division_of_subset: assumes "p division_of (\p)" "q \ p" shows "q division_of (\q)" + apply(rule division_ofI) proof- note as=division_ofD[OF assms(1)] + show "finite q" apply(rule finite_subset) using as(1) assms(2) by auto + { fix k assume "k \ q" hence kp:"k\p" using assms(2) by auto show "k\\q" using `k \ q` by auto + show "\a b. k = {a..b}" using as(4)[OF kp] by auto show "k \ {}" using as(3)[OF kp] by auto } + fix k1 k2 assume "k1 \ q" "k2 \ q" "k1 \ k2" hence *:"k1\p" "k2\p" "k1\k2" using assms(2) by auto + show "interior k1 \ interior k2 = {}" using as(5)[OF *] by auto qed auto + +lemma division_of_union_self[intro]: "p division_of s \ p division_of (\p)" unfolding division_of_def by auto + +lemma division_of_content_0: assumes "content {a..b} = 0" "d division_of {a..b}" shows "\k\d. content k = 0" + unfolding forall_in_division[OF assms(2)] apply(rule,rule,rule) apply(drule division_ofD(2)[OF assms(2)]) + apply(drule content_subset) unfolding assms(1) proof- case goal1 thus ?case using content_pos_le[of a b] by auto qed + +lemma division_inter: assumes "p1 division_of s1" "p2 division_of (s2::(real^'a) set)" + shows "{k1 \ k2 | k1 k2 .k1 \ p1 \ k2 \ p2 \ k1 \ k2 \ {}} division_of (s1 \ s2)" (is "?A' division_of _") proof- +let ?A = "{s. s \ (\(k1,k2). k1 \ k2) ` (p1 \ p2) \ s \ {}}" have *:"?A' = ?A" by auto +show ?thesis unfolding * proof(rule division_ofI) have "?A \ (\(x, y). x \ y) ` (p1 \ p2)" by auto + moreover have "finite (p1 \ p2)" using assms unfolding division_of_def by auto ultimately show "finite ?A" by auto + have *:"\s. \{x\s. x \ {}} = \s" by auto show "\?A = s1 \ s2" apply(rule set_ext) unfolding * and Union_image_eq UN_iff + using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto + { fix k assume "k\?A" then obtain k1 k2 where k:"k = k1 \ k2" "k1\p1" "k2\p2" "k\{}" by auto thus "k \ {}" by auto + show "k \ s1 \ s2" using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)] unfolding k by auto + guess a1 using division_ofD(4)[OF assms(1) k(2)] .. then guess b1 .. note ab1=this + guess a2 using division_ofD(4)[OF assms(2) k(3)] .. then guess b2 .. note ab2=this + show "\a b. k = {a..b}" unfolding k ab1 ab2 unfolding inter_interval by auto } fix k1 k2 + assume "k1\?A" then obtain x1 y1 where k1:"k1 = x1 \ y1" "x1\p1" "y1\p2" "k1\{}" by auto + assume "k2\?A" then obtain x2 y2 where k2:"k2 = x2 \ y2" "x2\p1" "y2\p2" "k2\{}" by auto + assume "k1 \ k2" hence th:"x1\x2 \ y1\y2" unfolding k1 k2 by auto + have *:"(interior x1 \ interior x2 = {} \ interior y1 \ interior y2 = {}) \ + interior(x1 \ y1) \ interior(x1) \ interior(x1 \ y1) \ interior(y1) \ + interior(x2 \ y2) \ interior(x2) \ interior(x2 \ y2) \ interior(y2) + \ interior(x1 \ y1) \ interior(x2 \ y2) = {}" by auto + show "interior k1 \ interior k2 = {}" unfolding k1 k2 apply(rule *) defer apply(rule_tac[1-4] subset_interior) + using division_ofD(5)[OF assms(1) k1(2) k2(2)] + using division_ofD(5)[OF assms(2) k1(3) k2(3)] using th by auto qed qed + +lemma division_inter_1: assumes "d division_of i" "{a..b::real^'n} \ i" + shows "{ {a..b} \ k |k. k \ d \ {a..b} \ k \ {} } division_of {a..b}" proof(cases "{a..b} = {}") + case True show ?thesis unfolding True and division_of_trivial by auto next + have *:"{a..b} \ i = {a..b}" using assms(2) by auto + case False show ?thesis using division_inter[OF division_of_self[OF False] assms(1)] unfolding * by auto qed + +lemma elementary_inter: assumes "p1 division_of s" "p2 division_of (t::(real^'n) set)" + shows "\p. p division_of (s \ t)" + by(rule,rule division_inter[OF assms]) + +lemma elementary_inters: assumes "finite f" "f\{}" "\s\f. \p. p division_of (s::(real^'n) set)" + shows "\p. p division_of (\ f)" using assms apply-proof(induct f rule:finite_induct) +case (insert x f) show ?case proof(cases "f={}") + case True thus ?thesis unfolding True using insert by auto next + case False guess p using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] .. + moreover guess px using insert(5)[rule_format,OF insertI1] .. ultimately + show ?thesis unfolding Inter_insert apply(rule_tac elementary_inter) by assumption+ qed qed auto + +lemma division_disjoint_union: + assumes "p1 division_of s1" "p2 division_of s2" "interior s1 \ interior s2 = {}" + shows "(p1 \ p2) division_of (s1 \ s2)" proof(rule division_ofI) + note d1 = division_ofD[OF assms(1)] and d2 = division_ofD[OF assms(2)] + show "finite (p1 \ p2)" using d1(1) d2(1) by auto + show "\(p1 \ p2) = s1 \ s2" using d1(6) d2(6) by auto + { fix k1 k2 assume as:"k1 \ p1 \ p2" "k2 \ p1 \ p2" "k1 \ k2" moreover let ?g="interior k1 \ interior k2 = {}" + { assume as:"k1\p1" "k2\p2" have ?g using subset_interior[OF d1(2)[OF as(1)]] subset_interior[OF d2(2)[OF as(2)]] + using assms(3) by blast } moreover + { assume as:"k1\p2" "k2\p1" have ?g using subset_interior[OF d1(2)[OF as(2)]] subset_interior[OF d2(2)[OF as(1)]] + using assms(3) by blast} ultimately + show ?g using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto } + fix k assume k:"k \ p1 \ p2" show "k \ s1 \ s2" using k d1(2) d2(2) by auto + show "k \ {}" using k d1(3) d2(3) by auto show "\a b. k = {a..b}" using k d1(4) d2(4) by auto qed + +lemma partial_division_extend_1: + assumes "{c..d} \ {a..b::real^'n}" "{c..d} \ {}" + obtains p where "p division_of {a..b}" "{c..d} \ p" +proof- def n \ "CARD('n)" have n:"1 \ n" "0 < n" "n \ 0" unfolding n_def by auto + guess \ using ex_bij_betw_nat_finite_1[OF finite_UNIV[where 'a='n]] .. note \=this + def \' \ "inv_into {1..n} \" + have \':"bij_betw \' UNIV {1..n}" using bij_betw_inv_into[OF \] unfolding \'_def n_def by auto + hence \'i:"\i. \' i \ {1..n}" unfolding bij_betw_def by auto + have \\'[simp]:"\i. \ (\' i) = i" unfolding \'_def apply(rule f_inv_into_f) unfolding n_def using \ unfolding bij_betw_def by auto + have \'\[simp]:"\i. i\{1..n} \ \' (\ i) = i" unfolding \'_def apply(rule inv_into_f_eq) using \ unfolding n_def bij_betw_def by auto + have "{c..d} \ {}" using assms by auto + let ?p1 = "\l. {(\ i. if \' i < l then c$i else a$i) .. (\ i. if \' i < l then d$i else if \' i = l then c$\ l else b$i)}" + let ?p2 = "\l. {(\ i. if \' i < l then c$i else if \' i = l then d$\ l else a$i) .. (\ i. if \' i < l then d$i else b$i)}" + let ?p = "{?p1 l |l. l \ {1..n+1}} \ {?p2 l |l. l \ {1..n+1}}" + have abcd:"\i. a $ i \ c $ i \ c$i \ d$i \ d $ i \ b $ i" using assms unfolding subset_interval interval_eq_empty by(auto simp add:not_le not_less) + show ?thesis apply(rule that[of ?p]) apply(rule division_ofI) + proof- have "\i. \' i < Suc n" + proof(rule ccontr,unfold not_less) fix i assume "Suc n \ \' i" + hence "\' i \ {1..n}" by auto thus False using \' unfolding bij_betw_def by auto + qed hence "c = (\ i. if \' i < Suc n then c $ i else a $ i)" + "d = (\ i. if \' i < Suc n then d $ i else if \' i = n + 1 then c $ \ (n + 1) else b $ i)" + unfolding Cart_eq Cart_lambda_beta using \' unfolding bij_betw_def by auto + thus cdp:"{c..d} \ ?p" apply-apply(rule UnI1) unfolding mem_Collect_eq apply(rule_tac x="n + 1" in exI) by auto + have "\l. l\{1..n+1} \ ?p1 l \ {a..b}" "\l. l\{1..n+1} \ ?p2 l \ {a..b}" + unfolding subset_eq apply(rule_tac[!] ballI,rule_tac[!] ccontr) + proof- fix l assume l:"l\{1..n+1}" fix x assume "x\{a..b}" + then guess i unfolding mem_interval not_all .. note i=this + show "x \ ?p1 l \ False" "x \ ?p2 l \ False" unfolding mem_interval apply(erule_tac[!] x=i in allE) + apply(case_tac[!] "\' i < l", case_tac[!] "\' i = l") using abcd[of i] i by auto + qed moreover have "\x. x \ {a..b} \ x \ \?p" + proof- fix x assume x:"x\{a..b}" + { presume "x\{c..d} \ x \ \?p" thus "x \ \?p" using cdp by blast } + let ?M = "{i. i\{1..n+1} \ \ (c $ \ i \ x $ \ i \ x $ \ i \ d $ \ i)}" + assume "x\{c..d}" then guess i0 unfolding mem_interval not_all .. + hence "\' i0 \ ?M" using \' unfolding bij_betw_def by(auto intro!:le_SucI) + hence M:"finite ?M" "?M \ {}" by auto + def l \ "Min ?M" note l = Min_less_iff[OF M,unfolded l_def[symmetric]] Min_in[OF M,unfolded mem_Collect_eq l_def[symmetric]] + Min_gr_iff[OF M,unfolded l_def[symmetric]] + have "x\?p1 l \ x\?p2 l" using l(2)[THEN conjunct2] unfolding de_Morgan_conj not_le + apply- apply(erule disjE) apply(rule disjI1) defer apply(rule disjI2) + proof- assume as:"x $ \ l < c $ \ l" + show "x \ ?p1 l" unfolding mem_interval Cart_lambda_beta + proof case goal1 have "\' i \ {1..n}" using \' unfolding bij_betw_def not_le by auto + thus ?case using as x[unfolded mem_interval,rule_format,of i] + apply auto using l(3)[of "\' i"] by(auto elim!:ballE[where x="\' i"]) + qed + next assume as:"x $ \ l > d $ \ l" + show "x \ ?p2 l" unfolding mem_interval Cart_lambda_beta + proof case goal1 have "\' i \ {1..n}" using \' unfolding bij_betw_def not_le by auto + thus ?case using as x[unfolded mem_interval,rule_format,of i] + apply auto using l(3)[of "\' i"] by(auto elim!:ballE[where x="\' i"]) + qed qed + thus "x \ \?p" using l(2) by blast + qed ultimately show "\?p = {a..b}" apply-apply(rule) defer apply(rule) by(assumption,blast) + + show "finite ?p" by auto + fix k assume k:"k\?p" then obtain l where l:"k = ?p1 l \ k = ?p2 l" "l \ {1..n + 1}" by auto + show "k\{a..b}" apply(rule,unfold mem_interval,rule,rule) + proof- fix i::'n and x assume "x \ k" moreover have "\' i < l \ \' i = l \ \' i > l" by auto + ultimately show "a$i \ x$i" "x$i \ b$i" using abcd[of i] using l by(auto elim:disjE elim!:allE[where x=i] simp add:vector_le_def) + qed have "\l. ?p1 l \ {}" "\l. ?p2 l \ {}" unfolding interval_eq_empty not_ex apply(rule_tac[!] allI) + proof- case goal1 thus ?case using abcd[of x] by auto + next case goal2 thus ?case using abcd[of x] by auto + qed thus "k \ {}" using k by auto + show "\a b. k = {a..b}" using k by auto + fix k' assume k':"k' \ ?p" "k \ k'" then obtain l' where l':"k' = ?p1 l' \ k' = ?p2 l'" "l' \ {1..n + 1}" by auto + { fix k k' l l' + assume k:"k\?p" and l:"k = ?p1 l \ k = ?p2 l" "l \ {1..n + 1}" + assume k':"k' \ ?p" "k \ k'" and l':"k' = ?p1 l' \ k' = ?p2 l'" "l' \ {1..n + 1}" + assume "l \ l'" fix x + have "x \ interior k \ interior k'" + proof(rule,cases "l' = n+1") assume x:"x \ interior k \ interior k'" + case True hence "\i. \' i < l'" using \'i by(auto simp add:less_Suc_eq_le) + hence k':"k' = {c..d}" using l'(1) \'i by(auto simp add:Cart_nth_inverse) + have ln:"l < n + 1" + proof(rule ccontr) case goal1 hence l2:"l = n+1" using l by auto + hence "\i. \' i < l" using \'i by(auto simp add:less_Suc_eq_le) + hence "k = {c..d}" using l(1) \'i by(auto simp add:Cart_nth_inverse) + thus False using `k\k'` k' by auto + qed have **:"\' (\ l) = l" using \'\[of l] using l ln by auto + have "x $ \ l < c $ \ l \ d $ \ l < x $ \ l" using l(1) apply- + proof(erule disjE) + assume as:"k = ?p1 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] + show ?thesis using *[of "\ l"] using ln unfolding Cart_lambda_beta ** by auto + next assume as:"k = ?p2 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] + show ?thesis using *[of "\ l"] using ln unfolding Cart_lambda_beta ** by auto + qed thus False using x unfolding k' unfolding Int_iff interior_closed_interval mem_interval + by(auto elim!:allE[where x="\ l"]) + next case False hence "l < n + 1" using l'(2) using `l\l'` by auto + hence ln:"l \ {1..n}" "l' \ {1..n}" using l l' False by auto + note \l = \'\[OF ln(1)] \'\[OF ln(2)] + assume x:"x \ interior k \ interior k'" + show False using l(1) l'(1) apply- + proof(erule_tac[!] disjE)+ + assume as:"k = ?p1 l" "k' = ?p1 l'" + note * = x[unfolded as Int_iff interior_closed_interval mem_interval] + have "l \ l'" using k'(2)[unfolded as] by auto + thus False using * by(smt Cart_lambda_beta \l) + next assume as:"k = ?p2 l" "k' = ?p2 l'" + note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] + have "l \ l'" apply(rule) using k'(2)[unfolded as] by auto + thus False using *[of "\ l"] *[of "\ l'"] + unfolding Cart_lambda_beta \l using `l \ l'` by auto + next assume as:"k = ?p1 l" "k' = ?p2 l'" + note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] + show False using *[of "\ l"] *[of "\ l'"] + unfolding Cart_lambda_beta \l using `l \ l'` using abcd[of "\ l'"] by smt + next assume as:"k = ?p2 l" "k' = ?p1 l'" + note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] + show False using *[of "\ l"] *[of "\ l'"] + unfolding Cart_lambda_beta \l using `l \ l'` using abcd[of "\ l'"] by smt + qed qed } + from this[OF k l k' l'] this[OF k'(1) l' k _ l] have "\x. x \ interior k \ interior k'" + apply - apply(cases "l' \ l") using k'(2) by auto + thus "interior k \ interior k' = {}" by auto +qed qed + +lemma partial_division_extend_interval: assumes "p division_of (\p)" "(\p) \ {a..b}" + obtains q where "p \ q" "q division_of {a..b::real^'n}" proof(cases "p = {}") + case True guess q apply(rule elementary_interval[of a b]) . + thus ?thesis apply- apply(rule that[of q]) unfolding True by auto next + case False note p = division_ofD[OF assms(1)] + have *:"\k\p. \q. q division_of {a..b} \ k\q" proof case goal1 + guess c using p(4)[OF goal1] .. then guess d .. note cd_ = this + have *:"{c..d} \ {a..b}" "{c..d} \ {}" using p(2,3)[OF goal1, unfolded cd_] using assms(2) by auto + guess q apply(rule partial_division_extend_1[OF *]) . thus ?case unfolding cd_ by auto qed + guess q using bchoice[OF *] .. note q = conjunctD2[OF this[rule_format]] + have "\x. x\p \ \d. d division_of \(q x - {x})" apply(rule,rule_tac p="q x" in division_of_subset) proof- + fix x assume x:"x\p" show "q x division_of \q x" apply-apply(rule division_ofI) + using division_ofD[OF q(1)[OF x]] by auto show "q x - {x} \ q x" by auto qed + hence "\d. d division_of \ ((\i. \(q i - {i})) ` p)" apply- apply(rule elementary_inters) + apply(rule finite_imageI[OF p(1)]) unfolding image_is_empty apply(rule False) by auto + then guess d .. note d = this + show ?thesis apply(rule that[of "d \ p"]) proof- + have *:"\s f t. s \ {} \ (\i\s. f i \ i = t) \ t = \ (f ` s) \ (\s)" by auto + have *:"{a..b} = \ (\i. \(q i - {i})) ` p \ \p" apply(rule *[OF False]) proof fix i assume i:"i\p" + show "\(q i - {i}) \ i = {a..b}" using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto qed + show "d \ p division_of {a..b}" unfolding * apply(rule division_disjoint_union[OF d assms(1)]) + apply(rule inter_interior_unions_intervals) apply(rule p open_interior ballI)+ proof(assumption,rule) + fix k assume k:"k\p" have *:"\u t s. u \ s \ s \ t = {} \ u \ t = {}" by auto + show "interior (\(\i. \(q i - {i})) ` p) \ interior k = {}" apply(rule *[of _ "interior (\(q k - {k}))"]) + defer apply(subst Int_commute) apply(rule inter_interior_unions_intervals) proof- note qk=division_ofD[OF q(1)[OF k]] + show "finite (q k - {k})" "open (interior k)" "\t\q k - {k}. \a b. t = {a..b}" using qk by auto + show "\t\q k - {k}. interior k \ interior t = {}" using qk(5) using q(2)[OF k] by auto + have *:"\x s. x \ s \ \s \ x" by auto show "interior (\(\i. \(q i - {i})) ` p) \ interior (\(q k - {k}))" + apply(rule subset_interior *)+ using k by auto qed qed qed auto qed + +lemma elementary_bounded[dest]: "p division_of s \ bounded (s::(real^'n) set)" + unfolding division_of_def by(metis bounded_Union bounded_interval) + +lemma elementary_subset_interval: "p division_of s \ \a b. s \ {a..b::real^'n}" + by(meson elementary_bounded bounded_subset_closed_interval) + +lemma division_union_intervals_exists: assumes "{a..b::real^'n} \ {}" + obtains p where "(insert {a..b} p) division_of ({a..b} \ {c..d})" proof(cases "{c..d} = {}") + case True show ?thesis apply(rule that[of "{}"]) unfolding True using assms by auto next + case False note false=this show ?thesis proof(cases "{a..b} \ {c..d} = {}") + have *:"\a b. {a,b} = {a} \ {b}" by auto + case True show ?thesis apply(rule that[of "{{c..d}}"]) unfolding * apply(rule division_disjoint_union) + using false True assms using interior_subset by auto next + case False obtain u v where uv:"{a..b} \ {c..d} = {u..v}" unfolding inter_interval by auto + have *:"{u..v} \ {c..d}" using uv by auto + guess p apply(rule partial_division_extend_1[OF * False[unfolded uv]]) . note p=this division_ofD[OF this(1)] + have *:"{a..b} \ {c..d} = {a..b} \ \(p - {{u..v}})" "\x s. insert x s = {x} \ s" using p(8) unfolding uv[THEN sym] by auto + show thesis apply(rule that[of "p - {{u..v}}"]) unfolding *(1) apply(subst *(2)) apply(rule division_disjoint_union) + apply(rule,rule assms) apply(rule division_of_subset[of p]) apply(rule division_of_union_self[OF p(1)]) defer + unfolding interior_inter[THEN sym] proof- + have *:"\cd p uv ab. p \ cd \ ab \ cd = uv \ ab \ p = uv \ p" by auto + have "interior ({a..b} \ \(p - {{u..v}})) = interior({u..v} \ \(p - {{u..v}}))" + apply(rule arg_cong[of _ _ interior]) apply(rule *[OF _ uv]) using p(8) by auto + also have "\ = {}" unfolding interior_inter apply(rule inter_interior_unions_intervals) using p(6) p(7)[OF p(2)] p(3) by auto + finally show "interior ({a..b} \ \(p - {{u..v}})) = {}" by assumption qed auto qed qed + +lemma division_of_unions: assumes "finite f" "\p. p\f \ p division_of (\p)" + "\k1 k2. \k1 \ \f; k2 \ \f; k1 \ k2\ \ interior k1 \ interior k2 = {}" + shows "\f division_of \\f" apply(rule division_ofI) prefer 5 apply(rule assms(3)|assumption)+ + apply(rule finite_Union assms(1))+ prefer 3 apply(erule UnionE) apply(rule_tac s=X in division_ofD(3)[OF assms(2)]) + using division_ofD[OF assms(2)] by auto + +lemma elementary_union_interval: assumes "p division_of \p" + obtains q where "q division_of ({a..b::real^'n} \ \p)" proof- + note assm=division_ofD[OF assms] + have lem1:"\f s. \\ (f ` s) = \(\x.\(f x)) ` s" by auto + have lem2:"\f s. f \ {} \ \{s \ t |t. t \ f} = s \ \f" by auto +{ presume "p={} \ thesis" "{a..b} = {} \ thesis" "{a..b} \ {} \ interior {a..b} = {} \ thesis" + "p\{} \ interior {a..b}\{} \ {a..b} \ {} \ thesis" + thus thesis by auto +next assume as:"p={}" guess p apply(rule elementary_interval[of a b]) . + thus thesis apply(rule_tac that[of p]) unfolding as by auto +next assume as:"{a..b}={}" show thesis apply(rule that) unfolding as using assms by auto +next assume as:"interior {a..b} = {}" "{a..b} \ {}" + show thesis apply(rule that[of "insert {a..b} p"],rule division_ofI) + unfolding finite_insert apply(rule assm(1)) unfolding Union_insert + using assm(2-4) as apply- by(fastsimp dest: assm(5))+ +next assume as:"p \ {}" "interior {a..b} \ {}" "{a..b}\{}" + have "\k\p. \q. (insert {a..b} q) division_of ({a..b} \ k)" proof case goal1 + from assm(4)[OF this] guess c .. then guess d .. + thus ?case apply-apply(rule division_union_intervals_exists[OF as(3),of c d]) by auto + qed from bchoice[OF this] guess q .. note q=division_ofD[OF this[rule_format]] + let ?D = "\{insert {a..b} (q k) | k. k \ p}" + show thesis apply(rule that[of "?D"]) proof(rule division_ofI) + have *:"{insert {a..b} (q k) |k. k \ p} = (\k. insert {a..b} (q k)) ` p" by auto + show "finite ?D" apply(rule finite_Union) unfolding * apply(rule finite_imageI) using assm(1) q(1) by auto + show "\?D = {a..b} \ \p" unfolding * lem1 unfolding lem2[OF as(1), of "{a..b}",THEN sym] + using q(6) by auto + fix k assume k:"k\?D" thus " k \ {a..b} \ \p" using q(2) by auto + show "k \ {}" using q(3) k by auto show "\a b. k = {a..b}" using q(4) k by auto + fix k' assume k':"k'\?D" "k\k'" + obtain x where x: "k \insert {a..b} (q x)" "x\p" using k by auto + obtain x' where x':"k'\insert {a..b} (q x')" "x'\p" using k' by auto + show "interior k \ interior k' = {}" proof(cases "x=x'") + case True show ?thesis apply(rule q(5)) using x x' k' unfolding True by auto + next case False + { presume "k = {a..b} \ ?thesis" "k' = {a..b} \ ?thesis" + "k \ {a..b} \ k' \ {a..b} \ ?thesis" + thus ?thesis by auto } + { assume as':"k = {a..b}" show ?thesis apply(rule q(5)) using x' k'(2) unfolding as' by auto } + { assume as':"k' = {a..b}" show ?thesis apply(rule q(5)) using x k'(2) unfolding as' by auto } + assume as':"k \ {a..b}" "k' \ {a..b}" + guess c using q(4)[OF x(2,1)] .. then guess d .. note c_d=this + have "interior k \ interior {a..b} = {}" apply(rule q(5)) using x k'(2) using as' by auto + hence "interior k \ interior x" apply- + apply(rule interior_subset_union_intervals[OF c_d _ as(2) q(2)[OF x(2,1)]]) by auto moreover + guess c using q(4)[OF x'(2,1)] .. then guess d .. note c_d=this + have "interior k' \ interior {a..b} = {}" apply(rule q(5)) using x' k'(2) using as' by auto + hence "interior k' \ interior x'" apply- + apply(rule interior_subset_union_intervals[OF c_d _ as(2) q(2)[OF x'(2,1)]]) by auto + ultimately show ?thesis using assm(5)[OF x(2) x'(2) False] by auto + qed qed } qed + +lemma elementary_unions_intervals: + assumes "finite f" "\s. s \ f \ \a b. s = {a..b::real^'n}" + obtains p where "p division_of (\f)" proof- + have "\p. p division_of (\f)" proof(induct_tac f rule:finite_subset_induct) + show "\p. p division_of \{}" using elementary_empty by auto + fix x F assume as:"finite F" "x \ F" "\p. p division_of \F" "x\f" + from this(3) guess p .. note p=this + from assms(2)[OF as(4)] guess a .. then guess b .. note ab=this + have *:"\F = \p" using division_ofD[OF p] by auto + show "\p. p division_of \insert x F" using elementary_union_interval[OF p[unfolded *], of a b] + unfolding Union_insert ab * by auto + qed(insert assms,auto) thus ?thesis apply-apply(erule exE,rule that) by auto qed + +lemma elementary_union: assumes "ps division_of s" "pt division_of (t::(real^'n) set)" + obtains p where "p division_of (s \ t)" +proof- have "s \ t = \ps \ \pt" using assms unfolding division_of_def by auto + hence *:"\(ps \ pt) = s \ t" by auto + show ?thesis apply-apply(rule elementary_unions_intervals[of "ps\pt"]) + unfolding * prefer 3 apply(rule_tac p=p in that) + using assms[unfolded division_of_def] by auto qed + +lemma partial_division_extend: fixes t::"(real^'n) set" + assumes "p division_of s" "q division_of t" "s \ t" + obtains r where "p \ r" "r division_of t" proof- + note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)] + obtain a b where ab:"t\{a..b}" using elementary_subset_interval[OF assms(2)] by auto + guess r1 apply(rule partial_division_extend_interval) apply(rule assms(1)[unfolded divp(6)[THEN sym]]) + apply(rule subset_trans) by(rule ab assms[unfolded divp(6)[THEN sym]])+ note r1 = this division_ofD[OF this(2)] + guess p' apply(rule elementary_unions_intervals[of "r1 - p"]) using r1(3,6) by auto + then obtain r2 where r2:"r2 division_of (\(r1 - p)) \ (\q)" + apply- apply(drule elementary_inter[OF _ assms(2)[unfolded divq(6)[THEN sym]]]) by auto + { fix x assume x:"x\t" "x\s" + hence "x\\r1" unfolding r1 using ab by auto + then guess r unfolding Union_iff .. note r=this moreover + have "r \ p" proof assume "r\p" hence "x\s" using divp(2) r by auto + thus False using x by auto qed + ultimately have "x\\(r1 - p)" by auto } + hence *:"t = \p \ (\(r1 - p) \ \q)" unfolding divp divq using assms(3) by auto + show ?thesis apply(rule that[of "p \ r2"]) unfolding * defer apply(rule division_disjoint_union) + unfolding divp(6) apply(rule assms r2)+ + proof- have "interior s \ interior (\(r1-p)) = {}" + proof(rule inter_interior_unions_intervals) + show "finite (r1 - p)" "open (interior s)" "\t\r1-p. \a b. t = {a..b}" using r1 by auto + have *:"\s. (\x. x \ s \ False) \ s = {}" by auto + show "\t\r1-p. interior s \ interior t = {}" proof(rule) + fix m x assume as:"m\r1-p" + have "interior m \ interior (\p) = {}" proof(rule inter_interior_unions_intervals) + show "finite p" "open (interior m)" "\t\p. \a b. t = {a..b}" using divp by auto + show "\t\p. interior m \ interior t = {}" apply(rule, rule r1(7)) using as using r1 by auto + qed thus "interior s \ interior m = {}" unfolding divp by auto + qed qed + thus "interior s \ interior (\(r1-p) \ (\q)) = {}" using interior_subset by auto + qed auto qed + +subsection {* Tagged (partial) divisions. *} + +definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) where + "(s tagged_partial_division_of i) \ + finite s \ + (\x k. (x,k) \ s \ x \ k \ k \ i \ (\a b. k = {a..b})) \ + (\x1 k1 x2 k2. (x1,k1) \ s \ (x2,k2) \ s \ ((x1,k1) \ (x2,k2)) + \ (interior(k1) \ interior(k2) = {}))" + +lemma tagged_partial_division_ofD[dest]: assumes "s tagged_partial_division_of i" + shows "finite s" "\x k. (x,k) \ s \ x \ k" "\x k. (x,k) \ s \ k \ i" + "\x k. (x,k) \ s \ \a b. k = {a..b}" + "\x1 k1 x2 k2. (x1,k1) \ s \ (x2,k2) \ s \ (x1,k1) \ (x2,k2) \ interior(k1) \ interior(k2) = {}" + using assms unfolding tagged_partial_division_of_def apply- by blast+ + +definition tagged_division_of (infixr "tagged'_division'_of" 40) where + "(s tagged_division_of i) \ + (s tagged_partial_division_of i) \ (\{k. \x. (x,k) \ s} = i)" + +lemma tagged_division_of_finite[dest]: "s tagged_division_of i \ finite s" + unfolding tagged_division_of_def tagged_partial_division_of_def by auto + +lemma tagged_division_of: + "(s tagged_division_of i) \ + finite s \ + (\x k. (x,k) \ s + \ x \ k \ k \ i \ (\a b. k = {a..b})) \ + (\x1 k1 x2 k2. (x1,k1) \ s \ (x2,k2) \ s \ ~((x1,k1) = (x2,k2)) + \ (interior(k1) \ interior(k2) = {})) \ + (\{k. \x. (x,k) \ s} = i)" + unfolding tagged_division_of_def tagged_partial_division_of_def by auto + +lemma tagged_division_ofI: assumes + "finite s" "\x k. (x,k) \ s \ x \ k" "\x k. (x,k) \ s \ k \ i" "\x k. (x,k) \ s \ \a b. k = {a..b}" + "\x1 k1 x2 k2. (x1,k1) \ s \ (x2,k2) \ s \ ~((x1,k1) = (x2,k2)) \ (interior(k1) \ interior(k2) = {})" + "(\{k. \x. (x,k) \ s} = i)" + shows "s tagged_division_of i" + unfolding tagged_division_of apply(rule) defer apply rule + apply(rule allI impI conjI assms)+ apply assumption + apply(rule, rule assms, assumption) apply(rule assms, assumption) + using assms(1,5-) apply- by blast+ + +lemma tagged_division_ofD[dest]: assumes "s tagged_division_of i" + shows "finite s" "\x k. (x,k) \ s \ x \ k" "\x k. (x,k) \ s \ k \ i" "\x k. (x,k) \ s \ \a b. k = {a..b}" + "\x1 k1 x2 k2. (x1,k1) \ s \ (x2,k2) \ s \ ~((x1,k1) = (x2,k2)) \ (interior(k1) \ interior(k2) = {})" + "(\{k. \x. (x,k) \ s} = i)" using assms unfolding tagged_division_of apply- by blast+ + +lemma division_of_tagged_division: assumes"s tagged_division_of i" shows "(snd ` s) division_of i" +proof(rule division_ofI) note assm=tagged_division_ofD[OF assms] + show "\snd ` s = i" "finite (snd ` s)" using assm by auto + fix k assume k:"k \ snd ` s" then obtain xk where xk:"(xk, k) \ s" by auto + thus "k \ i" "k \ {}" "\a b. k = {a..b}" using assm apply- by fastsimp+ + fix k' assume k':"k' \ snd ` s" "k \ k'" from this(1) obtain xk' where xk':"(xk', k') \ s" by auto + thus "interior k \ interior k' = {}" apply-apply(rule assm(5)) apply(rule xk xk')+ using k' by auto +qed + +lemma partial_division_of_tagged_division: assumes "s tagged_partial_division_of i" + shows "(snd ` s) division_of \(snd ` s)" +proof(rule division_ofI) note assm=tagged_partial_division_ofD[OF assms] + show "finite (snd ` s)" "\snd ` s = \snd ` s" using assm by auto + fix k assume k:"k \ snd ` s" then obtain xk where xk:"(xk, k) \ s" by auto + thus "k\{}" "\a b. k = {a..b}" "k \ \snd ` s" using assm by auto + fix k' assume k':"k' \ snd ` s" "k \ k'" from this(1) obtain xk' where xk':"(xk', k') \ s" by auto + thus "interior k \ interior k' = {}" apply-apply(rule assm(5)) apply(rule xk xk')+ using k' by auto +qed + +lemma tagged_partial_division_subset: assumes "s tagged_partial_division_of i" "t \ s" + shows "t tagged_partial_division_of i" + using assms unfolding tagged_partial_division_of_def using finite_subset[OF assms(2)] by blast + +lemma setsum_over_tagged_division_lemma: fixes d::"(real^'m) set \ 'a::real_normed_vector" + assumes "p tagged_division_of i" "\u v. {u..v} \ {} \ content {u..v} = 0 \ d {u..v} = 0" + shows "setsum (\(x,k). d k) p = setsum d (snd ` p)" +proof- note assm=tagged_division_ofD[OF assms(1)] + have *:"(\(x,k). d k) = d \ snd" unfolding o_def apply(rule ext) by auto + show ?thesis unfolding * apply(subst eq_commute) proof(rule setsum_reindex_nonzero) + show "finite p" using assm by auto + fix x y assume as:"x\p" "y\p" "x\y" "snd x = snd y" + obtain a b where ab:"snd x = {a..b}" using assm(4)[of "fst x" "snd x"] as(1) by auto + have "(fst x, snd y) \ p" "(fst x, snd y) \ y" unfolding as(4)[THEN sym] using as(1-3) by auto + hence "interior (snd x) \ interior (snd y) = {}" apply-apply(rule assm(5)[of "fst x" _ "fst y"]) using as by auto + hence "content {a..b} = 0" unfolding as(4)[THEN sym] ab content_eq_0_interior by auto + hence "d {a..b} = 0" apply-apply(rule assms(2)) using assm(2)[of "fst x" "snd x"] as(1) unfolding ab[THEN sym] by auto + thus "d (snd x) = 0" unfolding ab by auto qed qed + +lemma tag_in_interval: "p tagged_division_of i \ (x,k) \ p \ x \ i" by auto + +lemma tagged_division_of_empty: "{} tagged_division_of {}" + unfolding tagged_division_of by auto + +lemma tagged_partial_division_of_trivial[simp]: + "p tagged_partial_division_of {} \ p = {}" + unfolding tagged_partial_division_of_def by auto + +lemma tagged_division_of_trivial[simp]: + "p tagged_division_of {} \ p = {}" + unfolding tagged_division_of by auto + +lemma tagged_division_of_self: + "x \ {a..b} \ {(x,{a..b})} tagged_division_of {a..b}" + apply(rule tagged_division_ofI) by auto + +lemma tagged_division_union: + assumes "p1 tagged_division_of s1" "p2 tagged_division_of s2" "interior s1 \ interior s2 = {}" + shows "(p1 \ p2) tagged_division_of (s1 \ s2)" +proof(rule tagged_division_ofI) note p1=tagged_division_ofD[OF assms(1)] and p2=tagged_division_ofD[OF assms(2)] + show "finite (p1 \ p2)" using p1(1) p2(1) by auto + show "\{k. \x. (x, k) \ p1 \ p2} = s1 \ s2" using p1(6) p2(6) by blast + fix x k assume xk:"(x,k)\p1\p2" show "x\k" "\a b. k = {a..b}" using xk p1(2,4) p2(2,4) by auto + show "k\s1\s2" using xk p1(3) p2(3) by blast + fix x' k' assume xk':"(x',k')\p1\p2" "(x,k) \ (x',k')" + have *:"\a b. a\ s1 \ b\ s2 \ interior a \ interior b = {}" using assms(3) subset_interior by blast + show "interior k \ interior k' = {}" apply(cases "(x,k)\p1", case_tac[!] "(x',k')\p1") + apply(rule p1(5)) prefer 4 apply(rule *) prefer 6 apply(subst Int_commute,rule *) prefer 8 apply(rule p2(5)) + using p1(3) p2(3) using xk xk' by auto qed + +lemma tagged_division_unions: + assumes "finite iset" "\i\iset. (pfn(i) tagged_division_of i)" + "\i1 \ iset. \i2 \ iset. ~(i1 = i2) \ (interior(i1) \ interior(i2) = {})" + shows "\(pfn ` iset) tagged_division_of (\iset)" +proof(rule tagged_division_ofI) + note assm = tagged_division_ofD[OF assms(2)[rule_format]] + show "finite (\pfn ` iset)" apply(rule finite_Union) using assms by auto + have "\{k. \x. (x, k) \ \pfn ` iset} = \(\i. \{k. \x. (x, k) \ pfn i}) ` iset" by blast + also have "\ = \iset" using assm(6) by auto + finally show "\{k. \x. (x, k) \ \pfn ` iset} = \iset" . + fix x k assume xk:"(x,k)\\pfn ` iset" then obtain i where i:"i \ iset" "(x, k) \ pfn i" by auto + show "x\k" "\a b. k = {a..b}" "k \ \iset" using assm(2-4)[OF i] using i(1) by auto + fix x' k' assume xk':"(x',k')\\pfn ` iset" "(x, k) \ (x', k')" then obtain i' where i':"i' \ iset" "(x', k') \ pfn i'" by auto + have *:"\a b. i\i' \ a\ i \ b\ i' \ interior a \ interior b = {}" using i(1) i'(1) + using assms(3)[rule_format] subset_interior by blast + show "interior k \ interior k' = {}" apply(cases "i=i'") + using assm(5)[OF i _ xk'(2)] i'(2) using assm(3)[OF i] assm(3)[OF i'] defer apply-apply(rule *) by auto +qed + +lemma tagged_partial_division_of_union_self: + assumes "p tagged_partial_division_of s" shows "p tagged_division_of (\(snd ` p))" + apply(rule tagged_division_ofI) using tagged_partial_division_ofD[OF assms] by auto + +lemma tagged_division_of_union_self: assumes "p tagged_division_of s" + shows "p tagged_division_of (\(snd ` p))" + apply(rule tagged_division_ofI) using tagged_division_ofD[OF assms] by auto + +subsection {* Fine-ness of a partition w.r.t. a gauge. *} + +definition fine (infixr "fine" 46) where + "d fine s \ (\(x,k) \ s. k \ d(x))" + +lemma fineI: assumes "\x k. (x,k) \ s \ k \ d x" + shows "d fine s" using assms unfolding fine_def by auto + +lemma fineD[dest]: assumes "d fine s" + shows "\x k. (x,k) \ s \ k \ d x" using assms unfolding fine_def by auto + +lemma fine_inter: "(\x. d1 x \ d2 x) fine p \ d1 fine p \ d2 fine p" + unfolding fine_def by auto + +lemma fine_inters: + "(\x. \ {f d x | d. d \ s}) fine p \ (\d\s. (f d) fine p)" + unfolding fine_def by blast + +lemma fine_union: + "d fine p1 \ d fine p2 \ d fine (p1 \ p2)" + unfolding fine_def by blast + +lemma fine_unions:"(\p. p \ ps \ d fine p) \ d fine (\ps)" + unfolding fine_def by auto + +lemma fine_subset: "p \ q \ d fine q \ d fine p" + unfolding fine_def by blast + +subsection {* Gauge integral. Define on compact intervals first, then use a limit. *} + +definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46) where + "(f has_integral_compact_interval y) i \ + (\e>0. \d. gauge d \ + (\p. p tagged_division_of i \ d fine p + \ norm(setsum (\(x,k). content k *\<^sub>R f x) p - y) < e))" + +definition has_integral (infixr "has'_integral" 46) where +"((f::(real^'n \ 'b::real_normed_vector)) has_integral y) i \ + if (\a b. i = {a..b}) then (f has_integral_compact_interval y) i + else (\e>0. \B>0. \a b. ball 0 B \ {a..b} + \ (\z. ((\x. if x \ i then f x else 0) has_integral_compact_interval z) {a..b} \ + norm(z - y) < e))" + +lemma has_integral: + "(f has_integral y) ({a..b}) \ + (\e>0. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p + \ norm(setsum (\(x,k). content(k) *\<^sub>R f x) p - y) < e))" + unfolding has_integral_def has_integral_compact_interval_def by auto + +lemma has_integralD[dest]: assumes + "(f has_integral y) ({a..b})" "e>0" + obtains d where "gauge d" "\p. p tagged_division_of {a..b} \ d fine p + \ norm(setsum (\(x,k). content(k) *\<^sub>R f(x)) p - y) < e" + using assms unfolding has_integral by auto + +lemma has_integral_alt: + "(f has_integral y) i \ + (if (\a b. i = {a..b}) then (f has_integral y) i + else (\e>0. \B>0. \a b. ball 0 B \ {a..b} + \ (\z. ((\x. if x \ i then f(x) else 0) + has_integral z) ({a..b}) \ + norm(z - y) < e)))" + unfolding has_integral unfolding has_integral_compact_interval_def has_integral_def by auto + +lemma has_integral_altD: + assumes "(f has_integral y) i" "\ (\a b. i = {a..b})" "e>0" + obtains B where "B>0" "\a b. ball 0 B \ {a..b}\ (\z. ((\x. if x \ i then f(x) else 0) has_integral z) ({a..b}) \ norm(z - y) < e)" + using assms unfolding has_integral unfolding has_integral_compact_interval_def has_integral_def by auto + +definition integrable_on (infixr "integrable'_on" 46) where + "(f integrable_on i) \ \y. (f has_integral y) i" + +definition "integral i f \ SOME y. (f has_integral y) i" + +lemma integrable_integral[dest]: + "f integrable_on i \ (f has_integral (integral i f)) i" + unfolding integrable_on_def integral_def by(rule someI_ex) + +lemma has_integral_integrable[intro]: "(f has_integral i) s \ f integrable_on s" + unfolding integrable_on_def by auto + +lemma has_integral_integral:"f integrable_on s \ (f has_integral (integral s f)) s" + by auto + +lemma setsum_content_null: + assumes "content({a..b}) = 0" "p tagged_division_of {a..b}" + shows "setsum (\(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)" +proof(rule setsum_0',rule) fix y assume y:"y\p" + obtain x k where xk:"y = (x,k)" using surj_pair[of y] by blast + note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]] + from this(2) guess c .. then guess d .. note c_d=this + have "(\(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x" unfolding xk by auto + also have "\ = 0" using content_subset[OF assm(1)[unfolded c_d]] content_pos_le[of c d] + unfolding assms(1) c_d by auto + finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . +qed + +subsection {* Some basic combining lemmas. *} + +lemma tagged_division_unions_exists: + assumes "finite iset" "\i \ iset. \p. p tagged_division_of i \ d fine p" + "\i1\iset. \i2\iset. ~(i1 = i2) \ (interior(i1) \ interior(i2) = {})" "(\iset = i)" + obtains p where "p tagged_division_of i" "d fine p" +proof- guess pfn using bchoice[OF assms(2)] .. note pfn = conjunctD2[OF this[rule_format]] + show thesis apply(rule_tac p="\(pfn ` iset)" in that) unfolding assms(4)[THEN sym] + apply(rule tagged_division_unions[OF assms(1) _ assms(3)]) defer + apply(rule fine_unions) using pfn by auto +qed + +subsection {* The set we're concerned with must be closed. *} + +lemma division_of_closed: "s division_of i \ closed (i::(real^'n) set)" + unfolding division_of_def by(fastsimp intro!: closed_Union closed_interval) + +subsection {* General bisection principle for intervals; might be useful elsewhere. *} + +lemma interval_bisection_step: + assumes "P {}" "(\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P(s \ t))" "~(P {a..b::real^'n})" + obtains c d where "~(P{c..d})" + "\i. a$i \ c$i \ c$i \ d$i \ d$i \ b$i \ 2 * (d$i - c$i) \ b$i - a$i" +proof- have "{a..b} \ {}" using assms(1,3) by auto + note ab=this[unfolded interval_eq_empty not_ex not_less] + { fix f have "finite f \ + (\s\f. P s) \ + (\s\f. \a b. s = {a..b}) \ + (\s\f.\t\f. ~(s = t) \ interior(s) \ interior(t) = {}) \ P(\f)" + proof(induct f rule:finite_induct) + case empty show ?case using assms(1) by auto + next case (insert x f) show ?case unfolding Union_insert apply(rule assms(2)[rule_format]) + apply rule defer apply rule defer apply(rule inter_interior_unions_intervals) + using insert by auto + qed } note * = this + let ?A = "{{c..d} | c d. \i. (c$i = a$i) \ (d$i = (a$i + b$i) / 2) \ (c$i = (a$i + b$i) / 2) \ (d$i = b$i)}" + let ?PP = "\c d. \i. a$i \ c$i \ c$i \ d$i \ d$i \ b$i \ 2 * (d$i - c$i) \ b$i - a$i" + { presume "\c d. ?PP c d \ P {c..d} \ False" + thus thesis unfolding atomize_not not_all apply-apply(erule exE)+ apply(rule_tac c=x and d=xa in that) by auto } + assume as:"\c d. ?PP c d \ P {c..d}" + have "P (\ ?A)" proof(rule *, rule_tac[2-] ballI, rule_tac[4] ballI, rule_tac[4] impI) + let ?B = "(\s.{(\ i. if i \ s then a$i else (a$i + b$i) / 2) .. + (\ i. if i \ s then (a$i + b$i) / 2 else b$i)}) ` {s. s \ UNIV}" + have "?A \ ?B" proof case goal1 + then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+) note c_d=this[rule_format] + have *:"\a b c d. a = c \ b = d \ {a..b} = {c..d}" by auto + show "x\?B" unfolding image_iff apply(rule_tac x="{i. c$i = a$i}" in bexI) + unfolding c_d apply(rule * ) unfolding Cart_eq cond_component Cart_lambda_beta + proof(rule_tac[1-2] allI) fix i show "c $ i = (if i \ {i. c $ i = a $ i} then a $ i else (a $ i + b $ i) / 2)" + "d $ i = (if i \ {i. c $ i = a $ i} then (a $ i + b $ i) / 2 else b $ i)" + using c_d(2)[of i] ab[THEN spec[where x=i]] by(auto simp add:field_simps) + qed auto qed + thus "finite ?A" apply(rule finite_subset[of _ ?B]) by auto + fix s assume "s\?A" then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+) + note c_d=this[rule_format] + show "P s" unfolding c_d apply(rule as[rule_format]) proof- case goal1 show ?case + using c_d(2)[of i] using ab[THEN spec[where x=i]] by auto qed + show "\a b. s = {a..b}" unfolding c_d by auto + fix t assume "t\?A" then guess e unfolding mem_Collect_eq .. then guess f apply- by(erule exE,(erule conjE)+) + note e_f=this[rule_format] + assume "s \ t" hence "\ (c = e \ d = f)" unfolding c_d e_f by auto + then obtain i where "c$i \ e$i \ d$i \ f$i" unfolding de_Morgan_conj Cart_eq by auto + hence i:"c$i \ e$i" "d$i \ f$i" apply- apply(erule_tac[!] disjE) + proof- assume "c$i \ e$i" thus "d$i \ f$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp + next assume "d$i \ f$i" thus "c$i \ e$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp + qed have *:"\s t. (\a. a\s \ a\t \ False) \ s \ t = {}" by auto + show "interior s \ interior t = {}" unfolding e_f c_d interior_closed_interval proof(rule *) + fix x assume "x\{c<..{e<.. ?A = {a..b}" proof(rule set_ext,rule) + fix x assume "x\\?A" then guess Y unfolding Union_iff .. + from this(1) guess c unfolding mem_Collect_eq .. then guess d .. + note c_d = this[THEN conjunct2,rule_format] `x\Y`[unfolded this[THEN conjunct1]] + show "x\{a..b}" unfolding mem_interval proof + fix i show "a $ i \ x $ i \ x $ i \ b $ i" + using c_d(1)[of i] c_d(2)[unfolded mem_interval,THEN spec[where x=i]] by auto qed + next fix x assume x:"x\{a..b}" + have "\i. \c d. (c = a$i \ d = (a$i + b$i) / 2 \ c = (a$i + b$i) / 2 \ d = b$i) \ c\x$i \ x$i \ d" + (is "\i. \c d. ?P i c d") unfolding mem_interval proof fix i + have "?P i (a$i) ((a $ i + b $ i) / 2) \ ?P i ((a $ i + b $ i) / 2) (b$i)" + using x[unfolded mem_interval,THEN spec[where x=i]] by auto thus "\c d. ?P i c d" by blast + qed thus "x\\?A" unfolding Union_iff lambda_skolem unfolding Bex_def mem_Collect_eq + apply-apply(erule exE)+ apply(rule_tac x="{xa..xaa}" in exI) unfolding mem_interval by auto + qed finally show False using assms by auto qed + +lemma interval_bisection: + assumes "P {}" "(\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P(s \ t))" "\ P {a..b::real^'n}" + obtains x where "x \ {a..b}" "\e>0. \c d. x \ {c..d} \ {c..d} \ ball x e \ {c..d} \ {a..b} \ ~P({c..d})" +proof- + have "\x. \y. \ P {fst x..snd x} \ (\ P {fst y..snd y} \ (\i. fst x$i \ fst y$i \ fst y$i \ snd y$i \ snd y$i \ snd x$i \ + 2 * (snd y$i - fst y$i) \ snd x$i - fst x$i))" proof case goal1 thus ?case proof- + presume "\ P {fst x..snd x} \ ?thesis" + thus ?thesis apply(cases "P {fst x..snd x}") by auto + next assume as:"\ P {fst x..snd x}" from interval_bisection_step[of P, OF assms(1-2) as] guess c d . + thus ?thesis apply- apply(rule_tac x="(c,d)" in exI) by auto + qed qed then guess f apply-apply(drule choice) by(erule exE) note f=this + def AB \ "\n. (f ^^ n) (a,b)" def A \ "\n. fst(AB n)" and B \ "\n. snd(AB n)" note ab_def = this AB_def + have "A 0 = a" "B 0 = b" "\n. \ P {A(Suc n)..B(Suc n)} \ + (\i. A(n)$i \ A(Suc n)$i \ A(Suc n)$i \ B(Suc n)$i \ B(Suc n)$i \ B(n)$i \ + 2 * (B(Suc n)$i - A(Suc n)$i) \ B(n)$i - A(n)$i)" (is "\n. ?P n") + proof- show "A 0 = a" "B 0 = b" unfolding ab_def by auto + case goal3 note S = ab_def funpow.simps o_def id_apply show ?case + proof(induct n) case 0 thus ?case unfolding S apply(rule f[rule_format]) using assms(3) by auto + next case (Suc n) show ?case unfolding S apply(rule f[rule_format]) using Suc unfolding S by auto + qed qed note AB = this(1-2) conjunctD2[OF this(3),rule_format] + + have interv:"\e. 0 < e \ \n. \x\{A n..B n}. \y\{A n..B n}. dist x y < e" + proof- case goal1 guess n using real_arch_pow2[of "(setsum (\i. b$i - a$i) UNIV) / e"] .. note n=this + show ?case apply(rule_tac x=n in exI) proof(rule,rule) + fix x y assume xy:"x\{A n..B n}" "y\{A n..B n}" + have "dist x y \ setsum (\i. abs((x - y)$i)) UNIV" unfolding vector_dist_norm by(rule norm_le_l1) + also have "\ \ setsum (\i. B n$i - A n$i) UNIV" + proof(rule setsum_mono) fix i show "\(x - y) $ i\ \ B n $ i - A n $ i" + using xy[unfolded mem_interval,THEN spec[where x=i]] + unfolding vector_minus_component by auto qed + also have "\ \ setsum (\i. b$i - a$i) UNIV / 2^n" unfolding setsum_divide_distrib + proof(rule setsum_mono) case goal1 thus ?case + proof(induct n) case 0 thus ?case unfolding AB by auto + next case (Suc n) have "B (Suc n) $ i - A (Suc n) $ i \ (B n $ i - A n $ i) / 2" using AB(4)[of n i] by auto + also have "\ \ (b $ i - a $ i) / 2 ^ Suc n" using Suc by(auto simp add:field_simps) finally show ?case . + qed qed + also have "\ < e" using n using goal1 by(auto simp add:field_simps) finally show "dist x y < e" . + qed qed + { fix n m ::nat assume "m \ n" then guess d unfolding le_Suc_ex_iff .. note d=this + have "{A n..B n} \ {A m..B m}" unfolding d + proof(induct d) case 0 thus ?case by auto + next case (Suc d) show ?case apply(rule subset_trans[OF _ Suc]) + apply(rule) unfolding mem_interval apply(rule,erule_tac x=i in allE) + proof- case goal1 thus ?case using AB(4)[of "m + d" i] by(auto simp add:field_simps) + qed qed } note ABsubset = this + have "\a. \n. a\{A n..B n}" apply(rule decreasing_closed_nest[rule_format,OF closed_interval _ ABsubset interv]) + proof- fix n show "{A n..B n} \ {}" apply(cases "0{a..b}" using x0[of 0] unfolding AB . + fix e assume "0 < (e::real)" from interv[OF this] guess n .. note n=this + show "\c d. x0 \ {c..d} \ {c..d} \ ball x0 e \ {c..d} \ {a..b} \ \ P {c..d}" + apply(rule_tac x="A n" in exI,rule_tac x="B n" in exI) apply(rule,rule x0) apply rule defer + proof show "\ P {A n..B n}" apply(cases "0 ball x0 e" using n using x0[of n] by auto + show "{A n..B n} \ {a..b}" unfolding AB(1-2)[symmetric] apply(rule ABsubset) by auto + qed qed qed + +subsection {* Cousin's lemma. *} + +lemma fine_division_exists: assumes "gauge g" + obtains p where "p tagged_division_of {a..b::real^'n}" "g fine p" +proof- presume "\ (\p. p tagged_division_of {a..b} \ g fine p) \ False" + then guess p unfolding atomize_not not_not .. thus thesis apply-apply(rule that[of p]) by auto +next assume as:"\ (\p. p tagged_division_of {a..b} \ g fine p)" + guess x apply(rule interval_bisection[of "\s. \p. p tagged_division_of s \ g fine p",rule_format,OF _ _ as]) + apply(rule_tac x="{}" in exI) defer apply(erule conjE exE)+ + proof- show "{} tagged_division_of {} \ g fine {}" unfolding fine_def by auto + fix s t p p' assume "p tagged_division_of s" "g fine p" "p' tagged_division_of t" "g fine p'" "interior s \ interior t = {}" + thus "\p. p tagged_division_of s \ t \ g fine p" apply-apply(rule_tac x="p \ p'" in exI) apply rule + apply(rule tagged_division_union) prefer 4 apply(rule fine_union) by auto + qed note x=this + obtain e where e:"e>0" "ball x e \ g x" using gaugeD[OF assms, of x] unfolding open_contains_ball by auto + from x(2)[OF e(1)] guess c d apply-apply(erule exE conjE)+ . note c_d = this + have "g fine {(x, {c..d})}" unfolding fine_def using e using c_d(2) by auto + thus False using tagged_division_of_self[OF c_d(1)] using c_d by auto qed + +subsection {* Basic theorems about integrals. *} + +lemma has_integral_unique: fixes f::"real^'n \ 'a::real_normed_vector" + assumes "(f has_integral k1) i" "(f has_integral k2) i" shows "k1 = k2" +proof(rule ccontr) let ?e = "norm(k1 - k2) / 2" assume as:"k1 \ k2" hence e:"?e > 0" by auto + have lem:"\f::real^'n \ 'a. \ a b k1 k2. + (f has_integral k1) ({a..b}) \ (f has_integral k2) ({a..b}) \ k1 \ k2 \ False" + proof- case goal1 let ?e = "norm(k1 - k2) / 2" from goal1(3) have e:"?e > 0" by auto + guess d1 by(rule has_integralD[OF goal1(1) e]) note d1=this + guess d2 by(rule has_integralD[OF goal1(2) e]) note d2=this + guess p by(rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)],of a b]) note p=this + let ?c = "(\(x, k)\p. content k *\<^sub>R f x)" have "norm (k1 - k2) \ norm (?c - k2) + norm (?c - k1)" + using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] by(auto simp add:group_simps norm_minus_commute) + also have "\ < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" + apply(rule add_strict_mono) apply(rule_tac[!] d2(2) d1(2)) using p unfolding fine_def by auto + finally show False by auto + qed { presume "\ (\a b. i = {a..b}) \ False" + thus False apply-apply(cases "\a b. i = {a..b}") + using assms by(auto simp add:has_integral intro:lem[OF _ _ as]) } + assume as:"\ (\a b. i = {a..b})" + guess B1 by(rule has_integral_altD[OF assms(1) as,OF e]) note B1=this[rule_format] + guess B2 by(rule has_integral_altD[OF assms(2) as,OF e]) note B2=this[rule_format] + have "\a b::real^'n. ball 0 B1 \ ball 0 B2 \ {a..b}" apply(rule bounded_subset_closed_interval) + using bounded_Un bounded_ball by auto then guess a b apply-by(erule exE)+ + note ab=conjunctD2[OF this[unfolded Un_subset_iff]] + guess w using B1(2)[OF ab(1)] .. note w=conjunctD2[OF this] + guess z using B2(2)[OF ab(2)] .. note z=conjunctD2[OF this] + have "z = w" using lem[OF w(1) z(1)] by auto + hence "norm (k1 - k2) \ norm (z - k2) + norm (w - k1)" + using norm_triangle_ineq4[of "k1 - w" "k2 - z"] by(auto simp add: norm_minus_commute) + also have "\ < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" apply(rule add_strict_mono) by(rule_tac[!] z(2) w(2)) + finally show False by auto qed + +lemma integral_unique[intro]: + "(f has_integral y) k \ integral k f = y" + unfolding integral_def apply(rule some_equality) by(auto intro: has_integral_unique) + +lemma has_integral_is_0: fixes f::"real^'n \ 'a::real_normed_vector" + assumes "\x\s. f x = 0" shows "(f has_integral 0) s" +proof- have lem:"\a b. \f::real^'n \ 'a. + (\x\{a..b}. f(x) = 0) \ (f has_integral 0) ({a..b})" unfolding has_integral + proof(rule,rule) fix a b e and f::"real^'n \ 'a" + assume as:"\x\{a..b}. f x = 0" "0 < (e::real)" + show "\d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e)" + apply(rule_tac x="\x. ball x 1" in exI) apply(rule,rule gaugeI) unfolding centre_in_ball defer apply(rule open_ball) + proof(rule,rule,erule conjE) case goal1 + have "(\(x, k)\p. content k *\<^sub>R f x) = 0" proof(rule setsum_0',rule) + fix x assume x:"x\p" have "f (fst x) = 0" using tagged_division_ofD(2-3)[OF goal1(1), of "fst x" "snd x"] using as x by auto + thus "(\(x, k). content k *\<^sub>R f x) x = 0" apply(subst surjective_pairing[of x]) unfolding split_conv by auto + qed thus ?case using as by auto + qed auto qed { presume "\ (\a b. s = {a..b}) \ ?thesis" + thus ?thesis apply-apply(cases "\a b. s = {a..b}") + using assms by(auto simp add:has_integral intro:lem) } + have *:"(\x. if x \ s then f x else 0) = (\x. 0)" apply(rule ext) using assms by auto + assume "\ (\a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P * + apply(rule,rule,rule_tac x=1 in exI,rule) defer apply(rule,rule,rule) + proof- fix e::real and a b assume "e>0" + thus "\z. ((\x::real^'n. 0::'a) has_integral z) {a..b} \ norm (z - 0) < e" + apply(rule_tac x=0 in exI) apply(rule,rule lem) by auto + qed auto qed + +lemma has_integral_0[simp]: "((\x::real^'n. 0) has_integral 0) s" + apply(rule has_integral_is_0) by auto + +lemma has_integral_0_eq[simp]: "((\x. 0) has_integral i) s \ i = 0" + using has_integral_unique[OF has_integral_0] by auto + +lemma has_integral_linear: fixes f::"real^'n \ 'a::real_normed_vector" + assumes "(f has_integral y) s" "bounded_linear h" shows "((h o f) has_integral ((h y))) s" +proof- interpret bounded_linear h using assms(2) . from pos_bounded guess B .. note B=conjunctD2[OF this,rule_format] + have lem:"\f::real^'n \ 'a. \ y a b. + (f has_integral y) ({a..b}) \ ((h o f) has_integral h(y)) ({a..b})" + proof(subst has_integral,rule,rule) case goal1 + from pos_bounded guess B .. note B=conjunctD2[OF this,rule_format] + have *:"e / B > 0" apply(rule divide_pos_pos) using goal1(2) B by auto + guess g using has_integralD[OF goal1(1) *] . note g=this + show ?case apply(rule_tac x=g in exI) apply(rule,rule g(1)) + proof(rule,rule,erule conjE) fix p assume as:"p tagged_division_of {a..b}" "g fine p" + have *:"\x k. h ((\(x, k). content k *\<^sub>R f x) x) = (\(x, k). h (content k *\<^sub>R f x)) x" by auto + have "(\(x, k)\p. content k *\<^sub>R (h \ f) x) = setsum (h \ (\(x, k). content k *\<^sub>R f x)) p" + unfolding o_def unfolding scaleR[THEN sym] * by simp + also have "\ = h (\(x, k)\p. content k *\<^sub>R f x)" using setsum[of "\(x,k). content k *\<^sub>R f x" p] using as by auto + finally have *:"(\(x, k)\p. content k *\<^sub>R (h \ f) x) = h (\(x, k)\p. content k *\<^sub>R f x)" . + show "norm ((\(x, k)\p. content k *\<^sub>R (h \ f) x) - h y) < e" unfolding * diff[THEN sym] + apply(rule le_less_trans[OF B(2)]) using g(2)[OF as] B(1) by(auto simp add:field_simps) + qed qed { presume "\ (\a b. s = {a..b}) \ ?thesis" + thus ?thesis apply-apply(cases "\a b. s = {a..b}") using assms by(auto simp add:has_integral intro!:lem) } + assume as:"\ (\a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P + proof(rule,rule) fix e::real assume e:"0B>0. \a b. ball 0 B \ {a..b} \ (\z. ((\x. if x \ s then (h \ f) x else 0) has_integral z) {a..b} \ norm (z - h y) < e)" + apply(rule_tac x=M in exI) apply(rule,rule M(1)) + proof(rule,rule,rule) case goal1 guess z using M(2)[OF goal1(1)] .. note z=conjunctD2[OF this] + have *:"(\x. if x \ s then (h \ f) x else 0) = h \ (\x. if x \ s then f x else 0)" + unfolding o_def apply(rule ext) using zero by auto + show ?case apply(rule_tac x="h z" in exI,rule) unfolding * apply(rule lem[OF z(1)]) unfolding diff[THEN sym] + apply(rule le_less_trans[OF B(2)]) using B(1) z(2) by(auto simp add:field_simps) + qed qed qed + +lemma has_integral_cmul: + shows "(f has_integral k) s \ ((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) s" + unfolding o_def[THEN sym] apply(rule has_integral_linear,assumption) + by(rule scaleR.bounded_linear_right) + +lemma has_integral_neg: + shows "(f has_integral k) s \ ((\x. -(f x)) has_integral (-k)) s" + apply(drule_tac c="-1" in has_integral_cmul) by auto + +lemma has_integral_add: fixes f::"real^'n \ 'a::real_normed_vector" + assumes "(f has_integral k) s" "(g has_integral l) s" + shows "((\x. f x + g x) has_integral (k + l)) s" +proof- have lem:"\f g::real^'n \ 'a. \a b k l. + (f has_integral k) ({a..b}) \ (g has_integral l) ({a..b}) \ + ((\x. f(x) + g(x)) has_integral (k + l)) ({a..b})" proof- case goal1 + show ?case unfolding has_integral proof(rule,rule) fix e::real assume e:"e>0" hence *:"e/2>0" by auto + guess d1 using has_integralD[OF goal1(1) *] . note d1=this + guess d2 using has_integralD[OF goal1(2) *] . note d2=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 + g x)) - (k + l)) < e)" + apply(rule_tac x="\x. (d1 x) \ (d2 x)" in exI) apply(rule,rule gauge_inter[OF d1(1) d2(1)]) + proof(rule,rule,erule conjE) fix p assume as:"p tagged_division_of {a..b}" "(\x. d1 x \ d2 x) fine p" + have *:"(\(x, k)\p. content k *\<^sub>R (f x + g x)) = (\(x, k)\p. content k *\<^sub>R f x) + (\(x, k)\p. content k *\<^sub>R g x)" + unfolding scaleR_right_distrib setsum_addf[of "\(x,k). content k *\<^sub>R f x" "\(x,k). content k *\<^sub>R g x" p,THEN sym] + by(rule setsum_cong2,auto) + have "norm ((\(x, k)\p. content k *\<^sub>R (f x + g x)) - (k + l)) = norm (((\(x, k)\p. content k *\<^sub>R f x) - k) + ((\(x, k)\p. content k *\<^sub>R g x) - l))" + unfolding * by(auto simp add:group_simps) also let ?res = "\" + from as have *:"d1 fine p" "d2 fine p" unfolding fine_inter by auto + have "?res < e/2 + e/2" apply(rule le_less_trans[OF norm_triangle_ineq]) + apply(rule add_strict_mono) using d1(2)[OF as(1) *(1)] and d2(2)[OF as(1) *(2)] by auto + finally show "norm ((\(x, k)\p. content k *\<^sub>R (f x + g x)) - (k + l)) < e" by auto + qed qed qed { presume "\ (\a b. s = {a..b}) \ ?thesis" + thus ?thesis apply-apply(cases "\a b. s = {a..b}") using assms by(auto simp add:has_integral intro!:lem) } + assume as:"\ (\a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P + proof(rule,rule) case goal1 hence *:"e/2 > 0" by auto + from has_integral_altD[OF assms(1) as *] guess B1 . note B1=this[rule_format] + from has_integral_altD[OF assms(2) as *] guess B2 . note B2=this[rule_format] + show ?case apply(rule_tac x="max B1 B2" in exI) apply(rule,rule min_max.less_supI1,rule B1) + proof(rule,rule,rule) fix a b assume "ball 0 (max B1 B2) \ {a..b::real^'n}" + hence *:"ball 0 B1 \ {a..b::real^'n}" "ball 0 B2 \ {a..b::real^'n}" by auto + guess w using B1(2)[OF *(1)] .. note w=conjunctD2[OF this] + guess z using B2(2)[OF *(2)] .. note z=conjunctD2[OF this] + have *:"\x. (if x \ s then f x + g x else 0) = (if x \ s then f x else 0) + (if x \ s then g x else 0)" by auto + show "\z. ((\x. if x \ s then f x + g x else 0) has_integral z) {a..b} \ norm (z - (k + l)) < e" + apply(rule_tac x="w + z" in exI) apply(rule,rule lem[OF w(1) z(1), unfolded *[THEN sym]]) + using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2) by(auto simp add:field_simps) + qed qed qed + +lemma has_integral_sub: + shows "(f has_integral k) s \ (g has_integral l) s \ ((\x. f(x) - g(x)) has_integral (k - l)) s" + using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding group_simps by auto + +lemma integral_0: "integral s (\x::real^'n. 0::real^'m) = 0" + by(rule integral_unique has_integral_0)+ + +lemma integral_add: + shows "f integrable_on s \ g integrable_on s \ + integral s (\x. f x + g x) = integral s f + integral s g" + apply(rule integral_unique) apply(drule integrable_integral)+ + apply(rule has_integral_add) by assumption+ + +lemma integral_cmul: + shows "f integrable_on s \ integral s (\x. c *\<^sub>R f x) = c *\<^sub>R integral s f" + apply(rule integral_unique) apply(drule integrable_integral)+ + apply(rule has_integral_cmul) by assumption+ + +lemma integral_neg: + shows "f integrable_on s \ integral s (\x. - f x) = - integral s f" + apply(rule integral_unique) apply(drule integrable_integral)+ + apply(rule has_integral_neg) by assumption+ + +lemma integral_sub: + shows "f integrable_on s \ g integrable_on s \ integral s (\x. f x - g x) = integral s f - integral s g" + apply(rule integral_unique) apply(drule integrable_integral)+ + apply(rule has_integral_sub) by assumption+ + +lemma integrable_0: "(\x. 0) integrable_on s" + unfolding integrable_on_def using has_integral_0 by auto + +lemma integrable_add: + shows "f integrable_on s \ g integrable_on s \ (\x. f x + g x) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_add) + +lemma integrable_cmul: + shows "f integrable_on s \ (\x. c *\<^sub>R f(x)) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_cmul) + +lemma integrable_neg: + shows "f integrable_on s \ (\x. -f(x)) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_neg) + +lemma integrable_sub: + shows "f integrable_on s \ g integrable_on s \ (\x. f x - g x) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_sub) + +lemma integrable_linear: + shows "f integrable_on s \ bounded_linear h \ (h o f) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_linear) + +lemma integral_linear: + shows "f integrable_on s \ bounded_linear h \ integral s (h o f) = h(integral s f)" + apply(rule has_integral_unique) defer unfolding has_integral_integral + apply(drule has_integral_linear,assumption,assumption) unfolding has_integral_integral[THEN sym] + apply(rule integrable_linear) by assumption+ + +lemma has_integral_setsum: + assumes "finite t" "\a\t. ((f a) has_integral (i a)) s" + shows "((\x. setsum (\a. f a x) t) has_integral (setsum i t)) s" +proof(insert assms(1) subset_refl[of t],induct rule:finite_subset_induct) + case (insert x F) show ?case unfolding setsum_insert[OF insert(1,3)] + apply(rule has_integral_add) using insert assms by auto +qed auto + +lemma integral_setsum: + shows "finite t \ \a\t. (f a) integrable_on s \ + integral s (\x. setsum (\a. f a x) t) = setsum (\a. integral s (f a)) t" + apply(rule integral_unique) apply(rule has_integral_setsum) + using integrable_integral by auto + +lemma integrable_setsum: + shows "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] by auto + +lemma has_integral_eq: + assumes "\x\s. f x = g x" "(f has_integral k) s" shows "(g has_integral k) s" + using has_integral_sub[OF assms(2), of "\x. f x - g x" 0] + using has_integral_is_0[of s "\x. f x - g x"] using assms(1) by auto + +lemma integrable_eq: + shows "\x\s. f x = g x \ f integrable_on s \ g integrable_on s" + unfolding integrable_on_def using has_integral_eq[of s f g] by auto + +lemma has_integral_eq_eq: + shows "\x\s. f x = g x \ ((f has_integral k) s \ (g has_integral k) s)" + using has_integral_eq[of s f g] has_integral_eq[of s g f] by auto + +lemma has_integral_null[dest]: + assumes "content({a..b}) = 0" shows "(f has_integral 0) ({a..b})" + unfolding has_integral apply(rule,rule,rule_tac x="\x. ball x 1" in exI,rule) defer +proof(rule,rule,erule conjE) fix e::real assume e:"e>0" thus "gauge (\x. ball x 1)" by auto + fix p assume p:"p tagged_division_of {a..b}" (*"(\x. ball x 1) fine p"*) + have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) = 0" unfolding norm_eq_zero diff_0_right + using setsum_content_null[OF assms(1) p, of f] . + thus "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e" using e by auto qed + +lemma has_integral_null_eq[simp]: + shows "content({a..b}) = 0 \ ((f has_integral i) ({a..b}) \ i = 0)" + apply rule apply(rule has_integral_unique,assumption) + apply(drule has_integral_null,assumption) + apply(drule has_integral_null) by auto + +lemma integral_null[dest]: shows "content({a..b}) = 0 \ integral({a..b}) f = 0" + by(rule integral_unique,drule has_integral_null) + +lemma integrable_on_null[dest]: shows "content({a..b}) = 0 \ f integrable_on {a..b}" + unfolding integrable_on_def apply(drule has_integral_null) by auto + +lemma has_integral_empty[intro]: shows "(f has_integral 0) {}" + unfolding empty_as_interval apply(rule has_integral_null) + using content_empty unfolding empty_as_interval . + +lemma has_integral_empty_eq[simp]: shows "(f has_integral i) {} \ i = 0" + apply(rule,rule has_integral_unique,assumption) by auto + +lemma integrable_on_empty[intro]: shows "f integrable_on {}" unfolding integrable_on_def by auto + +lemma integral_empty[simp]: shows "integral {} f = 0" + apply(rule integral_unique) using has_integral_empty . + +lemma has_integral_refl[intro]: shows "(f has_integral 0) {a..a}" + apply(rule has_integral_null) unfolding content_eq_0_interior + unfolding interior_closed_interval using interval_sing by auto + +lemma integrable_on_refl[intro]: shows "f integrable_on {a..a}" unfolding integrable_on_def by auto + +lemma integral_refl: shows "integral {a..a} f = 0" apply(rule integral_unique) by auto + +subsection {* Cauchy-type criterion for integrability. *} + +lemma integrable_cauchy: fixes f::"real^'n \ '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 + 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 + 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 "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 vector_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 + 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 + 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 "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 + 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 + then guess y unfolding convergent_eq_cauchy[THEN sym] .. note y=this[unfolded Lim_sequentially,rule_format] + 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 + 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 + using N2[rule_format,unfolded vector_dist_norm,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 + +subsection {* Additivity of integral on abutting intervals. *} + +lemma interval_split: + "{a..b::real^'n} \ {x. x$k \ c} = {a .. (\ i. if i = k then min (b$k) c else b$i)}" + "{a..b} \ {x. x$k \ c} = {(\ i. if i = k then max (a$k) c else a$i) .. b}" + apply(rule_tac[!] set_ext) unfolding Int_iff mem_interval mem_Collect_eq + unfolding Cart_lambda_beta by auto + +lemma content_split: + "content {a..b::real^'n} = content({a..b} \ {x. x$k \ c}) + content({a..b} \ {x. x$k >= c})" +proof- note simps = interval_split content_closed_interval_cases Cart_lambda_beta vector_le_def + { presume "a\b \ ?thesis" thus ?thesis apply(cases "a\b") unfolding simps by auto } + have *:"UNIV = insert k (UNIV - {k})" "\x. finite (UNIV-{x::'n})" "\x. x\UNIV-{x}" by auto + have *:"\X Y Z. (\i\UNIV. Z i (if i = k then X else Y i)) = Z k X * (\i\UNIV-{k}. Z i (Y i))" + "(\i\UNIV. b$i - a$i) = (\i\UNIV-{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 "\ a $ k \ c \ \ c \ b $ k \ False" + unfolding not_le using as[unfolded vector_le_def,rule_format,of k] by auto + ultimately show ?thesis + unfolding simps unfolding *(1)[of "\i x. b$i - x"] *(1)[of "\i x. x - a$i"] *(2) by(auto) +qed + +lemma division_split_left_inj: + assumes "d division_of i" "k1 \ d" "k2 \ d" "k1 \ k2" + "k1 \ {x::real^'n. x$k \ c} = k2 \ {x. x$k \ c}" + shows "content(k1 \ {x. x$k \ c}) = 0" +proof- note d=division_ofD[OF assms(1)] + have *:"\a b::real^'n. \ c k. (content({a..b} \ {x. x$k \ c}) = 0 \ interior({a..b} \ {x. x$k \ c}) = {})" + unfolding interval_split 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 + +lemma division_split_right_inj: + assumes "d division_of i" "k1 \ d" "k2 \ d" "k1 \ k2" + "k1 \ {x::real^'n. x$k \ c} = k2 \ {x. x$k \ c}" + shows "content(k1 \ {x. x$k \ c}) = 0" +proof- note d=division_ofD[OF assms(1)] + have *:"\a b::real^'n. \ c k. (content({a..b} \ {x. x$k >= c}) = 0 \ interior({a..b} \ {x. x$k >= c}) = {})" + unfolding interval_split 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 + +lemma tagged_division_split_left_inj: + assumes "d tagged_division_of i" "(x1,k1) \ d" "(x2,k2) \ d" "k1 \ k2" "k1 \ {x. x$k \ c} = k2 \ {x. x$k \ c}" + shows "content(k1 \ {x. x$k \ c}) = 0" +proof- have *:"\a b c. (a,b) \ c \ b \ snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto + show ?thesis apply(rule division_split_left_inj[OF division_of_tagged_division[OF assms(1)]]) + apply(rule_tac[1-2] *) using assms(2-) by auto qed + +lemma tagged_division_split_right_inj: + assumes "d tagged_division_of i" "(x1,k1) \ d" "(x2,k2) \ d" "k1 \ k2" "k1 \ {x. x$k \ c} = k2 \ {x. x$k \ c}" + shows "content(k1 \ {x. x$k \ c}) = 0" +proof- have *:"\a b c. (a,b) \ c \ b \ snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto + show ?thesis apply(rule division_split_right_inj[OF division_of_tagged_division[OF assms(1)]]) + apply(rule_tac[1-2] *) using assms(2-) by auto qed + +lemma division_split: + assumes "p division_of {a..b::real^'n}" + shows "{l \ {x. x$k \ c} | l. l \ p \ ~(l \ {x. x$k \ c} = {})} division_of ({a..b} \ {x. x$k \ c})" (is "?p1 division_of ?I1") and + "{l \ {x. x$k \ c} | l. l \ p \ ~(l \ {x. x$k \ c} = {})} division_of ({a..b} \ {x. x$k \ c})" (is "?p2 division_of ?I2") +proof(rule_tac[!] division_ofI) note p=division_ofD[OF assms] + show "finite ?p1" "finite ?p2" using p(1) by auto show "\?p1 = ?I1" "\?p2 = ?I2" unfolding p(6)[THEN sym] by auto + { fix k assume "k\?p1" then guess l unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l=this + guess u v using p(4)[OF l(2)] apply-by(erule exE)+ note uv=this + show "k\?I1" "k \ {}" "\a b. k = {a..b}" unfolding l + using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split) by auto + fix k' assume "k'\?p1" then guess l' unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l'=this + assume "k\k'" thus "interior k \ interior k' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto } + { fix k assume "k\?p2" then guess l unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l=this + guess u v using p(4)[OF l(2)] apply-by(erule exE)+ note uv=this + show "k\?I2" "k \ {}" "\a b. k = {a..b}" unfolding l + using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split) by auto + fix k' assume "k'\?p2" then guess l' unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l'=this + assume "k\k'" thus "interior k \ interior k' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto } +qed + +lemma has_integral_split: fixes f::"real^'n \ 'a::real_normed_vector" + assumes "(f has_integral i) ({a..b} \ {x. x$k \ c})" "(f has_integral j) ({a..b} \ {x. x$k \ c})" + shows "(f has_integral (i + j)) ({a..b})" +proof(unfold has_integral,rule,rule) case goal1 hence e:"e/2>0" by auto + guess d1 using has_integralD[OF assms(1)[unfolded interval_split] e] . note d1=this[unfolded interval_split[THEN sym]] + guess d2 using has_integralD[OF assms(2)[unfolded interval_split] e] . note d2=this[unfolded interval_split[THEN sym]] + 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,rule) defer apply(rule,rule,(erule conjE)+) + proof- 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)] + 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" + proof- fix x kk assume as:"(x,kk)\p" + show "~(kk \ {x. x$k \ c} = {}) \ x$k \ c" + proof(rule ccontr) case goal1 + from this(2)[unfolded not_le] have "kk \ ball x \x $ k - c\" + using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto + hence "\y. y \ ball x \x $ k - c\ \ {x. x $ k \ c}" using goal1(1) by blast + then guess y .. hence "\x $ k - y $ k\ < \x $ k - c\" "y$k \ c" apply-apply(rule le_less_trans) + using component_le_norm[of "x - y" k,unfolded vector_minus_component] by(auto simp add:vector_dist_norm) + thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps) + qed + show "~(kk \ {x. x$k \ c} = {}) \ x$k \ c" + proof(rule ccontr) case goal1 + from this(2)[unfolded not_le] have "kk \ ball x \x $ k - c\" + using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto + hence "\y. y \ ball x \x $ k - c\ \ {x. x $ k \ c}" using goal1(1) by blast + then guess y .. hence "\x $ k - y $ k\ < \x $ k - c\" "y$k \ c" apply-apply(rule le_less_trans) + using component_le_norm[of "x - y" k,unfolded vector_minus_component] by(auto simp add:vector_dist_norm) + thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps) + qed + qed + + have lem1: "\f P Q. (\x k. (x,k) \ {(x,f k) | x k. P x k} \ Q x k) \ (\x k. P x k \ Q x (f k))" by auto + have lem2: "\f s P f. finite s \ finite {(x,f k) | x k. (x,k) \ s \ P x k}" + proof- case goal1 thus ?case apply-apply(rule finite_subset[of _ "(\(x,k). (x,f k)) ` s"]) by auto qed + have lem3: "\g::(real ^ 'n \ bool) \ real ^ 'n \ bool. finite p \ + setsum (\(x,k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \ p \ ~(g k = {})} + = setsum (\(x,k). content k *\<^sub>R f x) ((\(x,k). (x,g k)) ` p)" + apply(rule setsum_mono_zero_left) prefer 3 + proof fix g::"(real ^ 'n \ bool) \ real ^ 'n \ bool" and i::"(real^'n) \ ((real^'n) set)" + assume "i \ (\(x, k). (x, g k)) ` p - {(x, g k) |x k. (x, k) \ p \ g k \ {}}" + then obtain x k where xk:"i=(x,g k)" "(x,k)\p" "(x,g k) \ {(x, g k) |x k. (x, k) \ p \ g k \ {}}" by auto + have "content (g k) = 0" using xk using content_empty by auto + thus "(\(x, k). content k *\<^sub>R f x) i = 0" unfolding xk split_conv by auto + qed auto + have lem4:"\g. (\(x,l). content (g l) *\<^sub>R f x) = (\(x,l). content l *\<^sub>R f x) o (\(x,l). (x,g l))" apply(rule ext) by auto + + let ?M1 = "{(x,kk \ {x. x$k \ c}) |x kk. (x,kk) \ p \ kk \ {x. x$k \ c} \ {}}" + have "norm ((\(x, k)\?M1. content k *\<^sub>R f x) - i) < e/2" apply(rule d1(2),rule tagged_division_ofI) + apply(rule lem2 p(3))+ prefer 6 apply(rule fineI) + proof- show "\{k. \x. (x, k) \ ?M1} = {a..b} \ {x. x$k \ c}" unfolding p(8)[THEN sym] by auto + fix x l assume xl:"(x,l)\?M1" + then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) . note xl'=this + have "l' \ d1 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto + thus "l \ d1 x" unfolding xl' by auto + show "x\l" "l \ {a..b} \ {x. x $ k \ c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4) + using lem0(1)[OF xl'(3-4)] by auto + show "\a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[where c=c and k=k]) + fix y r let ?goal = "interior l \ interior r = {}" assume yr:"(y,r)\?M1" + then guess y' r' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) . note yr'=this + assume as:"(x,l) \ (y,r)" show "interior l \ interior r = {}" + proof(cases "l' = r' \ x' = y'") + case False thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + next case True hence "l' \ r'" using as unfolding xl' yr' by auto + thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + qed qed moreover + + let ?M2 = "{(x,kk \ {x. x$k \ c}) |x kk. (x,kk) \ p \ kk \ {x. x$k \ c} \ {}}" + have "norm ((\(x, k)\?M2. content k *\<^sub>R f x) - j) < e/2" apply(rule d2(2),rule tagged_division_ofI) + apply(rule lem2 p(3))+ prefer 6 apply(rule fineI) + proof- show "\{k. \x. (x, k) \ ?M2} = {a..b} \ {x. x$k \ c}" unfolding p(8)[THEN sym] by auto + fix x l assume xl:"(x,l)\?M2" + then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) . note xl'=this + have "l' \ d2 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto + thus "l \ d2 x" unfolding xl' by auto + show "x\l" "l \ {a..b} \ {x. x $ k \ c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4) + using lem0(2)[OF xl'(3-4)] by auto + show "\a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[where c=c and k=k]) + fix y r let ?goal = "interior l \ interior r = {}" assume yr:"(y,r)\?M2" + then guess y' r' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) . note yr'=this + assume as:"(x,l) \ (y,r)" show "interior l \ interior r = {}" + proof(cases "l' = r' \ x' = y'") + case False thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + next case True hence "l' \ r'" using as unfolding xl' yr' by auto + thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + qed qed ultimately + + have "norm (((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j)) < e/2 + e/2" + apply- apply(rule norm_triangle_lt) by auto + also { have *:"\x y. x = (0::real) \ x *\<^sub>R (y::'a) = 0" using scaleR_zero_left by auto + have "((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j) + = (\(x, k)\?M1. content k *\<^sub>R f x) + (\(x, k)\?M2. content k *\<^sub>R f x) - (i + j)" by auto + also have "\ = (\(x, ka)\p. content (ka \ {x. x $ k \ c}) *\<^sub>R f x) + (\(x, ka)\p. content (ka \ {x. c \ x $ k}) *\<^sub>R f x) - (i + j)" + unfolding lem3[OF p(3)] apply(subst setsum_reindex_nonzero[OF p(3)]) defer apply(subst setsum_reindex_nonzero[OF p(3)]) + defer unfolding lem4[THEN sym] apply(rule refl) unfolding split_paired_all split_conv apply(rule_tac[!] *) + proof- case goal1 thus ?case apply- apply(rule tagged_division_split_left_inj [OF p(1), of a b aa ba]) by auto + next case goal2 thus ?case apply- apply(rule tagged_division_split_right_inj[OF p(1), of a b aa ba]) by auto + qed also note setsum_addf[THEN sym] + 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 *\<^sub>R f x) x" unfolding split_paired_all split_conv + proof- fix a b assume "(a,b) \ p" from p(6)[OF this] guess u v apply-by(erule exE)+ note uv=this + thus "content (b \ {x. x $ k \ c}) *\<^sub>R f a + content (b \ {x. c \ x $ k}) *\<^sub>R f a = content b *\<^sub>R f a" + unfolding scaleR_left_distrib[THEN sym] unfolding uv content_split[of u v k c] by auto + qed note setsum_cong2[OF this] + finally have "(\(x, k)\{(x, kk \ {x. x $ k \ c}) |x kk. (x, kk) \ p \ kk \ {x. x $ k \ c} \ {}}. content k *\<^sub>R f x) - i + + ((\(x, k)\{(x, kk \ {x. c \ x $ k}) |x kk. (x, kk) \ p \ kk \ {x. c \ x $ k} \ {}}. content k *\<^sub>R f x) - j) = + (\(x, ka)\p. content ka *\<^sub>R f x) - (i + j)" by auto } + finally show "norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed + +subsection {* A sort of converse, integrability on subintervals. *} + +lemma tagged_division_union_interval: + assumes "p1 tagged_division_of ({a..b} \ {x::real^'n. x$k \ (c::real)})" "p2 tagged_division_of ({a..b} \ {x. x$k \ c})" + 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]) + unfolding interval_split interior_closed_interval + by(auto simp add: vector_less_def Cart_lambda_beta elim!:allE[where x=k]) qed + +lemma has_integral_separate_sides: fixes f::"real^'m \ 'a::real_normed_vector" + assumes "(f has_integral i) ({a..b})" "e>0" + 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] . 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 + 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 fastsimp + moreover have "interior {x. x $ k = c} = {}" + proof(rule ccontr) case goal1 then obtain x where x:"x\interior {x. x$k = c}" by auto + then guess e unfolding mem_interior .. note e=this + have x:"x$k = c" using x interior_subset by fastsimp + have *:"\i. \(x - (x + (\ i. if i = k then e / 2 else 0))) $ i\ = (if i = k then e/2 else 0)" using e by auto + have "x + (\ i. if i = k then e/2 else 0) \ ball x e" unfolding mem_ball vector_dist_norm + apply(rule le_less_trans[OF norm_le_l1]) unfolding * + unfolding setsum_delta[OF finite_UNIV] using e by auto + hence "x + (\ i. if i = k then e/2 else 0) \ {x. x$k = c}" using e by auto + thus False unfolding mem_Collect_eq using e x by auto + qed ultimately have "content b = 0" unfolding uv content_eq_0_interior apply-apply(drule subset_interior) by auto + thus "content b *\<^sub>R f a = 0" by auto + qed auto + also have "\ < e" by(rule 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::"real^'n \ 'a::{real_normed_vector,complete_space}" assumes "f integrable_on {a..b}" + 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 unfolding integrable_on_def .. note y=this + def b' \ "(\ i. if i = k then min (b$k) c else b$i)::real^'n" + and a' \ "(\ i. if i = k then max (a$k) c else a$i)::real^'n" + show ?t1 ?t2 unfolding interval_split integrable_cauchy unfolding interval_split[THEN sym] + proof(rule_tac[!] allI impI)+ fix e::real assume "e>0" hence "e/2>0" by auto + from has_integral_separate_sides[OF y this,of k 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 \ + 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 "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]] + using as unfolding interval_split b'_def[symmetric] a'_def[symmetric] + using p using assms by(auto simp add:group_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" + 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 b'_def[symmetric] a'_def[symmetric] + using p using assms by(auto simp add:group_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) \ ((real^'n) set \ 'a) \ bool" where + "operative opp f \ + (\a b. content {a..b} = 0 \ f {a..b} = neutral(opp)) \ + (\a b c k. f({a..b}) = + opp (f({a..b} \ {x. x$k \ c})) + (f({a..b} \ {x. x$k \ c})))" + +lemma operativeD[dest]: assumes "operative opp f" + shows "\a b. content {a..b} = 0 \ f {a..b} = neutral(opp)" + "\a b c k. 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" + unfolding operative_def by auto + +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 + +subsection {* Using additivity of lifted function to encode definedness. *} + +lemma forall_option: "(\x. P x) \ P None \ (\x. P(Some x))" + by (metis map_of.simps option.nchotomy) + +lemma exists_option: + "(\x. P x) \ P None \ (\x. P(Some x))" + by (metis map_of.simps 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" + apply(induct v) by 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)" + +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" + unfolding monoidal_def using assms by fastsimp + +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)" + using assms unfolding monoidal_def apply- by metis+ + +lemma monoidal_simps[simp]: assumes "monoidal opp" + shows "opp (neutral opp) a = a" "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 +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) 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 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 fun_left_comm_monoidal[intro]: assumes "monoidal opp" shows "fun_left_comm opp" + unfolding fun_left_comm_def using monoidal_ac[OF assms] by auto + +lemma support_clauses: + "\f g s. support opp f {} = {}" + "\f g s. support opp f (insert x s) = (if f(x) = neutral opp then support opp f s else insert x (support opp f s))" + "\f g s. support opp f (s - {x}) = (support opp f s) - {x}" + "\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 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" + 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 + show ?thesis unfolding iterate_def if_P[OF True] * by auto +next case False note x=this + note * = fun_left_comm.fun_left_comm_apply[OF fun_left_comm_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 fun_left_comm.fold_insert[OF * finite_support]) + using `finite s` unfolding support_def using False x by auto 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 + +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 operative_content[intro]: "operative (op +) content" + unfolding operative_def content_split[THEN sym] neutral_add by auto + +lemma neutral_monoid[simp]: "neutral ((op +)::('a::comm_monoid_add) \ 'a \ 'a) = 0" + unfolding neutral_def apply(rule some_equality) defer + apply(erule_tac x=0 in allE) by auto + +lemma monoidal_monoid[intro]: + shows "monoidal ((op +)::('a::comm_monoid_add) \ 'a \ 'a)" + unfolding monoidal_def neutral_monoid by(auto simp add: group_simps) + +lemma operative_integral: fixes f::"real^'n \ 'a::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)+ +proof- fix a b c k 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] + unfolding if_P[OF integrable_split(1)[OF True]] if_P[OF integrable_split(2)[OF True]] + unfolding lifted.simps option.inject apply(rule integral_unique) apply(rule has_integral_split) + apply(rule_tac[!] integrable_integral integrable_split)+ using True by assumption+ + 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) 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::real^'n} = 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 + +subsection {* Points of division of a partition. *} + +definition "division_points (k::(real^'n) set) d = + {(j,x). (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: 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 ` UNIV)" + unfolding division_points_def by auto + show ?thesis unfolding * using assm by auto qed + +lemma division_points_subset: + assumes "d division_of {a..b}" "\i. a$i < b$i" "a$k < c" "c < b$k" + 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. a$i \ b$i" "\i. a$i \ (\ i. if i = k then min (b $ k) c else b $ i) $ i" + "\i. (\ i. if i = k then max (a $ k) c else a $ i) $ i \ b$i" "min (b $ k) c = c" "max (a $ k) c = c" + using assms using less_imp_le by auto + show ?t1 unfolding division_points_def interval_split[of a b] + unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] Cart_lambda_beta unfolding * + unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ unfolding mem_Collect_eq 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)" + "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} \ {}" + from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this + have *:"\i. u $ i \ (\ i. if i = k then min (v $ k) c else v $ i) $ i" using as(6) unfolding l interval_split interval_ne_empty as . + have **:"\i. u$i \ v$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto + show "a $ fst x < snd x \ snd x < b $ fst x \ (\i\d. interval_lowerbound i $ fst x = snd x \ interval_upperbound i $ fst x = snd x)" + using as(1-3,5) unfolding l interval_split interval_ne_empty as interval_bounds[OF *] Cart_lambda_beta apply- + apply(rule,assumption,rule) defer apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **] + apply(case_tac[!] "fst x = k") using assms by auto + qed + show ?t2 unfolding division_points_def interval_split[of a b] + unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] Cart_lambda_beta unfolding * + unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ unfolding mem_Collect_eq 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" "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} \ {}" + from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this + have *:"\i. (\ i. if i = k then max (u $ k) c else u $ i) $ i \ v $ i" using as(6) unfolding l interval_split interval_ne_empty as . + have **:"\i. u$i \ v$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto + show "a $ fst x < snd x \ snd x < b $ fst x \ (\i\d. interval_lowerbound i $ fst x = snd x \ interval_upperbound i $ fst x = snd x)" + using as(1-3,5) unfolding l interval_split interval_ne_empty as interval_bounds[OF *] Cart_lambda_beta apply- + apply rule defer apply(rule,assumption) apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **] + apply(case_tac[!] "fst x = k") using assms by auto qed qed + +lemma division_points_psubset: + assumes "d division_of {a..b}" "\i. a$i < b$i" "a$k < c" "c < b$k" + "l \ d" "interval_lowerbound l$k = c \ interval_upperbound l$k = c" + 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. 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. u$i \ v$i" "\i. 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 apply(subst interval_bounds) prefer 3 apply(subst interval_bounds) + unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab 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] + apply (auto simp add:interval_bounds) unfolding * by auto + thus "?D1 \ ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) 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 apply(subst interval_bounds) prefer 3 apply(subst interval_bounds) + unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab 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] + apply (auto simp add:interval_bounds) unfolding * by auto + thus "?D2 \ ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) by auto qed + +subsection {* Preservation by divisions and tagged divisions. *} + +lemma support_support[simp]:"support opp f (support opp f s) = support opp f s" + unfolding support_def by auto + +lemma iterate_support[simp]: "iterate opp (support opp f s) f = iterate opp s f" + unfolding iterate_def support_support by auto + +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" + 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)[THEN sym]) + 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 + show ?thesis + apply(cases "finite (support opp g (f ` s))") + apply(subst (1) iterate_support[THEN sym],subst (2) iterate_support[THEN sym]) + 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. *) +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" + 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" + unfolding support_def using assms(3) by auto + show ?thesis unfolding * + apply(subst iterate_support[THEN sym]) unfolding support_clauses + apply(subst iterate_image[OF assms(1)]) defer + apply(subst(2) iterate_support[THEN sym]) apply(subst **) + unfolding inj_on_def using assms(3,4) unfolding support_def by auto 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[THEN sym]) + 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" + shows "iterate opp s f = iterate opp s g" +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" + case True note support_subset[of opp f s] + thus ?thesis apply- apply(subst iterate_support[THEN sym],subst(2) iterate_support[THEN sym]) 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::"(real^'n) set \ 'a" + assumes "monoidal opp" "operative opp f" "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) + 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 } + assume "content {a..b} \ 0" note ab = this[unfolded content_lt_nz[THEN sym] content_pos_lt_eq] + hence ab':"\i. 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. 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) apply(rule) + proof- fix u v j assume as:"{u..v} \ d" note division_ofD(3)[OF goal1(4) this] + hence uv:"\i. u$i \ v$i" "u$j \ v$j" unfolding interval_ne_empty by auto + have *:"\p r Q. p \ r \ (\x\d. Q x) \ p \ r \ (Q {u..v})" using as 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 by auto + 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) by auto + qed have "(1/2) *\<^sub>R (a+b) \ {a..b}" unfolding mem_interval using ab by(auto intro!:less_imp_le) + note this[unfolded division_ofD(6)[OF goal1(4),THEN sym] 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 + 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 Cart_eq + proof(rule_tac[!] allI) fix j 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] by auto + 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" unfolding Cart_eq 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 exI) 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 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} \ {}}" + def cb \ "(\ i. if i = k then c else b$i)" and ca \ "(\ i. if i = k then c else a$i)" + 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 apply(rule_tac[!] goal1(1)[rule_format]) + using division_split[OF goal1(4), where k=k and c=c] + unfolding interval_split d1_def[symmetric] d2_def[symmetric] unfolding goal1(2) Suc_le_mono + using goal1(2-3) using division_points_finite[OF goal1(4)] by auto + have "f {a..b} = opp (iterate opp d1 f) (iterate opp d2 f)" (is "_ = ?prev") + unfolding * apply(rule operativeD(2)) using goal1(3) . + 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[THEN sym] 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 + apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym] apply(rule division_split_left_inj) + apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule 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[THEN sym] 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 + apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym] apply(rule division_split_right_inj) + apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as)+ + 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) . + 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[THEN sym]) using goal1 by auto + 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[THEN sym,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)[THEN sym] uv by auto + qed also have "\ = f {a..b}" + using operative_division[OF assms(1-2) division_of_tagged_division[OF assms(3)]] . + finally show ?thesis . qed + +subsection {* Additivity of content. *} + +lemma setsum_iterate: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 + thus ?thesis unfolding * setsum_def iterate_def fold_image_def fold'_def + unfolding neutral_monoid . qed + +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,THEN sym] + apply(subst setsum_iterate) using assms by auto + +lemma additive_content_tagged_division: + assumes "d tagged_division_of {a..b}" + shows "setsum (\(x,l). content l) d = content({a..b})" + unfolding operative_tagged_division[OF monoidal_monoid operative_content assms,THEN sym] + apply(subst setsum_iterate) using assms by auto + +subsection {* Finally, the integral of a constant\ *} + +lemma has_integral_const[intro]: + "((\x. c) has_integral (content({a..b::real^'n}) *\<^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[THEN sym, unfolded o_def]) + defer apply(subst additive_content_tagged_division[unfolded split_def]) apply assumption by auto + +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 setsum_norm) defer unfolding norm_scaleR setsum_left_distrib[THEN sym] + apply(rule order_trans[OF mult_left_mono],rule assms,rule setsum_abs_ge_zero) + apply(subst real_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 setsum_norm) defer 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[THEN sym] apply(subst real_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),THEN sym] split_def by auto + 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(insert assms,auto) 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[THEN sym] apply(rule setsum_cong2) unfolding scaleR.diff_right by auto + +lemma has_integral_bound: fixes f::"real^'n \ 'a::real_normed_vector" + assumes "0 \ B" "(f has_integral i) ({a..b})" "\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[THEN sym]) 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] + 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 + +subsection {* Similar theorems about relationship among components. *} + +lemma rsum_component_le: fixes f::"real^'n \ real^'m" + assumes "p tagged_division_of {a..b}" "\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 setsum_component apply(rule setsum_mono) + apply(rule mp) defer apply assumption apply(induct_tac x,rule) unfolding split_conv +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 Cart_nth.scaleR real_scaleR_def apply(rule mult_left_mono) + defer apply(rule content_pos_le,rule assms(2)[rule_format]) using assm by auto qed + +lemma has_integral_component_le: fixes f::"real^'n \ real^'m" + assumes "(f has_integral i) s" "(g has_integral j) s" "\x\s. (f x)$k \ (g x)$k" + shows "i$k \ j$k" +proof- have lem:"\a b g i j. \f::real^'n \ real^'m. (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 hence *:"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] + 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 component_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 unfolding Cart_nth.diff using rsum_component_le[OF p(1) goal1(3)] by smt + 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 + note has_integral_altD[OF _ as this] from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format] + have "bounded (ball 0 B1 \ ball (0::real^'n) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+ + from bounded_subset_closed_interval[OF this] guess a b apply- by(erule 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 smt(*SMTSMT*) + 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 Cart_nth.diff by(rule *) qed + +lemma integral_component_le: fixes f::"real^'n \ real^'m" + assumes "f integrable_on s" "g integrable_on s" "\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_dest_vec1_le: fixes f::"real^'n \ real^1" + assumes "(f has_integral i) s" "(g has_integral j) s" "\x\s. f x \ g x" + shows "dest_vec1 i \ dest_vec1 j" apply(rule has_integral_component_le[OF assms(1-2)]) + using assms(3) unfolding vector_le_def by auto + +lemma integral_dest_vec1_le: fixes f::"real^'n \ real^1" + assumes "f integrable_on s" "g integrable_on s" "\x\s. f x \ g x" + shows "dest_vec1(integral s f) \ dest_vec1(integral s g)" + apply(rule has_integral_dest_vec1_le) apply(rule_tac[1-2] integrable_integral) using assms by auto + +lemma has_integral_component_pos: fixes f::"real^'n \ real^'m" + assumes "(f has_integral i) s" "\x\s. 0 \ (f x)$k" shows "0 \ i$k" + using has_integral_component_le[OF has_integral_0 assms(1)] using assms(2) by auto + +lemma integral_component_pos: fixes f::"real^'n \ real^'m" + assumes "f integrable_on s" "\x\s. 0 \ (f x)$k" shows "0 \ (integral s f)$k" + apply(rule has_integral_component_pos) using assms by auto + +lemma has_integral_dest_vec1_pos: fixes f::"real^'n \ real^1" + assumes "(f has_integral i) s" "\x\s. 0 \ f x" shows "0 \ i" + using has_integral_component_pos[OF assms(1), of 1] + using assms(2) unfolding vector_le_def by auto + +lemma integral_dest_vec1_pos: fixes f::"real^'n \ real^1" + assumes "f integrable_on s" "\x\s. 0 \ f x" shows "0 \ integral s f" + apply(rule has_integral_dest_vec1_pos) using assms by auto + +lemma has_integral_component_neg: fixes f::"real^'n \ real^'m" + assumes "(f has_integral i) s" "\x\s. (f x)$k \ 0" shows "i$k \ 0" + using has_integral_component_le[OF assms(1) has_integral_0] assms(2) by auto + +lemma has_integral_dest_vec1_neg: fixes f::"real^'n \ real^1" + assumes "(f has_integral i) s" "\x\s. f x \ 0" shows "i \ 0" + using has_integral_component_neg[OF assms(1),of 1] using assms(2) by auto + +lemma has_integral_component_lbound: + assumes "(f has_integral i) {a..b}" "\x\{a..b}. B \ f(x)$k" shows "B * content {a..b} \ i$k" + using has_integral_component_le[OF has_integral_const assms(1),of "(\ i. B)" k] assms(2) + unfolding Cart_lambda_beta vector_scaleR_component by(auto simp add:field_simps) + +lemma has_integral_component_ubound: + assumes "(f has_integral i) {a..b}" "\x\{a..b}. f x$k \ B" + shows "i$k \ B * content({a..b})" + using has_integral_component_le[OF assms(1) has_integral_const, of k "vec B"] + unfolding vec_component Cart_nth.scaleR using assms(2) by(auto simp add:field_simps) + +lemma integral_component_lbound: + assumes "f integrable_on {a..b}" "\x\{a..b}. B \ f(x)$k" + 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: + assumes "f integrable_on {a..b}" "\x\{a..b}. f(x)$k \ B" + shows "(integral({a..b}) f)$k \ B * content({a..b})" + apply(rule has_integral_component_ubound) using assms unfolding has_integral_integral by auto + +subsection {* Uniform limit of integrable functions is integrable. *} + +lemma real_arch_invD: + "0 < (e::real) \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" + by(subst(asm) real_arch_inv) + +lemma integrable_uniform_limit: fixes f::"real^'n \ 'a::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 + 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] + 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)" + using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] + using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:group_simps) + also have "\ < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps) + finally show ?case . + qed + show ?case unfolding vector_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 real_divide_def 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:group_simps simp add:norm_minus_commute) + qed qed qed + from this[unfolded convergent_eq_cauchy[THEN sym]] 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 + from s[unfolded Lim_sequentially,rule_format,OF 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 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)" + using norm_triangle_ineq[of "sf - sg" "sg - s"] + using norm_triangle_ineq[of "sg - i" " i - s"] by(auto simp add:group_simps) + also have "\ < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_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,unfolded vector_dist_norm],auto) + qed qed qed qed + +subsection {* Negligible sets. *} + +definition "indicator s \ (\x. if x \ s then 1 else (0::real))" + +lemma dest_vec1_indicator: + "indicator s x = (if x \ s then 1 else 0)" unfolding indicator_def by auto + +lemma indicator_pos_le[intro]: "0 \ (indicator s x)" unfolding indicator_def by auto + +lemma indicator_le_1[intro]: "(indicator s x) \ 1" unfolding indicator_def by auto + +lemma dest_vec1_indicator_abs_le_1: "abs(indicator s x) \ 1" + unfolding indicator_def by auto + +definition "negligible (s::(real^'n) set) \ (\a b. ((indicator s) has_integral 0) {a..b})" + +lemma indicator_simps[simp]:"x\s \ indicator s x = 1" "x\s \ indicator s x = 0" + unfolding indicator_def by auto + +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" + 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: shows "{a..b} \ {x . abs(x$k - c) \ (e::real)} = + {(\ i. if i = k then max (a$k) (c - e) else a$i) .. (\ i. if i = k then min (b$k) (c + e) else b$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 by(rule refl) qed + +lemma division_doublesplit: assumes "p division_of {a..b::real^'n}" + 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})" +proof- have *:"\x c. abs(x - c) \ e \ x \ c - e \ x \ c + e" by auto + have **:"\p q p' q'. p division_of q \ p = p' \ q = q' \ p' division_of q'" by auto + note division_split(1)[OF assms, where c="c+e" and k=k,unfolded interval_split] + note division_split(2)[OF this, where c="c-e" and k=k] + thus ?thesis apply(rule **) unfolding interval_doublesplit unfolding * unfolding interval_split interval_doublesplit + apply(rule set_ext) unfolding mem_Collect_eq apply rule apply(erule conjE exE)+ apply(rule_tac x=la in exI) defer + apply(erule conjE exE)+ apply(rule_tac x="l \ {x. c + e \ x $ k}" in exI) apply rule defer apply rule + apply(rule_tac x=l in exI) by blast+ qed + +lemma content_doublesplit: assumes "0 < e" + obtains d where "0 < d" "content({a..b} \ {x. abs(x$k - c) \ d}) < e" +proof(cases "content {a..b} = 0") + case True show ?thesis apply(rule that[of 1]) defer unfolding interval_doublesplit + apply(rule le_less_trans[OF content_subset]) defer apply(subst True) + unfolding interval_doublesplit[THEN sym] using assms by auto +next case False def d \ "e / 3 / setprod (\i. b$i - a$i) (UNIV - {k})" + note False[unfolded content_eq_0 not_ex not_le, rule_format] + hence prod0:"0 < setprod (\i. b$i - a$i) (UNIV - {k})" apply-apply(rule setprod_pos) by smt + hence "d > 0" unfolding d_def using assms by(auto simp add:field_simps) thus ?thesis + proof(rule that[of d]) have *:"UNIV = insert k (UNIV - {k})" by auto + have **:"{a..b} \ {x. \x $ k - c\ \ d} \ {} \ + (\i\UNIV - {k}. interval_upperbound ({a..b} \ {x. \x $ k - c\ \ d}) $ i - interval_lowerbound ({a..b} \ {x. \x $ k - c\ \ d}) $ i) + = (\i\UNIV - {k}. b$i - a$i)" apply(rule setprod_cong,rule refl) + unfolding interval_doublesplit interval_eq_empty not_ex not_less unfolding interval_bounds by auto + show "content ({a..b} \ {x. \x $ k - c\ \ d}) < e" apply(cases) unfolding content_def apply(subst if_P,assumption,rule assms) + unfolding if_not_P apply(subst *) apply(subst setprod_insert) unfolding ** + unfolding interval_doublesplit interval_eq_empty not_ex not_less unfolding interval_bounds unfolding Cart_lambda_beta if_P[OF refl] + proof- have "(min (b $ k) (c + d) - max (a $ k) (c - d)) \ 2 * d" by auto + also have "... < e / (\i\UNIV - {k}. b $ i - a $ i)" unfolding d_def using assms prod0 by(auto simp add:field_simps) + finally show "(min (b $ k) (c + d) - max (a $ k) (c - d)) * (\i\UNIV - {k}. b $ i - a $ i) < e" + unfolding pos_less_divide_eq[OF prod0] . qed auto qed qed + +lemma negligible_standard_hyperplane[intro]: "negligible {x. x$k = (c::real)}" + unfolding negligible_def has_integral apply(rule,rule,rule,rule) +proof- case goal1 from content_doublesplit[OF this,of a b k c] guess d . note d=this let ?i = "indicator {x. x$k = c}" + show ?case apply(rule_tac x="\x. ball x d" in exI) apply(rule,rule gauge_ball,rule d) + proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \ (\x. ball x d) fine p" + have *:"(\(x, ka)\p. content ka *\<^sub>R ?i x) = (\(x, ka)\p. content (ka \ {x. abs(x$k - c) \ d}) *\<^sub>R ?i x)" + apply(rule setsum_cong2) unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv + apply(cases,rule disjI1,assumption,rule disjI2) + proof- fix x l assume as:"(x,l)\p" "?i x \ 0" hence xk:"x$k = c" unfolding indicator_def apply-by(rule ccontr,auto) + show "content l = content (l \ {x. \x $ k - c\ \ d})" apply(rule arg_cong[where f=content]) + apply(rule set_ext,rule,rule) unfolding mem_Collect_eq + proof- fix y assume y:"y\l" note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv] + note this[unfolded subset_eq mem_ball vector_dist_norm,rule_format,OF y] note le_less_trans[OF component_le_norm[of _ k] this] + thus "\y $ k - c\ \ d" unfolding Cart_nth.diff xk by auto + qed auto qed + note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]] + show "norm ((\(x, ka)\p. content ka *\<^sub>R ?i x) - 0) < e" unfolding diff_0_right * unfolding real_scaleR_def real_norm_def + apply(subst abs_of_nonneg) apply(rule setsum_nonneg,rule) unfolding split_paired_all split_conv + apply(rule mult_nonneg_nonneg) apply(drule p'(4)) apply(erule exE)+ apply(rule_tac b=b in back_subst) + prefer 2 apply(subst(asm) eq_commute) apply assumption + apply(subst interval_doublesplit) apply(rule content_pos_le) apply(rule indicator_pos_le) + proof- have "(\(x, ka)\p. content (ka \ {x. \x $ k - c\ \ d}) * ?i x) \ (\(x, ka)\p. content (ka \ {x. \x $ k - c\ \ d}))" + apply(rule setsum_mono) unfolding split_paired_all split_conv + apply(rule mult_right_le_one_le) apply(drule p'(4)) by(auto simp add:interval_doublesplit intro!:content_pos_le) + also have "... < e" apply(subst setsum_over_tagged_division_lemma[OF p[THEN conjunct1]]) + proof- case goal1 have "content ({u..v} \ {x. \x $ k - c\ \ d}) \ content {u..v}" + unfolding interval_doublesplit apply(rule content_subset) unfolding interval_doublesplit[THEN sym] by auto + thus ?case unfolding goal1 unfolding interval_doublesplit using content_pos_le by smt + next have *:"setsum content {l \ {x. \x $ k - c\ \ d} |l. l \ snd ` p \ l \ {x. \x $ k - c\ \ d} \ {}} \ 0" + apply(rule setsum_nonneg,rule) unfolding mem_Collect_eq image_iff apply(erule exE bexE conjE)+ unfolding split_paired_all + proof- fix x l a b assume as:"x = l \ {x. \x $ k - c\ \ d}" "(a, b) \ p" "l = snd (a, b)" + guess u v using p'(4)[OF as(2)] apply-by(erule exE)+ note * = this + show "content x \ 0" unfolding as snd_conv * interval_doublesplit by(rule content_pos_le) + qed have **:"norm (1::real) \ 1" by auto note division_doublesplit[OF p'',unfolded interval_doublesplit] + note dsum_bound[OF this **,unfolded interval_doublesplit[THEN sym]] + note this[unfolded real_scaleR_def real_norm_def class_semiring.semiring_rules, of k c d] note le_less_trans[OF this d(2)] + from this[unfolded abs_of_nonneg[OF *]] show "(\ka\snd ` p. content (ka \ {x. \x $ k - c\ \ d})) < e" + apply(subst vsum_nonzero_image_lemma[of "snd ` p" content "{}", unfolded o_def,THEN sym]) + apply(rule finite_imageI p' content_empty)+ unfolding forall_in_division[OF p''] + proof(rule,rule,rule,rule,rule,rule,rule,erule conjE) fix m n u v + assume as:"{m..n} \ snd ` p" "{u..v} \ snd ` p" "{m..n} \ {u..v}" "{m..n} \ {x. \x $ k - c\ \ d} = {u..v} \ {x. \x $ k - c\ \ d}" + have "({m..n} \ {x. \x $ k - c\ \ d}) \ ({u..v} \ {x. \x $ k - c\ \ d}) \ {m..n} \ {u..v}" by blast + note subset_interior[OF this, unfolded division_ofD(5)[OF p'' as(1-3)] interior_inter[of "{m..n}"]] + hence "interior ({m..n} \ {x. \x $ k - c\ \ d}) = {}" unfolding as Int_absorb by auto + thus "content ({m..n} \ {x. \x $ k - c\ \ d}) = 0" unfolding interval_doublesplit content_eq_0_interior[THEN sym] . + qed qed + finally show "(\(x, ka)\p. content (ka \ {x. \x $ k - c\ \ d}) * ?i x) < e" . + qed qed qed + +subsection {* A technical lemma about "refinement" of division. *} + +lemma tagged_division_finer: fixes p::"((real^'n) \ ((real^'n) set)) set" + assumes "p tagged_division_of {a..b}" "gauge d" + obtains q where "q tagged_division_of {a..b}" "d fine q" "\(x,k) \ p. k \ d(x) \ (x,k) \ q" +proof- + let ?P = "\p. p tagged_partial_division_of {a..b} \ gauge d \ + (\q. q tagged_division_of (\{k. \x. (x,k) \ p}) \ d fine q \ + (\(x,k) \ p. k \ d(x) \ (x,k) \ q))" + { have *:"finite p" "p tagged_partial_division_of {a..b}" using assms(1) unfolding tagged_division_of_def by auto + presume "\p. finite p \ ?P p" from this[rule_format,OF * assms(2)] guess q .. note q=this + thus ?thesis apply-apply(rule that[of q]) unfolding tagged_division_ofD[OF assms(1)] by auto + } fix p::"((real^'n) \ ((real^'n) set)) set" assume as:"finite p" + show "?P p" apply(rule,rule) using as proof(induct p) + case empty show ?case apply(rule_tac x="{}" in exI) unfolding fine_def by auto + next case (insert xk p) guess x k using surj_pair[of xk] apply- by(erule exE)+ note xk=this + note tagged_partial_division_subset[OF insert(4) subset_insertI] + from insert(3)[OF this insert(5)] guess q1 .. note q1 = conjunctD3[OF this] + have *:"\{l. \y. (y,l) \ insert xk p} = k \ \{l. \y. (y,l) \ p}" unfolding xk by auto + note p = tagged_partial_division_ofD[OF insert(4)] + from p(4)[unfolded xk, OF insertI1] guess u v apply-by(erule exE)+ note uv=this + + have "finite {k. \x. (x, k) \ p}" + apply(rule finite_subset[of _ "snd ` p"],rule) unfolding subset_eq image_iff mem_Collect_eq + apply(erule exE,rule_tac x="(xa,x)" in bexI) using p by auto + hence int:"interior {u..v} \ interior (\{k. \x. (x, k) \ p}) = {}" + apply(rule inter_interior_unions_intervals) apply(rule open_interior) apply(rule_tac[!] ballI) + unfolding mem_Collect_eq apply(erule_tac[!] exE) apply(drule p(4)[OF insertI2],assumption) + apply(rule p(5)) unfolding uv xk apply(rule insertI1,rule insertI2) apply assumption + using insert(2) unfolding uv xk by auto + + show ?case proof(cases "{u..v} \ d x") + case True thus ?thesis apply(rule_tac x="{(x,{u..v})} \ q1" in exI) apply rule + unfolding * uv apply(rule tagged_division_union,rule tagged_division_of_self) + apply(rule p[unfolded xk uv] insertI1)+ apply(rule q1,rule int) + apply(rule,rule fine_union,subst fine_def) defer apply(rule q1) + unfolding Ball_def split_paired_All split_conv apply(rule,rule,rule,rule) + apply(erule insertE) defer apply(rule UnI2) apply(drule q1(3)[rule_format]) unfolding xk uv by auto + next case False from fine_division_exists[OF assms(2), of u v] guess q2 . note q2=this + show ?thesis apply(rule_tac x="q2 \ q1" in exI) + apply rule unfolding * uv apply(rule tagged_division_union q2 q1 int fine_union)+ + unfolding Ball_def split_paired_All split_conv apply rule apply(rule fine_union) + apply(rule q1 q2)+ apply(rule,rule,rule,rule) apply(erule insertE) + apply(rule UnI2) defer apply(drule q1(3)[rule_format])using False unfolding xk uv by auto + qed qed qed + +subsection {* Hence the main theorem about negligible sets. *} + +lemma finite_product_dependent: assumes "finite s" "\x. x\s\ finite (t x)" + shows "finite {(i, j) |i j. i \ s \ j \ t i}" using assms +proof(induct) case (insert x s) + have *:"{(i, j) |i j. i \ insert x s \ j \ t i} = (\y. (x,y)) ` (t x) \ {(i, j) |i j. i \ s \ j \ t i}" by auto + show ?case unfolding * apply(rule finite_UnI) using insert by auto qed auto + +lemma sum_sum_product: assumes "finite s" "\i\s. finite (t i)" + shows "setsum (\i. setsum (x i) (t i)::real) s = setsum (\(i,j). x i j) {(i,j) | i j. i \ s \ j \ t i}" using assms +proof(induct) case (insert a s) + have *:"{(i, j) |i j. i \ insert a s \ j \ t i} = (\y. (a,y)) ` (t a) \ {(i, j) |i j. i \ s \ j \ t i}" by auto + show ?case unfolding * apply(subst setsum_Un_disjoint) unfolding setsum_insert[OF insert(1-2)] + prefer 4 apply(subst insert(3)) unfolding add_right_cancel + proof- show "setsum (x a) (t a) = (\(xa, y)\Pair a ` t a. x xa y)" apply(subst setsum_reindex) unfolding inj_on_def by auto + show "finite {(i, j) |i j. i \ s \ j \ t i}" apply(rule finite_product_dependent) using insert by auto + qed(insert insert, auto) qed auto + +lemma has_integral_negligible: fixes f::"real^'n \ 'a::real_normed_vector" + assumes "negligible s" "\x\(t - s). f x = 0" + shows "(f has_integral 0) t" +proof- presume P:"\f::real^'n \ 'a. \a b. (\x. ~(x \ s) \ f x = 0) \ (f has_integral 0) ({a..b})" + let ?f = "(\x. if x \ t then f x else 0)" + show ?thesis apply(rule_tac f="?f" in has_integral_eq) apply(rule) unfolding if_P apply(rule refl) + apply(subst has_integral_alt) apply(cases,subst if_P,assumption) unfolding if_not_P + proof- assume "\a b. t = {a..b}" then guess a b apply-by(erule exE)+ note t = this + show "(?f has_integral 0) t" unfolding t apply(rule P) using assms(2) unfolding t by auto + next show "\e>0. \B>0. \a b. ball 0 B \ {a..b} \ (\z. ((\x. if x \ t then ?f x else 0) has_integral z) {a..b} \ norm (z - 0) < e)" + apply(safe,rule_tac x=1 in exI,rule) apply(rule zero_less_one,safe) apply(rule_tac x=0 in exI) + apply(rule,rule P) using assms(2) by auto + qed +next fix f::"real^'n \ 'a" and a b::"real^'n" assume assm:"\x. x \ s \ f x = 0" + show "(f has_integral 0) {a..b}" unfolding has_integral + proof(safe) case goal1 + hence "\n. e / 2 / ((real n+1) * (2 ^ n)) > 0" + apply-apply(rule divide_pos_pos) defer apply(rule mult_pos_pos) by(auto simp add:field_simps) + note assms(1)[unfolded negligible_def has_integral,rule_format,OF this,of a b] note allI[OF this,of "\x. x"] + from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format]] + show ?case apply(rule_tac x="\x. d (nat \norm (f x)\) x" in exI) + proof safe show "gauge (\x. d (nat \norm (f x)\) x)" using d(1) unfolding gauge_def by auto + fix p assume as:"p tagged_division_of {a..b}" "(\x. d (nat \norm (f x)\) x) fine p" + let ?goal = "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e" + { presume "p\{} \ ?goal" thus ?goal apply(cases "p={}") using goal1 by auto } + assume as':"p \ {}" from real_arch_simple[of "Sup((\(x,k). norm(f x)) ` p)"] guess N .. + hence N:"\x\(\(x, k). norm (f x)) ` p. x \ real N" apply(subst(asm) Sup_finite_le_iff) using as as' by auto + have "\i. \q. q tagged_division_of {a..b} \ (d i) fine q \ (\(x, k)\p. k \ (d i) x \ (x, k) \ q)" + apply(rule,rule tagged_division_finer[OF as(1) d(1)]) by auto + from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]] + have *:"\i. (\(x, k)\q i. content k *\<^sub>R indicator s x) \ 0" apply(rule setsum_nonneg,safe) + unfolding real_scaleR_def apply(rule mult_nonneg_nonneg) apply(drule tagged_division_ofD(4)[OF q(1)]) by auto + have **:"\f g s t. finite s \ finite t \ (\(x,y) \ t. (0::real) \ g(x,y)) \ (\y\s. \x. (x,y) \ t \ f(y) \ g(x,y)) \ setsum f s \ setsum g t" + proof- case goal1 thus ?case apply-apply(rule setsum_le_included[of s t g snd f]) prefer 4 + apply safe apply(erule_tac x=x in ballE) apply(erule exE) apply(rule_tac x="(xa,x)" in bexI) by auto qed + have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) \ setsum (\i. (real i + 1) * + norm(setsum (\(x,k). content k *\<^sub>R indicator s x) (q i))) {0..N+1}" + unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right + apply(rule order_trans,rule setsum_norm) defer apply(subst sum_sum_product) prefer 3 + proof(rule **,safe) show "finite {(i, j) |i j. i \ {0..N + 1} \ j \ q i}" apply(rule finite_product_dependent) using q by auto + fix i a b assume as'':"(a,b) \ q i" show "0 \ (real i + 1) * (content b *\<^sub>R indicator s a)" + unfolding real_scaleR_def apply(rule mult_nonneg_nonneg) defer apply(rule mult_nonneg_nonneg) + using tagged_division_ofD(4)[OF q(1) as''] by auto + next fix i::nat show "finite (q i)" using q by auto + next fix x k assume xk:"(x,k) \ p" def n \ "nat \norm (f x)\" + have *:"norm (f x) \ (\(x, k). norm (f x)) ` p" using xk by auto + have nfx:"real n \ norm(f x)" "norm(f x) \ real n + 1" unfolding n_def by auto + hence "n \ {0..N + 1}" using N[rule_format,OF *] by auto + moreover note as(2)[unfolded fine_def,rule_format,OF xk,unfolded split_conv] + note q(3)[rule_format,OF xk,unfolded split_conv,rule_format,OF this] note this[unfolded n_def[symmetric]] + moreover have "norm (content k *\<^sub>R f x) \ (real n + 1) * (content k * indicator s x)" + proof(cases "x\s") case False thus ?thesis using assm by auto + next case True have *:"content k \ 0" using tagged_division_ofD(4)[OF as(1) xk] by auto + moreover have "content k * norm (f x) \ content k * (real n + 1)" apply(rule mult_mono) using nfx * by auto + ultimately show ?thesis unfolding abs_mult using nfx True by(auto simp add:field_simps) + qed ultimately show "\y. (y, x, k) \ {(i, j) |i j. i \ {0..N + 1} \ j \ q i} \ norm (content k *\<^sub>R f x) \ (real y + 1) * (content k *\<^sub>R indicator s x)" + apply(rule_tac x=n in exI,safe) apply(rule_tac x=n in exI,rule_tac x="(x,k)" in exI,safe) by auto + qed(insert as, auto) + also have "... \ setsum (\i. e / 2 / 2 ^ i) {0..N+1}" apply(rule setsum_mono) + proof- case goal1 thus ?case apply(subst mult_commute, subst pos_le_divide_eq[THEN sym]) + using d(2)[rule_format,of "q i" i] using q[rule_format] by(auto simp add:field_simps) + qed also have "... < e * inverse 2 * 2" unfolding real_divide_def setsum_right_distrib[THEN sym] + apply(rule mult_strict_left_mono) unfolding power_inverse atLeastLessThanSuc_atLeastAtMost[THEN sym] + apply(subst sumr_geometric) using goal1 by auto + finally show "?goal" by auto qed qed qed + +lemma has_integral_spike: fixes f::"real^'n \ 'a::real_normed_vector" + assumes "negligible s" "(\x\(t - s). g x = f x)" "(f has_integral y) t" + shows "(g has_integral y) t" +proof- { fix a b::"real^'n" and f g ::"real^'n \ 'a" and y::'a + assume as:"\x \ {a..b} - s. g x = f x" "(f has_integral y) {a..b}" + have "((\x. f x + (g x - f x)) has_integral (y + 0)) {a..b}" apply(rule has_integral_add[OF as(2)]) + apply(rule has_integral_negligible[OF assms(1)]) using as by auto + hence "(g has_integral y) {a..b}" by auto } note * = this + show ?thesis apply(subst has_integral_alt) using assms(2-) apply-apply(rule cond_cases,safe) + apply(rule *, assumption+) apply(subst(asm) has_integral_alt) unfolding if_not_P + apply(erule_tac x=e in allE,safe,rule_tac x=B in exI,safe) apply(erule_tac x=a in allE,erule_tac x=b in allE,safe) + apply(rule_tac x=z in exI,safe) apply(rule *[where fa2="\x. if x\t then f x else 0"]) by auto qed + +lemma has_integral_spike_eq: + assumes "negligible s" "\x\(t - s). g x = f x" + shows "((f has_integral y) t \ (g has_integral y) t)" + apply rule apply(rule_tac[!] has_integral_spike[OF assms(1)]) using assms(2) by auto + +lemma integrable_spike: assumes "negligible s" "\x\(t - s). g x = f x" "f integrable_on t" + shows "g integrable_on t" + using assms unfolding integrable_on_def apply-apply(erule exE) + apply(rule,rule has_integral_spike) by fastsimp+ + +lemma integral_spike: assumes "negligible s" "\x\(t - s). g x = f x" + shows "integral t f = integral t g" + unfolding integral_def using has_integral_spike_eq[OF assms] by auto + +subsection {* Some other trivialities about negligible sets. *} + +lemma negligible_subset[intro]: assumes "negligible s" "t \ s" shows "negligible t" unfolding negligible_def +proof(safe) case goal1 show ?case using assms(1)[unfolded negligible_def,rule_format,of a b] + apply-apply(rule has_integral_spike[OF assms(1)]) defer apply assumption + using assms(2) unfolding indicator_def by auto qed + +lemma negligible_diff[intro?]: assumes "negligible s" shows "negligible(s - t)" using assms by auto + +lemma negligible_inter: assumes "negligible s \ negligible t" shows "negligible(s \ t)" using assms by auto + +lemma negligible_union: assumes "negligible s" "negligible t" shows "negligible (s \ t)" unfolding negligible_def +proof safe case goal1 note assm = assms[unfolded negligible_def,rule_format,of a b] + thus ?case apply(subst has_integral_spike_eq[OF assms(2)]) + defer apply assumption unfolding indicator_def by auto qed + +lemma negligible_union_eq[simp]: "negligible (s \ t) \ (negligible s \ negligible t)" + using negligible_union by auto + +lemma negligible_sing[intro]: "negligible {a::real^'n}" +proof- guess x using UNIV_witness[where 'a='n] .. + show ?thesis using negligible_standard_hyperplane[of x "a$x"] by auto qed + +lemma negligible_insert[simp]: "negligible(insert a s) \ negligible s" + apply(subst insert_is_Un) unfolding negligible_union_eq by auto + +lemma negligible_empty[intro]: "negligible {}" by auto + +lemma negligible_finite[intro]: assumes "finite s" shows "negligible s" + using assms apply(induct s) by auto + +lemma negligible_unions[intro]: assumes "finite s" "\t\s. negligible t" shows "negligible(\s)" + using assms by(induct,auto) + +lemma negligible: "negligible s \ (\t::(real^'n) set. (indicator s has_integral 0) t)" + apply safe defer apply(subst negligible_def) +proof- fix t::"(real^'n) set" assume as:"negligible s" + have *:"(\x. if x \ s \ t then 1 else 0) = (\x. if x\t then if x\s then 1 else 0 else 0)" by(rule ext,auto) + show "(indicator s has_integral 0) t" apply(subst has_integral_alt) + apply(cases,subst if_P,assumption) unfolding if_not_P apply(safe,rule as[unfolded negligible_def,rule_format]) + apply(rule_tac x=1 in exI) apply(safe,rule zero_less_one) apply(rule_tac x=0 in exI) + using negligible_subset[OF as,of "s \ t"] unfolding negligible_def indicator_def unfolding * by auto qed auto + +subsection {* Finite case of the spike theorem is quite commonly needed. *} + +lemma has_integral_spike_finite: assumes "finite s" "\x\t-s. g x = f x" + "(f has_integral y) t" shows "(g has_integral y) t" + apply(rule has_integral_spike) using assms by auto + +lemma has_integral_spike_finite_eq: assumes "finite s" "\x\t-s. g x = f x" + shows "((f has_integral y) t \ (g has_integral y) t)" + apply rule apply(rule_tac[!] has_integral_spike_finite) using assms by auto + +lemma integrable_spike_finite: + assumes "finite s" "\x\t-s. g x = f x" "f integrable_on t" shows "g integrable_on t" + using assms unfolding integrable_on_def apply safe apply(rule_tac x=y in exI) + apply(rule has_integral_spike_finite) by auto + +subsection {* In particular, the boundary of an interval is negligible. *} + +lemma negligible_frontier_interval: "negligible({a..b} - {a<.. ?A" apply rule unfolding Diff_iff mem_interval not_all + apply(erule conjE exE)+ apply(rule_tac X="{x. x $ xa = a $ xa} \ {x. x $ xa = b $ xa}" in UnionI) + apply(erule_tac[!] x=xa in allE) by auto + thus ?thesis apply-apply(rule negligible_subset[of ?A]) apply(rule negligible_unions[OF finite_imageI]) by auto qed + +lemma has_integral_spike_interior: + assumes "\x\{a<..x\{a<.. (g has_integral y) ({a..b}))" + apply rule apply(rule_tac[!] has_integral_spike_interior) using assms by auto + +lemma integrable_spike_interior: assumes "\x\{a<.. = True" + unfolding neutral_def apply(rule some_equality) by auto + +lemma monoidal_and[intro]: "monoidal op \" unfolding monoidal_def by auto + +lemma iterate_and[simp]: assumes "finite s" shows "(iterate op \) s p \ (\x\s. p x)" using assms +apply induct unfolding iterate_insert[OF monoidal_and] by auto + +lemma operative_division_and: assumes "operative op \ P" "d division_of {a..b}" + shows "(\i\d. P i) \ P {a..b}" + using operative_division[OF monoidal_and assms] division_of_finite[OF assms(2)] by auto + +lemma operative_approximable: assumes "0 \ e" fixes f::"real^'n \ 'a::banach" + shows "operative op \ (\i. \g. (\x\i. norm (f x - g (x::real^'n)) \ e) \ g integrable_on i)" unfolding operative_def neutral_and +proof safe fix a b::"real^'n" { assume "content {a..b} = 0" + thus "\g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" + apply(rule_tac x=f in exI) using assms by(auto intro!:integrable_on_null) } + { fix c k g assume as:"\x\{a..b}. norm (f x - g x) \ e" "g integrable_on {a..b}" + show "\g. (\x\{a..b} \ {x. x $ k \ c}. norm (f x - g x) \ e) \ g integrable_on {a..b} \ {x. x $ k \ c}" + "\g. (\x\{a..b} \ {x. c \ x $ k}. norm (f x - g x) \ e) \ g integrable_on {a..b} \ {x. c \ x $ k}" + apply(rule_tac[!] x=g in exI) using as(1) integrable_split[OF as(2)] by auto } + fix c k g1 g2 assume as:"\x\{a..b} \ {x. x $ k \ c}. norm (f x - g1 x) \ e" "g1 integrable_on {a..b} \ {x. x $ k \ c}" + "\x\{a..b} \ {x. c \ x $ k}. norm (f x - g2 x) \ e" "g2 integrable_on {a..b} \ {x. c \ x $ k}" + let ?g = "\x. if x$k = c then f x else if x$k \ c then g1 x else g2 x" + show "\g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" apply(rule_tac x="?g" in exI) + proof safe case goal1 thus ?case apply- apply(cases "x$k=c", case_tac "x$k < c") using as assms by auto + next case goal2 presume "?g integrable_on {a..b} \ {x. x $ k \ c}" "?g integrable_on {a..b} \ {x. x $ k \ c}" + then guess h1 h2 unfolding integrable_on_def by auto from has_integral_split[OF this] + show ?case unfolding integrable_on_def by auto + next show "?g integrable_on {a..b} \ {x. x $ k \ c}" "?g integrable_on {a..b} \ {x. x $ k \ c}" + apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) using as(2,4) by auto qed qed + +lemma approximable_on_division: fixes f::"real^'n \ 'a::banach" + assumes "0 \ e" "d division_of {a..b}" "\i\d. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" + obtains g where "\x\{a..b}. norm (f x - g x) \ e" "g integrable_on {a..b}" +proof- note * = operative_division[OF monoidal_and operative_approximable[OF assms(1)] assms(2)] + note this[unfolded iterate_and[OF division_of_finite[OF assms(2)]]] from assms(3)[unfolded this[of f]] + guess g .. thus thesis apply-apply(rule that[of g]) by auto qed + +lemma integrable_continuous: fixes f::"real^'n \ 'a::banach" + assumes "continuous_on {a..b} f" shows "f integrable_on {a..b}" +proof(rule integrable_uniform_limit,safe) fix e::real assume e:"0 < e" + from compact_uniformly_continuous[OF assms compact_interval,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d .. + note d=conjunctD2[OF this,rule_format] + from fine_division_exists[OF gauge_ball[OF d(1)], of a b] guess p . note p=this + note p' = tagged_division_ofD[OF p(1)] + have *:"\i\snd ` p. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" + proof(safe,unfold snd_conv) fix x l assume as:"(x,l) \ p" + from p'(4)[OF this] guess a b apply-by(erule exE)+ note l=this + show "\g. (\x\l. norm (f x - g x) \ e) \ g integrable_on l" apply(rule_tac x="\y. f x" in exI) + proof safe show "(\y. f x) integrable_on l" unfolding integrable_on_def l by(rule,rule has_integral_const) + fix y assume y:"y\l" note fineD[OF p(2) as,unfolded subset_eq,rule_format,OF this] + note d(2)[OF _ _ this[unfolded mem_ball]] + thus "norm (f y - f x) \ e" using y p'(2-3)[OF as] unfolding vector_dist_norm l norm_minus_commute by fastsimp qed qed + from e have "0 \ e" by auto from approximable_on_division[OF this division_of_tagged_division[OF p(1)] *] guess g . + thus "\g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" by auto qed + +subsection {* Specialization of additivity to one dimension. *} + +lemma operative_1_lt: assumes "monoidal opp" + shows "operative opp f \ ((\a b. b \ a \ f {a..b::real^1} = neutral opp) \ + (\a b c. a < c \ c < b \ opp (f{a..c})(f{c..b}) = f {a..b}))" + unfolding operative_def content_eq_0_1 forall_1 vector_le_def vector_less_def +proof safe fix a b c::"real^1" assume as:"\a b c. f {a..b} = opp (f ({a..b} \ {x. x $ 1 \ c})) (f ({a..b} \ {x. c \ x $ 1}))" "a $ 1 < c $ 1" "c $ 1 < b $ 1" + from this(2-) have "{a..b} \ {x. x $ 1 \ c $ 1} = {a..c}" "{a..b} \ {x. x $ 1 \ c $ 1} = {c..b}" by auto + thus "opp (f {a..c}) (f {c..b}) = f {a..b}" unfolding as(1)[rule_format,of a b "c$1"] by auto +next fix a b::"real^1" and c::real + assume as:"\a b. b $ 1 \ a $ 1 \ f {a..b} = neutral opp" "\a b c. a $ 1 < c $ 1 \ c $ 1 < b $ 1 \ opp (f {a..c}) (f {c..b}) = f {a..b}" + show "f {a..b} = opp (f ({a..b} \ {x. x $ 1 \ c})) (f ({a..b} \ {x. c \ x $ 1}))" + proof(cases "c \ {a$1 .. b$1}") + case False hence "c c>b$1" by auto + thus ?thesis apply-apply(erule disjE) + proof- assume "c {x. x $ 1 \ c} = {1..0}" "{a..b} \ {x. c \ x $ 1} = {a..b}" by auto + show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto + next assume "b$1 {x. x $ 1 \ c} = {a..b}" "{a..b} \ {x. c \ x $ 1} = {1..0}" by auto + show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto + qed + next case True hence *:"min (b $ 1) c = c" "max (a $ 1) c = c" by auto + show ?thesis unfolding interval_split num1_eq_iff if_True * vec_def[THEN sym] + proof(cases "c = a$1 \ c = b$1") + case False thus "f {a..b} = opp (f {a..vec1 c}) (f {vec1 c..b})" + apply-apply(subst as(2)[rule_format]) using True by auto + next case True thus "f {a..b} = opp (f {a..vec1 c}) (f {vec1 c..b})" apply- + proof(erule disjE) assume "c=a$1" hence *:"a = vec1 c" unfolding Cart_eq by auto + hence "f {a..vec1 c} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto + thus ?thesis using assms unfolding * by auto + next assume "c=b$1" hence *:"b = vec1 c" unfolding Cart_eq by auto + hence "f {vec1 c..b} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto + thus ?thesis using assms unfolding * by auto qed qed qed qed + +lemma operative_1_le: assumes "monoidal opp" + shows "operative opp f \ ((\a b. b \ a \ f {a..b::real^1} = neutral opp) \ + (\a b c. a \ c \ c \ b \ opp (f{a..c})(f{c..b}) = f {a..b}))" +unfolding operative_1_lt[OF assms] +proof safe fix a b c::"real^1" assume as:"\a b c. a \ c \ c \ b \ opp (f {a..c}) (f {c..b}) = f {a..b}" "a < c" "c < b" + show "opp (f {a..c}) (f {c..b}) = f {a..b}" apply(rule as(1)[rule_format]) using as(2-) unfolding vector_le_def vector_less_def by auto +next fix a b c ::"real^1" + assume "\a b. b \ a \ f {a..b} = neutral opp" "\a b c. a < c \ c < b \ opp (f {a..c}) (f {c..b}) = f {a..b}" "a \ c" "c \ b" + note as = this[rule_format] + show "opp (f {a..c}) (f {c..b}) = f {a..b}" + proof(cases "c = a \ c = b") + case False thus ?thesis apply-apply(subst as(2)) using as(3-) unfolding vector_le_def vector_less_def Cart_eq by(auto simp del:dest_vec1_eq) + next case True thus ?thesis apply- + proof(erule disjE) assume *:"c=a" hence "f {a..c} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto + thus ?thesis using assms unfolding * by auto + next assume *:"c=b" hence "f {c..b} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto + thus ?thesis using assms unfolding * by auto qed qed qed + +subsection {* Special case of additivity we need for the FCT. *} + +lemma additive_tagged_division_1: fixes f::"real^1 \ 'a::real_normed_vector" + assumes "dest_vec1 a \ dest_vec1 b" "p tagged_division_of {a..b}" + shows "setsum (\(x,k). f(interval_upperbound k) - f(interval_lowerbound k)) p = f b - f a" +proof- let ?f = "(\k::(real^1) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))" + have *:"operative op + ?f" unfolding operative_1_lt[OF monoidal_monoid] interval_eq_empty_1 + by(auto simp add:not_less interval_bound_1 vector_less_def) + have **:"{a..b} \ {}" using assms(1) by auto note operative_tagged_division[OF monoidal_monoid * assms(2)] + note * = this[unfolded if_not_P[OF **] interval_bound_1[OF assms(1)],THEN sym ] + show ?thesis unfolding * apply(subst setsum_iterate[THEN sym]) defer + apply(rule setsum_cong2) unfolding split_paired_all split_conv using assms(2) by auto qed + +subsection {* A useful lemma allowing us to factor out the content size. *} + +lemma has_integral_factor_content: + "(f has_integral i) {a..b} \ (\e>0. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p + \ norm (setsum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a..b}))" +proof(cases "content {a..b} = 0") + case True show ?thesis unfolding has_integral_null_eq[OF True] apply safe + apply(rule,rule,rule gauge_trivial,safe) unfolding setsum_content_null[OF True] True defer + apply(erule_tac x=1 in allE,safe) defer apply(rule fine_division_exists[of _ a b],assumption) + apply(erule_tac x=p in allE) unfolding setsum_content_null[OF True] by auto +next case False note F = this[unfolded content_lt_nz[THEN sym]] + let ?P = "\e opp. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" + show ?thesis apply(subst has_integral) + proof safe fix e::real assume e:"e>0" + { assume "\e>0. ?P e op <" thus "?P (e * content {a..b}) op \" apply(erule_tac x="e * content {a..b}" in allE) + apply(erule impE) defer apply(erule exE,rule_tac x=d in exI) + using F e by(auto simp add:field_simps intro:mult_pos_pos) } + { assume "\e>0. ?P (e * content {a..b}) op \" thus "?P e op <" apply(erule_tac x="e / 2 / content {a..b}" in allE) + apply(erule impE) defer apply(erule exE,rule_tac x=d in exI) + using F e by(auto simp add:field_simps intro:mult_pos_pos) } qed qed + +subsection {* Fundamental theorem of calculus. *} + +lemma fundamental_theorem_of_calculus: fixes f::"real^1 \ 'a::banach" + assumes "a \ b" "\x\{a..b}. ((f o vec1) has_vector_derivative f'(vec1 x)) (at x within {a..b})" + shows "(f' has_integral (f(vec1 b) - f(vec1 a))) ({vec1 a..vec1 b})" +unfolding has_integral_factor_content +proof safe fix e::real assume e:"e>0" have ab:"dest_vec1 (vec1 a) \ dest_vec1 (vec1 b)" using assms(1) by auto + note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt] + have *:"\P Q. \x\{a..b}. P x \ (\e>0. \d>0. Q x e d) \ \x. \(d::real)>0. x\{a..b} \ Q x e d" using e by blast + note this[OF assm,unfolded gauge_existence_lemma] from choice[OF this,unfolded Ball_def[symmetric]] + guess d .. note d=conjunctD2[OF this[rule_format],rule_format] + show "\d. gauge d \ (\p. p tagged_division_of {vec1 a..vec1 b} \ d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f (vec1 b) - f (vec1 a))) \ e * content {vec1 a..vec1 b})" + apply(rule_tac x="\x. ball x (d (dest_vec1 x))" in exI,safe) + apply(rule gauge_ball_dependent,rule,rule d(1)) + proof- fix p assume as:"p tagged_division_of {vec1 a..vec1 b}" "(\x. ball x (d (dest_vec1 x))) fine p" + show "norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f (vec1 b) - f (vec1 a))) \ e * content {vec1 a..vec1 b}" + unfolding content_1[OF ab] additive_tagged_division_1[OF ab as(1),of f,THEN sym] + unfolding vector_minus_component[THEN sym] additive_tagged_division_1[OF ab as(1),of "\x. x",THEN sym] + apply(subst dest_vec1_setsum) unfolding setsum_right_distrib defer unfolding setsum_subtractf[THEN sym] + proof(rule setsum_norm_le,safe) fix x k assume "(x,k)\p" + note xk = tagged_division_ofD(2-4)[OF as(1) this] from this(3) guess u v apply-by(erule exE)+ note k=this + have *:"dest_vec1 u \ dest_vec1 v" using xk unfolding k by auto + have ball:"\xa\k. xa \ ball x (d (dest_vec1 x))" using as(2)[unfolded fine_def,rule_format,OF `(x,k)\p`,unfolded split_conv subset_eq] . + have "norm ((v$1 - u$1) *\<^sub>R f' x - (f v - f u)) \ norm (f u - f x - (u$1 - x$1) *\<^sub>R f' x) + norm (f v - f x - (v$1 - x$1) *\<^sub>R f' x)" + apply(rule order_trans[OF _ norm_triangle_ineq4]) apply(rule eq_refl) apply(rule arg_cong[where f=norm]) + unfolding scaleR.diff_left by(auto simp add:group_simps) + also have "... \ e * norm (dest_vec1 u - dest_vec1 x) + e * norm (dest_vec1 v - dest_vec1 x)" + apply(rule add_mono) apply(rule d(2)[of "x$1" "u$1",unfolded o_def vec1_dest_vec1]) prefer 4 + apply(rule d(2)[of "x$1" "v$1",unfolded o_def vec1_dest_vec1]) + using ball[rule_format,of u] ball[rule_format,of v] + using xk(1-2) unfolding k subset_eq by(auto simp add:vector_dist_norm norm_real) + also have "... \ e * dest_vec1 (interval_upperbound k - interval_lowerbound k)" + unfolding k interval_bound_1[OF *] using xk(1) unfolding k by(auto simp add:vector_dist_norm norm_real field_simps) + finally show "norm (content k *\<^sub>R f' x - (f (interval_upperbound k) - f (interval_lowerbound k))) \ + e * dest_vec1 (interval_upperbound k - interval_lowerbound k)" unfolding k interval_bound_1[OF *] content_1[OF *] . + qed(insert as, auto) qed qed + +subsection {* Attempt a systematic general set of "offset" results for components. *} + +lemma gauge_modify: + assumes "(\s. open s \ open {x. f(x) \ s})" "gauge d" + shows "gauge (\x y. d (f x) (f y))" + using assms unfolding gauge_def apply safe defer apply(erule_tac x="f x" in allE) + apply(erule_tac x="d (f x)" in allE) unfolding mem_def Collect_def by auto + +subsection {* Only need trivial subintervals if the interval itself is trivial. *} + +lemma division_of_nontrivial: fixes s::"(real^'n) set set" + assumes "s division_of {a..b}" "content({a..b}) \ 0" + shows "{k. k \ s \ content k \ 0} division_of {a..b}" using assms(1) apply- +proof(induct "card s" arbitrary:s rule:nat_less_induct) + fix s::"(real^'n) set set" assume assm:"s division_of {a..b}" + "\mx. m = card x \ x division_of {a..b} \ {k \ x. content k \ 0} division_of {a..b}" + note s = division_ofD[OF assm(1)] let ?thesis = "{k \ s. content k \ 0} division_of {a..b}" + { presume *:"{k \ s. content k \ 0} \ s \ ?thesis" + show ?thesis apply cases defer apply(rule *,assumption) using assm(1) by auto } + assume noteq:"{k \ s. content k \ 0} \ s" + then obtain k where k:"k\s" "content k = 0" by auto + from s(4)[OF k(1)] guess c d apply-by(erule exE)+ note k=k this + from k have "card s > 0" unfolding card_gt_0_iff using assm(1) by auto + hence card:"card (s - {k}) < card s" using assm(1) k(1) apply(subst card_Diff_singleton_if) by auto + have *:"closed (\(s - {k}))" apply(rule closed_Union) defer apply rule apply(drule DiffD1,drule s(4)) + apply safe apply(rule closed_interval) using assm(1) by auto + have "k \ \(s - {k})" apply safe apply(rule *[unfolded closed_limpt,rule_format]) unfolding islimpt_approachable + proof safe fix x and e::real assume as:"x\k" "e>0" + from k(2)[unfolded k content_eq_0] guess i .. + hence i:"c$i = d$i" using s(3)[OF k(1),unfolded k] unfolding interval_ne_empty by smt + hence xi:"x$i = d$i" using as unfolding k mem_interval by smt + def y \ "(\ j. if j = i then if c$i \ (a$i + b$i) / 2 then c$i + min e (b$i - c$i) / 2 else c$i - min e (c$i - a$i) / 2 else x$j)" + show "\x'\\(s - {k}). x' \ x \ dist x' x < e" apply(rule_tac x=y in bexI) + proof have "d \ {c..d}" using s(3)[OF k(1)] unfolding k interval_eq_empty mem_interval by(fastsimp simp add: not_less) + 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 Cart_lambda_beta 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] by smt+ + thus "y \ x" unfolding Cart_eq by auto + have *:"UNIV = insert i (UNIV - {i})" by auto + have "norm (y - x) < e + setsum (\i. 0) (UNIV::'n set)" apply(rule le_less_trans[OF norm_le_l1]) + apply(subst *,subst setsum_insert) prefer 3 apply(rule add_less_le_mono) + proof- show "\(y - x) $ i\ < e" unfolding y_def Cart_lambda_beta vector_minus_component if_P[OF refl] + apply(cases) apply(subst if_P,assumption) unfolding if_not_P unfolding i xi using di as(2) by auto + show "(\i\UNIV - {i}. \(y - x) $ i\) \ (\i\UNIV. 0)" unfolding y_def by auto + qed auto thus "dist y x < e" unfolding vector_dist_norm by auto + have "y\k" unfolding k mem_interval apply rule apply(erule_tac x=i in allE) using xyi unfolding k i xi by auto + moreover have "y \ \s" unfolding s mem_interval + proof note simps = y_def Cart_lambda_beta if_not_P + fix j::'n show "a $ j \ y $ j \ y $ j \ b $ j" + proof(cases "j = i") case False have "x \ {a..b}" using s(2)[OF k(1)] as(1) by auto + thus ?thesis unfolding simps if_not_P[OF False] unfolding mem_interval by auto + next case True note T = this show ?thesis + proof(cases "c $ i \ (a $ i + b $ i) / 2") + case True show ?thesis unfolding simps if_P[OF T] if_P[OF True] unfolding i + using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps) + next case False thus ?thesis unfolding simps if_P[OF T] if_not_P[OF False] unfolding i + using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps) + qed qed qed + ultimately show "y \ \(s - {k})" by auto + qed qed hence "\(s - {k}) = {a..b}" unfolding s(6)[THEN sym] by auto + hence "{ka \ s - {k}. content ka \ 0} division_of {a..b}" apply-apply(rule assm(2)[rule_format,OF card refl]) + apply(rule division_ofI) defer apply(rule_tac[1-4] s) using assm(1) by auto + moreover have "{ka \ s - {k}. content ka \ 0} = {k \ s. content k \ 0}" using k by auto ultimately show ?thesis by auto qed + +subsection {* Integrabibility on subintervals. *} + +lemma operative_integrable: fixes f::"real^'n \ 'a::banach" shows + "operative op \ (\i. f integrable_on i)" + unfolding operative_def neutral_and apply safe apply(subst integrable_on_def) + unfolding has_integral_null_eq apply(rule,rule refl) apply(rule,assumption)+ + unfolding integrable_on_def by(auto intro: has_integral_split) + +lemma integrable_subinterval: fixes f::"real^'n \ 'a::banach" + assumes "f integrable_on {a..b}" "{c..d} \ {a..b}" shows "f integrable_on {c..d}" + apply(cases "{c..d} = {}") defer apply(rule partial_division_extend_1[OF assms(2)],assumption) + using operative_division_and[OF operative_integrable,THEN sym,of _ _ _ f] assms(1) by auto + +subsection {* Combining adjacent intervals in 1 dimension. *} + +lemma has_integral_combine: assumes "(a::real^1) \ c" "c \ b" + "(f has_integral i) {a..c}" "(f has_integral (j::'a::banach)) {c..b}" + shows "(f has_integral (i + j)) {a..b}" +proof- note operative_integral[of f, unfolded operative_1_le[OF monoidal_lifted[OF monoidal_monoid]]] + note conjunctD2[OF this,rule_format] note * = this(2)[OF conjI[OF assms(1-2)],unfolded if_P[OF assms(3)]] + hence "f integrable_on {a..b}" apply- apply(rule ccontr) apply(subst(asm) if_P) defer + apply(subst(asm) if_P) using assms(3-) by auto + with * show ?thesis apply-apply(subst(asm) if_P) defer apply(subst(asm) if_P) defer apply(subst(asm) if_P) + unfolding lifted.simps using assms(3-) by(auto simp add: integrable_on_def integral_unique) qed + +lemma integral_combine: fixes f::"real^1 \ 'a::banach" + assumes "a \ c" "c \ b" "f integrable_on ({a..b})" + shows "integral {a..c} f + integral {c..b} f = integral({a..b}) f" + apply(rule integral_unique[THEN sym]) apply(rule has_integral_combine[OF assms(1-2)]) + apply(rule_tac[!] integrable_integral integrable_subinterval[OF assms(3)])+ using assms(1-2) by auto + +lemma integrable_combine: fixes f::"real^1 \ 'a::banach" + assumes "a \ c" "c \ b" "f integrable_on {a..c}" "f integrable_on {c..b}" + shows "f integrable_on {a..b}" using assms unfolding integrable_on_def by(fastsimp intro!:has_integral_combine) + +subsection {* Reduce integrability to "local" integrability. *} + +lemma integrable_on_little_subintervals: fixes f::"real^'n \ 'a::banach" + assumes "\x\{a..b}. \d>0. \u v. x \ {u..v} \ {u..v} \ ball x d \ {u..v} \ {a..b} \ f integrable_on {u..v}" + shows "f integrable_on {a..b}" +proof- have "\x. \d. x\{a..b} \ d>0 \ (\u v. x \ {u..v} \ {u..v} \ ball x d \ {u..v} \ {a..b} \ f integrable_on {u..v})" + using assms by auto note this[unfolded gauge_existence_lemma] from choice[OF this] guess d .. note d=this[rule_format] + guess p apply(rule fine_division_exists[OF gauge_ball_dependent,of d a b]) using d by auto note p=this(1-2) + note division_of_tagged_division[OF this(1)] note * = operative_division_and[OF operative_integrable,OF this,THEN sym,of f] + show ?thesis unfolding * apply safe unfolding snd_conv + proof- fix x k assume "(x,k) \ p" note tagged_division_ofD(2-4)[OF p(1) this] fineD[OF p(2) this] + thus "f integrable_on k" apply safe apply(rule d[THEN conjunct2,rule_format,of x]) by auto qed qed + +subsection {* Second FCT or existence of antiderivative. *} + +lemma integrable_const[intro]:"(\x. c) integrable_on {a..b}" + unfolding integrable_on_def by(rule,rule has_integral_const) + +lemma integral_has_vector_derivative: fixes f::"real \ 'a::banach" + assumes "continuous_on {a..b} f" "x \ {a..b}" + shows "((\u. integral {vec a..vec u} (f o dest_vec1)) has_vector_derivative f(x)) (at x within {a..b})" + unfolding has_vector_derivative_def has_derivative_within_alt +apply safe apply(rule scaleR.bounded_linear_left) +proof- fix e::real assume e:"e>0" + note compact_uniformly_continuous[OF assms(1) compact_real_interval,unfolded uniformly_continuous_on_def] + from this[rule_format,OF e] guess d apply-by(erule conjE exE)+ note d=this[rule_format] + let ?I = "\a b. integral {vec1 a..vec1 b} (f \ dest_vec1)" + show "\d>0. \y\{a..b}. norm (y - x) < d \ norm (?I a y - ?I a x - (y - x) *\<^sub>R f x) \ e * norm (y - x)" + proof(rule,rule,rule d,safe) case goal1 show ?case proof(cases "y < x") + case False have "f \ dest_vec1 integrable_on {vec1 a..vec1 y}" apply(rule integrable_subinterval,rule integrable_continuous) + apply(rule continuous_on_o_dest_vec1 assms)+ unfolding not_less using assms(2) goal1 by auto + hence *:"?I a y - ?I a x = ?I x y" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine) + using False unfolding not_less using assms(2) goal1 by auto + have **:"norm (y - x) = content {vec1 x..vec1 y}" apply(subst content_1) using False unfolding not_less by auto + show ?thesis unfolding ** apply(rule has_integral_bound[where f="(\u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def + defer apply(rule has_integral_sub) apply(rule integrable_integral) + apply(rule integrable_subinterval,rule integrable_continuous) apply(rule continuous_on_o_dest_vec1[unfolded o_def] assms)+ + proof- show "{vec1 x..vec1 y} \ {vec1 a..vec1 b}" using goal1 assms(2) by auto + have *:"y - x = norm(y - x)" using False by auto + show "((\xa. f x) has_integral (y - x) *\<^sub>R f x) {vec1 x..vec1 y}" apply(subst *) unfolding ** by auto + show "\xa\{vec1 x..vec1 y}. norm (f (dest_vec1 xa) - f x) \ e" apply safe apply(rule less_imp_le) + apply(rule d(2)[unfolded vector_dist_norm]) using assms(2) using goal1 by auto + qed(insert e,auto) + next case True have "f \ dest_vec1 integrable_on {vec1 a..vec1 x}" apply(rule integrable_subinterval,rule integrable_continuous) + apply(rule continuous_on_o_dest_vec1 assms)+ unfolding not_less using assms(2) goal1 by auto + hence *:"?I a x - ?I a y = ?I y x" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine) + using True using assms(2) goal1 by auto + have **:"norm (y - x) = content {vec1 y..vec1 x}" apply(subst content_1) using True unfolding not_less by auto + have ***:"\fy fx c::'a. fx - fy - (y - x) *\<^sub>R c = -(fy - fx - (x - y) *\<^sub>R c)" unfolding scaleR_left.diff by auto + show ?thesis apply(subst ***) unfolding norm_minus_cancel ** + apply(rule has_integral_bound[where f="(\u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def + defer apply(rule has_integral_sub) apply(subst minus_minus[THEN sym]) unfolding minus_minus + apply(rule integrable_integral) apply(rule integrable_subinterval,rule integrable_continuous) + apply(rule continuous_on_o_dest_vec1[unfolded o_def] assms)+ + proof- show "{vec1 y..vec1 x} \ {vec1 a..vec1 b}" using goal1 assms(2) by auto + have *:"x - y = norm(y - x)" using True by auto + show "((\xa. f x) has_integral (x - y) *\<^sub>R f x) {vec1 y..vec1 x}" apply(subst *) unfolding ** by auto + show "\xa\{vec1 y..vec1 x}. norm (f (dest_vec1 xa) - f x) \ e" apply safe apply(rule less_imp_le) + apply(rule d(2)[unfolded vector_dist_norm]) using assms(2) using goal1 by auto + qed(insert e,auto) qed qed qed + +lemma integral_has_vector_derivative': fixes f::"real^1 \ 'a::banach" + assumes "continuous_on {a..b} f" "x \ {a..b}" + shows "((\u. (integral {a..vec u} f)) has_vector_derivative f x) (at (x$1) within {a$1..b$1})" + using integral_has_vector_derivative[OF continuous_on_o_vec1[OF assms(1)], of "x$1"] + unfolding o_def vec1_dest_vec1 using assms(2) by auto + +lemma antiderivative_continuous: assumes "continuous_on {a..b::real} f" + obtains g where "\x\ {a..b}. (g has_vector_derivative (f(x)::_::banach)) (at x within {a..b})" + apply(rule that,rule) using integral_has_vector_derivative[OF assms] by auto + +subsection {* Combined fundamental theorem of calculus. *} + +lemma antiderivative_integral_continuous: fixes f::"real \ 'a::banach" assumes "continuous_on {a..b} f" + obtains g where "\u\{a..b}. \v \ {a..b}. u \ v \ ((f o dest_vec1) has_integral (g v - g u)) {vec u..vec v}" +proof- from antiderivative_continuous[OF assms] guess g . note g=this + show ?thesis apply(rule that[of g]) + proof safe case goal1 have "\x\{u..v}. (g has_vector_derivative f x) (at x within {u..v})" + apply(rule,rule has_vector_derivative_within_subset) apply(rule g[rule_format]) using goal1(1-2) by auto + thus ?case using fundamental_theorem_of_calculus[OF goal1(3),of "g o dest_vec1" "f o dest_vec1"] + unfolding o_def vec1_dest_vec1 by auto qed qed + +subsection {* General "twiddling" for interval-to-interval function image. *} + +lemma has_integral_twiddle: + assumes "0 < r" "\x. h(g x) = x" "\x. g(h x) = x" "\x. continuous (at x) g" + "\u v. \w z. g ` {u..v} = {w..z}" + "\u v. \w z. h ` {u..v} = {w..z}" + "\u v. content(g ` {u..v}) = r * content {u..v}" + "(f has_integral i) {a..b}" + shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` {a..b})" +proof- { presume *:"{a..b} \ {} \ ?thesis" + show ?thesis apply cases defer apply(rule *,assumption) + proof- case goal1 thus ?thesis unfolding goal1 assms(8)[unfolded goal1 has_integral_empty_eq] by auto qed } + assume "{a..b} \ {}" from assms(6)[rule_format,of a b] guess w z apply-by(erule exE)+ note wz=this + have inj:"inj g" "inj h" unfolding inj_on_def apply safe apply(rule_tac[!] ccontr) + using assms(2) apply(erule_tac x=x in allE) using assms(2) apply(erule_tac x=y in allE) defer + using assms(3) apply(erule_tac x=x in allE) using assms(3) apply(erule_tac x=y in allE) by auto + show ?thesis unfolding has_integral_def has_integral_compact_interval_def apply(subst if_P) apply(rule,rule,rule wz) + proof safe fix e::real assume e:"e>0" hence "e * r > 0" using assms(1) by(rule mult_pos_pos) + from assms(8)[unfolded has_integral,rule_format,OF this] guess d apply-by(erule exE conjE)+ note d=this[rule_format] + def d' \ "\x y. d (g x) (g y)" have d':"\x. d' x = {y. g y \ (d (g x))}" unfolding d'_def by(auto simp add:mem_def) + show "\d. gauge d \ (\p. p tagged_division_of h ` {a..b} \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" + proof(rule_tac x=d' in exI,safe) show "gauge d'" using d(1) unfolding gauge_def d' using continuous_open_preimage_univ[OF assms(4)] by auto + fix p assume as:"p tagged_division_of h ` {a..b}" "d' fine p" note p = tagged_division_ofD[OF as(1)] + have "(\(x, k). (g x, g ` k)) ` p tagged_division_of {a..b} \ d fine (\(x, k). (g x, g ` k)) ` p" unfolding tagged_division_of + proof safe show "finite ((\(x, k). (g x, g ` k)) ` p)" using as by auto + show "d fine (\(x, k). (g x, g ` k)) ` p" using as(2) unfolding fine_def d' by auto + fix x k assume xk[intro]:"(x,k) \ p" show "g x \ g ` k" using p(2)[OF xk] by auto + show "\u v. g ` k = {u..v}" using p(4)[OF xk] using assms(5-6) by auto + { fix y assume "y \ k" thus "g y \ {a..b}" "g y \ {a..b}" using p(3)[OF xk,unfolded subset_eq,rule_format,of "h (g y)"] + using assms(2)[rule_format,of y] unfolding inj_image_mem_iff[OF inj(2)] by auto } + fix x' k' assume xk':"(x',k') \ p" fix z assume "z \ interior (g ` k)" "z \ interior (g ` k')" + hence *:"interior (g ` k) \ interior (g ` k') \ {}" by auto + have same:"(x, k) = (x', k')" apply-apply(rule ccontr,drule p(5)[OF xk xk']) + proof- assume as:"interior k \ interior k' = {}" from nonempty_witness[OF *] guess z . + hence "z \ g ` (interior k \ interior k')" using interior_image_subset[OF assms(4) inj(1)] + unfolding image_Int[OF inj(1)] by auto thus False using as by blast + qed thus "g x = g x'" by auto + { fix z assume "z \ k" thus "g z \ g ` k'" using same by auto } + { fix z assume "z \ k'" thus "g z \ g ` k" using same by auto } + next fix x assume "x \ {a..b}" hence "h x \ \{k. \x. (x, k) \ p}" using p(6) by auto + then guess X unfolding Union_iff .. note X=this from this(1) guess y unfolding mem_Collect_eq .. + thus "x \ \{k. \x. (x, k) \ (\(x, k). (g x, g ` k)) ` p}" apply- + apply(rule_tac X="g ` X" in UnionI) defer apply(rule_tac x="h x" in image_eqI) + using X(2) assms(3)[rule_format,of x] by auto + qed note ** = d(2)[OF this] have *:"inj_on (\(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastsimp + have "(\(x, k)\(\(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\(x, k)\p. content k *\<^sub>R f (g x)) - i" (is "?l = _") unfolding group_simps add_left_cancel + unfolding setsum_reindex[OF *] apply(subst scaleR_right.setsum) defer apply(rule setsum_cong2) unfolding o_def split_paired_all split_conv + apply(drule p(4)) apply safe unfolding assms(7)[rule_format] using p by auto + also have "... = r *\<^sub>R ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") unfolding scaleR.diff_right scaleR.scaleR_left[THEN sym] + unfolding real_scaleR_def using assms(1) by auto finally have *:"?l = ?r" . + show "norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" using ** unfolding * unfolding norm_scaleR + using assms(1) by(auto simp add:field_simps) qed qed qed + +subsection {* Special case of a basic affine transformation. *} + +lemma interval_image_affinity_interval: shows "\u v. (\x. m *\<^sub>R (x::real^'n) + c) ` {a..b} = {u..v}" + unfolding image_affinity_interval by auto + +lemmas Cart_simps = Cart_nth.add Cart_nth.minus Cart_nth.zero Cart_nth.diff Cart_nth.scaleR real_scaleR_def Cart_lambda_beta + Cart_eq vector_le_def vector_less_def + +lemma setprod_cong2: assumes "\x. x \ A \ f x = g x" shows "setprod f A = setprod g A" + apply(rule setprod_cong) using assms by auto + +lemma content_image_affinity_interval: + "content((\x::real^'n. m *\<^sub>R x + c) ` {a..b}) = (abs m) ^ CARD('n) * content {a..b}" (is "?l = ?r") +proof- { presume *:"{a..b}\{} \ ?thesis" show ?thesis apply(cases,rule *,assumption) + unfolding not_not using content_empty by auto } + assume as:"{a..b}\{}" show ?thesis proof(cases "m \ 0") + case True show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_P[OF True] + unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') + defer apply(subst setprod_constant[THEN sym]) apply(rule finite_UNIV) unfolding setprod_timesf[THEN sym] + apply(rule setprod_cong2) using True as unfolding interval_ne_empty Cart_simps not_le + by(auto simp add:field_simps intro:mult_left_mono) + next case False show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_not_P[OF False] + unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') + defer apply(subst setprod_constant[THEN sym]) apply(rule finite_UNIV) unfolding setprod_timesf[THEN sym] + apply(rule setprod_cong2) using False as unfolding interval_ne_empty Cart_simps not_le + by(auto simp add:field_simps mult_le_cancel_left_neg) qed qed + +lemma has_integral_affinity: assumes "(f has_integral i) {a..b::real^'n}" "m \ 0" + shows "((\x. f(m *\<^sub>R x + c)) has_integral ((1 / (abs(m) ^ CARD('n::finite))) *\<^sub>R i)) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` {a..b})" + apply(rule has_integral_twiddle,safe) unfolding Cart_eq Cart_simps apply(rule zero_less_power) + defer apply(insert assms(2), simp add:field_simps) apply(insert assms(2), simp add:field_simps) + apply(rule continuous_intros)+ apply(rule interval_image_affinity_interval)+ apply(rule content_image_affinity_interval) using assms by auto + +lemma integrable_affinity: assumes "f integrable_on {a..b}" "m \ 0" + shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` {a..b})" + using assms unfolding integrable_on_def apply safe apply(drule has_integral_affinity) by auto + +subsection {* Special case of stretching coordinate axes separately. *} + +lemma image_stretch_interval: + "(\x. \ k. m k * x$k) ` {a..b::real^'n} = + (if {a..b} = {} then {} else {(\ k. min (m(k) * a$k) (m(k) * b$k)) .. (\ k. max (m(k) * a$k) (m(k) * b$k))})" (is "?l = ?r") +proof(cases "{a..b}={}") case True thus ?thesis unfolding True by auto +next have *:"\P Q. (\i. P i) \ (\i. Q i) \ (\i. P i \ Q i)" by auto + case False note ab = this[unfolded interval_ne_empty] + show ?thesis apply-apply(rule set_ext) + proof- fix x::"real^'n" have **:"\P Q. (\i. P i = Q i) \ (\i. P i) = (\i. Q i)" by auto + show "x \ ?l \ x \ ?r" unfolding if_not_P[OF False] + unfolding image_iff mem_interval Bex_def Cart_simps Cart_eq * + unfolding lambda_skolem[THEN sym,of "\ i xa. (a $ i \ xa \ xa \ b $ i) \ x $ i = m i * xa"] + proof(rule **,rule) fix i::'n show "(\xa. (a $ i \ xa \ xa \ b $ i) \ x $ i = m i * xa) = + (min (m i * a $ i) (m i * b $ i) \ x $ i \ x $ i \ max (m i * a $ i) (m i * b $ i))" + proof(cases "m i = 0") case True thus ?thesis using ab by auto + next case False hence "0 < m i \ 0 > m i" by auto thus ?thesis apply- + proof(erule disjE) assume as:"0 < m i" hence *:"min (m i * a $ i) (m i * b $ i) = m i * a $ i" + "max (m i * a $ i) (m i * b $ i) = m i * b $ i" using ab unfolding min_def max_def by auto + show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$i" in exI) + using as by(auto simp add:field_simps) + next assume as:"0 > m i" hence *:"max (m i * a $ i) (m i * b $ i) = m i * a $ i" + "min (m i * a $ i) (m i * b $ i) = m i * b $ i" using ab as unfolding min_def max_def + by(auto simp add:field_simps mult_le_cancel_left_neg intro:real_le_antisym) + show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$i" in exI) + using as by(auto simp add:field_simps) qed qed qed qed qed + +lemma interval_image_stretch_interval: "\u v. (\x. \ k. m k * x$k) ` {a..b::real^'n} = {u..v}" + unfolding image_stretch_interval by auto + +lemma content_image_stretch_interval: + "content((\x::real^'n. \ k. m k * x$k) ` {a..b}) = abs(setprod m UNIV) * content({a..b})" +proof(cases "{a..b} = {}") case True thus ?thesis + unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto +next case False hence "(\x. \ k. m k * x $ k) ` {a..b} \ {}" by auto + thus ?thesis using False unfolding content_def image_stretch_interval apply- unfolding interval_bounds' if_not_P + unfolding abs_setprod setprod_timesf[THEN sym] apply(rule setprod_cong2) unfolding Cart_lambda_beta + proof- fix i::'n have "(m i < 0 \ m i > 0) \ m i = 0" by auto + thus "max (m i * a $ i) (m i * b $ i) - min (m i * a $ i) (m i * b $ i) = \m i\ * (b $ i - a $ i)" + apply-apply(erule disjE)+ unfolding min_def max_def using False[unfolded interval_ne_empty,rule_format,of i] + by(auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos) qed qed + +lemma has_integral_stretch: assumes "(f has_integral i) {a..b}" "\k. ~(m k = 0)" + shows "((\x. f(\ k. m k * x$k)) has_integral + ((1/(abs(setprod m UNIV))) *\<^sub>R i)) ((\x. \ k. 1/(m k) * x$k) ` {a..b})" + apply(rule has_integral_twiddle) unfolding zero_less_abs_iff content_image_stretch_interval + unfolding image_stretch_interval empty_as_interval Cart_eq using assms +proof- show "\x. continuous (at x) (\x. \ k. m k * x $ k)" + apply(rule,rule linear_continuous_at) unfolding linear_linear + unfolding linear_def Cart_simps Cart_eq by(auto simp add:field_simps) qed auto + +lemma integrable_stretch: + assumes "f integrable_on {a..b}" "\k. ~(m k = 0)" + shows "(\x. f(\ k. m k * x$k)) integrable_on ((\x. \ k. 1/(m k) * x$k) ` {a..b})" + using assms unfolding integrable_on_def apply-apply(erule exE) apply(drule has_integral_stretch) by auto + +subsection {* even more special cases. *} + +lemma uminus_interval_vector[simp]:"uminus ` {a..b} = {-b .. -a::real^'n}" + apply(rule set_ext,rule) defer unfolding image_iff + apply(rule_tac x="-x" in bexI) by(auto simp add:vector_le_def minus_le_iff le_minus_iff) + +lemma has_integral_reflect_lemma[intro]: assumes "(f has_integral i) {a..b}" + shows "((\x. f(-x)) has_integral i) {-b .. -a}" + using has_integral_affinity[OF assms, of "-1" 0] by auto + +lemma has_integral_reflect[simp]: "((\x. f(-x)) has_integral i) {-b..-a} \ (f has_integral i) ({a..b})" + apply rule apply(drule_tac[!] has_integral_reflect_lemma) by auto + +lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on {-b..-a} \ f integrable_on {a..b}" + unfolding integrable_on_def by auto + +lemma integral_reflect[simp]: "integral {-b..-a} (\x. f(-x)) = integral ({a..b}) f" + unfolding integral_def by auto + +subsection {* Stronger form of FCT; quite a tedious proof. *} + +(** move this **) +declare norm_triangle_ineq4[intro] + +lemma bgauge_existence_lemma: "(\x\s. \d::real. 0 < d \ q d x) \ (\x. \d>0. x\s \ q d x)" by(meson zero_less_one) + +lemma additive_tagged_division_1': fixes f::"real \ 'a::real_normed_vector" + assumes "a \ b" "p tagged_division_of {vec1 a..vec1 b}" + shows "setsum (\(x,k). f (dest_vec1 (interval_upperbound k)) - f(dest_vec1 (interval_lowerbound k))) p = f b - f a" + using additive_tagged_division_1[OF _ assms(2), of "f o dest_vec1"] + unfolding o_def vec1_dest_vec1 using assms(1) by auto + +lemma split_minus[simp]:"(\(x, k). ?f x k) x - (\(x, k). ?g x k) x = (\(x, k). ?f x k - ?g x k) x" + unfolding split_def by(rule refl) + +lemma norm_triangle_le_sub: "norm x + norm y \ e \ norm (x - y) \ e" + apply(subst(asm)(2) norm_minus_cancel[THEN sym]) + apply(drule norm_triangle_le) by(auto simp add:group_simps) + +lemma fundamental_theorem_of_calculus_interior: + assumes"a \ b" "continuous_on {a..b} f" "\x\{a<.. ?thesis" + show ?thesis proof(cases,rule *,assumption) + assume "\ a < b" hence "a = b" using assms(1) by auto + hence *:"{vec a .. vec b} = {vec b}" "f b - f a = 0" apply(auto simp add: Cart_simps) by smt + show ?thesis unfolding *(2) apply(rule has_integral_null) unfolding content_eq_0_1 using * `a=b` by auto + qed } assume ab:"a < b" + let ?P = "\e. \d. gauge d \ (\p. p tagged_division_of {vec1 a..vec1 b} \ d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R (f' \ dest_vec1) x) - (f b - f a)) \ e * content {vec1 a..vec1 b})" + { presume "\e. e>0 \ ?P e" thus ?thesis unfolding has_integral_factor_content by auto } + fix e::real assume e:"e>0" + note assms(3)[unfolded has_vector_derivative_def has_derivative_at_alt ball_conj_distrib] + note conjunctD2[OF this] note bounded=this(1) and this(2) + from this(2) have "\x\{a<..d>0. \y. norm (y - x) < d \ norm (f y - f x - (y - x) *\<^sub>R f' x) \ e/2 * norm (y - x)" + apply-apply safe apply(erule_tac x=x in ballE,erule_tac x="e/2" in allE) using e by auto note this[unfolded bgauge_existence_lemma] + from choice[OF this] guess d .. note conjunctD2[OF this[rule_format]] note d = this[rule_format] + have "bounded (f ` {a..b})" apply(rule compact_imp_bounded compact_continuous_image)+ using compact_real_interval assms by auto + from this[unfolded bounded_pos] guess B .. note B = this[rule_format] + + have "\da. 0 < da \ (\c. a \ c \ {a..c} \ {a..b} \ {a..c} \ ball a da + \ norm(content {vec1 a..vec1 c} *\<^sub>R f' a - (f c - f a)) \ (e * (b - a)) / 4)" + proof- have "a\{a..b}" using ab by auto + note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] + note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" using e ab by(auto simp add:field_simps) + from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format] + have "\l. 0 < l \ norm(l *\<^sub>R f' a) \ (e * (b - a)) / 8" + proof(cases "f' a = 0") case True + thus ?thesis apply(rule_tac x=1 in exI) using ab e by(auto intro!:mult_nonneg_nonneg) + next case False thus ?thesis + apply(rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI) + using ab e by(auto simp add:field_simps) + qed then guess l .. note l = conjunctD2[OF this] + show ?thesis apply(rule_tac x="min k l" in exI) apply safe unfolding min_less_iff_conj apply(rule,(rule l k)+) + proof- fix c assume as:"a \ c" "{a..c} \ {a..b}" "{a..c} \ ball a (min k l)" + note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_interval] + have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \ norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" by(rule norm_triangle_ineq4) + also have "... \ e * (b - a) / 8 + e * (b - a) / 8" + proof(rule add_mono) case goal1 have "\c - a\ \ \l\" using as' by auto + thus ?case apply-apply(rule order_trans[OF _ l(2)]) unfolding norm_scaleR apply(rule mult_right_mono) by auto + next case goal2 show ?case apply(rule less_imp_le) apply(cases "a = c") defer + apply(rule k(2)[unfolded vector_dist_norm]) using as' e ab by(auto simp add:field_simps) + qed finally show "norm (content {vec1 a..vec1 c} *\<^sub>R f' a - (f c - f a)) \ e * (b - a) / 4" unfolding content_1'[OF as(1)] by auto + qed qed then guess da .. note da=conjunctD2[OF this,rule_format] + + have "\db>0. \c\b. {c..b} \ {a..b} \ {c..b} \ ball b db \ norm(content {vec1 c..vec1 b} *\<^sub>R f' b - (f b - f c)) \ (e * (b - a)) / 4" + proof- have "b\{a..b}" using ab by auto + note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] + note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" using e ab by(auto simp add:field_simps) + from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format] + have "\l. 0 < l \ norm(l *\<^sub>R f' b) \ (e * (b - a)) / 8" + proof(cases "f' b = 0") case True + thus ?thesis apply(rule_tac x=1 in exI) using ab e by(auto intro!:mult_nonneg_nonneg) + next case False thus ?thesis + apply(rule_tac x="(e * (b - a)) / 8 / norm (f' b)" in exI) + using ab e by(auto simp add:field_simps) + qed then guess l .. note l = conjunctD2[OF this] + show ?thesis apply(rule_tac x="min k l" in exI) apply safe unfolding min_less_iff_conj apply(rule,(rule l k)+) + proof- fix c assume as:"c \ b" "{c..b} \ {a..b}" "{c..b} \ ball b (min k l)" + note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_interval] + have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \ norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" by(rule norm_triangle_ineq4) + also have "... \ e * (b - a) / 8 + e * (b - a) / 8" + proof(rule add_mono) case goal1 have "\c - b\ \ \l\" using as' by auto + thus ?case apply-apply(rule order_trans[OF _ l(2)]) unfolding norm_scaleR apply(rule mult_right_mono) by auto + next case goal2 show ?case apply(rule less_imp_le) apply(cases "b = c") defer apply(subst norm_minus_commute) + apply(rule k(2)[unfolded vector_dist_norm]) using as' e ab by(auto simp add:field_simps) + qed finally show "norm (content {vec1 c..vec1 b} *\<^sub>R f' b - (f b - f c)) \ e * (b - a) / 4" unfolding content_1'[OF as(1)] by auto + qed qed then guess db .. note db=conjunctD2[OF this,rule_format] + + let ?d = "(\x. ball x (if x=vec1 a then da else if x=vec b then db else d (dest_vec1 x)))" + show "?P e" apply(rule_tac x="?d" in exI) + proof safe case goal1 show ?case apply(rule gauge_ball_dependent) using ab db(1) da(1) d(1) by auto + next case goal2 note as=this let ?A = "{t. fst t \ {vec1 a, vec1 b}}" note p = tagged_division_ofD[OF goal2(1)] + have pA:"p = (p \ ?A) \ (p - ?A)" "finite (p \ ?A)" "finite (p - ?A)" "(p \ ?A) \ (p - ?A) = {}" using goal2 by auto + note * = additive_tagged_division_1'[OF assms(1) goal2(1), THEN sym] + have **:"\n1 s1 n2 s2::real. n2 \ s2 / 2 \ n1 - s1 \ s2 / 2 \ n1 + n2 \ s1 + s2" by arith + show ?case unfolding content_1'[OF assms(1)] and *[of "\x. x"] *[of f] setsum_subtractf[THEN sym] split_minus + unfolding setsum_right_distrib apply(subst(2) pA,subst pA) unfolding setsum_Un_disjoint[OF pA(2-)] + proof(rule norm_triangle_le,rule **) + case goal1 show ?case apply(rule order_trans,rule setsum_norm_le) apply(rule pA) defer apply(subst divide.setsum) + proof(rule order_refl,safe,unfold not_le o_def split_conv fst_conv,rule ccontr) fix x k assume as:"(x,k) \ p" + "e * (dest_vec1 (interval_upperbound k) - dest_vec1 (interval_lowerbound k)) / 2 + < norm (content k *\<^sub>R f' (dest_vec1 x) - (f (dest_vec1 (interval_upperbound k)) - f (dest_vec1 (interval_lowerbound k))))" + from p(4)[OF this(1)] guess u v apply-by(erule exE)+ note k=this + hence "\i. u$i \ v$i" and uv:"{u,v}\{u..v}" using p(2)[OF as(1)] by auto note this(1) this(1)[unfolded forall_1] + note result = as(2)[unfolded k interval_bounds[OF this(1)] content_1[OF this(2)]] + + assume as':"x \ vec1 a" "x \ vec1 b" hence "x$1 \ {a<..R f' (x$1) - (f (v$1) - f (u$1))) = + norm ((f (u$1) - f (x$1) - (u$1 - x$1) *\<^sub>R f' (x$1)) - (f (v$1) - f (x$1) - (v$1 - x$1) *\<^sub>R f' (x$1)))" + apply(rule arg_cong[of _ _ norm]) unfolding scaleR_left.diff by auto + also have "... \ e / 2 * norm (u$1 - x$1) + e / 2 * norm (v$1 - x$1)" apply(rule norm_triangle_le_sub) + apply(rule add_mono) apply(rule_tac[!] *) using fineD[OF goal2(2) as(1)] as' unfolding k subset_eq + apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE) using uv by(auto simp add:dist_real) + also have "... \ e / 2 * norm (v$1 - u$1)" using p(2)[OF as(1)] unfolding k by(auto simp add:field_simps) + finally have "e * (dest_vec1 v - dest_vec1 u) / 2 < e * (dest_vec1 v - dest_vec1 u) / 2" + apply- apply(rule less_le_trans[OF result]) using uv by auto thus False by auto qed + + next have *:"\x s1 s2::real. 0 \ s1 \ x \ (s1 + s2) / 2 \ x - s1 \ s2 / 2" by auto + case goal2 show ?case apply(rule *) apply(rule setsum_nonneg) apply(rule,unfold split_paired_all split_conv) + defer unfolding setsum_Un_disjoint[OF pA(2-),THEN sym] pA(1)[THEN sym] unfolding setsum_right_distrib[THEN sym] + apply(subst additive_tagged_division_1[OF _ as(1)]) unfolding vec1_dest_vec1 apply(rule assms) + proof- fix x k assume "(x,k) \ p \ {t. fst t \ {vec1 a, vec1 b}}" note xk=IntD1[OF this] + from p(4)[OF this] guess u v apply-by(erule exE)+ note uv=this + with p(2)[OF xk] have "{u..v} \ {}" by auto + thus "0 \ e * ((interval_upperbound k)$1 - (interval_lowerbound k)$1)" + unfolding uv using e by(auto simp add:field_simps) + next have *:"\s f t e. setsum f s = setsum f t \ norm(setsum f t) \ e \ norm(setsum f s) \ e" by auto + show "norm (\(x, k)\p \ ?A. content k *\<^sub>R (f' \ dest_vec1) x - + (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) \ e * (b - a) / 2" + apply(rule *[where t="p \ {t. fst t \ {vec1 a, vec1 b} \ content(snd t) \ 0}"]) + apply(rule setsum_mono_zero_right[OF pA(2)]) defer apply(rule) unfolding split_paired_all split_conv o_def + proof- fix x k assume "(x,k) \ p \ {t. fst t \ {vec1 a, vec1 b}} - p \ {t. fst t \ {vec1 a, vec1 b} \ content (snd t) \ 0}" + hence xk:"(x,k)\p" "content k = 0" by auto from p(4)[OF xk(1)] guess u v apply-by(erule exE)+ note uv=this + have "k\{}" using p(2)[OF xk(1)] by auto hence *:"u = v" using xk unfolding uv content_eq_0_1 interval_eq_empty by auto + thus "content k *\<^sub>R (f' (x$1)) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1)) = 0" using xk unfolding uv by auto + next have *:"p \ {t. fst t \ {vec1 a, vec1 b} \ content(snd t) \ 0} = + {t. t\p \ fst t = vec1 a \ content(snd t) \ 0} \ {t. t\p \ fst t = vec1 b \ content(snd t) \ 0}" by blast + have **:"\s f. \e::real. (\x y. x \ s \ y \ s \ x = y) \ (\x. x \ s \ norm(f x) \ e) \ e>0 \ norm(setsum f s) \ e" + proof(case_tac "s={}") case goal2 then obtain x where "x\s" by auto hence *:"s = {x}" using goal2(1) by auto + thus ?case using `x\s` goal2(2) by auto + qed auto + case goal2 show ?case apply(subst *, subst setsum_Un_disjoint) prefer 4 apply(rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"]) + apply(rule norm_triangle_le,rule add_mono) apply(rule_tac[1-2] **) + proof- let ?B = "\x. {t \ p. fst t = vec1 x \ content (snd t) \ 0}" + have pa:"\k. (vec1 a, k) \ p \ \v. k = {vec1 a .. v} \ vec1 a \ v" + proof- case goal1 guess u v using p(4)[OF goal1] apply-by(erule exE)+ note uv=this + have *:"u \ v" using p(2)[OF goal1] unfolding uv by auto + have u:"u = vec1 a" proof(rule ccontr) have "u \ {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto + have "u \ vec1 a" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "u\vec1 a" ultimately + have "u > vec1 a" unfolding Cart_simps by auto + thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:Cart_simps) + qed thus ?case apply(rule_tac x=v in exI) unfolding uv using * by auto + qed + have pb:"\k. (vec1 b, k) \ p \ \v. k = {v .. vec1 b} \ vec1 b \ v" + proof- case goal1 guess u v using p(4)[OF goal1] apply-by(erule exE)+ note uv=this + have *:"u \ v" using p(2)[OF goal1] unfolding uv by auto + have u:"v = vec1 b" proof(rule ccontr) have "u \ {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto + have "v \ vec1 b" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "v\vec1 b" ultimately + have "v < vec1 b" unfolding Cart_simps by auto + thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:Cart_simps) + qed thus ?case apply(rule_tac x=u in exI) unfolding uv using * by auto + qed + + show "\x y. x \ ?B a \ y \ ?B a \ x = y" apply(rule,rule,rule,unfold split_paired_all) + unfolding mem_Collect_eq fst_conv snd_conv apply safe + proof- fix x k k' assume k:"(vec1 a, k) \ p" "(vec1 a, k') \ p" "content k \ 0" "content k' \ 0" + guess v using pa[OF k(1)] .. note v = conjunctD2[OF this] + guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "vec1 (min (v$1) (v'$1))" + have "{vec1 a <..< ?v} \ k \ k'" unfolding v v' by(auto simp add:Cart_simps) note subset_interior[OF this,unfolded interior_inter] + moreover have "vec1 ((a + ?v$1)/2) \ {vec1 a <..< ?v}" using k(3-) unfolding v v' content_eq_0_1 not_le by(auto simp add:Cart_simps) + ultimately have "vec1 ((a + ?v$1)/2) \ interior k \ interior k'" unfolding interior_open[OF open_interval] by auto + hence *:"k = k'" apply- apply(rule ccontr) using p(5)[OF k(1-2)] by auto + { assume "x\k" thus "x\k'" unfolding * . } { assume "x\k'" thus "x\k" unfolding * . } + qed + show "\x y. x \ ?B b \ y \ ?B b \ x = y" apply(rule,rule,rule,unfold split_paired_all) + unfolding mem_Collect_eq fst_conv snd_conv apply safe + proof- fix x k k' assume k:"(vec1 b, k) \ p" "(vec1 b, k') \ p" "content k \ 0" "content k' \ 0" + guess v using pb[OF k(1)] .. note v = conjunctD2[OF this] + guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "vec1 (max (v$1) (v'$1))" + have "{?v <..< vec1 b} \ k \ k'" unfolding v v' by(auto simp add:Cart_simps) note subset_interior[OF this,unfolded interior_inter] + moreover have "vec1 ((b + ?v$1)/2) \ {?v <..< vec1 b}" using k(3-) unfolding v v' content_eq_0_1 not_le by(auto simp add:Cart_simps) + ultimately have "vec1 ((b + ?v$1)/2) \ interior k \ interior k'" unfolding interior_open[OF open_interval] by auto + hence *:"k = k'" apply- apply(rule ccontr) using p(5)[OF k(1-2)] by auto + { assume "x\k" thus "x\k'" unfolding * . } { assume "x\k'" thus "x\k" unfolding * . } + qed + + let ?a = a and ?b = b (* a is something else while proofing the next theorem. *) + show "\x. x \ ?B a \ norm ((\(x, k). content k *\<^sub>R f' (x$1) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) x) + \ e * (b - a) / 4" apply safe unfolding fst_conv snd_conv apply safe unfolding vec1_dest_vec1 + proof- case goal1 guess v using pa[OF goal1(1)] .. note v = conjunctD2[OF this] + have "vec1 ?a\{vec1 ?a..v}" using v(2) by auto hence "dest_vec1 v \ ?b" using p(3)[OF goal1(1)] unfolding subset_eq v by auto + moreover have "{?a..dest_vec1 v} \ ball ?a da" using fineD[OF as(2) goal1(1)] + apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe apply(erule_tac x="vec1 x" in ballE) + by(auto simp add:Cart_simps subset_eq dist_real v dist_real_def) ultimately + show ?case unfolding v unfolding interval_bounds[OF v(2)[unfolded v vector_le_def]] vec1_dest_vec1 apply- + apply(rule da(2)[of "v$1",unfolded vec1_dest_vec1]) + using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0_1 by auto + qed + show "\x. x \ ?B b \ norm ((\(x, k). content k *\<^sub>R f' (x$1) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) x) + \ e * (b - a) / 4" apply safe unfolding fst_conv snd_conv apply safe unfolding vec1_dest_vec1 + proof- case goal1 guess v using pb[OF goal1(1)] .. note v = conjunctD2[OF this] + have "vec1 ?b\{v..vec1 ?b}" using v(2) by auto hence "dest_vec1 v \ ?a" using p(3)[OF goal1(1)] unfolding subset_eq v by auto + moreover have "{dest_vec1 v..?b} \ ball ?b db" using fineD[OF as(2) goal1(1)] + apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe apply(erule_tac x="vec1 x" in ballE) using ab + by(auto simp add:Cart_simps subset_eq dist_real v dist_real_def) ultimately + show ?case unfolding v unfolding interval_bounds[OF v(2)[unfolded v vector_le_def]] vec1_dest_vec1 apply- + apply(rule db(2)[of "v$1",unfolded vec1_dest_vec1]) + using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0_1 by auto + qed + qed(insert p(1) ab e, auto simp add:field_simps) qed auto qed qed qed qed + +subsection {* Stronger form with finite number of exceptional points. *} + +lemma fundamental_theorem_of_calculus_interior_strong: fixes f::"real \ 'a::banach" + assumes"finite s" "a \ b" "continuous_on {a..b} f" + "\x\{a<..{a<.. vec1 c" "vec1 c \ vec1 b" by auto + thus ?thesis apply(subst *) apply(rule has_integral_combine) apply assumption+ + apply(rule_tac[!] Suc(1)[OF cs(3)]) using Suc(3) unfolding cs + proof- show "continuous_on {a..c} f" "continuous_on {c..b} f" + apply(rule_tac[!] continuous_on_subset[OF Suc(5)]) using True by auto + let ?P = "\i j. \x\{i<.. 'a::banach" + assumes "finite s" "a \ b" "continuous_on {a..b} f" + "\x\{a..b} - s. (f has_vector_derivative f'(x)) (at x)" + shows "((f' o dest_vec1) has_integral (f(b) - f(a))) {vec1 a..vec1 b}" + apply(rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f']) + using assms(4) by auto + +end diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy --- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,5 @@ theory Multivariate_Analysis -imports Determinants Derivative +imports Determinants Integration_MV begin end diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Wed Feb 17 21:13:40 2010 +0100 @@ -3179,6 +3179,23 @@ (\e>0. \d>0. \x\s. \ x'\s. dist x' x < d --> dist (f x') (f x) < e)" + +text{* Lifting and dropping *} + +lemma continuous_on_o_dest_vec1: fixes f::"real \ 'a::real_normed_vector" + assumes "continuous_on {a..b::real} f" shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)" + using assms unfolding continuous_on_def apply safe + apply(erule_tac x="x$1" in ballE,erule_tac x=e in allE) apply safe + apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real + apply(erule_tac x="dest_vec1 x'" in ballE) by(auto simp add:vector_le_def) + +lemma continuous_on_o_vec1: fixes f::"real^1 \ 'a::real_normed_vector" + assumes "continuous_on {a..b} f" shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)" + using assms unfolding continuous_on_def apply safe + apply(erule_tac x="vec x" in ballE,erule_tac x=e in allE) apply safe + apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real + apply(erule_tac x="vec1 x'" in ballE) by(auto simp add:vector_le_def) + text{* Some simple consequential lemmas. *} lemma uniformly_continuous_imp_continuous: @@ -3708,6 +3725,17 @@ shows "\x. continuous (at x) f \ closed s \ closed (f -` s)" unfolding vimage_def by (rule continuous_closed_preimage_univ) +lemma interior_image_subset: fixes f::"_::metric_space \ _::metric_space" + assumes "\x. continuous (at x) f" "inj f" + shows "interior (f ` s) \ f ` (interior s)" + apply rule unfolding interior_def mem_Collect_eq image_iff apply safe +proof- fix x T assume as:"open T" "x \ T" "T \ f ` s" + hence "x \ f ` s" by auto then guess y unfolding image_iff .. note y=this + thus "\xa\{x. \T. open T \ x \ T \ T \ s}. x = f xa" apply(rule_tac x=y in bexI) using assms as + apply safe apply(rule_tac x="{x. f x \ T}" in exI) apply(safe,rule continuous_open_preimage_univ) + proof- fix x assume "f x \ T" hence "f x \ f ` s" using as by auto + thus "x \ s" unfolding inj_image_mem_iff[OF assms(2)] . qed auto qed + text{* Equality of continuous functions on closure and related results. *} lemma continuous_closed_in_preimage_constant: diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/SetInterval.thy Wed Feb 17 21:13:40 2010 +0100 @@ -970,6 +970,27 @@ finally show ?thesis . qed +lemma setsum_le_included: + fixes f :: "'a \ 'b::{comm_monoid_add,ordered_ab_semigroup_add_imp_le}" + assumes "finite s" "finite t" + and "\y\t. 0 \ g y" "(\x\s. \y\t. i y = x \ f x \ g y)" + shows "setsum f s \ setsum g t" +proof - + have "setsum f s \ setsum (\y. setsum g {x. x\t \ i x = y}) s" + proof (rule setsum_mono) + fix y assume "y \ s" + with assms obtain z where z: "z \ t" "y = i z" "f y \ g z" by auto + with assms show "f y \ setsum g {x \ t. i x = y}" (is "?A y \ ?B y") + using order_trans[of "?A (i z)" "setsum g {z}" "?B (i z)", intro] + by (auto intro!: setsum_mono2) + qed + also have "... \ setsum (\y. setsum g {x. x\t \ i x = y}) (i ` t)" + using assms(2-4) by (auto intro!: setsum_mono2 setsum_nonneg) + also have "... \ setsum g t" + using assms by (auto simp: setsum_image_gen[symmetric]) + finally show ?thesis . +qed + lemma setsum_multicount_gen: assumes "finite s" "finite t" "\j\t. (card {i\s. R i j} = k j)" shows "setsum (\i. (card {j\t. R i j})) s = setsum k t" (is "?l = ?r") diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Tools/primrec.ML --- a/src/HOL/Tools/primrec.ML Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Tools/primrec.ML Wed Feb 17 21:13:40 2010 +0100 @@ -8,16 +8,16 @@ signature PRIMREC = sig val add_primrec: (binding * typ option * mixfix) list -> - (Attrib.binding * term) list -> local_theory -> thm list * local_theory + (Attrib.binding * term) list -> local_theory -> (term list * thm list) * local_theory val add_primrec_cmd: (binding * string option * mixfix) list -> - (Attrib.binding * string) list -> local_theory -> thm list * local_theory + (Attrib.binding * string) list -> local_theory -> (term list * thm list) * local_theory val add_primrec_global: (binding * typ option * mixfix) list -> - (Attrib.binding * term) list -> theory -> thm list * theory + (Attrib.binding * term) list -> theory -> (term list * thm list) * theory val add_primrec_overloaded: (string * (string * typ) * bool) list -> (binding * typ option * mixfix) list -> - (Attrib.binding * term) list -> theory -> thm list * theory + (Attrib.binding * term) list -> theory -> (term list * thm list) * theory val add_primrec_simple: ((binding * typ) * mixfix) list -> term list -> - local_theory -> (string * thm list list) * local_theory + local_theory -> (string * (term list * thm list)) * local_theory end; structure Primrec : PRIMREC = @@ -244,7 +244,7 @@ val rewrites = rec_rewrites' @ map (snd o snd) defs; fun tac _ = EVERY [rewrite_goals_tac rewrites, rtac refl 1]; val _ = message ("Proving equations for primrec function(s) " ^ commas_quote names); - in map (fn eq => [Goal.prove lthy frees [] eq tac]) eqs end; + in map (fn eq => Goal.prove lthy frees [] eq tac) eqs end; in ((prefix, (fs, defs)), prove) end handle PrimrecError (msg, some_eqn) => error ("Primrec definition error:\n" ^ msg ^ (case some_eqn @@ -260,7 +260,7 @@ in lthy |> fold_map Local_Theory.define defs - |-> (fn defs => `(fn lthy => (prefix, prove lthy defs))) + |-> (fn defs => `(fn lthy => (prefix, (map fst defs, prove lthy defs)))) end; local @@ -285,10 +285,10 @@ in lthy |> add_primrec_simple fixes (map snd spec) - |-> (fn (prefix, simps) => fold_map Local_Theory.note (attr_bindings prefix ~~ simps) + |-> (fn (prefix, (ts, simps)) => fold_map Local_Theory.note (attr_bindings prefix ~~ map single simps) #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps') + #>> (fn (_, simps'') => (ts, simps'')) ##> (Spec_Rules.add Spec_Rules.Equational (specs_of simps')))) - |>> snd end; in @@ -301,16 +301,16 @@ fun add_primrec_global fixes specs thy = let val lthy = Theory_Target.init NONE thy; - val (simps, lthy') = add_primrec fixes specs lthy; + val ((ts, simps), lthy') = add_primrec fixes specs lthy; val simps' = ProofContext.export lthy' lthy simps; - in (simps', Local_Theory.exit_global lthy') end; + in ((ts, simps'), Local_Theory.exit_global lthy') end; fun add_primrec_overloaded ops fixes specs thy = let val lthy = Theory_Target.overloading ops thy; - val (simps, lthy') = add_primrec fixes specs lthy; + val ((ts, simps), lthy') = add_primrec fixes specs lthy; val simps' = ProofContext.export lthy' lthy simps; - in (simps', Local_Theory.exit_global lthy') end; + in ((ts, simps'), Local_Theory.exit_global lthy') end; (* outer syntax *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOL/Tools/quickcheck_generators.ML --- a/src/HOL/Tools/quickcheck_generators.ML Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOL/Tools/quickcheck_generators.ML Wed Feb 17 21:13:40 2010 +0100 @@ -139,7 +139,7 @@ val eqs0 = [subst_v @{term "0::code_numeral"} eq, subst_v (@{term "Suc_code_numeral"} $ t_k) eq]; val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0; - val ((_, eqs2), lthy') = Primrec.add_primrec_simple + val ((_, (_, eqs2)), lthy') = Primrec.add_primrec_simple [((Binding.conceal (Binding.name random_aux), T), NoSyn)] eqs1 lthy; val cT_random_aux = inst pt_random_aux; val cT_rhs = inst pt_rhs; @@ -148,7 +148,7 @@ [(cT_random_aux, cert t_random_aux), (cT_rhs, cert t_rhs)]); val tac = ALLGOALS (rtac rule) THEN ALLGOALS (simp_tac rew_ss) - THEN (ALLGOALS (ProofContext.fact_tac (flat eqs2))) + THEN (ALLGOALS (ProofContext.fact_tac eqs2)) val simp = Skip_Proof.prove lthy' [v] [] eq (K tac); in (simp, lthy') end; diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/Cfun.thy --- a/src/HOLCF/Cfun.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/Cfun.thy Wed Feb 17 21:13:40 2010 +0100 @@ -521,7 +521,7 @@ text {* results about strictify *} lemma cont_strictify1: "cont (\f. if x = \ then \ else f\x)" -by (simp add: cont_if) +by simp lemma monofun_strictify2: "monofun (\x. if x = \ then \ else f\x)" apply (rule monofunI) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/Deflation.thy --- a/src/HOLCF/Deflation.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/Deflation.thy Wed Feb 17 21:13:40 2010 +0100 @@ -316,16 +316,14 @@ subsection {* Uniqueness of ep-pairs *} lemma ep_pair_unique_e_lemma: - assumes "ep_pair e1 p" and "ep_pair e2 p" + assumes 1: "ep_pair e1 p" and 2: "ep_pair e2 p" shows "e1 \ e2" proof (rule below_cfun_ext) - interpret e1: ep_pair e1 p by fact - interpret e2: ep_pair e2 p by fact fix x have "e1\(p\(e2\x)) \ e2\x" - by (rule e1.e_p_below) + by (rule ep_pair.e_p_below [OF 1]) thus "e1\x \ e2\x" - by (simp only: e2.e_inverse) + by (simp only: ep_pair.e_inverse [OF 2]) qed lemma ep_pair_unique_e: @@ -333,18 +331,16 @@ by (fast intro: below_antisym elim: ep_pair_unique_e_lemma) lemma ep_pair_unique_p_lemma: - assumes "ep_pair e p1" and "ep_pair e p2" + assumes 1: "ep_pair e p1" and 2: "ep_pair e p2" shows "p1 \ p2" proof (rule below_cfun_ext) - interpret p1: ep_pair e p1 by fact - interpret p2: ep_pair e p2 by fact fix x have "e\(p1\x) \ x" - by (rule p1.e_p_below) + by (rule ep_pair.e_p_below [OF 1]) hence "p2\(e\(p1\x)) \ p2\x" by (rule monofun_cfun_arg) thus "p1\x \ p2\x" - by (simp only: p2.e_inverse) + by (simp only: ep_pair.e_inverse [OF 2]) qed lemma ep_pair_unique_p: diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/FOCUS/Buffer_adm.thy --- a/src/HOLCF/FOCUS/Buffer_adm.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/FOCUS/Buffer_adm.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/FOCUS/Buffer_adm.thy - ID: $Id$ Author: David von Oheimb, TU Muenchen *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/FOCUS/FOCUS.thy --- a/src/HOLCF/FOCUS/FOCUS.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/FOCUS/FOCUS.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/FOCUS/FOCUS.thy - ID: $Id$ Author: David von Oheimb, TU Muenchen *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/FOCUS/Stream_adm.thy --- a/src/HOLCF/FOCUS/Stream_adm.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/FOCUS/Stream_adm.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/ex/Stream_adm.thy - ID: $Id$ Author: David von Oheimb, TU Muenchen *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IMP/Denotational.thy --- a/src/HOLCF/IMP/Denotational.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IMP/Denotational.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IMP/Denotational.thy - ID: $Id$ Author: Tobias Nipkow and Robert Sandner, TUM Copyright 1996 TUM *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IMP/HoareEx.thy --- a/src/HOLCF/IMP/HoareEx.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IMP/HoareEx.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IMP/HoareEx.thy - ID: $Id$ Author: Tobias Nipkow, TUM Copyright 1997 TUM *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IMP/README.html --- a/src/HOLCF/IMP/README.html Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IMP/README.html Wed Feb 17 21:13:40 2010 +0100 @@ -1,7 +1,5 @@ - - diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IMP/ROOT.ML --- a/src/HOLCF/IMP/ROOT.ML Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IMP/ROOT.ML Wed Feb 17 21:13:40 2010 +0100 @@ -1,3 +1,1 @@ -(* $Id$ *) - use_thys ["HoareEx"]; diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Abschannel.thy --- a/src/HOLCF/IOA/ABP/Abschannel.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Abschannel.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Abschannel.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Abschannel_finite.thy --- a/src/HOLCF/IOA/ABP/Abschannel_finite.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Abschannel_finite.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Abschannels.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Action.thy --- a/src/HOLCF/IOA/ABP/Action.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Action.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Action.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Check.ML --- a/src/HOLCF/IOA/ABP/Check.ML Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Check.ML Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Check.ML - ID: $Id$ Author: Olaf Mueller The Model Checker. diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Correctness.thy --- a/src/HOLCF/IOA/ABP/Correctness.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Correctness.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Correctness.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Env.thy --- a/src/HOLCF/IOA/ABP/Env.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Env.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Impl.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Impl.thy --- a/src/HOLCF/IOA/ABP/Impl.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Impl.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Impl.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Impl_finite.thy --- a/src/HOLCF/IOA/ABP/Impl_finite.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Impl_finite.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Impl.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Lemmas.thy --- a/src/HOLCF/IOA/ABP/Lemmas.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Lemmas.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Lemmas.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Packet.thy --- a/src/HOLCF/IOA/ABP/Packet.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Packet.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Packet.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Receiver.thy --- a/src/HOLCF/IOA/ABP/Receiver.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Receiver.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Receiver.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Sender.thy --- a/src/HOLCF/IOA/ABP/Sender.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Sender.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Sender.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/ABP/Spec.thy --- a/src/HOLCF/IOA/ABP/Spec.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/ABP/Spec.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Spec.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/Modelcheck/Cockpit.thy --- a/src/HOLCF/IOA/Modelcheck/Cockpit.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/Modelcheck/Cockpit.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,6 +1,3 @@ - -(* $Id$ *) - theory Cockpit imports MuIOAOracle begin diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/Modelcheck/ROOT.ML --- a/src/HOLCF/IOA/Modelcheck/ROOT.ML Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/Modelcheck/ROOT.ML Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/Modelcheck/ROOT.ML - ID: $Id$ Author: Olaf Mueller and Tobias Hamberger, TU Muenchen Modelchecker setup for I/O automata. diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Abschannel.thy --- a/src/HOLCF/IOA/NTP/Abschannel.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Abschannel.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Abschannel.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Action.thy --- a/src/HOLCF/IOA/NTP/Action.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Action.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Action.thy - ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Correctness.thy --- a/src/HOLCF/IOA/NTP/Correctness.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Correctness.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Correctness.thy - ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Impl.thy --- a/src/HOLCF/IOA/NTP/Impl.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Impl.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Impl.thy - ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Lemmas.thy --- a/src/HOLCF/IOA/NTP/Lemmas.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Lemmas.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Lemmas.thy - ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Multiset.thy --- a/src/HOLCF/IOA/NTP/Multiset.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Multiset.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Multiset.thy - ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Packet.thy --- a/src/HOLCF/IOA/NTP/Packet.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Packet.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Packet.thy - ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Receiver.thy --- a/src/HOLCF/IOA/NTP/Receiver.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Receiver.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Receiver.thy - ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Sender.thy --- a/src/HOLCF/IOA/NTP/Sender.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Sender.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Sender.thy - ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/NTP/Spec.thy --- a/src/HOLCF/IOA/NTP/Spec.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/NTP/Spec.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/NTP/Spec.thy - ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/README.html --- a/src/HOLCF/IOA/README.html Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/README.html Wed Feb 17 21:13:40 2010 +0100 @@ -1,7 +1,5 @@ - - diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/Storage/Action.thy --- a/src/HOLCF/IOA/Storage/Action.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/Storage/Action.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/ABP/Action.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/Storage/Correctness.thy --- a/src/HOLCF/IOA/Storage/Correctness.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/Storage/Correctness.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/example/Correctness.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/Storage/Impl.thy --- a/src/HOLCF/IOA/Storage/Impl.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/Storage/Impl.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/example/Spec.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/Storage/Spec.thy --- a/src/HOLCF/IOA/Storage/Spec.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/Storage/Spec.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/example/Spec.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/Asig.thy --- a/src/HOLCF/IOA/meta_theory/Asig.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/Asig.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IOA/meta_theory/Asig.thy - ID: $Id$ Author: Olaf Müller, Tobias Nipkow & Konrad Slind *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/Automata.thy --- a/src/HOLCF/IOA/meta_theory/Automata.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/Automata.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/Automata.thy - ID: $Id$ Author: Olaf Müller, Konrad Slind, Tobias Nipkow *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/CompoExecs.thy --- a/src/HOLCF/IOA/meta_theory/CompoExecs.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/CompoExecs.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/CompoExecs.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/CompoScheds.thy --- a/src/HOLCF/IOA/meta_theory/CompoScheds.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/CompoScheds.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/CompoScheds.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/CompoTraces.thy --- a/src/HOLCF/IOA/meta_theory/CompoTraces.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/CompoTraces.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/CompoTraces.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/Compositionality.thy --- a/src/HOLCF/IOA/meta_theory/Compositionality.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/Compositionality.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/Compositionality.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/Deadlock.thy --- a/src/HOLCF/IOA/meta_theory/Deadlock.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/Deadlock.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/Deadlock.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/IOA.thy --- a/src/HOLCF/IOA/meta_theory/IOA.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/IOA.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/IOA.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/LiveIOA.thy --- a/src/HOLCF/IOA/meta_theory/LiveIOA.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/LiveIOA.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/LiveIOA.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/Pred.thy --- a/src/HOLCF/IOA/meta_theory/Pred.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/Pred.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/Pred.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/RefCorrectness.thy --- a/src/HOLCF/IOA/meta_theory/RefCorrectness.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/RefCorrectness.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/RefCorrectness.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/RefMappings.thy --- a/src/HOLCF/IOA/meta_theory/RefMappings.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/RefMappings.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/RefMappings.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/Seq.thy --- a/src/HOLCF/IOA/meta_theory/Seq.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/Seq.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/Seq.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/ShortExecutions.thy --- a/src/HOLCF/IOA/meta_theory/ShortExecutions.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/ShortExecutions.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/ShortExecutions.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/SimCorrectness.thy --- a/src/HOLCF/IOA/meta_theory/SimCorrectness.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/SimCorrectness.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/SimCorrectness.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/Simulations.thy --- a/src/HOLCF/IOA/meta_theory/Simulations.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/Simulations.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/Simulations.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/TL.thy --- a/src/HOLCF/IOA/meta_theory/TL.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/TL.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/TLS.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IOA/meta_theory/TLS.thy --- a/src/HOLCF/IOA/meta_theory/TLS.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IOA/meta_theory/TLS.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/IOA/meta_theory/TLS.thy - ID: $Id$ Author: Olaf Müller *) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/IsaMakefile --- a/src/HOLCF/IsaMakefile Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/IsaMakefile Wed Feb 17 21:13:40 2010 +0100 @@ -95,6 +95,7 @@ ex/Dagstuhl.thy \ ex/Dnat.thy \ ex/Domain_ex.thy \ + ex/Domain_Proofs.thy \ ex/Fix2.thy \ ex/Fixrec_ex.thy \ ex/Focus_ex.thy \ @@ -103,6 +104,7 @@ ex/New_Domain.thy \ ex/Powerdomain_ex.thy \ ex/Stream.thy \ + ex/Strict_Fun.thy \ ex/ROOT.ML @$(ISABELLE_TOOL) usedir $(OUT)/HOLCF ex diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/README.html --- a/src/HOLCF/README.html Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/README.html Wed Feb 17 21:13:40 2010 +0100 @@ -1,7 +1,5 @@ - - diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/Dagstuhl.thy --- a/src/HOLCF/ex/Dagstuhl.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/Dagstuhl.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Dagstuhl imports Stream begin @@ -56,10 +54,10 @@ lemma wir_moel: "YS = YYS" apply (rule stream.take_lemmas) apply (induct_tac n) - apply (simp (no_asm) add: stream.rews) + apply (simp (no_asm)) apply (subst YS_def2) apply (subst YYS_def2) - apply (simp add: stream.rews) + apply simp apply (rule lemma5 [symmetric, THEN subst]) apply (rule refl) done diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/Dnat.thy --- a/src/HOLCF/ex/Dnat.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/Dnat.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/Dnat.thy - ID: $Id$ Author: Franz Regensburger Theory for the domain of natural numbers dnat = one ++ dnat @@ -34,18 +33,18 @@ lemma iterator1: "iterator $ UU $ f $ x = UU" apply (subst iterator_def2) - apply (simp add: dnat.rews) + apply simp done lemma iterator2: "iterator $ dzero $ f $ x = x" apply (subst iterator_def2) - apply (simp add: dnat.rews) + apply simp done lemma iterator3: "n ~= UU ==> iterator $ (dsucc $ n) $ f $ x = f $ (iterator $ n $ f $ x)" apply (rule trans) apply (subst iterator_def2) - apply (simp add: dnat.rews) + apply simp apply (rule refl) done @@ -59,13 +58,13 @@ apply (rule_tac x = y in dnat.casedist) apply simp apply simp - apply (simp add: dnat.dist_les) + apply simp apply (rule allI) apply (rule_tac x = y in dnat.casedist) apply (fast intro!: UU_I) apply (thin_tac "ALL y. d << y --> d = UU | d = y") - apply (simp add: dnat.dist_les) - apply (simp (no_asm_simp) add: dnat.rews dnat.injects dnat.inverts) + apply simp + apply (simp (no_asm_simp)) apply (drule_tac x="da" in spec) apply simp done diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/Fix2.thy --- a/src/HOLCF/ex/Fix2.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/Fix2.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/ex/Fix2.thy - ID: $Id$ Author: Franz Regensburger Show that fix is the unique least fixed-point operator. diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/Focus_ex.thy --- a/src/HOLCF/ex/Focus_ex.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/Focus_ex.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - (* Specification of the following loop back device diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/Hoare.thy --- a/src/HOLCF/ex/Hoare.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/Hoare.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/ex/hoare.thy - ID: $Id$ Author: Franz Regensburger Theory for an example by C.A.R. Hoare diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/Loop.thy --- a/src/HOLCF/ex/Loop.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/Loop.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/ex/Loop.thy - ID: $Id$ Author: Franz Regensburger *) @@ -115,7 +114,7 @@ and t = "g$ (iterate n$ (step$b$g) $x) " in ssubst) prefer 2 apply (assumption) apply (simp add: step_def2) -apply (simp del: iterate_Suc add: loop_lemma2) +apply (drule (1) loop_lemma2, simp) done lemma loop_lemma4 [rule_format]: diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/Powerdomain_ex.thy --- a/src/HOLCF/ex/Powerdomain_ex.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/Powerdomain_ex.thy Wed Feb 17 21:13:40 2010 +0100 @@ -14,8 +14,6 @@ domain ordering = LT | EQ | GT -declare ordering.rews [simp] - definition compare :: "int lift \ int lift \ ordering" where "compare = (FLIFT x y. if x < y then LT else if x = y then EQ else GT)" diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/ROOT.ML --- a/src/HOLCF/ex/ROOT.ML Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/ROOT.ML Wed Feb 17 21:13:40 2010 +0100 @@ -5,4 +5,5 @@ use_thys ["Dnat", "Stream", "Dagstuhl", "Focus_ex", "Fix2", "Hoare", "Loop", "Fixrec_ex", "Powerdomain_ex", "Domain_ex", "Domain_Proofs", + "Strict_Fun", "New_Domain"]; diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/Stream.thy --- a/src/HOLCF/ex/Stream.thy Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/Stream.thy Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/ex/Stream.thy - ID: $Id$ Author: Franz Regensburger, David von Oheimb, Borislav Gajanovic *) @@ -54,8 +53,6 @@ | \ \ s1)" -declare stream.rews [simp add] - (* ----------------------------------------------------------------------- *) (* theorems about scons *) (* ----------------------------------------------------------------------- *) @@ -149,13 +146,13 @@ apply (insert stream.reach [of s], erule subst) back apply (simp add: fix_def2 stream.take_def) apply (insert contlub_cfun_fun [of "%i. iterate i$stream_copy$UU" s,THEN sym]) -by (simp add: chain_iterate) +by simp lemma chain_stream_take: "chain (%i. stream_take i$s)" apply (rule chainI) apply (rule monofun_cfun_fun) apply (simp add: stream.take_def del: iterate_Suc) -by (rule chainE, simp add: chain_iterate) +by (rule chainE, simp) lemma stream_take_prefix [simp]: "stream_take n$s << s" apply (insert stream_reach2 [of s]) diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/Strict_Fun.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/ex/Strict_Fun.thy Wed Feb 17 21:13:40 2010 +0100 @@ -0,0 +1,239 @@ +(* Title: HOLCF/ex/Strict_Fun.thy + Author: Brian Huffman +*) + +header {* The Strict Function Type *} + +theory Strict_Fun +imports HOLCF +begin + +pcpodef (open) ('a, 'b) sfun (infixr "->!" 0) + = "{f :: 'a \ 'b. f\\ = \}" +by simp_all + +syntax (xsymbols) + sfun :: "type \ type \ type" (infixr "\!" 0) + +text {* TODO: Define nice syntax for abstraction, application. *} + +definition + sfun_abs :: "('a \ 'b) \ ('a \! 'b)" +where + "sfun_abs = (\ f. Abs_sfun (strictify\f))" + +definition + sfun_rep :: "('a \! 'b) \ 'a \ 'b" +where + "sfun_rep = (\ f. Rep_sfun f)" + +lemma sfun_rep_beta: "sfun_rep\f = Rep_sfun f" + unfolding sfun_rep_def by (simp add: cont_Rep_sfun) + +lemma sfun_rep_strict1 [simp]: "sfun_rep\\ = \" + unfolding sfun_rep_beta by (rule Rep_sfun_strict) + +lemma sfun_rep_strict2 [simp]: "sfun_rep\f\\ = \" + unfolding sfun_rep_beta by (rule Rep_sfun [simplified]) + +lemma strictify_cancel: "f\\ = \ \ strictify\f = f" + by (simp add: expand_cfun_eq strictify_conv_if) + +lemma sfun_abs_sfun_rep: "sfun_abs\(sfun_rep\f) = f" + unfolding sfun_abs_def sfun_rep_def + apply (simp add: cont_Abs_sfun cont_Rep_sfun) + apply (simp add: Rep_sfun_inject [symmetric] Abs_sfun_inverse) + apply (simp add: expand_cfun_eq strictify_conv_if) + apply (simp add: Rep_sfun [simplified]) + done + +lemma sfun_rep_sfun_abs [simp]: "sfun_rep\(sfun_abs\f) = strictify\f" + unfolding sfun_abs_def sfun_rep_def + apply (simp add: cont_Abs_sfun cont_Rep_sfun) + apply (simp add: Abs_sfun_inverse) + done + +lemma ep_pair_sfun: "ep_pair sfun_rep sfun_abs" +apply default +apply (rule sfun_abs_sfun_rep) +apply (simp add: expand_cfun_below strictify_conv_if) +done + +interpretation sfun: ep_pair sfun_rep sfun_abs + by (rule ep_pair_sfun) + +subsection {* Map functional for strict function space *} + +definition + sfun_map :: "('b \ 'a) \ ('c \ 'd) \ ('a \! 'c) \ ('b \! 'd)" +where + "sfun_map = (\ a b. sfun_abs oo cfun_map\a\b oo sfun_rep)" + +lemma sfun_map_ID: "sfun_map\ID\ID = ID" + unfolding sfun_map_def + by (simp add: cfun_map_ID expand_cfun_eq) + +lemma sfun_map_map: + assumes "f2\\ = \" and "g2\\ = \" shows + "sfun_map\f1\g1\(sfun_map\f2\g2\p) = + sfun_map\(\ x. f2\(f1\x))\(\ x. g1\(g2\x))\p" +unfolding sfun_map_def +by (simp add: expand_cfun_eq strictify_cancel assms cfun_map_map) + +lemma ep_pair_sfun_map: + assumes 1: "ep_pair e1 p1" + assumes 2: "ep_pair e2 p2" + shows "ep_pair (sfun_map\p1\e2) (sfun_map\e1\p2)" +proof + interpret e1p1: pcpo_ep_pair e1 p1 + unfolding pcpo_ep_pair_def by fact + interpret e2p2: pcpo_ep_pair e2 p2 + unfolding pcpo_ep_pair_def by fact + fix f show "sfun_map\e1\p2\(sfun_map\p1\e2\f) = f" + unfolding sfun_map_def + apply (simp add: sfun.e_eq_iff [symmetric] strictify_cancel) + apply (rule ep_pair.e_inverse) + apply (rule ep_pair_cfun_map [OF 1 2]) + done + fix g show "sfun_map\p1\e2\(sfun_map\e1\p2\g) \ g" + unfolding sfun_map_def + apply (simp add: sfun.e_below_iff [symmetric] strictify_cancel) + apply (rule ep_pair.e_p_below) + apply (rule ep_pair_cfun_map [OF 1 2]) + done +qed + +lemma deflation_sfun_map: + assumes 1: "deflation d1" + assumes 2: "deflation d2" + shows "deflation (sfun_map\d1\d2)" +apply (simp add: sfun_map_def) +apply (rule deflation.intro) +apply simp +apply (subst strictify_cancel) +apply (simp add: cfun_map_def deflation_strict 1 2) +apply (simp add: cfun_map_def deflation.idem 1 2) +apply (simp add: sfun.e_below_iff [symmetric]) +apply (subst strictify_cancel) +apply (simp add: cfun_map_def deflation_strict 1 2) +apply (rule deflation.below) +apply (rule deflation_cfun_map [OF 1 2]) +done + +lemma finite_deflation_sfun_map: + assumes 1: "finite_deflation d1" + assumes 2: "finite_deflation d2" + shows "finite_deflation (sfun_map\d1\d2)" +proof (intro finite_deflation.intro finite_deflation_axioms.intro) + interpret d1: finite_deflation d1 by fact + interpret d2: finite_deflation d2 by fact + have "deflation d1" and "deflation d2" by fact+ + thus "deflation (sfun_map\d1\d2)" by (rule deflation_sfun_map) + from 1 2 have "finite_deflation (cfun_map\d1\d2)" + by (rule finite_deflation_cfun_map) + then have "finite {f. cfun_map\d1\d2\f = f}" + by (rule finite_deflation.finite_fixes) + moreover have "inj (\f. sfun_rep\f)" + by (rule inj_onI, simp) + ultimately have "finite ((\f. sfun_rep\f) -` {f. cfun_map\d1\d2\f = f})" + by (rule finite_vimageI) + then show "finite {f. sfun_map\d1\d2\f = f}" + unfolding sfun_map_def sfun.e_eq_iff [symmetric] + by (simp add: strictify_cancel + deflation_strict `deflation d1` `deflation d2`) +qed + +subsection {* Strict function space is bifinite *} + +instantiation sfun :: (bifinite, bifinite) bifinite +begin + +definition + "approx = (\i. sfun_map\(approx i)\(approx i))" + +instance proof + show "chain (approx :: nat \ ('a \! 'b) \ ('a \! 'b))" + unfolding approx_sfun_def by simp +next + fix x :: "'a \! 'b" + show "(\i. approx i\x) = x" + unfolding approx_sfun_def + by (simp add: lub_distribs sfun_map_ID [unfolded ID_def]) +next + fix i :: nat and x :: "'a \! 'b" + show "approx i\(approx i\x) = approx i\x" + unfolding approx_sfun_def + by (intro deflation.idem deflation_sfun_map deflation_approx) +next + fix i :: nat + show "finite {x::'a \! 'b. approx i\x = x}" + unfolding approx_sfun_def + by (intro finite_deflation.finite_fixes + finite_deflation_sfun_map + finite_deflation_approx) +qed + +end + +subsection {* Strict function space is representable *} + +instantiation sfun :: (rep, rep) rep +begin + +definition + "emb = udom_emb oo sfun_map\prj\emb" + +definition + "prj = sfun_map\emb\prj oo udom_prj" + +instance +apply (default, unfold emb_sfun_def prj_sfun_def) +apply (rule ep_pair_comp) +apply (rule ep_pair_sfun_map) +apply (rule ep_pair_emb_prj) +apply (rule ep_pair_emb_prj) +apply (rule ep_pair_udom) +done + +end + +text {* + A deflation constructor lets us configure the domain package to work + with the strict function space type constructor. +*} + +definition + sfun_defl :: "TypeRep \ TypeRep \ TypeRep" +where + "sfun_defl = TypeRep_fun2 sfun_map" + +lemma cast_sfun_defl: + "cast\(sfun_defl\A\B) = udom_emb oo sfun_map\(cast\A)\(cast\B) oo udom_prj" +unfolding sfun_defl_def +apply (rule cast_TypeRep_fun2) +apply (erule (1) finite_deflation_sfun_map) +done + +lemma REP_sfun: "REP('a::rep \! 'b::rep) = sfun_defl\REP('a)\REP('b)" +apply (rule cast_eq_imp_eq, rule ext_cfun) +apply (simp add: cast_REP cast_sfun_defl) +apply (simp only: prj_sfun_def emb_sfun_def) +apply (simp add: sfun_map_def cfun_map_def strictify_cancel) +done + +lemma isodefl_sfun: + "isodefl d1 t1 \ isodefl d2 t2 \ + isodefl (sfun_map\d1\d2) (sfun_defl\t1\t2)" +apply (rule isodeflI) +apply (simp add: cast_sfun_defl cast_isodefl) +apply (simp add: emb_sfun_def prj_sfun_def) +apply (simp add: sfun_map_map deflation_strict [OF isodefl_imp_deflation]) +done + +setup {* + Domain_Isomorphism.add_type_constructor + (@{type_name "sfun"}, @{term sfun_defl}, @{const_name sfun_map}, + @{thm REP_sfun}, @{thm isodefl_sfun}, @{thm sfun_map_ID}) +*} + +end diff -r 69fa4c39dab2 -r a815c3f4eef2 src/HOLCF/ex/hoare.txt --- a/src/HOLCF/ex/hoare.txt Wed Feb 17 20:50:14 2010 +0100 +++ b/src/HOLCF/ex/hoare.txt Wed Feb 17 21:13:40 2010 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - Proves about loops and tail-recursive functions ===============================================