# HG changeset patch # User huffman # Date 1272298885 25200 # Node ID 06475a1547cbb1ac1386ac65199978dc6b37e5ce # Parent 1debc8e29f6a7f58f7e461fa1f46f2a1db7a93a6 fix lots of looping simp calls and other warnings diff -r 1debc8e29f6a -r 06475a1547cb src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Sun Apr 25 23:22:29 2010 -0700 +++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Apr 26 09:21:25 2010 -0700 @@ -22,8 +22,6 @@ imports Convex_Euclidean_Space begin -declare norm_scaleR[simp] - lemma brouwer_compactness_lemma: assumes "compact s" "continuous_on s f" "\ (\x\s. (f x = (0::real^'n)))" obtains d where "0 < d" "\x\s. d \ norm(f x)" proof(cases "s={}") case False @@ -131,7 +129,7 @@ lemma image_lemma_2: assumes "finite s" "finite t" "card s = card t" "f ` s \ t" "f ` s \ t" "b \ t" shows "(card {s'. \a\s. (s' = s - {a}) \ f ` s' = t - {b}} = 0) \ (card {s'. \a\s. (s' = s - {a}) \ f ` s' = t - {b}} = 2)" proof(cases "{a\s. f ` (s - {a}) = t - {b}} = {}") - case True thus ?thesis apply-apply(rule disjI1, rule image_lemma_0) using assms(1) by(auto simp add:card_0_eq) + case True thus ?thesis apply-apply(rule disjI1, rule image_lemma_0) using assms(1) by auto next let ?M = "{a\s. f ` (s - {a}) = t - {b}}" case False then obtain a where "a\?M" by auto hence a:"a\s" "f ` (s - {a}) = t - {b}" by auto have "f a \ t - {b}" using a and assms by auto @@ -1691,11 +1689,11 @@ path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \ path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \ path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2) - by(auto simp add: pathstart_join pathfinish_join path_image_join path_image_linepath path_join path_linepath) + by(auto simp add: path_image_join path_linepath) have abab: "{a..b} \ {?a..?b}" by(auto simp add:vector_le_def forall_2 vector_2) guess z apply(rule fashoda[of ?P1 ?P2 ?a ?b]) unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 proof- - show "path ?P1" "path ?P2" using assms by(auto simp add: pathstart_join pathfinish_join path_join) + show "path ?P1" "path ?P2" using assms by auto have "path_image ?P1 \ {?a .. ?b}" unfolding P1P2 path_image_linepath apply(rule Un_least)+ defer 3 apply(rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format]) unfolding mem_interval forall_2 vector_2 using ab startfin abab assms(3) diff -r 1debc8e29f6a -r 06475a1547cb src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Sun Apr 25 23:22:29 2010 -0700 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Apr 26 09:21:25 2010 -0700 @@ -24,7 +24,7 @@ lemma dest_vec1_simps[simp]: fixes a::"real^1" shows "a$1 = 0 \ a = 0" (*"a \ 1 \ dest_vec1 a \ 1" "0 \ a \ 0 \ dest_vec1 a"*) "a \ b \ dest_vec1 a \ dest_vec1 b" "dest_vec1 (1::real^1) = 1" - by(auto simp add:vector_component_simps forall_1 Cart_eq) + by(auto simp add: vector_le_def Cart_eq) lemma norm_not_0:"(x::real^'n)\0 \ norm x \ 0" by auto @@ -50,7 +50,7 @@ lemma mem_interval_1: fixes x :: "real^1" shows "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b)" "(x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" -by(simp_all add: Cart_eq vector_less_def vector_le_def forall_1) +by(simp_all add: Cart_eq vector_less_def vector_le_def) lemma image_smult_interval:"(\x. m *\<^sub>R (x::real^'n)) ` {a..b} = (if {a..b} = {} then {} else if 0 \ m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R a})" @@ -66,8 +66,7 @@ apply(rule_tac [!] allI)apply(rule_tac [!] impI) apply(rule_tac[2] x="vec1 x" in exI)apply(rule_tac[4] x="vec1 x" in exI) apply(rule_tac[6] x="vec1 x" in exI)apply(rule_tac[8] x="vec1 x" in exI) - by (auto simp add: vector_less_def vector_le_def forall_1 - vec1_dest_vec1[unfolded One_nat_def]) + by (auto simp add: vector_less_def vector_le_def) lemma dest_vec1_setsum: assumes "finite S" shows " dest_vec1 (setsum f S) = setsum (\x. dest_vec1 (f x)) S" @@ -90,7 +89,7 @@ using one_le_card_finite by auto lemma real_dimindex_ge_1:"real (CARD('n::finite)) \ 1" - by(metis dimindex_ge_1 linorder_not_less real_eq_of_nat real_le_trans real_of_nat_1 real_of_nat_le_iff) + by(metis dimindex_ge_1 real_eq_of_nat real_of_nat_1 real_of_nat_le_iff) lemma real_dimindex_gt_0:"real (CARD('n::finite)) > 0" apply(rule less_le_trans[OF _ real_dimindex_ge_1]) by auto @@ -480,8 +479,8 @@ next fix t u assume asm:"\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s" "finite (t::'a set)" (*"finite t" "t \ s" "\x\t. (0::real) \ u x" "setsum u t = 1"*) - from this(2) have "\u. t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ (\x\t. u x *\<^sub>R x) \ s" apply(induct_tac t rule:finite_induct) - prefer 3 apply (rule,rule) apply(erule conjE)+ proof- + from this(2) have "\u. t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ (\x\t. u x *\<^sub>R x) \ s" apply(induct t rule:finite_induct) + prefer 2 apply (rule,rule) apply(erule conjE)+ proof- fix x f u assume ind:"\u. f \ s \ (\x\f. 0 \ u x) \ setsum u f = 1 \ (\x\f. u x *\<^sub>R x) \ s" assume as:"finite f" "x \ f" "insert x f \ s" "\x\insert x f. 0 \ u x" "setsum u (insert x f) = (1::real)" show "(\x\insert x f. u x *\<^sub>R x) \ s" proof(cases "u x = 1") @@ -776,7 +775,7 @@ lemma convex_cball: fixes x :: "'a::real_normed_vector" shows "convex(cball x e)" -proof(auto simp add: convex_def Ball_def mem_cball) +proof(auto simp add: convex_def Ball_def) fix y z assume yz:"dist x y \ e" "dist x z \ e" fix u v ::real assume uv:" 0 \ u" "0 \ v" "u + v = 1" have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ u * dist x y + v * dist x z" using uv yz @@ -1144,7 +1143,7 @@ hence "x + y \ s" using `?lhs`[unfolded convex_def, THEN conjunct1] apply(erule_tac x="2*\<^sub>R x" in ballE) apply(erule_tac x="2*\<^sub>R y" in ballE) apply(erule_tac x="1/2" in allE) apply simp apply(erule_tac x="1/2" in allE) by auto } - thus ?thesis unfolding convex_def cone_def by auto + thus ?thesis unfolding convex_def cone_def by blast qed lemma affine_dependent_biggerset: fixes s::"(real^'n) set" @@ -1259,7 +1258,7 @@ fixes s :: "'a::real_normed_vector set" assumes "open s" shows "open(convex hull s)" - unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(10) + unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(10) proof(rule, rule) fix a assume "\sa u. finite sa \ sa \ s \ (\x\sa. 0 \ u x) \ setsum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = a" then obtain t u where obt:"finite t" "t\s" "\x\t. 0 \ u x" "setsum u t = 1" "(\v\t. u v *\<^sub>R v) = a" by auto @@ -1279,7 +1278,7 @@ hence "Min i \ b x" unfolding i_def apply(rule_tac Min_le) using obt(1) by auto hence "x + (y - a) \ cball x (b x)" using y unfolding mem_cball dist_norm by auto moreover from `x\t` have "x\s" using obt(2) by auto - ultimately have "x + (y - a) \ s" using y and b[THEN bspec[where x=x]] unfolding subset_eq by auto } + ultimately have "x + (y - a) \ s" using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast } moreover have *:"inj_on (\v. v + (y - a)) t" unfolding inj_on_def by auto have "(\v\(\v. v + (y - a)) ` t. u (v - (y - a))) = 1" @@ -1349,7 +1348,7 @@ show ?thesis unfolding caratheodory[of s] proof(induct ("CARD('n) + 1")) have *:"{x.\sa. finite sa \ sa \ s \ card sa \ 0 \ x \ convex hull sa} = {}" - using compact_empty by (auto simp add: convex_hull_empty) + using compact_empty by auto case 0 thus ?case unfolding * by simp next case (Suc n) @@ -1359,11 +1358,11 @@ fix x assume "\t. finite t \ t \ s \ card t \ Suc n \ x \ convex hull t" then obtain t where t:"finite t" "t \ s" "card t \ Suc n" "x \ convex hull t" by auto show "x\s" proof(cases "card t = 0") - case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by(simp add: convex_hull_empty) + case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by simp next case False hence "card t = Suc 0" using t(3) `n=0` by auto then obtain a where "t = {a}" unfolding card_Suc_eq by auto - thus ?thesis using t(2,4) by (simp add: convex_hull_singleton) + thus ?thesis using t(2,4) by simp qed next fix x assume "x\s" @@ -1398,7 +1397,7 @@ show ?P proof(cases "u={}") case True hence "x=a" using t(4)[unfolded au] by auto show ?P unfolding `x=a` apply(rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI) - using t and `n\0` unfolding au by(auto intro!: exI[where x="{a}"] simp add: convex_hull_singleton) + using t and `n\0` unfolding au by(auto intro!: exI[where x="{a}"]) next case False obtain ux vx b where obt:"ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" using t(4)[unfolded au convex_hull_insert[OF False]] by auto @@ -1411,7 +1410,7 @@ qed thus ?thesis using compact_convex_combinations[OF assms Suc] by simp qed - qed + qed qed lemma finite_imp_compact_convex_hull: @@ -1851,7 +1850,7 @@ lemma convex_hull_scaling: "convex hull ((\x. c *\<^sub>R x) ` s) = (\x. c *\<^sub>R x) ` (convex hull s)" apply(cases "c=0") defer apply(rule convex_hull_bilemma[rule_format, of _ _ inverse]) apply(rule convex_hull_scaling_lemma) - unfolding image_image scaleR_scaleR by(auto simp add:image_constant_conv convex_hull_eq_empty) + unfolding image_image scaleR_scaleR by(auto simp add:image_constant_conv) lemma convex_hull_affinity: "convex hull ((\x. a + c *\<^sub>R x) ` s) = (\x. a + c *\<^sub>R x) ` (convex hull s)" @@ -2313,7 +2312,7 @@ hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: ring_simps) hence "u * a + v * b \ b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) } ultimately show "u *\<^sub>R x + v *\<^sub>R y \ s" apply- apply(rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) - using as(3-) dimindex_ge_1 apply- by(auto simp add: vector_component) qed + using as(3-) dimindex_ge_1 by auto qed lemma is_interval_connected: fixes s :: "(real ^ _) set" @@ -2336,7 +2335,7 @@ hence *:"dest_vec1 a < dest_vec1 x" "dest_vec1 x < dest_vec1 b" apply(rule_tac [!] ccontr) unfolding not_less by auto let ?halfl = "{z. inner (basis 1) z < dest_vec1 x} " and ?halfr = "{z. inner (basis 1) z > dest_vec1 x} " { fix y assume "y \ s" have "y \ ?halfr \ ?halfl" apply(rule ccontr) - using as(6) `y\s` by (auto simp add: inner_vector_def dest_vec1_eq) } + using as(6) `y\s` by (auto simp add: inner_vector_def) } moreover have "a\?halfl" "b\?halfr" using * by (auto simp add: inner_vector_def) hence "?halfl \ s \ {}" "?halfr \ s \ {}" using as(2-3) by auto ultimately show False apply(rule_tac notE[OF as(1)[unfolded connected_def]]) @@ -2374,7 +2373,7 @@ shows "\x\{a..b}. (f x)$k = y" apply(subst neg_equal_iff_equal[THEN sym]) unfolding vector_uminus_component[THEN sym] apply(rule ivt_increasing_component_on_1) using assms using continuous_on_neg - by(auto simp add:vector_uminus_component) + by auto lemma ivt_decreasing_component_1: fixes f::"real^1 \ real^'n" shows "dest_vec1 a \ dest_vec1 b \ \x\{a .. b}. continuous (at x) f @@ -2415,7 +2414,7 @@ unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff defer apply(rule_tac x=j in bexI) using i' by auto have i01:"x$i \ 1" "x$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] using i'(2) `x$i \ 0` - by(auto simp add: Cart_lambda_beta) + by auto show ?thesis proof(cases "x$i=1") case True have "\j\{i. x$i \ 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof- fix j assume "x $ j \ 0" "x $ j \ 1" @@ -2424,21 +2423,21 @@ hence "x$j \ x$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto thus False using True Suc(2) j by(auto simp add: vector_le_def elim!:ballE[where x=j]) qed thus "x\convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) - by(auto simp add: Cart_lambda_beta) + by auto next let ?y = "\j. if x$j = 0 then 0 else (x$j - x$i) / (1 - x$i)" case False hence *:"x = x$i *\<^sub>R (\ j. if x$j = 0 then 0 else 1) + (1 - x$i) *\<^sub>R (\ j. ?y j)" unfolding Cart_eq - by(auto simp add: Cart_lambda_beta vector_add_component vector_smult_component vector_minus_component field_simps) + by(auto simp add: field_simps) { fix j have "x$j \ 0 \ 0 \ (x $ j - x $ i) / (1 - x $ i)" "(x $ j - x $ i) / (1 - x $ i) \ 1" apply(rule_tac divide_nonneg_pos) using i(1)[of j] using False i01 - using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps Cart_lambda_beta) + using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps) hence "0 \ ?y j \ ?y j \ 1" by auto } - moreover have "i\{j. x$j \ 0} - {j. ((\ j. ?y j)::real^'n) $ j \ 0}" using i01 by(auto simp add: Cart_lambda_beta) + moreover have "i\{j. x$j \ 0} - {j. ((\ j. ?y j)::real^'n) $ j \ 0}" using i01 by auto hence "{j. x$j \ 0} \ {j. ((\ j. ?y j)::real^'n) $ j \ 0}" by auto - hence **:"{j. ((\ j. ?y j)::real^'n) $ j \ 0} \ {j. x$j \ 0}" apply - apply rule by(auto simp add: Cart_lambda_beta) + hence **:"{j. ((\ j. ?y j)::real^'n) $ j \ 0} \ {j. x$j \ 0}" apply - apply rule by auto have "card {j. ((\ j. ?y j)::real^'n) $ j \ 0} \ n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format]) apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1)) - unfolding mem_interval using i01 Suc(3) by (auto simp add: Cart_lambda_beta) + unfolding mem_interval using i01 Suc(3) by auto qed qed qed } note * = this show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule apply(rule_tac n2="CARD('n)" in *) prefer 3 apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule @@ -2453,7 +2452,7 @@ prefer 3 apply(rule unit_interval_convex_hull) apply rule unfolding mem_Collect_eq proof- fix x::"real^'n" assume as:"\i. x $ i = 0 \ x $ i = 1" show "x \ (\s. \ i. if i \ s then 1 else 0) ` UNIV" apply(rule image_eqI[where x="{i. x$i = 1}"]) - unfolding Cart_eq using as by(auto simp add:Cart_lambda_beta) qed auto + unfolding Cart_eq using as by auto qed auto subsection {* Hence any cube (could do any nonempty interval). *} @@ -2464,23 +2463,23 @@ unfolding image_iff defer apply(erule bexE) proof- fix y assume as:"y\{x - ?d .. x + ?d}" { fix i::'n have "x $ i \ d + y $ i" "y $ i \ d + x $ i" using as[unfolded mem_interval, THEN spec[where x=i]] - by(auto simp add: vector_component) + by auto hence "1 \ inverse d * (x $ i - y $ i)" "1 \ inverse d * (y $ i - x $ i)" apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[THEN sym] - using assms by(auto simp add: field_simps right_inverse) + using assms by(auto simp add: field_simps) hence "inverse d * (x $ i * 2) \ 2 + inverse d * (y $ i * 2)" "inverse d * (y $ i * 2) \ 2 + inverse d * (x $ i * 2)" by(auto simp add:field_simps) } hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ {0..1}" unfolding mem_interval using assms - by(auto simp add: Cart_eq vector_component_simps field_simps) + by(auto simp add: Cart_eq field_simps) thus "\z\{0..1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI) - using assms by(auto simp add: Cart_eq vector_le_def Cart_lambda_beta) + using assms by(auto simp add: Cart_eq vector_le_def) next fix y z assume as:"z\{0..1}" "y = x - ?d + (2*d) *\<^sub>R z" have "\i. 0 \ d * z $ i \ d * z $ i \ d" using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE) apply rule apply(rule mult_nonneg_nonneg) prefer 3 apply(rule mult_right_le_one_le) - using assms by(auto simp add: vector_component_simps Cart_eq) + using assms by(auto simp add: Cart_eq) thus "y \ {x - ?d..x + ?d}" unfolding as(2) mem_interval apply- apply rule using as(1)[unfolded mem_interval] - apply(erule_tac x=i in allE) using assms by(auto simp add: vector_component_simps Cart_eq) qed + apply(erule_tac x=i in allE) using assms by(auto simp add: Cart_eq) qed obtain s where "finite s" "{0..1::real^'n} = convex hull s" using unit_cube_convex_hull by auto thus ?thesis apply(rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` s"]) unfolding * and convex_hull_affinity by auto qed @@ -2570,8 +2569,8 @@ have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) let ?d = "(\ i. d)::real^'n" obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto - have "x\{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by(auto simp add:vector_component_simps) - hence "c\{}" using c by(auto simp add:convex_hull_empty) + have "x\{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by auto + hence "c\{}" using c by auto def k \ "Max (f ` c)" have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)]) apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof @@ -2579,7 +2578,7 @@ have e:"e = setsum (\i. d) (UNIV::'n set)" unfolding setsum_constant d_def using dimge1 by (metis eq_divide_imp mult_frac_num real_dimindex_gt_0 real_eq_of_nat real_less_def real_mult_commute) show "dist x z \ e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono) - using z[unfolded mem_interval] apply(erule_tac x=i in allE) by(auto simp add:field_simps vector_component_simps) qed + using z[unfolded mem_interval] apply(erule_tac x=i in allE) by auto qed hence k:"\y\{x - ?d..x + ?d}. f y \ k" unfolding c(2) apply(rule_tac convex_on_convex_hull_bound) apply assumption unfolding k_def apply(rule, rule Max_ge) using c(1) by auto have "d \ e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 using real_dimindex_ge_1 by auto @@ -2588,9 +2587,9 @@ hence "\y\cball x d. abs (f y) \ k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof fix y assume y:"y\cball x d" { fix i::'n have "x $ i - d \ y $ i" "y $ i \ x $ i + d" - using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by(auto simp add: vector_component) } + using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by auto } thus "f y \ k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm - by(auto simp add: vector_component_simps) qed + by auto qed hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous) apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball) apply force @@ -2716,17 +2715,17 @@ unfolding as(1) by(auto simp add:algebra_simps) show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" unfolding norm_minus_commute[of x a] * Cart_eq using as(2,3) - by(auto simp add: vector_component_simps field_simps) + by(auto simp add: field_simps) next assume as:"dist a b = dist a x + dist x b" have "norm (a - x) / norm (a - b) \ 1" unfolding divide_le_eq_1_pos[OF Fal2] unfolding as[unfolded dist_norm] norm_ge_zero by auto thus "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" apply(rule_tac x="dist a x / dist a b" in exI) unfolding dist_norm Cart_eq apply- apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4 proof rule fix i::'n have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i = ((norm (a - b) - norm (a - x)) * (a $ i) + norm (a - x) * (b $ i)) / norm (a - b)" - using Fal by(auto simp add:vector_component_simps field_simps) + using Fal by(auto simp add: field_simps) also have "\ = x$i" apply(rule divide_eq_imp[OF Fal]) unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq Cart_eq,rule_format, of i] - by(auto simp add:field_simps vector_component_simps) + by(auto simp add:field_simps) finally show "x $ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i" by auto qed(insert Fal2, auto) qed qed @@ -2735,7 +2734,7 @@ "between (b,a) (midpoint a b)" (is ?t2) proof- have *:"\x y z. x = (1/2::real) *\<^sub>R z \ y = (1/2) *\<^sub>R z \ norm z = norm x + norm y" by auto show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *) - by(auto simp add:field_simps Cart_eq vector_component_simps) qed + by(auto simp add:field_simps Cart_eq) qed lemma between_mem_convex_hull: "between (a,b) x \ x \ convex hull {a,b}" @@ -2754,7 +2753,7 @@ have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm unfolding norm_scaleR[THEN sym] apply(rule norm_eqI) using `e>0` - by(auto simp add:vector_component_simps Cart_eq field_simps) + by(auto simp add: Cart_eq field_simps) also have "\ = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:norm_eqI simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and `e>0` by(auto simp add:pos_divide_less_eq[OF `e>0`] real_mult_commute) @@ -2826,11 +2825,11 @@ fix x::"real^'n" and e assume "0xa. dist x xa < e \ (\x. 0 \ xa $ x) \ setsum (op $ xa) UNIV \ 1" show "(\xa. 0 < x $ xa) \ setsum (op $ x) UNIV < 1" apply(rule,rule) proof- fix i::'n show "0 < x $ i" using as[THEN spec[where x="x - (e / 2) *\<^sub>R basis i"]] and `e>0` - unfolding dist_norm by(auto simp add: norm_basis vector_component_simps basis_component elim:allE[where x=i]) + unfolding dist_norm by(auto simp add: norm_basis elim:allE[where x=i]) next guess a using UNIV_witness[where 'a='n] .. have **:"dist x (x + (e / 2) *\<^sub>R basis a) < e" using `e>0` and norm_basis[of a] - unfolding dist_norm by(auto simp add: vector_component_simps basis_component intro!: mult_strict_left_mono_comm) - have "\i. (x + (e / 2) *\<^sub>R basis a) $ i = x$i + (if i = a then e/2 else 0)" by(auto simp add:vector_component_simps) + unfolding dist_norm by(auto intro!: mult_strict_left_mono_comm) + have "\i. (x + (e / 2) *\<^sub>R basis a) $ i = x$i + (if i = a then e/2 else 0)" by auto hence *:"setsum (op $ (x + (e / 2) *\<^sub>R basis a)) UNIV = setsum (\i. x$i + (if a = i then e/2 else 0)) UNIV" by(rule_tac setsum_cong, auto) have "setsum (op $ x) UNIV < setsum (op $ (x + (e / 2) *\<^sub>R basis a)) UNIV" unfolding * setsum_addf using `0 setsum (\i. x$i + ?d) UNIV" proof(rule setsum_mono) fix i::'n have "abs (y$i - x$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i] - using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add:vector_component_simps norm_minus_commute) + using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add: norm_minus_commute) thus "y $ i \ x $ i + ?d" by auto qed also have "\ \ 1" unfolding setsum_addf setsum_constant card_enum real_eq_of_nat using dimindex_ge_1 by(auto simp add: Suc_le_eq) finally show "(\i. 0 \ y $ i) \ setsum (op $ y) UNIV \ 1" apply- proof(rule,rule) fix i::'n have "norm (x - y) < x$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] - using Min_gr_iff[of "op $ x ` dimset x"] dimindex_ge_1 by auto - thus "0 \ y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by(auto simp add: vector_component_simps) + by auto + thus "0 \ y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by auto qed auto qed auto qed lemma interior_std_simplex_nonempty: obtains a::"real^'n" where @@ -3040,9 +3039,6 @@ apply(erule_tac x="1-x" in ballE, erule_tac x="1-y" in ballE) unfolding mem_interval_1 by auto -(** move this **) -declare vector_scaleR_component[simp] - lemma simple_path_join_loop: assumes "injective_path g1" "injective_path g2" "pathfinish g2 = pathstart g1" "(path_image g1 \ path_image g2) \ {pathstart g1,pathstart g2}" @@ -3390,7 +3386,7 @@ hence "\ k \ \ (Suc k)" using \[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=k]] by auto hence **:"?basis k + ?basis (Suc k) \ {x. 0 < inner (?basis (Suc k)) x} \ (?A k)" "?basis k - ?basis (Suc k) \ {x. 0 > inner (?basis (Suc k)) x} \ ({x. 0 < inner (?basis (Suc k)) x} \ (?A k))" using d - by(auto simp add: inner_basis vector_component_simps intro!:bexI[where x=k]) + by(auto simp add: inner_basis intro!:bexI[where x=k]) show ?thesis unfolding * Un_assoc apply(rule path_connected_Un) defer apply(rule path_connected_Un) prefer 5 apply(rule_tac[1-2] convex_imp_path_connected, rule convex_halfspace_lt, rule convex_halfspace_gt) apply(rule Suc(1)) using d ** False by auto @@ -3404,7 +3400,7 @@ apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I) apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I) apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I) - using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps inner_basis) + using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add: inner_basis) qed qed auto qed note lem = this have ***:"\x::real^'n. (\i\{1..CARD('n)}. inner (basis (\ i)) x \ 0) \ (\i. inner (basis i) x \ 0)" @@ -3432,7 +3428,7 @@ apply(rule, rule continuous_at_within_inv[unfolded o_def inverse_eq_divide]) apply(rule continuous_at_within) apply(rule continuous_at_norm[unfolded o_def]) by auto thus ?thesis unfolding * ** using path_connected_punctured_universe[OF assms] - by(auto intro!: path_connected_continuous_image continuous_on_intros continuous_on_mul) qed + by(auto intro!: path_connected_continuous_image continuous_on_intros) qed lemma connected_sphere: "2 \ CARD('n) \ connected {x::real^'n. norm(x - a) = r}" using path_connected_sphere path_connected_imp_connected by auto diff -r 1debc8e29f6a -r 06475a1547cb src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Sun Apr 25 23:22:29 2010 -0700 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Mon Apr 26 09:21:25 2010 -0700 @@ -93,7 +93,7 @@ subsection {* Derivatives on real = Derivatives on @{typ "real^1"} *} -lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding vector_dist_norm by(auto simp add:vec1_dest_vec1_simps) +lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding vector_dist_norm by auto lemma bounded_linear_vec1_dest_vec1: fixes f::"real \ real" shows "linear (vec1 \ f \ dest_vec1) = bounded_linear f" (is "?l = ?r") proof- @@ -281,6 +281,8 @@ subsection {* differentiability. *} +no_notation Deriv.differentiable (infixl "differentiable" 60) + definition differentiable :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a net \ bool" (infixr "differentiable" 30) where "f differentiable net \ (\f'. (f has_derivative f') net)" @@ -754,11 +756,11 @@ lemma onorm_vec1: fixes f::"real \ real" shows "onorm (\x. vec1 (f (dest_vec1 x))) = onorm f" proof- have "\x::real^1. norm x = 1 \ x\{vec1 -1, vec1 (1::real)}" unfolding forall_vec1 by(auto simp add:Cart_eq) - hence 1:"{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by(auto simp add:norm_vec1) + hence 1:"{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by auto have 2:"{norm (vec1 (f (dest_vec1 x))) |x. norm x = 1} = (\x. norm (vec1 (f (dest_vec1 x)))) ` {x. norm x=1}" by auto have "\x::real. norm x = 1 \ x\{-1, 1}" by auto hence 3:"{x. norm x = 1} = {-1, (1::real)}" by auto have 4:"{norm (f x) |x. norm x = 1} = (\x. norm (f x)) ` {x. norm x=1}" by auto - show ?thesis unfolding onorm_def 1 2 3 4 by(simp add:Sup_finite_Max norm_vec1) qed + show ?thesis unfolding onorm_def 1 2 3 4 by(simp add:Sup_finite_Max) qed lemma differentiable_bound_real: fixes f::"real \ real" assumes "convex s" "\x\s. (f has_derivative f' x) (at x within s)" "\x\s. onorm(f' x) \ B" and x:"x\s" and y:"y\s" @@ -1013,7 +1015,7 @@ unfolding ph' * apply(rule diff_chain_within) defer apply(rule bounded_linear.has_derivative[OF assms(3)]) apply(rule has_derivative_intros) defer apply(rule has_derivative_sub[where g'="\x.0",unfolded diff_0_right]) apply(rule has_derivative_at_within) using assms(5) and `u\s` `a\s` - by(auto intro!: has_derivative_intros derivative_linear) + by(auto intro!: has_derivative_intros derivative_linear) have **:"bounded_linear (\x. f' u x - f' a x)" "bounded_linear (\x. f' a x - f' u x)" apply(rule_tac[!] bounded_linear_sub) apply(rule_tac[!] derivative_linear) using assms(5) `u\s` `a\s` by auto have "onorm (\v. v - g' (f' u v)) \ onorm g' * onorm (\w. f' a w - f' u w)" unfolding * apply(rule onorm_compose) diff -r 1debc8e29f6a -r 06475a1547cb src/HOL/Multivariate_Analysis/Determinants.thy --- a/src/HOL/Multivariate_Analysis/Determinants.thy Sun Apr 25 23:22:29 2010 -0700 +++ b/src/HOL/Multivariate_Analysis/Determinants.thy Mon Apr 26 09:21:25 2010 -0700 @@ -813,7 +813,7 @@ lemma cramer: fixes A ::"real^'n^'n" assumes d0: "det A \ 0" - shows "A *v x = b \ x = (\ k. det(\ i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)" + shows "A *v x = b \ x = (\ k. det(\ i j. if j=k then b$i else A$i$j) / det A)" proof- from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1" unfolding invertible_det_nz[symmetric] invertible_def by blast @@ -821,7 +821,7 @@ hence "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc) then have xe: "\x. A*v x = b" by blast {fix x assume x: "A *v x = b" - have "x = (\ k. det(\ i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)" + have "x = (\ k. det(\ i j. if j=k then b$i else A$i$j) / det A)" unfolding x[symmetric] using d0 by (simp add: Cart_eq cramer_lemma field_simps)} with xe show ?thesis by auto @@ -993,15 +993,15 @@ moreover {assume "x = 0" "y \ 0" then have "dist (?g x) (?g y) = dist x y" - apply (simp add: dist_norm norm_mul) + apply (simp add: dist_norm) apply (rule f1[rule_format]) - by(simp add: norm_mul field_simps)} + by(simp add: field_simps)} moreover {assume "x \ 0" "y = 0" then have "dist (?g x) (?g y) = dist x y" - apply (simp add: dist_norm norm_mul) + apply (simp add: dist_norm) apply (rule f1[rule_format]) - by(simp add: norm_mul field_simps)} + by(simp add: field_simps)} moreover {assume z: "x \ 0" "y \ 0" have th00: "x = norm x *s (inverse (norm x) *s x)" "y = norm y *s (inverse (norm y) *s y)" "norm x *s f ((inverse (norm x) *s x)) = norm x *s f (inverse (norm x) *s x)" @@ -1013,7 +1013,7 @@ "norm (f (inverse (norm x) *s x) - f (inverse (norm y) *s y)) = norm (inverse (norm x) *s x - inverse (norm y) *s y)" using z - by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_norm]) + by (auto simp add: vector_smult_assoc field_simps intro: f1[rule_format] fd1[rule_format, unfolded dist_norm]) from z th0[OF th00] have "dist (?g x) (?g y) = dist x y" by (simp add: dist_norm)} ultimately have "dist (?g x) (?g y) = dist x y" by blast} @@ -1047,7 +1047,7 @@ by (simp add: nat_number setprod_numseg mult_commute) lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" - by (simp add: det_def permutes_sing sign_id UNIV_1) + by (simp add: det_def sign_id UNIV_1) lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" proof- @@ -1057,7 +1057,7 @@ unfolding setsum_over_permutations_insert[OF f12] unfolding permutes_sing apply (simp add: sign_swap_id sign_id swap_id_eq) - by (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31)) + by (simp add: arith_simps(31)[symmetric] del: arith_simps(31)) qed lemma det_3: "det (A::'a::comm_ring_1^3^3) = @@ -1078,7 +1078,7 @@ unfolding permutes_sing apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) - apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31)) + apply (simp add: arith_simps(31)[symmetric] del: arith_simps(31)) by (simp add: ring_simps) qed diff -r 1debc8e29f6a -r 06475a1547cb src/HOL/Multivariate_Analysis/Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Sun Apr 25 23:22:29 2010 -0700 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Apr 26 09:21:25 2010 -0700 @@ -670,7 +670,7 @@ subsection{* The collapse of the general concepts to dimension one. *} lemma vector_one: "(x::'a ^1) = (\ i. (x$1))" - by (simp add: Cart_eq forall_1) + by (simp add: Cart_eq) lemma forall_one: "(\(x::'a ^1). P x) \ (\x. P(\ i. x))" apply auto @@ -775,8 +775,7 @@ lemma sqrt_even_pow2: assumes n: "even n" shows "sqrt(2 ^ n) = 2 ^ (n div 2)" proof- - from n obtain m where m: "n = 2*m" unfolding even_nat_equiv_def2 - by (auto simp add: nat_number) + from n obtain m where m: "n = 2*m" unfolding even_mult_two_ex .. from m have "sqrt(2 ^ n) = sqrt ((2 ^ m) ^ 2)" by (simp only: power_mult[symmetric] mult_commute) then show ?thesis using m by simp @@ -785,7 +784,7 @@ lemma real_div_sqrt: "0 <= x ==> x / sqrt(x) = sqrt(x)" apply (cases "x = 0", simp_all) using sqrt_divide_self_eq[of x] - apply (simp add: inverse_eq_divide real_sqrt_ge_0_iff field_simps) + apply (simp add: inverse_eq_divide field_simps) done text{* Hence derive more interesting properties of the norm. *} @@ -798,8 +797,8 @@ by (rule norm_zero) lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x" - by (simp add: norm_vector_def vector_component setL2_right_distrib - abs_mult cong: strong_setL2_cong) + by (simp add: norm_vector_def setL2_right_distrib abs_mult) + lemma norm_eq_0_dot: "(norm x = 0) \ (inner x x = (0::real))" by (simp add: norm_vector_def setL2_def power2_eq_square) lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero) @@ -866,10 +865,14 @@ by (auto simp add: norm_eq_sqrt_inner) lemma real_abs_le_square_iff: "\x\ \ \y\ \ (x::real)^2 \ y^2" -proof- - have "x^2 \ y^2 \ (x -y) * (y + x) \ 0" by (simp add: ring_simps power2_eq_square) - also have "\ \ \x\ \ \y\" apply (simp add: zero_compare_simps real_abs_def not_less) by arith -finally show ?thesis .. +proof + assume "\x\ \ \y\" + then have "\x\\ \ \y\\" by (rule power_mono, simp) + then show "x\ \ y\" by simp +next + assume "x\ \ y\" + then have "sqrt (x\) \ sqrt (y\)" by (rule real_sqrt_le_mono) + then show "\x\ \ \y\" by simp qed lemma norm_le_square: "norm(x) <= a \ 0 <= a \ x \ x <= a^2" @@ -1179,7 +1182,7 @@ assumes fS: "finite S" shows "setsum f S *s v = setsum (\x. f x *s v) S" proof(induct rule: finite_induct[OF fS]) - case 1 then show ?case by (simp add: vector_smult_lzero) + case 1 then show ?case by simp next case (2 x F) from "2.hyps" have "setsum f (insert x F) *s v = (f x + setsum f F) *s v" @@ -1398,7 +1401,7 @@ apply (subgoal_tac "vector [v$1] = v") apply simp apply (vector vector_def) - apply (simp add: forall_1) + apply simp done lemma forall_vector_2: "(\v::'a::zero^2. P v) \ (\x y. P(vector[x, y]))" @@ -1536,7 +1539,7 @@ unfolding norm_mul apply (simp only: mult_commute) apply (rule mult_mono) - by (auto simp add: ring_simps norm_ge_zero) } + by (auto simp add: ring_simps) } then have th: "\i\ ?S. norm ((x$i) *s f (basis i :: real ^'m)) \ norm (f (basis i)) * norm x" by metis from real_setsum_norm_le[OF fS, of "\i. (x$i) *s (f (basis i))", OF th] have "norm (f x) \ ?B * norm x" unfolding th0 setsum_left_distrib by metis} @@ -1553,9 +1556,9 @@ let ?K = "\B\ + 1" have Kp: "?K > 0" by arith {assume C: "B < 0" - have "norm (1::real ^ 'n) > 0" by (simp add: zero_less_norm_iff) + have "norm (1::real ^ 'n) > 0" by simp with C have "B * norm (1:: real ^ 'n) < 0" - by (simp add: zero_compare_simps) + by (simp add: mult_less_0_iff) with B[rule_format, of 1] norm_ge_zero[of "f 1"] have False by simp } then have Bp: "B \ 0" by ferrack @@ -1677,11 +1680,11 @@ apply (rule real_setsum_norm_le) using fN fM apply simp - apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] norm_mul ring_simps) + apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] ring_simps) apply (rule mult_mono) - apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm) + apply (auto simp add: zero_le_mult_iff component_le_norm) apply (rule mult_mono) - apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm) + apply (auto simp add: zero_le_mult_iff component_le_norm) done} then show ?thesis by metis qed @@ -1701,7 +1704,7 @@ have "B * norm x * norm y \ ?K * norm x * norm y" apply - apply (rule mult_right_mono, rule mult_right_mono) - by (auto simp add: norm_ge_zero) + by auto then have "norm (h x y) \ ?K * norm x * norm y" using B[rule_format, of x y] by simp} with Kp show ?thesis by blast @@ -2006,8 +2009,8 @@ have uv': "u = 0 \ v \ 0" using u v uv by arith have "a = a * (u + v)" unfolding uv by simp hence th: "u * a + v * a = a" by (simp add: ring_simps) - from xa u have "u \ 0 \ u*x < u*a" by (simp add: mult_compare_simps) - from ya v have "v \ 0 \ v * y < v * a" by (simp add: mult_compare_simps) + from xa u have "u \ 0 \ u*x < u*a" by (simp add: mult_strict_left_mono) + from ya v have "v \ 0 \ v * y < v * a" by (simp add: mult_strict_left_mono) from xa ya u v have "u * x + v * y < u * a + v * a" apply (cases "u = 0", simp_all add: uv') apply(rule mult_strict_left_mono) @@ -2048,7 +2051,7 @@ assumes x: "0 <= (x::real)" and y:"0 <= y" and z: "0 <= z" and xy: "x^2 <= y^2 + z^2" shows "x <= y + z" proof- - have "y^2 + z^2 \ y^2 + 2*y*z + z^2" using z y by (simp add: zero_compare_simps) + have "y^2 + z^2 \ y^2 + 2*y*z + z^2" using z y by (simp add: mult_nonneg_nonneg) with xy have th: "x ^2 \ (y+z)^2" by (simp add: power2_eq_square ring_simps) from y z have yz: "y + z \ 0" by arith from power2_le_imp_le[OF th yz] show ?thesis . @@ -2100,10 +2103,10 @@ {assume x0: "x \ 0" hence n0: "norm x \ 0" by (metis norm_eq_zero) let ?c = "1/ norm x" - have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul) + have "norm (?c*s x) = 1" using x0 by (simp add: n0) with H have "norm (f(?c*s x)) \ b" by blast hence "?c * norm (f x) \ b" - by (simp add: linear_cmul[OF lf] norm_mul) + by (simp add: linear_cmul[OF lf]) hence "norm (f x) \ b * norm x" using n0 norm_ge_zero[of x] by (auto simp add: field_simps)} ultimately have "norm (f x) \ b * norm x" by blast} @@ -2221,18 +2224,24 @@ where "dest_vec1 x \ (x$1)" lemma vec1_component[simp]: "(vec1 x)$1 = x" - by (simp add: ) - -lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y" - by (simp_all add: Cart_eq forall_1) - -lemma forall_vec1: "(\x. P x) \ (\x. P (vec1 x))" by (metis vec1_dest_vec1) - -lemma exists_vec1: "(\x. P x) \ (\x. P(vec1 x))" by (metis vec1_dest_vec1) - -lemma vec1_eq[simp]: "vec1 x = vec1 y \ x = y" by (metis vec1_dest_vec1) - -lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \ x = y" by (metis vec1_dest_vec1) + by simp + +lemma vec1_dest_vec1: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y" + by (simp_all add: Cart_eq) + +declare vec1_dest_vec1(1) [simp] + +lemma forall_vec1: "(\x. P x) \ (\x. P (vec1 x))" + by (metis vec1_dest_vec1(1)) + +lemma exists_vec1: "(\x. P x) \ (\x. P(vec1 x))" + by (metis vec1_dest_vec1(1)) + +lemma vec1_eq[simp]: "vec1 x = vec1 y \ x = y" + by (metis vec1_dest_vec1(2)) + +lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \ x = y" + by (metis vec1_dest_vec1(1)) lemma vec_in_image_vec: "vec x \ (vec ` S) \ x \ S" by auto @@ -2260,8 +2269,8 @@ lemma dest_vec1_sum: assumes fS: "finite S" shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S" apply (induct rule: finite_induct[OF fS]) - apply (simp add: dest_vec1_vec) - apply (auto simp add:vector_minus_component) + apply simp + apply auto done lemma norm_vec1: "norm(vec1 x) = abs(x)" @@ -2270,7 +2279,7 @@ lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)" by (simp only: dist_real vec1_component) lemma abs_dest_vec1: "norm x = \dest_vec1 x\" - by (metis vec1_dest_vec1 norm_vec1) + by (metis vec1_dest_vec1(1) norm_vec1) lemmas vec1_dest_vec1_simps = forall_vec1 vec_add[THEN sym] dist_vec1 vec_sub[THEN sym] vec1_dest_vec1 norm_vec1 vector_smult_component vec1_eq vec_cmul[THEN sym] smult_conv_scaleR[THEN sym] o_def dist_real_def norm_vec1 real_norm_def @@ -2298,7 +2307,7 @@ shows "f = (\x. vec1(row 1 (matrix f) \ x))" apply (rule ext) apply (subst matrix_works[OF lf, symmetric]) - apply (simp add: Cart_eq matrix_vector_mult_def row_def inner_vector_def mult_commute forall_1) + apply (simp add: Cart_eq matrix_vector_mult_def row_def inner_vector_def mult_commute) done lemma dest_vec1_eq_0: "dest_vec1 x = 0 \ x = 0" @@ -2366,13 +2375,13 @@ by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0) lemma pastecart_vec[simp]: "pastecart (vec x) (vec x) = vec x" - by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart) + by (simp add: pastecart_eq) lemma pastecart_add[simp]:"pastecart (x1::'a::{plus,times}^_) y1 + pastecart x2 y2 = pastecart (x1 + x2) (y1 + y2)" - by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart) + by (simp add: pastecart_eq) lemma pastecart_cmul[simp]: "pastecart (c *s (x1::'a::{plus,times}^_)) (c *s y1) = c *s pastecart x1 y1" - by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart) + by (simp add: pastecart_eq) lemma pastecart_neg[simp]: "pastecart (- (x::'a::ring_1^_)) (- y) = - pastecart x y" unfolding vector_sneg_minus1 pastecart_cmul .. @@ -2384,7 +2393,7 @@ fixes f:: "'d \ 'a::semiring_1^_" assumes fS: "finite S" shows "pastecart (setsum f S) (setsum g S) = setsum (\i. pastecart (f i) (g i)) S" - by (simp add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS] fstcart_pastecart sndcart_pastecart) + by (simp add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS]) lemma setsum_Plus: "\finite A; finite B\ \ @@ -2402,7 +2411,7 @@ lemma norm_fstcart: "norm(fstcart x) <= norm (x::real ^('n::finite + 'm::finite))" proof- have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))" - by (simp add: pastecart_fst_snd) + by simp have th1: "fstcart x \ fstcart x \ pastecart (fstcart x) (sndcart x) \ pastecart (fstcart x) (sndcart x)" by (simp add: inner_vector_def setsum_UNIV_sum pastecart_def setsum_nonneg) then show ?thesis @@ -2417,7 +2426,7 @@ lemma norm_sndcart: "norm(sndcart x) <= norm (x::real ^('n::finite + 'm::finite))" proof- have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))" - by (simp add: pastecart_fst_snd) + by simp have th1: "sndcart x \ sndcart x \ pastecart (fstcart x) (sndcart x) \ pastecart (fstcart x) (sndcart x)" by (simp add: inner_vector_def setsum_UNIV_sum pastecart_def setsum_nonneg) then show ?thesis @@ -2519,7 +2528,7 @@ lemma real_arch_inv: "0 < e \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" using reals_Archimedean - apply (auto simp add: field_simps inverse_positive_iff_positive) + apply (auto simp add: field_simps) apply (subgoal_tac "inverse (real n) > 0") apply arith apply simp @@ -2732,7 +2741,8 @@ "0 \ span S" "x\ span S \ y \ span S ==> x + y \ span S" "x \ span S \ c *s x \ span S" - by (metis span_def hull_subset subset_eq subspace_span subspace_def)+ + by (metis span_def hull_subset subset_eq) + (metis subspace_span subspace_def)+ lemma span_induct: assumes SP: "\x. x \ S ==> P x" and P: "subspace P" and x: "x \ span S" shows "P x" @@ -2830,7 +2840,7 @@ (* Individual closure properties. *) -lemma span_superset: "x \ S ==> x \ span S" by (metis span_clauses) +lemma span_superset: "x \ S ==> x \ span S" by (metis span_clauses(1)) lemma span_0: "0 \ span S" by (metis subspace_span subspace_0) @@ -2847,8 +2857,7 @@ by (metis subspace_span subspace_sub) lemma span_setsum: "finite A \ \x \ A. f x \ span S ==> setsum f A \ span S" - apply (rule subspace_setsum) - by (metis subspace_span subspace_setsum)+ + by (rule subspace_setsum, rule subspace_span) lemma span_add_eq: "(x::'a::ring_1^_) \ span S \ x + y \ span S \ y \ span S" apply (auto simp only: span_add span_sub) @@ -3110,7 +3119,7 @@ from fS SP aP have th0: "finite ?S" "?S \ P" "?v \ ?S" "?u ?v \ 0" by auto have s0: "setsum (\v. ?u v *s v) ?S = 0" using fS aS - apply (simp add: vector_smult_lneg vector_smult_lid setsum_clauses ring_simps ) + apply (simp add: vector_smult_lneg setsum_clauses ring_simps) apply (subst (2) ua[symmetric]) apply (rule setsum_cong2) by auto @@ -3479,7 +3488,7 @@ lemma basis_card_eq_dim: "B \ (V:: (real ^'n) set) \ V \ span B \ independent B \ finite B \ card B = dim V" - by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_mono independent_bound) + by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_bound) lemma dim_unique: "(B::(real ^'n) set) \ V \ V \ span B \ independent B \ card B = n \ dim V = n" by (metis basis_card_eq_dim) @@ -3669,7 +3678,7 @@ apply (subst Cy) using C(1) fth apply (simp only: setsum_clauses) unfolding smult_conv_scaleR - apply (auto simp add: inner_simps inner_eq_zero_iff inner_commute[of y a] dot_lsum[OF fth]) + apply (auto simp add: inner_simps inner_commute[of y a] dot_lsum[OF fth]) apply (rule setsum_0') apply clarsimp apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) @@ -3686,7 +3695,7 @@ using C(1) fth apply (simp only: setsum_clauses) unfolding smult_conv_scaleR apply (subst inner_commute[of x]) - apply (auto simp add: inner_simps inner_eq_zero_iff inner_commute[of x a] dot_rsum[OF fth]) + apply (auto simp add: inner_simps inner_commute[of x a] dot_rsum[OF fth]) apply (rule setsum_0') apply clarsimp apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) @@ -3759,10 +3768,10 @@ apply (subst B') using fB fth unfolding setsum_clauses(2)[OF fth] apply simp unfolding inner_simps smult_conv_scaleR - apply (clarsimp simp add: inner_simps inner_eq_zero_iff smult_conv_scaleR dot_lsum) + apply (clarsimp simp add: inner_simps smult_conv_scaleR dot_lsum) apply (rule setsum_0', rule ballI) unfolding inner_commute - by (auto simp add: x field_simps inner_eq_zero_iff intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])} + by (auto simp add: x field_simps intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])} then show "\x \ B. ?a \ x = 0" by blast qed with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"]) @@ -3915,7 +3924,7 @@ {assume xb: "x \ b" have h0: "0 = ?h x" apply (rule conjunct2[OF h, rule_format]) - apply (metis span_superset insertI1 xb x) + apply (metis span_superset x) apply simp apply (metis span_superset xb) done @@ -4262,8 +4271,7 @@ {fix y have "?P y" proof(rule span_induct_alt[of ?P "columns A"]) show "\x\real ^ 'm. setsum (\i. (x$i) *s column i A) ?U = 0" - apply (rule exI[where x=0]) - by (simp add: zero_index vector_smult_lzero) + by (rule exI[where x=0], simp) next fix c y1 y2 assume y1: "y1 \ columns A" and y2: "?P y2" from y1 obtain i where i: "i \ ?U" "y1 = column i A" @@ -4687,7 +4695,7 @@ hence d2: "(sqrt (real ?d))^2 = real ?d" by (auto intro: real_sqrt_pow2) have th: "sqrt (real ?d) * infnorm x \ 0" - by (simp add: zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le) + by (simp add: zero_le_mult_iff infnorm_pos_le) have th1: "x \ x \ (sqrt (real ?d) * infnorm x)^2" unfolding power_mult_distrib d2 unfolding real_of_nat_def inner_vector_def @@ -4861,4 +4869,3 @@ done end - \ No newline at end of file diff -r 1debc8e29f6a -r 06475a1547cb src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Sun Apr 25 23:22:29 2010 -0700 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Mon Apr 26 09:21:25 2010 -0700 @@ -933,7 +933,7 @@ lemma has_integral_vec1: assumes "(f has_integral k) {a..b}" shows "((\x. vec1 (f x)) has_integral (vec1 k)) {a..b}" proof- have *:"\p. (\(x, k)\p. content k *\<^sub>R vec1 (f x)) - vec1 k = vec1 ((\(x, k)\p. content k *\<^sub>R f x) - k)" - unfolding vec_sub Cart_eq by(auto simp add:vec1_dest_vec1_simps split_beta) + unfolding vec_sub Cart_eq by(auto simp add: split_beta) show ?thesis using assms unfolding has_integral apply safe apply(erule_tac x=e in allE,safe) apply(rule_tac x=d in exI,safe) apply(erule_tac x=p in allE,safe) unfolding * norm_vector_1 by auto qed @@ -1356,7 +1356,7 @@ 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 + using has_integral_eq[of s f g] has_integral_eq[of s g f] by rule auto lemma has_integral_null[dest]: assumes "content({a..b}) = 0" shows "(f has_integral 0) ({a..b})" @@ -1653,7 +1653,7 @@ 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 + by(auto simp add: vector_less_def 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" @@ -1743,11 +1743,11 @@ 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) + by (metis option.nchotomy) lemma exists_option: "(\x. P x) \ P None \ (\x. P(Some x))" - by (metis map_of.simps option.nchotomy) + by (metis option.nchotomy) fun lifted where "lifted (opp::'a\'a\'b) (Some x) (Some y) = Some(opp x y)" | @@ -1842,9 +1842,8 @@ 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 neutral_monoid: "neutral ((op +)::('a::comm_monoid_add) \ 'a \ 'a) = 0" + by (rule neutral_add) (* FIXME: duplicate *) lemma monoidal_monoid[intro]: shows "monoidal ((op +)::('a::comm_monoid_add) \ 'a \ 'a)" @@ -1941,7 +1940,7 @@ 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 + apply auto 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" @@ -1952,7 +1951,7 @@ 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 + apply auto 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. *} @@ -2254,7 +2253,7 @@ 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 + apply(rule mp) defer apply assumption unfolding split_paired_all apply 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 @@ -2903,7 +2902,7 @@ 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) + by(auto simp add:not_less 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 @@ -3340,7 +3339,7 @@ proof- { presume *:"a < b \ ?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 + hence *:"{vec a .. vec b} = {vec b}" "f b - f a = 0" by(auto simp add: Cart_eq vector_le_def order_antisym) 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 \ @@ -3422,7 +3421,7 @@ 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<.. 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 @@ -3669,7 +3668,7 @@ from indefinite_integral_continuous_left[OF assms(1) `a0`] guess d . note d=this show ?thesis apply(rule,rule,rule d,safe) apply(subst dist_commute) unfolding `x=b` vector_dist_norm apply(rule d(2)[rule_format]) unfolding norm_real by auto - next assume "a xb" and xr:"a\x" "x xb" and xr:"a\x" "x0`] guess d1 . note d1=this from indefinite_integral_continuous_right[OF assms(1) xr `e>0`] guess d2 . note d2=this show ?thesis apply(rule_tac x="min d1 d2" in exI) @@ -3726,7 +3725,7 @@ unfolding o_def using assms(5) defer apply-apply(rule) proof- fix t assume as:"t\{0..1} - {t. (1 - t) *\<^sub>R c + t *\<^sub>R x \ k}" have *:"c - t *\<^sub>R c + t *\<^sub>R x \ s - k" apply safe apply(rule conv[unfolded scaleR_simps]) - using `x\s` `c\s` as by(auto simp add:scaleR_simps) + using `x\s` `c\s` as by(auto simp add: algebra_simps) have "(f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x) has_derivative (\x. 0) \ (\z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) (at t within {0..1})" apply(rule diff_chain_within) apply(rule has_derivative_add) unfolding scaleR_simps apply(rule has_derivative_sub) apply(rule has_derivative_const) @@ -3949,7 +3948,7 @@ lemma has_integral_spike_set_eq: fixes f::"real^'n \ 'a::banach" assumes "negligible((s - t) \ (t - s))" shows "((f has_integral y) s \ (f has_integral y) t)" - unfolding has_integral_restrict_univ[THEN sym,of f] apply(rule has_integral_spike_eq[OF assms]) by auto + unfolding has_integral_restrict_univ[THEN sym,of f] apply(rule has_integral_spike_eq[OF assms]) by (safe, auto split: split_if_asm) lemma has_integral_spike_set[dest]: fixes f::"real^'n \ 'a::banach" assumes "negligible((s - t) \ (t - s))" "(f has_integral y) s" diff -r 1debc8e29f6a -r 06475a1547cb src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Sun Apr 25 23:22:29 2010 -0700 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Apr 26 09:21:25 2010 -0700 @@ -48,16 +48,17 @@ "\S T. openin U S \ openin U T \ openin U (S\T)" "\K. (\S \ K. openin U S) \ openin U (\K)" using openin[of U] unfolding istopology_def Collect_def mem_def - by (metis mem_def subset_eq)+ + unfolding subset_eq Ball_def mem_def by auto lemma openin_subset[intro]: "openin U S \ S \ topspace U" unfolding topspace_def by blast lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses) lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" - by (simp add: openin_clauses) - -lemma openin_Union[intro]: "(\S \K. openin U S) \ openin U (\ K)" by (simp add: openin_clauses) + using openin_clauses by simp + +lemma openin_Union[intro]: "(\S \K. openin U S) \ openin U (\ K)" + using openin_clauses by simp lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" using openin_Union[of "{S,T}" U] by auto @@ -946,7 +947,7 @@ by (metis frontier_def closure_closed Diff_subset) lemma frontier_empty[simp]: "frontier {} = {}" - by (simp add: frontier_def closure_empty) + by (simp add: frontier_def) lemma frontier_subset_eq: "frontier S \ S \ closed S" proof- @@ -954,7 +955,7 @@ hence "closure S \ S" using interior_subset unfolding frontier_def by auto hence "closed S" using closure_subset_eq by auto } - thus ?thesis using frontier_subset_closed[of S] by auto + thus ?thesis using frontier_subset_closed[of S] .. qed lemma frontier_complement: "frontier(- S) = frontier S" @@ -1588,7 +1589,7 @@ (* FIXME: Only one congruence rule for tendsto can be used at a time! *) -lemma Lim_cong_within[cong add]: +lemma Lim_cong_within(*[cong add]*): fixes a :: "'a::metric_space" fixes l :: "'b::metric_space" (* TODO: generalize *) shows "(\x. x \ a \ f x = g x) ==> ((\x. f x) ---> l) (at a within S) \ ((g ---> l) (at a within S))" @@ -1667,7 +1668,7 @@ { fix e::real assume "e>0" hence "\N::nat. \n::nat\N. inverse (real n) < e" using real_arch_inv[of e] apply auto apply(rule_tac x=n in exI) - by (metis not_le le_imp_inverse_le not_less real_of_nat_gt_zero_cancel_iff real_of_nat_less_iff xt1(7)) + by (metis le_imp_inverse_le not_less real_of_nat_gt_zero_cancel_iff real_of_nat_less_iff xt1(7)) } thus ?thesis unfolding Lim_sequentially dist_norm by simp qed @@ -1702,7 +1703,7 @@ apply (simp add: interior_def, safe) apply (force simp add: open_contains_cball) apply (rule_tac x="ball x e" in exI) - apply (simp add: open_ball centre_in_ball subset_trans [OF ball_subset_cball]) + apply (simp add: subset_trans [OF ball_subset_cball]) done lemma islimpt_ball: @@ -1877,14 +1878,14 @@ lemma frontier_ball: fixes a :: "'a::real_normed_vector" shows "0 < e ==> frontier(ball a e) = {x. dist a x = e}" - apply (simp add: frontier_def closure_ball interior_open open_ball order_less_imp_le) + apply (simp add: frontier_def closure_ball interior_open order_less_imp_le) apply (simp add: expand_set_eq) by arith lemma frontier_cball: fixes a :: "'a::{real_normed_vector, perfect_space}" shows "frontier(cball a e) = {x. dist a x = e}" - apply (simp add: frontier_def interior_cball closed_cball closure_closed order_less_imp_le) + apply (simp add: frontier_def interior_cball closed_cball order_less_imp_le) apply (simp add: expand_set_eq) by arith @@ -2004,9 +2005,10 @@ lemma bounded_ball[simp,intro]: "bounded(ball x e)" by (metis ball_subset_cball bounded_cball bounded_subset) -lemma finite_imp_bounded[intro]: assumes "finite S" shows "bounded S" +lemma finite_imp_bounded[intro]: + fixes S :: "'a::metric_space set" assumes "finite S" shows "bounded S" proof- - { fix a F assume as:"bounded F" + { fix a and F :: "'a set" assume as:"bounded F" then obtain x e where "\y\F. dist x y \ e" unfolding bounded_def by auto hence "\y\(insert a F). dist x y \ max e (dist x a)" by auto hence "bounded (insert a F)" unfolding bounded_def by (intro exI) @@ -2216,7 +2218,7 @@ apply (rule allI, rule impI, rule ext) apply (erule conjE) apply (induct_tac x) -apply (simp add: nat_rec_0) +apply simp apply (erule_tac x="n" in allE) apply (simp) done @@ -2648,7 +2650,8 @@ { fix x y assume "x\t" "y\t" "f x = f y" hence "x \ f x" "y \ f x \ y = x" using f[THEN bspec[where x=x]] and `t\s` by auto hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\t` and `t\s` by auto } - hence "infinite (f ` t)" using assms(2) using finite_imageD[unfolded inj_on_def, of f t] by auto + hence "inj_on f t" unfolding inj_on_def by simp + hence "infinite (f ` t)" using assms(2) using finite_imageD by auto moreover { fix x assume "x\t" "f x \ g" from g(3) assms(3) `x\t` obtain h where "h\g" and "x\h" by auto @@ -3143,7 +3146,7 @@ using `?lhs`[unfolded continuous_within Lim_within] by auto { fix y assume "y\f ` (ball x d \ s)" hence "y \ ball (f x) e" using d(2) unfolding dist_nz[THEN sym] - apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto + apply (auto simp add: dist_commute) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto } hence "\d>0. f ` (ball x d \ s) \ ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute) } thus ?rhs by auto @@ -3875,7 +3878,7 @@ fixes f :: "'a::metric_space \ 'b::metric_space" (* TODO: generalize *) assumes "connected s" "continuous_on s f" "open {x \ s. f x = a}" "\x \ s. f x = a" shows "\x \ s. f x = a" -using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by auto +using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by fast text{* Some arithmetical combinations (more to prove). *} @@ -4387,7 +4390,7 @@ using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\n. sndcart (x n)"], THEN spec[where x="sndcart l"]] unfolding Lim_sequentially by auto hence "l \ {pastecart x y |x y. x \ s \ y \ t}" apply- unfolding mem_Collect_eq apply(rule_tac x="fstcart l" in exI,rule_tac x="sndcart l" in exI) by auto } - thus ?thesis unfolding closed_sequential_limits by auto + thus ?thesis unfolding closed_sequential_limits by blast qed lemma compact_pastecart: @@ -4478,7 +4481,7 @@ have "{x - y | x y . x\s \ y\s} \ {}" using `s \ {}` by auto then obtain x where x:"x\{x - y |x y. x \ s \ y \ s}" "\y\{x - y |x y. x \ s \ y \ s}. norm y \ norm x" using compact_differences[OF assms(1) assms(1)] - using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\s \ y\s}" 0] by(auto simp add: norm_minus_cancel) + using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\s \ y\s}" 0] by auto from x(1) obtain a b where "a\s" "b\s" "x = a - b" by auto thus ?thesis using x(2)[unfolded `x = a - b`] by blast qed @@ -4499,7 +4502,7 @@ hence "norm (x - y) \ 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps) } note * = this { fix x y assume "x\s" "y\s" hence "s \ {}" by auto - have "norm(x - y) \ diameter s" unfolding diameter_def using `s\{}` *[OF `x\s` `y\s`] `x\s` `y\s` + have "norm(x - y) \ diameter s" unfolding diameter_def using `s\{}` *[OF `x\s` `y\s`] `x\s` `y\s` by simp (blast intro!: Sup_upper *) } moreover { fix d::real assume "d>0" "d < diameter s" @@ -4530,10 +4533,10 @@ proof- have b:"bounded s" using assms(1) by (rule compact_imp_bounded) then obtain x y where xys:"x\s" "y\s" and xy:"\u\s. \v\s. norm (u - v) \ norm (x - y)" using compact_sup_maxdistance[OF assms] by auto - hence "diameter s \ norm (x - y)" - by (force simp add: diameter_def intro!: Sup_least) + hence "diameter s \ norm (x - y)" + unfolding diameter_def by clarsimp (rule Sup_least, fast+) thus ?thesis - by (metis b diameter_bounded_bound order_antisym xys) + by (metis b diameter_bounded_bound order_antisym xys) qed text{* Related results with closure as the conclusion. *} @@ -4722,7 +4725,7 @@ lemma mem_interval_1: fixes x :: "real^1" shows "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b)" "(x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" -by(simp_all add: Cart_eq vector_less_def vector_le_def forall_1) +by(simp_all add: Cart_eq vector_less_def vector_le_def) lemma vec1_interval:fixes a::"real" shows "vec1 ` {a .. b} = {vec1 a .. vec1 b}" @@ -4748,7 +4751,7 @@ have "a$i < b$i" using as[THEN spec[where x=i]] by auto hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i" unfolding vector_smult_component and vector_add_component - by (auto simp add: less_divide_eq_number_of1) } + by auto } hence "{a <..< b} \ {}" using mem_interval(1)[of "?x" a b] by auto } ultimately show ?th1 by blast @@ -4763,7 +4766,7 @@ have "a$i \ b$i" using as[THEN spec[where x=i]] by auto hence "a$i \ ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \ b$i" unfolding vector_smult_component and vector_add_component - by (auto simp add: less_divide_eq_number_of1) } + by auto } hence "{a .. b} \ {}" using mem_interval(2)[of "?x" a b] by auto } ultimately show ?th2 by blast qed @@ -4826,13 +4829,13 @@ { fix j have "c $ j < ?x $ j \ ?x $ j < d $ j" unfolding Cart_lambda_beta apply(cases "j=i") using as(2)[THEN spec[where x=j]] - by (auto simp add: less_divide_eq_number_of1 as2) } + by (auto simp add: as2) } hence "?x\{c<..{a .. b}" unfolding mem_interval apply auto apply(rule_tac x=i in exI) using as(2)[THEN spec[where x=i]] and as2 - by (auto simp add: less_divide_eq_number_of1) + by auto ultimately have False using as by auto } hence "a$i \ c$i" by(rule ccontr)auto moreover @@ -4841,13 +4844,13 @@ { fix j have "d $ j > ?x $ j \ ?x $ j > c $ j" unfolding Cart_lambda_beta apply(cases "j=i") using as(2)[THEN spec[where x=j]] - by (auto simp add: less_divide_eq_number_of1 as2) } + by (auto simp add: as2) } hence "?x\{c<..{a .. b}" unfolding mem_interval apply auto apply(rule_tac x=i in exI) using as(2)[THEN spec[where x=i]] and as2 - by (auto simp add: less_divide_eq_number_of1) + by auto ultimately have False using as by auto } hence "b$i \ d$i" by(rule ccontr)auto ultimately @@ -4878,7 +4881,7 @@ lemma inter_interval: fixes a :: "'a::linorder^'n" shows "{a .. b} \ {c .. d} = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" unfolding expand_set_eq and Int_iff and mem_interval - by (auto simp add: less_divide_eq_number_of1 intro!: bexI) + by auto (* Moved interval_open_subset_closed a bit upwards *) @@ -4918,7 +4921,7 @@ using open_interval[of "vec1 a" "vec1 b"] unfolding open_contains_ball apply-apply(rule,erule_tac x="vec1 x" in ballE) apply(erule exE,rule_tac x=e in exI) unfolding subset_eq mem_ball apply(rule) defer apply(rule,erule conjE,erule_tac x="vec1 xa" in ballE) - by(auto simp add: vec1_dest_vec1_simps vector_less_def forall_1) + by(auto simp add: dist_vec1 dist_real_def vector_less_def) lemma closed_interval[intro]: fixes a :: "real^'n" shows "closed {a .. b}" proof- @@ -4999,7 +5002,7 @@ have "a $ i < ((1 / 2) *\<^sub>R (a + b)) $ i \ ((1 / 2) *\<^sub>R (a + b)) $ i < b $ i" using assms[unfolded interval_ne_empty, THEN spec[where x=i]] unfolding vector_smult_component and vector_add_component - by(auto simp add: less_divide_eq_number_of1) } + by auto } thus ?thesis unfolding mem_interval by auto qed @@ -5041,7 +5044,7 @@ x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)" by (auto simp add: algebra_simps) hence "f n < b" and "a < f n" using open_closed_interval_convex[OF open_interval_midpoint[OF assms] as *] unfolding f_def by auto - hence False using fn unfolding f_def using xc by(auto simp add: vector_mul_lcancel vector_ssub_ldistrib) } + hence False using fn unfolding f_def using xc by(auto simp add: vector_ssub_ldistrib) } moreover { assume "\ (f ---> x) sequentially" { fix e::real assume "e>0" @@ -5457,7 +5460,7 @@ then obtain y where y:"y\t" "g y = x" by auto then obtain x' where x':"x'\s" "f x' = y" using assms(3) by auto hence "x \ s" unfolding g_def using someI2[of "\b. b\s \ f b = y" x' "\x. x\s"] unfolding y(2)[THEN sym] and g_def by auto } - ultimately have "x\s \ x \ g ` t" by auto } + ultimately have "x\s \ x \ g ` t" .. } hence "g ` t = s" by auto ultimately show ?thesis unfolding homeomorphism_def homeomorphic_def @@ -5599,7 +5602,7 @@ let ?S' = "{x::real^'m. x\s \ norm x = norm a}" let ?S'' = "{x::real^'m. norm x = norm a}" - have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by (auto simp add: norm_minus_cancel) + have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by auto hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto moreover have "?S' = s \ ?S''" by auto ultimately have "compact ?S'" using closed_inter_compact[of s ?S''] using s(1) by auto @@ -5646,7 +5649,7 @@ lemma subspace_substandard: "subspace {x::real^_. (\i. P i \ x$i = 0)}" - unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE) + unfolding subspace_def by auto lemma closed_substandard: "closed {x::real^'n. \i. P i --> x$i = 0}" (is "closed ?A") @@ -5663,7 +5666,7 @@ then obtain B where BB:"B \ ?Bs" and B:"B = {x::real^'n. inner (basis i) x = 0}" by auto hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto } hence "x\?A" by auto } - ultimately have "x\?A \ x\ \?Bs" by auto } + ultimately have "x\?A \ x\ \?Bs" .. } hence "?A = \ ?Bs" by auto thus ?thesis by(auto simp add: closed_Inter closed_hyperplane) qed @@ -5839,23 +5842,23 @@ case False { fix y assume "a \ y" "y \ b" "m > 0" hence "m *\<^sub>R a + c \ m *\<^sub>R y + c" "m *\<^sub>R y + c \ m *\<^sub>R b + c" - unfolding vector_le_def by(auto simp add: vector_smult_component vector_add_component) + unfolding vector_le_def by auto } moreover { fix y assume "a \ y" "y \ b" "m < 0" hence "m *\<^sub>R b + c \ m *\<^sub>R y + c" "m *\<^sub>R y + c \ m *\<^sub>R a + c" - unfolding vector_le_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE) + unfolding vector_le_def by(auto simp add: mult_left_mono_neg) } moreover { fix y assume "m > 0" "m *\<^sub>R a + c \ y" "y \ m *\<^sub>R b + c" hence "y \ (\x. m *\<^sub>R x + c) ` {a..b}" unfolding image_iff Bex_def mem_interval vector_le_def - apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric] + apply(auto simp add: vector_smult_assoc pth_3[symmetric] intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"]) by(auto simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute diff_le_iff) } moreover { fix y assume "m *\<^sub>R b + c \ y" "y \ m *\<^sub>R a + c" "m < 0" hence "y \ (\x. m *\<^sub>R x + c) ` {a..b}" unfolding image_iff Bex_def mem_interval vector_le_def - apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric] + apply(auto simp add: vector_smult_assoc pth_3[symmetric] intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"]) by(auto simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute diff_le_iff) }