# HG changeset patch # User krauss # Date 1364020253 -3600 # Node ID 0a0c9a45d294459c6c60fb3f4dfd7c71a8bf753e # Parent 637aa1649ac785e05fdb350dfb5fc18e2a00b43d# Parent dc39d69774bb9524dbd732a871c6ab320101af07 merged diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Deriv.thy Sat Mar 23 07:30:53 2013 +0100 @@ -19,15 +19,6 @@ ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "DERIV f x :> D = ((%h. (f(x + h) - f x) / h) -- 0 --> D)" -primrec - Bolzano_bisect :: "(real \ real \ bool) \ real \ real \ nat \ real \ real" where - "Bolzano_bisect P a b 0 = (a, b)" - | "Bolzano_bisect P a b (Suc n) = - (let (x, y) = Bolzano_bisect P a b n - in if P (x, (x+y) / 2) then ((x+y)/2, y) - else (x, (x+y)/2))" - - subsection {* Derivatives *} lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --> D)" @@ -434,225 +425,75 @@ subsection {* Nested Intervals and Bisection *} -text{*Lemmas about nested intervals and proof by bisection (cf.Harrison). - All considerably tidied by lcp.*} - -lemma lemma_f_mono_add [rule_format (no_asm)]: "(\n. (f::nat=>real) n \ f (Suc n)) --> f m \ f(m + no)" -apply (induct "no") -apply (auto intro: order_trans) -done - -lemma f_inc_g_dec_Beq_f: "[| \n. f(n) \ f(Suc n); - \n. g(Suc n) \ g(n); - \n. f(n) \ g(n) |] - ==> Bseq (f :: nat \ real)" -apply (rule_tac k = "f 0" and K = "g 0" in BseqI2, rule allI) -apply (rule conjI) -apply (induct_tac "n") -apply (auto intro: order_trans) -apply (rule_tac y = "g n" in order_trans) -apply (induct_tac [2] "n") -apply (auto intro: order_trans) -done - -lemma f_inc_g_dec_Beq_g: "[| \n. f(n) \ f(Suc n); - \n. g(Suc n) \ g(n); - \n. f(n) \ g(n) |] - ==> Bseq (g :: nat \ real)" -apply (subst Bseq_minus_iff [symmetric]) -apply (rule_tac g = "%x. - (f x)" in f_inc_g_dec_Beq_f) -apply auto -done - -lemma f_inc_imp_le_lim: - fixes f :: "nat \ real" - shows "\\n. f n \ f (Suc n); convergent f\ \ f n \ lim f" - by (rule incseq_le, simp add: incseq_SucI, simp add: convergent_LIMSEQ_iff) - -lemma lim_uminus: - fixes g :: "nat \ 'a::real_normed_vector" - shows "convergent g ==> lim (%x. - g x) = - (lim g)" -apply (rule tendsto_minus [THEN limI]) -apply (simp add: convergent_LIMSEQ_iff) -done - -lemma g_dec_imp_lim_le: - fixes g :: "nat \ real" - shows "\\n. g (Suc n) \ g(n); convergent g\ \ lim g \ g n" - by (rule decseq_le, simp add: decseq_SucI, simp add: convergent_LIMSEQ_iff) - -lemma lemma_nest: "[| \n. f(n) \ f(Suc n); - \n. g(Suc n) \ g(n); - \n. f(n) \ g(n) |] - ==> \l m :: real. l \ m & ((\n. f(n) \ l) & f ----> l) & - ((\n. m \ g(n)) & g ----> m)" -apply (subgoal_tac "monoseq f & monoseq g") -prefer 2 apply (force simp add: LIMSEQ_iff monoseq_Suc) -apply (subgoal_tac "Bseq f & Bseq g") -prefer 2 apply (blast intro: f_inc_g_dec_Beq_f f_inc_g_dec_Beq_g) -apply (auto dest!: Bseq_monoseq_convergent simp add: convergent_LIMSEQ_iff) -apply (rule_tac x = "lim f" in exI) -apply (rule_tac x = "lim g" in exI) -apply (auto intro: LIMSEQ_le) -apply (auto simp add: f_inc_imp_le_lim g_dec_imp_lim_le convergent_LIMSEQ_iff) -done +lemma nested_sequence_unique: + assumes "\n. f n \ f (Suc n)" "\n. g (Suc n) \ g n" "\n. f n \ g n" "(\n. f n - g n) ----> 0" + shows "\l::real. ((\n. f n \ l) \ f ----> l) \ ((\n. l \ g n) \ g ----> l)" +proof - + have "incseq f" unfolding incseq_Suc_iff by fact + have "decseq g" unfolding decseq_Suc_iff by fact -lemma lemma_nest_unique: "[| \n. f(n) \ f(Suc n); - \n. g(Suc n) \ g(n); - \n. f(n) \ g(n); - (%n. f(n) - g(n)) ----> 0 |] - ==> \l::real. ((\n. f(n) \ l) & f ----> l) & - ((\n. l \ g(n)) & g ----> l)" -apply (drule lemma_nest, auto) -apply (subgoal_tac "l = m") -apply (drule_tac [2] f = f in tendsto_diff) -apply (auto intro: LIMSEQ_unique) -done - -text{*The universal quantifiers below are required for the declaration - of @{text Bolzano_nest_unique} below.*} - -lemma Bolzano_bisect_le: - "a \ b ==> \n. fst (Bolzano_bisect P a b n) \ snd (Bolzano_bisect P a b n)" -apply (rule allI) -apply (induct_tac "n") -apply (auto simp add: Let_def split_def) -done - -lemma Bolzano_bisect_fst_le_Suc: "a \ b ==> - \n. fst(Bolzano_bisect P a b n) \ fst (Bolzano_bisect P a b (Suc n))" -apply (rule allI) -apply (induct_tac "n") -apply (auto simp add: Bolzano_bisect_le Let_def split_def) -done - -lemma Bolzano_bisect_Suc_le_snd: "a \ b ==> - \n. snd(Bolzano_bisect P a b (Suc n)) \ snd (Bolzano_bisect P a b n)" -apply (rule allI) -apply (induct_tac "n") -apply (auto simp add: Bolzano_bisect_le Let_def split_def) -done - -lemma eq_divide_2_times_iff: "((x::real) = y / (2 * z)) = (2 * x = y/z)" -apply (auto) -apply (drule_tac f = "%u. (1/2) *u" in arg_cong) -apply (simp) -done - -lemma Bolzano_bisect_diff: - "a \ b ==> - snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = - (b-a) / (2 ^ n)" -apply (induct "n") -apply (auto simp add: eq_divide_2_times_iff add_divide_distrib Let_def split_def) -done - -lemmas Bolzano_nest_unique = - lemma_nest_unique - [OF Bolzano_bisect_fst_le_Suc Bolzano_bisect_Suc_le_snd Bolzano_bisect_le] - - -lemma not_P_Bolzano_bisect: - assumes P: "!!a b c. [| P(a,b); P(b,c); a \ b; b \ c|] ==> P(a,c)" - and notP: "~ P(a,b)" - and le: "a \ b" - shows "~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" -proof (induct n) - case 0 show ?case using notP by simp - next - case (Suc n) - thus ?case - by (auto simp del: surjective_pairing [symmetric] - simp add: Let_def split_def Bolzano_bisect_le [OF le] - P [of "fst (Bolzano_bisect P a b n)" _ "snd (Bolzano_bisect P a b n)"]) + { fix n + from `decseq g` have "g n \ g 0" by (rule decseqD) simp + with `\n. f n \ g n`[THEN spec, of n] have "f n \ g 0" by auto } + then obtain u where "f ----> u" "\i. f i \ u" + using incseq_convergent[OF `incseq f`] by auto + moreover + { fix n + from `incseq f` have "f 0 \ f n" by (rule incseqD) simp + with `\n. f n \ g n`[THEN spec, of n] have "f 0 \ g n" by simp } + then obtain l where "g ----> l" "\i. l \ g i" + using decseq_convergent[OF `decseq g`] by auto + moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF `f ----> u` `g ----> l`]] + ultimately show ?thesis by auto qed -(*Now we re-package P_prem as a formula*) -lemma not_P_Bolzano_bisect': - "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); - ~ P(a,b); a \ b |] ==> - \n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" -by (blast elim!: not_P_Bolzano_bisect [THEN [2] rev_notE]) - - +lemma Bolzano[consumes 1, case_names trans local]: + fixes P :: "real \ real \ bool" + assumes [arith]: "a \ b" + assumes trans: "\a b c. \P a b; P b c; a \ b; b \ c\ \ P a c" + assumes local: "\x. a \ x \ x \ b \ \d>0. \a b. a \ x \ x \ b \ b - a < d \ P a b" + shows "P a b" +proof - + def bisect \ "nat_rec (a, b) (\n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))" + def l \ "\n. fst (bisect n)" and u \ "\n. snd (bisect n)" + have l[simp]: "l 0 = a" "\n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)" + and u[simp]: "u 0 = b" "\n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)" + by (simp_all add: l_def u_def bisect_def split: prod.split) -lemma lemma_BOLZANO: - "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); - \x. \d::real. 0 < d & - (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)); - a \ b |] - ==> P(a,b)" -apply (rule Bolzano_nest_unique [where P=P, THEN exE], assumption+) -apply (rule tendsto_minus_cancel) -apply (simp (no_asm_simp) add: Bolzano_bisect_diff LIMSEQ_divide_realpow_zero) -apply (rule ccontr) -apply (drule not_P_Bolzano_bisect', assumption+) -apply (rename_tac "l") -apply (drule_tac x = l in spec, clarify) -apply (simp add: LIMSEQ_iff) -apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) -apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) -apply (drule real_less_half_sum, auto) -apply (drule_tac x = "fst (Bolzano_bisect P a b (no + noa))" in spec) -apply (drule_tac x = "snd (Bolzano_bisect P a b (no + noa))" in spec) -apply safe -apply (simp_all (no_asm_simp)) -apply (rule_tac y = "abs (fst (Bolzano_bisect P a b (no + noa)) - l) + abs (snd (Bolzano_bisect P a b (no + noa)) - l)" in order_le_less_trans) -apply (simp (no_asm_simp) add: abs_if) -apply (rule real_sum_of_halves [THEN subst]) -apply (rule add_strict_mono) -apply (simp_all add: diff_minus [symmetric]) -done - + { fix n have "l n \ u n" by (induct n) auto } note this[simp] -lemma lemma_BOLZANO2: "((\a b c. (a \ b & b \ c & P(a,b) & P(b,c)) --> P(a,c)) & - (\x. \d::real. 0 < d & - (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)))) - --> (\a b. a \ b --> P(a,b))" -apply clarify -apply (blast intro: lemma_BOLZANO) -done - - -subsection {* Intermediate Value Theorem *} - -text {*Prove Contrapositive by Bisection*} + have "\x. ((\n. l n \ x) \ l ----> x) \ ((\n. x \ u n) \ u ----> x)" + proof (safe intro!: nested_sequence_unique) + fix n show "l n \ l (Suc n)" "u (Suc n) \ u n" by (induct n) auto + next + { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) } + then show "(\n. l n - u n) ----> 0" by (simp add: LIMSEQ_divide_realpow_zero) + qed fact + then obtain x where x: "\n. l n \ x" "\n. x \ u n" and "l ----> x" "u ----> x" by auto + obtain d where "0 < d" and d: "\a b. a \ x \ x \ b \ b - a < d \ P a b" + using `l 0 \ x` `x \ u 0` local[of x] by auto -lemma IVT: "[| f(a::real) \ (y::real); y \ f(b); - a \ b; - (\x. a \ x & x \ b --> isCont f x) |] - ==> \x. a \ x & x \ b & f(x) = y" -apply (rule contrapos_pp, assumption) -apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> ~ (f (u) \ y & y \ f (v))" in lemma_BOLZANO2) -apply safe -apply simp_all -apply (simp add: isCont_iff LIM_eq) -apply (rule ccontr) -apply (subgoal_tac "a \ x & x \ b") - prefer 2 - apply simp - apply (drule_tac P = "%d. 0 ?P d" and x = 1 in spec, arith) -apply (drule_tac x = x in spec)+ -apply simp -apply (drule_tac P = "%r. ?P r --> (\s>0. ?Q r s) " and x = "\y - f x\" in spec) -apply safe -apply simp -apply (drule_tac x = s in spec, clarify) -apply (cut_tac x = "f x" and y = y in linorder_less_linear, safe) -apply (drule_tac x = "ba-x" in spec) -apply (simp_all add: abs_if) -apply (drule_tac x = "aa-x" in spec) -apply (case_tac "x \ aa", simp_all) -done - -lemma IVT2: "[| f(b::real) \ (y::real); y \ f(a); - a \ b; - (\x. a \ x & x \ b --> isCont f x) - |] ==> \x. a \ x & x \ b & f(x) = y" -apply (subgoal_tac "- f a \ -y & -y \ - f b", clarify) -apply (drule IVT [where f = "%x. - f x"], assumption) -apply simp_all -done + show "P a b" + proof (rule ccontr) + assume "\ P a b" + { fix n have "\ P (l n) (u n)" + proof (induct n) + case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto + qed (simp add: `\ P a b`) } + moreover + { have "eventually (\n. x - d / 2 < l n) sequentially" + using `0 < d` `l ----> x` by (intro order_tendstoD[of _ x]) auto + moreover have "eventually (\n. u n < x + d / 2) sequentially" + using `0 < d` `u ----> x` by (intro order_tendstoD[of _ x]) auto + ultimately have "eventually (\n. P (l n) (u n)) sequentially" + proof eventually_elim + fix n assume "x - d / 2 < l n" "u n < x + d / 2" + from add_strict_mono[OF this] have "u n - l n < d" by simp + with x show "P (l n) (u n)" by (rule d) + qed } + ultimately show False by simp + qed +qed (*HOL style here: object-level formulations*) lemma IVT_objl: "(f(a::real) \ (y::real) & y \ f(b) & a \ b & @@ -668,30 +509,65 @@ done +lemma compact_Icc[simp, intro]: "compact {a .. b::real}" +proof (cases "a \ b", rule compactI) + fix C assume C: "a \ b" "\t\C. open t" "{a..b} \ \C" + def T == "{a .. b}" + from C(1,3) show "\C'\C. finite C' \ {a..b} \ \C'" + proof (induct rule: Bolzano) + case (trans a b c) + then have *: "{a .. c} = {a .. b} \ {b .. c}" by auto + from trans obtain C1 C2 where "C1\C \ finite C1 \ {a..b} \ \C1" "C2\C \ finite C2 \ {b..c} \ \C2" + by (auto simp: *) + with trans show ?case + unfolding * by (intro exI[of _ "C1 \ C2"]) auto + next + case (local x) + then have "x \ \C" using C by auto + with C(2) obtain c where "x \ c" "open c" "c \ C" by auto + then obtain e where "0 < e" "{x - e <..< x + e} \ c" + by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff) + with `c \ C` show ?case + by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto + qed +qed simp + subsection {* Boundedness of continuous functions *} text{*By bisection, function continuous on closed interval is bounded above*} +lemma isCont_eq_Ub: + fixes f :: "real \ 'a::linorder_topology" + shows "a \ b \ \x::real. a \ x \ x \ b \ isCont f x \ + \M. (\x. a \ x \ x \ b \ f x \ M) \ (\x. a \ x \ x \ b \ f x = M)" + using continuous_attains_sup[of "{a .. b}" f] + apply (simp add: continuous_at_imp_continuous_on Ball_def) + apply safe + apply (rule_tac x="f x" in exI) + apply auto + done + +lemma isCont_eq_Lb: + fixes f :: "real \ 'a::linorder_topology" + shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ + \M. (\x. a \ x \ x \ b \ M \ f x) \ (\x. a \ x \ x \ b \ f x = M)" + using continuous_attains_inf[of "{a .. b}" f] + apply (simp add: continuous_at_imp_continuous_on Ball_def) + apply safe + apply (rule_tac x="f x" in exI) + apply auto + done + lemma isCont_bounded: - "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. \x::real. a \ x & x \ b --> f(x) \ M" -apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> (\M. \x. u \ x & x \ v --> f x \ M)" in lemma_BOLZANO2) -apply safe -apply simp_all -apply (rename_tac x xa ya M Ma) -apply (metis linorder_not_less order_le_less order_trans) -apply (case_tac "a \ x & x \ b") - prefer 2 - apply (rule_tac x = 1 in exI, force) -apply (simp add: LIM_eq isCont_iff) -apply (drule_tac x = x in spec, auto) -apply (erule_tac V = "\M. \x. a \ x & x \ b & ~ f x \ M" in thin_rl) -apply (drule_tac x = 1 in spec, auto) -apply (rule_tac x = s in exI, clarify) -apply (rule_tac x = "\f x\ + 1" in exI, clarify) -apply (drule_tac x = "xa-x" in spec) -apply (auto simp add: abs_ge_self) -done + fixes f :: "real \ 'a::linorder_topology" + shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. \x. a \ x \ x \ b \ f x \ M" + using isCont_eq_Ub[of a b f] by auto + +lemma isCont_has_Ub: + fixes f :: "real \ 'a::linorder_topology" + shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ + \M. (\x. a \ x \ x \ b \ f x \ M) \ (\N. N < M \ (\x. a \ x \ x \ b \ N < f x))" + using isCont_eq_Ub[of a b f] by auto text{*Refine the above to existence of least upper bound*} @@ -699,72 +575,6 @@ (\t. isLub UNIV S t)" by (blast intro: reals_complete) -lemma isCont_has_Ub: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. (\x::real. a \ x & x \ b --> f(x) \ M) & - (\N. N < M --> (\x. a \ x & x \ b & N < f(x)))" -apply (cut_tac S = "Collect (%y. \x. a \ x & x \ b & y = f x)" - in lemma_reals_complete) -apply auto -apply (drule isCont_bounded, assumption) -apply (auto simp add: isUb_def leastP_def isLub_def setge_def setle_def) -apply (rule exI, auto) -apply (auto dest!: spec simp add: linorder_not_less) -done - -text{*Now show that it attains its upper bound*} - -lemma isCont_eq_Ub: - assumes le: "a \ b" - and con: "\x::real. a \ x & x \ b --> isCont f x" - shows "\M::real. (\x. a \ x & x \ b --> f(x) \ M) & - (\x. a \ x & x \ b & f(x) = M)" -proof - - from isCont_has_Ub [OF le con] - obtain M where M1: "\x. a \ x \ x \ b \ f x \ M" - and M2: "!!N. N \x. a \ x \ x \ b \ N < f x" by blast - show ?thesis - proof (intro exI, intro conjI) - show " \x. a \ x \ x \ b \ f x \ M" by (rule M1) - show "\x. a \ x \ x \ b \ f x = M" - proof (rule ccontr) - assume "\ (\x. a \ x \ x \ b \ f x = M)" - with M1 have M3: "\x. a \ x & x \ b --> f x < M" - by (fastforce simp add: linorder_not_le [symmetric]) - hence "\x. a \ x & x \ b --> isCont (%x. inverse (M - f x)) x" - by (auto simp add: con) - from isCont_bounded [OF le this] - obtain k where k: "!!x. a \ x & x \ b --> inverse (M - f x) \ k" by auto - have Minv: "!!x. a \ x & x \ b --> 0 < inverse (M - f (x))" - by (simp add: M3 algebra_simps) - have "!!x. a \ x & x \ b --> inverse (M - f x) < k+1" using k - by (auto intro: order_le_less_trans [of _ k]) - with Minv - have "!!x. a \ x & x \ b --> inverse(k+1) < inverse(inverse(M - f x))" - by (intro strip less_imp_inverse_less, simp_all) - hence invlt: "!!x. a \ x & x \ b --> inverse(k+1) < M - f x" - by simp - have "M - inverse (k+1) < M" using k [of a] Minv [of a] le - by (simp, arith) - from M2 [OF this] - obtain x where ax: "a \ x & x \ b & M - inverse(k+1) < f x" .. - thus False using invlt [of x] by force - qed - qed -qed - - -text{*Same theorem for lower bound*} - -lemma isCont_eq_Lb: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. (\x::real. a \ x & x \ b --> M \ f(x)) & - (\x. a \ x & x \ b & f(x) = M)" -apply (subgoal_tac "\x. a \ x & x \ b --> isCont (%x. - (f x)) x") -prefer 2 apply (blast intro: isCont_minus) -apply (drule_tac f = "(%x. - (f x))" in isCont_eq_Ub) -apply safe -apply auto -done - text{*Another version.*} @@ -1021,16 +831,7 @@ lemma lemma_MVT: "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)" -proof cases - assume "a=b" thus ?thesis by simp -next - assume "a\b" - hence ba: "b-a \ 0" by arith - show ?thesis - by (rule real_mult_left_cancel [OF ba, THEN iffD1], - simp add: right_diff_distrib, - simp add: left_diff_distrib) -qed + by (cases "a = b") (simp_all add: field_simps) theorem MVT: assumes lt: "a < b" @@ -1280,152 +1081,47 @@ by simp qed -subsection {* Continuous injective functions *} - -text{*Dull lemma: an continuous injection on an interval must have a -strict maximum at an end point, not in the middle.*} - -lemma lemma_isCont_inj: - fixes f :: "real \ real" - assumes d: "0 < d" - and inj [rule_format]: "\z. \z-x\ \ d --> g(f z) = z" - and cont: "\z. \z-x\ \ d --> isCont f z" - shows "\z. \z-x\ \ d & f x < f z" -proof (rule ccontr) - assume "~ (\z. \z-x\ \ d & f x < f z)" - hence all [rule_format]: "\z. \z - x\ \ d --> f z \ f x" by auto - show False - proof (cases rule: linorder_le_cases [of "f(x-d)" "f(x+d)"]) - case le - from d cont all [of "x+d"] - have flef: "f(x+d) \ f x" - and xlex: "x - d \ x" - and cont': "\z. x - d \ z \ z \ x \ isCont f z" - by (auto simp add: abs_if) - from IVT [OF le flef xlex cont'] - obtain x' where "x-d \ x'" "x' \ x" "f x' = f(x+d)" by blast - moreover - hence "g(f x') = g (f(x+d))" by simp - ultimately show False using d inj [of x'] inj [of "x+d"] - by (simp add: abs_le_iff) - next - case ge - from d cont all [of "x-d"] - have flef: "f(x-d) \ f x" - and xlex: "x \ x+d" - and cont': "\z. x \ z \ z \ x+d \ isCont f z" - by (auto simp add: abs_if) - from IVT2 [OF ge flef xlex cont'] - obtain x' where "x \ x'" "x' \ x+d" "f x' = f(x-d)" by blast - moreover - hence "g(f x') = g (f(x-d))" by simp - ultimately show False using d inj [of x'] inj [of "x-d"] - by (simp add: abs_le_iff) - qed -qed - - -text{*Similar version for lower bound.*} - -lemma lemma_isCont_inj2: - fixes f g :: "real \ real" - shows "[|0 < d; \z. \z-x\ \ d --> g(f z) = z; - \z. \z-x\ \ d --> isCont f z |] - ==> \z. \z-x\ \ d & f z < f x" -apply (insert lemma_isCont_inj - [where f = "%x. - f x" and g = "%y. g(-y)" and x = x and d = d]) -apply (simp add: linorder_not_le) -done - -text{*Show there's an interval surrounding @{term "f(x)"} in -@{text "f[[x - d, x + d]]"} .*} - -lemma isCont_inj_range: - fixes f :: "real \ real" - assumes d: "0 < d" - and inj: "\z. \z-x\ \ d --> g(f z) = z" - and cont: "\z. \z-x\ \ d --> isCont f z" - shows "\e>0. \y. \y - f x\ \ e --> (\z. \z-x\ \ d & f z = y)" -proof - - have "x-d \ x+d" "\z. x-d \ z \ z \ x+d \ isCont f z" using cont d - by (auto simp add: abs_le_iff) - from isCont_Lb_Ub [OF this] - obtain L M - where all1 [rule_format]: "\z. x-d \ z \ z \ x+d \ L \ f z \ f z \ M" - and all2 [rule_format]: - "\y. L \ y \ y \ M \ (\z. x-d \ z \ z \ x+d \ f z = y)" - by auto - with d have "L \ f x & f x \ M" by simp - moreover have "L \ f x" - proof - - from lemma_isCont_inj2 [OF d inj cont] - obtain u where "\u - x\ \ d" "f u < f x" by auto - thus ?thesis using all1 [of u] by arith - qed - moreover have "f x \ M" - proof - - from lemma_isCont_inj [OF d inj cont] - obtain u where "\u - x\ \ d" "f x < f u" by auto - thus ?thesis using all1 [of u] by arith - qed - ultimately have "L < f x & f x < M" by arith - hence "0 < f x - L" "0 < M - f x" by arith+ - from real_lbound_gt_zero [OF this] - obtain e where e: "0 < e" "e < f x - L" "e < M - f x" by auto - thus ?thesis - proof (intro exI conjI) - show "0y. \y - f x\ \ e \ (\z. \z - x\ \ d \ f z = y)" - proof (intro strip) - fix y::real - assume "\y - f x\ \ e" - with e have "L \ y \ y \ M" by arith - from all2 [OF this] - obtain z where "x - d \ z" "z \ x + d" "f z = y" by blast - thus "\z. \z - x\ \ d \ f z = y" - by (force simp add: abs_le_iff) - qed - qed -qed - - text{*Continuity of inverse function*} lemma isCont_inverse_function: fixes f g :: "real \ real" assumes d: "0 < d" - and inj: "\z. \z-x\ \ d --> g(f z) = z" - and cont: "\z. \z-x\ \ d --> isCont f z" + and inj: "\z. \z-x\ \ d \ g (f z) = z" + and cont: "\z. \z-x\ \ d \ isCont f z" shows "isCont g (f x)" -proof (simp add: isCont_iff LIM_eq) - show "\r. 0 < r \ - (\s>0. \z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r)" - proof (intro strip) - fix r::real - assume r: "0 e < d" by blast - with inj cont - have e_simps: "\z. \z-x\ \ e --> g (f z) = z" - "\z. \z-x\ \ e --> isCont f z" by auto - from isCont_inj_range [OF e this] - obtain e' where e': "0 < e'" - and all: "\y. \y - f x\ \ e' \ (\z. \z - x\ \ e \ f z = y)" - by blast - show "\s>0. \z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r" - proof (intro exI conjI) - show "0z. z \ 0 \ \z\ < e' \ \g (f x + z) - g (f x)\ < r" - proof (intro strip) - fix z::real - assume z: "z \ 0 \ \z\ < e'" - with e e_lt e_simps all [rule_format, of "f x + z"] - show "\g (f x + z) - g (f x)\ < r" by force - qed - qed - qed +proof - + let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + d}" + + have f: "continuous_on ?D f" + using cont by (intro continuous_at_imp_continuous_on ballI) auto + then have g: "continuous_on (f`?D) g" + using inj by (intro continuous_on_inv) auto + + from d f have "{min ?A ?B <..< max ?A ?B} \ f ` ?D" + by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max) + with g have "continuous_on {min ?A ?B <..< max ?A ?B} g" + by (rule continuous_on_subset) + moreover + have "(?A < f x \ f x < ?B) \ (?B < f x \ f x < ?A)" + using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto + then have "f x \ {min ?A ?B <..< max ?A ?B}" + by auto + ultimately + show ?thesis + by (simp add: continuous_on_eq_continuous_at) qed +lemma isCont_inverse_function2: + fixes f g :: "real \ real" shows + "\a < x; x < b; + \z. a \ z \ z \ b \ g (f z) = z; + \z. a \ z \ z \ b \ isCont f z\ + \ isCont g (f x)" +apply (rule isCont_inverse_function + [where f=f and d="min (x - a) (b - x)"]) +apply (simp_all add: abs_le_iff) +done + text {* Derivative of inverse function *} lemma DERIV_inverse_function: @@ -1475,7 +1171,6 @@ using neq by (rule tendsto_inverse) qed - subsection {* Generalized Mean Value Theorem *} theorem GMVT: @@ -1545,21 +1240,6 @@ ==> isCont g (f x)" by (rule isCont_inverse_function) -lemma isCont_inv_fun_inv: - fixes f g :: "real \ real" - shows "[| 0 < d; - \z. \z - x\ \ d --> g(f(z)) = z; - \z. \z - x\ \ d --> isCont f z |] - ==> \e. 0 < e & - (\y. 0 < \y - f(x)\ & \y - f(x)\ < e --> f(g(y)) = y)" -apply (drule isCont_inj_range) -prefer 2 apply (assumption, assumption, auto) -apply (rule_tac x = e in exI, auto) -apply (rotate_tac 2) -apply (drule_tac x = y in spec, auto) -done - - text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*} lemma LIM_fun_gt_zero: "[| f -- c --> (l::real); 0 < l |] diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Library/Product_Vector.thy --- a/src/HOL/Library/Product_Vector.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Library/Product_Vector.thy Sat Mar 23 07:30:53 2013 +0100 @@ -117,15 +117,6 @@ by (simp add: closed_vimage_fst closed_vimage_snd closed_Int) qed -lemma openI: (* TODO: move *) - assumes "\x. x \ S \ \T. open T \ x \ T \ T \ S" - shows "open S" -proof - - have "open (\{T. open T \ T \ S})" by auto - moreover have "\{T. open T \ T \ S} = S" by (auto dest!: assms) - ultimately show "open S" by simp -qed - lemma subset_fst_imageI: "A \ B \ S \ y \ B \ A \ fst ` S" unfolding image_def subset_eq by force @@ -202,15 +193,23 @@ (simp add: subsetD [OF `A \ B \ S`]) qed +lemma continuous_fst[continuous_intros]: "continuous F f \ continuous F (\x. fst (f x))" + unfolding continuous_def by (rule tendsto_fst) + +lemma continuous_snd[continuous_intros]: "continuous F f \ continuous F (\x. snd (f x))" + unfolding continuous_def by (rule tendsto_snd) + +lemma continuous_Pair[continuous_intros]: "continuous F f \ continuous F g \ continuous F (\x. (f x, g x))" + unfolding continuous_def by (rule tendsto_Pair) + lemma isCont_fst [simp]: "isCont f a \ isCont (\x. fst (f x)) a" - unfolding isCont_def by (rule tendsto_fst) + by (fact continuous_fst) lemma isCont_snd [simp]: "isCont f a \ isCont (\x. snd (f x)) a" - unfolding isCont_def by (rule tendsto_snd) + by (fact continuous_snd) -lemma isCont_Pair [simp]: - "\isCont f a; isCont g a\ \ isCont (\x. (f x, g x)) a" - unfolding isCont_def by (rule tendsto_Pair) +lemma isCont_Pair [simp]: "\isCont f a; isCont g a\ \ isCont (\x. (f x, g x)) a" + by (fact continuous_Pair) subsubsection {* Separation axioms *} diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Lim.thy --- a/src/HOL/Lim.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Lim.thy Sat Mar 23 07:30:53 2013 +0100 @@ -10,38 +10,8 @@ imports SEQ begin -text{*Standard Definitions*} - -abbreviation - LIM :: "['a::topological_space \ 'b::topological_space, 'a, 'b] \ bool" - ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where - "f -- a --> L \ (f ---> L) (at a)" - -definition - isCont :: "['a::topological_space \ 'b::topological_space, 'a] \ bool" where - "isCont f a = (f -- a --> (f a))" - -definition - isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where - "isUCont f = (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" - subsection {* Limits of Functions *} -lemma LIM_def: "f -- a --> L = - (\r > 0. \s > 0. \x. x \ a & dist x a < s - --> dist (f x) L < r)" -unfolding tendsto_iff eventually_at .. - -lemma metric_LIM_I: - "(\r. 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r) - \ f -- a --> L" -by (simp add: LIM_def) - -lemma metric_LIM_D: - "\f -- a --> L; 0 < r\ - \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r" -by (simp add: LIM_def) - lemma LIM_eq: fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" shows "f -- a --> L = @@ -81,8 +51,6 @@ shows "(\h. f (a + h)) -- 0 --> L \ f -- a --> L" by (drule_tac k="- a" in LIM_offset, simp) -lemma LIM_cong_limit: "\ f -- x --> L ; K = L \ \ f -- x --> K" by simp - lemma LIM_zero: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "(f ---> l) F \ ((\x. f x - l) ---> 0) F" @@ -98,13 +66,6 @@ shows "((\x. f x - l) ---> 0) F = (f ---> l) F" unfolding tendsto_iff dist_norm by simp -lemma metric_LIM_imp_LIM: - assumes f: "f -- a --> l" - assumes le: "\x. x \ a \ dist (g x) m \ dist (f x) l" - shows "g -- a --> m" - by (rule metric_tendsto_imp_tendsto [OF f], - auto simp add: eventually_at_topological le) - lemma LIM_imp_LIM: fixes f :: "'a::topological_space \ 'b::real_normed_vector" fixes g :: "'a::topological_space \ 'c::real_normed_vector" @@ -114,48 +75,6 @@ by (rule metric_LIM_imp_LIM [OF f], simp add: dist_norm le) -lemma LIM_const_not_eq: - fixes a :: "'a::perfect_space" - fixes k L :: "'b::t2_space" - shows "k \ L \ \ (\x. k) -- a --> L" - by (simp add: tendsto_const_iff) - -lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] - -lemma LIM_const_eq: - fixes a :: "'a::perfect_space" - fixes k L :: "'b::t2_space" - shows "(\x. k) -- a --> L \ k = L" - by (simp add: tendsto_const_iff) - -lemma LIM_unique: - fixes a :: "'a::perfect_space" - fixes L M :: "'b::t2_space" - shows "\f -- a --> L; f -- a --> M\ \ L = M" - using at_neq_bot by (rule tendsto_unique) - -text{*Limits are equal for functions equal except at limit point*} -lemma LIM_equal: - "[| \x. x \ a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)" -unfolding tendsto_def eventually_at_topological by simp - -lemma LIM_cong: - "\a = b; \x. x \ b \ f x = g x; l = m\ - \ ((\x. f x) -- a --> l) = ((\x. g x) -- b --> m)" -by (simp add: LIM_equal) - -lemma metric_LIM_equal2: - assumes 1: "0 < R" - assumes 2: "\x. \x \ a; dist x a < R\ \ f x = g x" - shows "g -- a --> l \ f -- a --> l" -apply (rule topological_tendstoI) -apply (drule (2) topological_tendstoD) -apply (simp add: eventually_at, safe) -apply (rule_tac x="min d R" in exI, safe) -apply (simp add: 1) -apply (simp add: 2) -done - lemma LIM_equal2: fixes f g :: "'a::real_normed_vector \ 'b::topological_space" assumes 1: "0 < R" @@ -163,21 +82,6 @@ shows "g -- a --> l \ f -- a --> l" by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm) -lemma LIM_compose_eventually: - assumes f: "f -- a --> b" - assumes g: "g -- b --> c" - assumes inj: "eventually (\x. f x \ b) (at a)" - shows "(\x. g (f x)) -- a --> c" - using g f inj by (rule tendsto_compose_eventually) - -lemma metric_LIM_compose2: - assumes f: "f -- a --> b" - assumes g: "g -- b --> c" - assumes inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ b" - shows "(\x. g (f x)) -- a --> c" - using g f inj [folded eventually_at] - by (rule tendsto_compose_eventually) - lemma LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f: "f -- a --> b" @@ -186,16 +90,13 @@ shows "(\x. g (f x)) -- a --> c" by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) -lemma LIM_o: "\g -- l --> g l; f -- a --> l\ \ (g \ f) -- a --> g l" - unfolding o_def by (rule tendsto_compose) - lemma real_LIM_sandwich_zero: fixes f g :: "'a::topological_space \ real" assumes f: "f -- a --> 0" assumes 1: "\x. x \ a \ 0 \ g x" assumes 2: "\x. x \ a \ g x \ f x" shows "g -- a --> 0" -proof (rule LIM_imp_LIM [OF f]) +proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *) fix x assume x: "x \ a" have "norm (g x - 0) = g x" by (simp add: 1 x) also have "g x \ f x" by (rule 2 [OF x]) @@ -217,63 +118,6 @@ shows "isCont f x = (\h. f (x + h)) -- 0 --> f x" by (simp add: isCont_def LIM_isCont_iff) -lemma isCont_ident [simp]: "isCont (\x. x) a" - unfolding isCont_def by (rule tendsto_ident_at) - -lemma isCont_const [simp]: "isCont (\x. k) a" - unfolding isCont_def by (rule tendsto_const) - -lemma isCont_norm [simp]: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "isCont f a \ isCont (\x. norm (f x)) a" - unfolding isCont_def by (rule tendsto_norm) - -lemma isCont_rabs [simp]: - fixes f :: "'a::topological_space \ real" - shows "isCont f a \ isCont (\x. \f x\) a" - unfolding isCont_def by (rule tendsto_rabs) - -lemma isCont_add [simp]: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "\isCont f a; isCont g a\ \ isCont (\x. f x + g x) a" - unfolding isCont_def by (rule tendsto_add) - -lemma isCont_minus [simp]: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "isCont f a \ isCont (\x. - f x) a" - unfolding isCont_def by (rule tendsto_minus) - -lemma isCont_diff [simp]: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "\isCont f a; isCont g a\ \ isCont (\x. f x - g x) a" - unfolding isCont_def by (rule tendsto_diff) - -lemma isCont_mult [simp]: - fixes f g :: "'a::topological_space \ 'b::real_normed_algebra" - shows "\isCont f a; isCont g a\ \ isCont (\x. f x * g x) a" - unfolding isCont_def by (rule tendsto_mult) - -lemma isCont_inverse [simp]: - fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" - shows "\isCont f a; f a \ 0\ \ isCont (\x. inverse (f x)) a" - unfolding isCont_def by (rule tendsto_inverse) - -lemma isCont_divide [simp]: - fixes f g :: "'a::topological_space \ 'b::real_normed_field" - shows "\isCont f a; isCont g a; g a \ 0\ \ isCont (\x. f x / g x) a" - unfolding isCont_def by (rule tendsto_divide) - -lemma isCont_tendsto_compose: - "\isCont g l; (f ---> l) F\ \ ((\x. g (f x)) ---> g l) F" - unfolding isCont_def by (rule tendsto_compose) - -lemma metric_isCont_LIM_compose2: - assumes f [unfolded isCont_def]: "isCont f a" - assumes g: "g -- f a --> l" - assumes inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ f a" - shows "(\x. g (f x)) -- a --> l" -by (rule metric_LIM_compose2 [OF f g inj]) - lemma isCont_LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f [unfolded isCont_def]: "isCont f a" @@ -282,41 +126,60 @@ shows "(\x. g (f x)) -- a --> l" by (rule LIM_compose2 [OF f g inj]) -lemma isCont_o2: "\isCont f a; isCont g (f a)\ \ isCont (\x. g (f x)) a" - unfolding isCont_def by (rule tendsto_compose) + +lemma isCont_norm [simp]: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + shows "isCont f a \ isCont (\x. norm (f x)) a" + by (fact continuous_norm) + +lemma isCont_rabs [simp]: + fixes f :: "'a::t2_space \ real" + shows "isCont f a \ isCont (\x. \f x\) a" + by (fact continuous_rabs) -lemma isCont_o: "\isCont f a; isCont g (f a)\ \ isCont (g o f) a" - unfolding o_def by (rule isCont_o2) +lemma isCont_add [simp]: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + shows "\isCont f a; isCont g a\ \ isCont (\x. f x + g x) a" + by (fact continuous_add) + +lemma isCont_minus [simp]: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + shows "isCont f a \ isCont (\x. - f x) a" + by (fact continuous_minus) + +lemma isCont_diff [simp]: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + shows "\isCont f a; isCont g a\ \ isCont (\x. f x - g x) a" + by (fact continuous_diff) + +lemma isCont_mult [simp]: + fixes f g :: "'a::t2_space \ 'b::real_normed_algebra" + shows "\isCont f a; isCont g a\ \ isCont (\x. f x * g x) a" + by (fact continuous_mult) lemma (in bounded_linear) isCont: "isCont g a \ isCont (\x. f (g x)) a" - unfolding isCont_def by (rule tendsto) + by (fact continuous) lemma (in bounded_bilinear) isCont: "\isCont f a; isCont g a\ \ isCont (\x. f x ** g x) a" - unfolding isCont_def by (rule tendsto) + by (fact continuous) -lemmas isCont_scaleR [simp] = +lemmas isCont_scaleR [simp] = bounded_bilinear.isCont [OF bounded_bilinear_scaleR] lemmas isCont_of_real [simp] = bounded_linear.isCont [OF bounded_linear_of_real] lemma isCont_power [simp]: - fixes f :: "'a::topological_space \ 'b::{power,real_normed_algebra}" + fixes f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" shows "isCont f a \ isCont (\x. f x ^ n) a" - unfolding isCont_def by (rule tendsto_power) - -lemma isCont_sgn [simp]: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "\isCont f a; f a \ 0\ \ isCont (\x. sgn (f x)) a" - unfolding isCont_def by (rule tendsto_sgn) + by (fact continuous_power) lemma isCont_setsum [simp]: - fixes f :: "'a \ 'b::topological_space \ 'c::real_normed_vector" - fixes A :: "'a set" + fixes f :: "'a \ 'b::t2_space \ 'c::real_normed_vector" shows "\i\A. isCont (f i) a \ isCont (\x. \i\A. f i x) a" - unfolding isCont_def by (simp add: tendsto_setsum) + by (auto intro: continuous_setsum) lemmas isCont_intros = isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus @@ -325,18 +188,6 @@ subsection {* Uniform Continuity *} -lemma isUCont_isCont: "isUCont f ==> isCont f x" -by (simp add: isUCont_def isCont_def LIM_def, force) - -lemma isUCont_Cauchy: - "\isUCont f; Cauchy X\ \ Cauchy (\n. f (X n))" -unfolding isUCont_def -apply (rule metric_CauchyI) -apply (drule_tac x=e in spec, safe) -apply (drule_tac e=s in metric_CauchyD, safe) -apply (rule_tac x=M in exI, simp) -done - lemma (in bounded_linear) isUCont: "isUCont f" unfolding isUCont_def dist_norm proof (intro allI impI) @@ -360,58 +211,6 @@ by (rule isUCont [THEN isUCont_Cauchy]) -subsection {* Relation of LIM and LIMSEQ *} - -lemma sequentially_imp_eventually_within: - fixes a :: "'a::metric_space" - assumes "\f. (\n. f n \ s \ f n \ a) \ f ----> a \ - eventually (\n. P (f n)) sequentially" - shows "eventually P (at a within s)" -proof (rule ccontr) - let ?I = "\n. inverse (real (Suc n))" - def F \ "\n::nat. SOME x. x \ s \ x \ a \ dist x a < ?I n \ \ P x" - assume "\ eventually P (at a within s)" - hence P: "\d>0. \x. x \ s \ x \ a \ dist x a < d \ \ P x" - unfolding Limits.eventually_within Limits.eventually_at by fast - hence "\n. \x. x \ s \ x \ a \ dist x a < ?I n \ \ P x" by simp - hence F: "\n. F n \ s \ F n \ a \ dist (F n) a < ?I n \ \ P (F n)" - unfolding F_def by (rule someI_ex) - hence F0: "\n. F n \ s" and F1: "\n. F n \ a" - and F2: "\n. dist (F n) a < ?I n" and F3: "\n. \ P (F n)" - by fast+ - from LIMSEQ_inverse_real_of_nat have "F ----> a" - by (rule metric_tendsto_imp_tendsto, - simp add: dist_norm F2 less_imp_le) - hence "eventually (\n. P (F n)) sequentially" - using assms F0 F1 by simp - thus "False" by (simp add: F3) -qed - -lemma sequentially_imp_eventually_at: - fixes a :: "'a::metric_space" - assumes "\f. (\n. f n \ a) \ f ----> a \ - eventually (\n. P (f n)) sequentially" - shows "eventually P (at a)" - using assms sequentially_imp_eventually_within [where s=UNIV] by simp - -lemma LIMSEQ_SEQ_conv1: - fixes f :: "'a::topological_space \ 'b::topological_space" - assumes f: "f -- a --> l" - shows "\S. (\n. S n \ a) \ S ----> a \ (\n. f (S n)) ----> l" - using tendsto_compose_eventually [OF f, where F=sequentially] by simp - -lemma LIMSEQ_SEQ_conv2: - fixes f :: "'a::metric_space \ 'b::topological_space" - assumes "\S. (\n. S n \ a) \ S ----> a \ (\n. f (S n)) ----> l" - shows "f -- a --> l" - using assms unfolding tendsto_def [where l=l] - by (simp add: sequentially_imp_eventually_at) - -lemma LIMSEQ_SEQ_conv: - "(\S. (\n. S n \ a) \ S ----> (a::'a::metric_space) \ (\n. X (S n)) ----> L) = - (X -- a --> (L::'b::topological_space))" - using LIMSEQ_SEQ_conv2 LIMSEQ_SEQ_conv1 .. - lemma LIM_less_bound: fixes f :: "real \ real" assumes ev: "b < x" "\ x' \ { b <..< x}. 0 \ f x'" and "isCont f x" diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Limits.thy Sat Mar 23 07:30:53 2013 +0100 @@ -8,513 +8,9 @@ imports RealVector begin -subsection {* Filters *} - -text {* - This definition also allows non-proper filters. -*} - -locale is_filter = - fixes F :: "('a \ bool) \ bool" - assumes True: "F (\x. True)" - assumes conj: "F (\x. P x) \ F (\x. Q x) \ F (\x. P x \ Q x)" - assumes mono: "\x. P x \ Q x \ F (\x. P x) \ F (\x. Q x)" - -typedef 'a filter = "{F :: ('a \ bool) \ bool. is_filter F}" -proof - show "(\x. True) \ ?filter" by (auto intro: is_filter.intro) -qed - -lemma is_filter_Rep_filter: "is_filter (Rep_filter F)" - using Rep_filter [of F] by simp - -lemma Abs_filter_inverse': - assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F" - using assms by (simp add: Abs_filter_inverse) - - -subsection {* Eventually *} - -definition eventually :: "('a \ bool) \ 'a filter \ bool" - where "eventually P F \ Rep_filter F P" - -lemma eventually_Abs_filter: - assumes "is_filter F" shows "eventually P (Abs_filter F) = F P" - unfolding eventually_def using assms by (simp add: Abs_filter_inverse) - -lemma filter_eq_iff: - shows "F = F' \ (\P. eventually P F = eventually P F')" - unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def .. - -lemma eventually_True [simp]: "eventually (\x. True) F" - unfolding eventually_def - by (rule is_filter.True [OF is_filter_Rep_filter]) - -lemma always_eventually: "\x. P x \ eventually P F" -proof - - assume "\x. P x" hence "P = (\x. True)" by (simp add: ext) - thus "eventually P F" by simp -qed - -lemma eventually_mono: - "(\x. P x \ Q x) \ eventually P F \ eventually Q F" - unfolding eventually_def - by (rule is_filter.mono [OF is_filter_Rep_filter]) - -lemma eventually_conj: - assumes P: "eventually (\x. P x) F" - assumes Q: "eventually (\x. Q x) F" - shows "eventually (\x. P x \ Q x) F" - using assms unfolding eventually_def - by (rule is_filter.conj [OF is_filter_Rep_filter]) - -lemma eventually_Ball_finite: - assumes "finite A" and "\y\A. eventually (\x. P x y) net" - shows "eventually (\x. \y\A. P x y) net" -using assms by (induct set: finite, simp, simp add: eventually_conj) - -lemma eventually_all_finite: - fixes P :: "'a \ 'b::finite \ bool" - assumes "\y. eventually (\x. P x y) net" - shows "eventually (\x. \y. P x y) net" -using eventually_Ball_finite [of UNIV P] assms by simp - -lemma eventually_mp: - assumes "eventually (\x. P x \ Q x) F" - assumes "eventually (\x. P x) F" - shows "eventually (\x. Q x) F" -proof (rule eventually_mono) - show "\x. (P x \ Q x) \ P x \ Q x" by simp - show "eventually (\x. (P x \ Q x) \ P x) F" - using assms by (rule eventually_conj) -qed - -lemma eventually_rev_mp: - assumes "eventually (\x. P x) F" - assumes "eventually (\x. P x \ Q x) F" - shows "eventually (\x. Q x) F" -using assms(2) assms(1) by (rule eventually_mp) - -lemma eventually_conj_iff: - "eventually (\x. P x \ Q x) F \ eventually P F \ eventually Q F" - by (auto intro: eventually_conj elim: eventually_rev_mp) - -lemma eventually_elim1: - assumes "eventually (\i. P i) F" - assumes "\i. P i \ Q i" - shows "eventually (\i. Q i) F" - using assms by (auto elim!: eventually_rev_mp) - -lemma eventually_elim2: - assumes "eventually (\i. P i) F" - assumes "eventually (\i. Q i) F" - assumes "\i. P i \ Q i \ R i" - shows "eventually (\i. R i) F" - using assms by (auto elim!: eventually_rev_mp) - -lemma eventually_subst: - assumes "eventually (\n. P n = Q n) F" - shows "eventually P F = eventually Q F" (is "?L = ?R") -proof - - from assms have "eventually (\x. P x \ Q x) F" - and "eventually (\x. Q x \ P x) F" - by (auto elim: eventually_elim1) - then show ?thesis by (auto elim: eventually_elim2) -qed - -ML {* - fun eventually_elim_tac ctxt thms thm = - let - val thy = Proof_Context.theory_of ctxt - val mp_thms = thms RL [@{thm eventually_rev_mp}] - val raw_elim_thm = - (@{thm allI} RS @{thm always_eventually}) - |> fold (fn thm1 => fn thm2 => thm2 RS thm1) mp_thms - |> fold (fn _ => fn thm => @{thm impI} RS thm) thms - val cases_prop = prop_of (raw_elim_thm RS thm) - val cases = (Rule_Cases.make_common (thy, cases_prop) [(("elim", []), [])]) - in - CASES cases (rtac raw_elim_thm 1) thm - end -*} - -method_setup eventually_elim = {* - Scan.succeed (fn ctxt => METHOD_CASES (eventually_elim_tac ctxt)) -*} "elimination of eventually quantifiers" - - -subsection {* Finer-than relation *} - -text {* @{term "F \ F'"} means that filter @{term F} is finer than -filter @{term F'}. *} - -instantiation filter :: (type) complete_lattice -begin - -definition le_filter_def: - "F \ F' \ (\P. eventually P F' \ eventually P F)" - -definition - "(F :: 'a filter) < F' \ F \ F' \ \ F' \ F" - -definition - "top = Abs_filter (\P. \x. P x)" - -definition - "bot = Abs_filter (\P. True)" - -definition - "sup F F' = Abs_filter (\P. eventually P F \ eventually P F')" - -definition - "inf F F' = Abs_filter - (\P. \Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" - -definition - "Sup S = Abs_filter (\P. \F\S. eventually P F)" - -definition - "Inf S = Sup {F::'a filter. \F'\S. F \ F'}" - -lemma eventually_top [simp]: "eventually P top \ (\x. P x)" - unfolding top_filter_def - by (rule eventually_Abs_filter, rule is_filter.intro, auto) - -lemma eventually_bot [simp]: "eventually P bot" - unfolding bot_filter_def - by (subst eventually_Abs_filter, rule is_filter.intro, auto) - -lemma eventually_sup: - "eventually P (sup F F') \ eventually P F \ eventually P F'" - unfolding sup_filter_def - by (rule eventually_Abs_filter, rule is_filter.intro) - (auto elim!: eventually_rev_mp) - -lemma eventually_inf: - "eventually P (inf F F') \ - (\Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" - unfolding inf_filter_def - apply (rule eventually_Abs_filter, rule is_filter.intro) - apply (fast intro: eventually_True) - apply clarify - apply (intro exI conjI) - apply (erule (1) eventually_conj) - apply (erule (1) eventually_conj) - apply simp - apply auto - done - -lemma eventually_Sup: - "eventually P (Sup S) \ (\F\S. eventually P F)" - unfolding Sup_filter_def - apply (rule eventually_Abs_filter, rule is_filter.intro) - apply (auto intro: eventually_conj elim!: eventually_rev_mp) - done - -instance proof - fix F F' F'' :: "'a filter" and S :: "'a filter set" - { show "F < F' \ F \ F' \ \ F' \ F" - by (rule less_filter_def) } - { show "F \ F" - unfolding le_filter_def by simp } - { assume "F \ F'" and "F' \ F''" thus "F \ F''" - unfolding le_filter_def by simp } - { assume "F \ F'" and "F' \ F" thus "F = F'" - unfolding le_filter_def filter_eq_iff by fast } - { show "F \ top" - unfolding le_filter_def eventually_top by (simp add: always_eventually) } - { show "bot \ F" - unfolding le_filter_def by simp } - { show "F \ sup F F'" and "F' \ sup F F'" - unfolding le_filter_def eventually_sup by simp_all } - { assume "F \ F''" and "F' \ F''" thus "sup F F' \ F''" - unfolding le_filter_def eventually_sup by simp } - { show "inf F F' \ F" and "inf F F' \ F'" - unfolding le_filter_def eventually_inf by (auto intro: eventually_True) } - { assume "F \ F'" and "F \ F''" thus "F \ inf F' F''" - unfolding le_filter_def eventually_inf - by (auto elim!: eventually_mono intro: eventually_conj) } - { assume "F \ S" thus "F \ Sup S" - unfolding le_filter_def eventually_Sup by simp } - { assume "\F. F \ S \ F \ F'" thus "Sup S \ F'" - unfolding le_filter_def eventually_Sup by simp } - { assume "F'' \ S" thus "Inf S \ F''" - unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } - { assume "\F'. F' \ S \ F \ F'" thus "F \ Inf S" - unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } -qed - -end - -lemma filter_leD: - "F \ F' \ eventually P F' \ eventually P F" - unfolding le_filter_def by simp - -lemma filter_leI: - "(\P. eventually P F' \ eventually P F) \ F \ F'" - unfolding le_filter_def by simp - -lemma eventually_False: - "eventually (\x. False) F \ F = bot" - unfolding filter_eq_iff by (auto elim: eventually_rev_mp) - -abbreviation (input) trivial_limit :: "'a filter \ bool" - where "trivial_limit F \ F = bot" - -lemma trivial_limit_def: "trivial_limit F \ eventually (\x. False) F" - by (rule eventually_False [symmetric]) - -lemma eventually_const: "\ trivial_limit net \ eventually (\x. P) net \ P" - by (cases P) (simp_all add: eventually_False) - - -subsection {* Map function for filters *} - -definition filtermap :: "('a \ 'b) \ 'a filter \ 'b filter" - where "filtermap f F = Abs_filter (\P. eventually (\x. P (f x)) F)" - -lemma eventually_filtermap: - "eventually P (filtermap f F) = eventually (\x. P (f x)) F" - unfolding filtermap_def - apply (rule eventually_Abs_filter) - apply (rule is_filter.intro) - apply (auto elim!: eventually_rev_mp) - done - -lemma filtermap_ident: "filtermap (\x. x) F = F" - by (simp add: filter_eq_iff eventually_filtermap) - -lemma filtermap_filtermap: - "filtermap f (filtermap g F) = filtermap (\x. f (g x)) F" - by (simp add: filter_eq_iff eventually_filtermap) - -lemma filtermap_mono: "F \ F' \ filtermap f F \ filtermap f F'" - unfolding le_filter_def eventually_filtermap by simp - -lemma filtermap_bot [simp]: "filtermap f bot = bot" - by (simp add: filter_eq_iff eventually_filtermap) - -lemma filtermap_sup: "filtermap f (sup F1 F2) = sup (filtermap f F1) (filtermap f F2)" - by (auto simp: filter_eq_iff eventually_filtermap eventually_sup) - -subsection {* Order filters *} - -definition at_top :: "('a::order) filter" - where "at_top = Abs_filter (\P. \k. \n\k. P n)" - -lemma eventually_at_top_linorder: "eventually P at_top \ (\N::'a::linorder. \n\N. P n)" - unfolding at_top_def -proof (rule eventually_Abs_filter, rule is_filter.intro) - fix P Q :: "'a \ bool" - assume "\i. \n\i. P n" and "\j. \n\j. Q n" - then obtain i j where "\n\i. P n" and "\n\j. Q n" by auto - then have "\n\max i j. P n \ Q n" by simp - then show "\k. \n\k. P n \ Q n" .. -qed auto - -lemma eventually_ge_at_top: - "eventually (\x. (c::_::linorder) \ x) at_top" - unfolding eventually_at_top_linorder by auto - -lemma eventually_at_top_dense: "eventually P at_top \ (\N::'a::dense_linorder. \n>N. P n)" - unfolding eventually_at_top_linorder -proof safe - fix N assume "\n\N. P n" then show "\N. \n>N. P n" by (auto intro!: exI[of _ N]) -next - fix N assume "\n>N. P n" - moreover from gt_ex[of N] guess y .. - ultimately show "\N. \n\N. P n" by (auto intro!: exI[of _ y]) -qed - -lemma eventually_gt_at_top: - "eventually (\x. (c::_::dense_linorder) < x) at_top" - unfolding eventually_at_top_dense by auto - -definition at_bot :: "('a::order) filter" - where "at_bot = Abs_filter (\P. \k. \n\k. P n)" - -lemma eventually_at_bot_linorder: - fixes P :: "'a::linorder \ bool" shows "eventually P at_bot \ (\N. \n\N. P n)" - unfolding at_bot_def -proof (rule eventually_Abs_filter, rule is_filter.intro) - fix P Q :: "'a \ bool" - assume "\i. \n\i. P n" and "\j. \n\j. Q n" - then obtain i j where "\n\i. P n" and "\n\j. Q n" by auto - then have "\n\min i j. P n \ Q n" by simp - then show "\k. \n\k. P n \ Q n" .. -qed auto - -lemma eventually_le_at_bot: - "eventually (\x. x \ (c::_::linorder)) at_bot" - unfolding eventually_at_bot_linorder by auto - -lemma eventually_at_bot_dense: - fixes P :: "'a::dense_linorder \ bool" shows "eventually P at_bot \ (\N. \nn\N. P n" then show "\N. \nnN. \n\N. P n" by (auto intro!: exI[of _ y]) -qed - -lemma eventually_gt_at_bot: - "eventually (\x. x < (c::_::dense_linorder)) at_bot" - unfolding eventually_at_bot_dense by auto - -subsection {* Sequentially *} - -abbreviation sequentially :: "nat filter" - where "sequentially == at_top" - -lemma sequentially_def: "sequentially = Abs_filter (\P. \k. \n\k. P n)" - unfolding at_top_def by simp - -lemma eventually_sequentially: - "eventually P sequentially \ (\N. \n\N. P n)" - by (rule eventually_at_top_linorder) - -lemma sequentially_bot [simp, intro]: "sequentially \ bot" - unfolding filter_eq_iff eventually_sequentially by auto - -lemmas trivial_limit_sequentially = sequentially_bot - -lemma eventually_False_sequentially [simp]: - "\ eventually (\n. False) sequentially" - by (simp add: eventually_False) - -lemma le_sequentially: - "F \ sequentially \ (\N. eventually (\n. N \ n) F)" - unfolding le_filter_def eventually_sequentially - by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp) - -lemma eventually_sequentiallyI: - assumes "\x. c \ x \ P x" - shows "eventually P sequentially" -using assms by (auto simp: eventually_sequentially) - - -subsection {* Standard filters *} - -definition within :: "'a filter \ 'a set \ 'a filter" (infixr "within" 70) - where "F within S = Abs_filter (\P. eventually (\x. x \ S \ P x) F)" - -definition (in topological_space) nhds :: "'a \ 'a filter" - where "nhds a = Abs_filter (\P. \S. open S \ a \ S \ (\x\S. P x))" - -definition (in topological_space) at :: "'a \ 'a filter" - where "at a = nhds a within - {a}" - -abbreviation at_right :: "'a\{topological_space, order} \ 'a filter" where - "at_right x \ at x within {x <..}" - -abbreviation at_left :: "'a\{topological_space, order} \ 'a filter" where - "at_left x \ at x within {..< x}" - definition at_infinity :: "'a::real_normed_vector filter" where "at_infinity = Abs_filter (\P. \r. \x. r \ norm x \ P x)" -lemma eventually_within: - "eventually P (F within S) = eventually (\x. x \ S \ P x) F" - unfolding within_def - by (rule eventually_Abs_filter, rule is_filter.intro) - (auto elim!: eventually_rev_mp) - -lemma within_UNIV [simp]: "F within UNIV = F" - unfolding filter_eq_iff eventually_within by simp - -lemma within_empty [simp]: "F within {} = bot" - unfolding filter_eq_iff eventually_within by simp - -lemma within_within_eq: "(F within S) within T = F within (S \ T)" - by (auto simp: filter_eq_iff eventually_within elim: eventually_elim1) - -lemma at_within_eq: "at x within T = nhds x within (T - {x})" - unfolding at_def within_within_eq by (simp add: ac_simps Diff_eq) - -lemma within_le: "F within S \ F" - unfolding le_filter_def eventually_within by (auto elim: eventually_elim1) - -lemma le_withinI: "F \ F' \ eventually (\x. x \ S) F \ F \ F' within S" - unfolding le_filter_def eventually_within by (auto elim: eventually_elim2) - -lemma le_within_iff: "eventually (\x. x \ S) F \ F \ F' within S \ F \ F'" - by (blast intro: within_le le_withinI order_trans) - -lemma eventually_nhds: - "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" -unfolding nhds_def -proof (rule eventually_Abs_filter, rule is_filter.intro) - have "open UNIV \ a \ UNIV \ (\x\UNIV. True)" by simp - thus "\S. open S \ a \ S \ (\x\S. True)" .. -next - fix P Q - assume "\S. open S \ a \ S \ (\x\S. P x)" - and "\T. open T \ a \ T \ (\x\T. Q x)" - then obtain S T where - "open S \ a \ S \ (\x\S. P x)" - "open T \ a \ T \ (\x\T. Q x)" by auto - hence "open (S \ T) \ a \ S \ T \ (\x\(S \ T). P x \ Q x)" - by (simp add: open_Int) - thus "\S. open S \ a \ S \ (\x\S. P x \ Q x)" .. -qed auto - -lemma eventually_nhds_metric: - "eventually P (nhds a) \ (\d>0. \x. dist x a < d \ P x)" -unfolding eventually_nhds open_dist -apply safe -apply fast -apply (rule_tac x="{x. dist x a < d}" in exI, simp) -apply clarsimp -apply (rule_tac x="d - dist x a" in exI, clarsimp) -apply (simp only: less_diff_eq) -apply (erule le_less_trans [OF dist_triangle]) -done - -lemma nhds_neq_bot [simp]: "nhds a \ bot" - unfolding trivial_limit_def eventually_nhds by simp - -lemma eventually_at_topological: - "eventually P (at a) \ (\S. open S \ a \ S \ (\x\S. x \ a \ P x))" -unfolding at_def eventually_within eventually_nhds by simp - -lemma eventually_at: - fixes a :: "'a::metric_space" - shows "eventually P (at a) \ (\d>0. \x. x \ a \ dist x a < d \ P x)" -unfolding at_def eventually_within eventually_nhds_metric by auto - -lemma eventually_within_less: (* COPY FROM Topo/eventually_within *) - "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" - unfolding eventually_within eventually_at dist_nz by auto - -lemma eventually_within_le: (* COPY FROM Topo/eventually_within_le *) - "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" - unfolding eventually_within_less by auto (metis dense order_le_less_trans) - -lemma at_eq_bot_iff: "at a = bot \ open {a}" - unfolding trivial_limit_def eventually_at_topological - by (safe, case_tac "S = {a}", simp, fast, fast) - -lemma at_neq_bot [simp]: "at (a::'a::perfect_space) \ bot" - by (simp add: at_eq_bot_iff not_open_singleton) - -lemma trivial_limit_at_left_real [simp]: (* maybe generalize type *) - "\ trivial_limit (at_left (x::real))" - unfolding trivial_limit_def eventually_within_le - apply clarsimp - apply (rule_tac x="x - d/2" in bexI) - apply (auto simp: dist_real_def) - done - -lemma trivial_limit_at_right_real [simp]: (* maybe generalize type *) - "\ trivial_limit (at_right (x::real))" - unfolding trivial_limit_def eventually_within_le - apply clarsimp - apply (rule_tac x="x + d/2" in bexI) - apply (auto simp: dist_real_def) - done - lemma eventually_at_infinity: "eventually P at_infinity \ (\b. \x. b \ norm x \ P x)" unfolding at_infinity_def @@ -553,8 +49,18 @@ subsection {* Boundedness *} -definition Bfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" - where "Bfun f F = (\K>0. eventually (\x. norm (f x) \ K) F)" +lemma Bfun_def: + "Bfun f F \ (\K>0. eventually (\x. norm (f x) \ K) F)" + unfolding Bfun_metric_def norm_conv_dist +proof safe + fix y K assume "0 < K" and *: "eventually (\x. dist (f x) y \ K) F" + moreover have "eventually (\x. dist (f x) 0 \ dist (f x) y + dist 0 y) F" + by (intro always_eventually) (metis dist_commute dist_triangle) + with * have "eventually (\x. dist (f x) 0 \ K + dist 0 y) F" + by eventually_elim auto + with `0 < K` show "\K>0. eventually (\x. dist (f x) 0 \ K) F" + by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto +qed auto lemma BfunI: assumes K: "eventually (\x. norm (f x) \ K) F" shows "Bfun f F" @@ -571,7 +77,6 @@ obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) F" using assms unfolding Bfun_def by fast - subsection {* Convergence to Zero *} definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" @@ -722,269 +227,23 @@ lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] - -subsection {* Limits *} - -definition filterlim :: "('a \ 'b) \ 'b filter \ 'a filter \ bool" where - "filterlim f F2 F1 \ filtermap f F1 \ F2" - -syntax - "_LIM" :: "pttrns \ 'a \ 'b \ 'a \ bool" ("(3LIM (_)/ (_)./ (_) :> (_))" [1000, 10, 0, 10] 10) - -translations - "LIM x F1. f :> F2" == "CONST filterlim (%x. f) F2 F1" - -lemma filterlim_iff: - "(LIM x F1. f x :> F2) \ (\P. eventually P F2 \ eventually (\x. P (f x)) F1)" - unfolding filterlim_def le_filter_def eventually_filtermap .. - -lemma filterlim_compose: - "filterlim g F3 F2 \ filterlim f F2 F1 \ filterlim (\x. g (f x)) F3 F1" - unfolding filterlim_def filtermap_filtermap[symmetric] by (metis filtermap_mono order_trans) - -lemma filterlim_mono: - "filterlim f F2 F1 \ F2 \ F2' \ F1' \ F1 \ filterlim f F2' F1'" - unfolding filterlim_def by (metis filtermap_mono order_trans) - -lemma filterlim_ident: "LIM x F. x :> F" - by (simp add: filterlim_def filtermap_ident) - -lemma filterlim_cong: - "F1 = F1' \ F2 = F2' \ eventually (\x. f x = g x) F2 \ filterlim f F1 F2 = filterlim g F1' F2'" - by (auto simp: filterlim_def le_filter_def eventually_filtermap elim: eventually_elim2) - -lemma filterlim_within: - "(LIM x F1. f x :> F2 within S) \ (eventually (\x. f x \ S) F1 \ (LIM x F1. f x :> F2))" - unfolding filterlim_def -proof safe - assume "filtermap f F1 \ F2 within S" then show "eventually (\x. f x \ S) F1" - by (auto simp: le_filter_def eventually_filtermap eventually_within elim!: allE[of _ "\x. x \ S"]) -qed (auto intro: within_le order_trans simp: le_within_iff eventually_filtermap) - -lemma filterlim_filtermap: "filterlim f F1 (filtermap g F2) = filterlim (\x. f (g x)) F1 F2" - unfolding filterlim_def filtermap_filtermap .. - -lemma filterlim_sup: - "filterlim f F F1 \ filterlim f F F2 \ filterlim f F (sup F1 F2)" - unfolding filterlim_def filtermap_sup by auto - -lemma filterlim_Suc: "filterlim Suc sequentially sequentially" - by (simp add: filterlim_iff eventually_sequentially) (metis le_Suc_eq) - -abbreviation (in topological_space) - tendsto :: "('b \ 'a) \ 'a \ 'b filter \ bool" (infixr "--->" 55) where - "(f ---> l) F \ filterlim f (nhds l) F" - -lemma tendsto_eq_rhs: "(f ---> x) F \ x = y \ (f ---> y) F" - by simp - -ML {* - -structure Tendsto_Intros = Named_Thms -( - val name = @{binding tendsto_intros} - val description = "introduction rules for tendsto" -) - -*} - -setup {* - Tendsto_Intros.setup #> - Global_Theory.add_thms_dynamic (@{binding tendsto_eq_intros}, - map (fn thm => @{thm tendsto_eq_rhs} OF [thm]) o Tendsto_Intros.get o Context.proof_of); -*} - -lemma tendsto_def: "(f ---> l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" - unfolding filterlim_def -proof safe - fix S assume "open S" "l \ S" "filtermap f F \ nhds l" - then show "eventually (\x. f x \ S) F" - unfolding eventually_nhds eventually_filtermap le_filter_def - by (auto elim!: allE[of _ "\x. x \ S"] eventually_rev_mp) -qed (auto elim!: eventually_rev_mp simp: eventually_nhds eventually_filtermap le_filter_def) - -lemma filterlim_at: - "(LIM x F. f x :> at b) \ (eventually (\x. f x \ b) F \ (f ---> b) F)" - by (simp add: at_def filterlim_within) - -lemma tendsto_mono: "F \ F' \ (f ---> l) F' \ (f ---> l) F" - unfolding tendsto_def le_filter_def by fast - -lemma topological_tendstoI: - "(\S. open S \ l \ S \ eventually (\x. f x \ S) F) - \ (f ---> l) F" - unfolding tendsto_def by auto - -lemma topological_tendstoD: - "(f ---> l) F \ open S \ l \ S \ eventually (\x. f x \ S) F" - unfolding tendsto_def by auto - -lemma tendstoI: - assumes "\e. 0 < e \ eventually (\x. dist (f x) l < e) F" - shows "(f ---> l) F" - apply (rule topological_tendstoI) - apply (simp add: open_dist) - apply (drule (1) bspec, clarify) - apply (drule assms) - apply (erule eventually_elim1, simp) - done - -lemma tendstoD: - "(f ---> l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" - apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD) - apply (clarsimp simp add: open_dist) - apply (rule_tac x="e - dist x l" in exI, clarsimp) - apply (simp only: less_diff_eq) - apply (erule le_less_trans [OF dist_triangle]) - apply simp - apply simp - done - -lemma tendsto_iff: - "(f ---> l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" - using tendstoI tendstoD by fast - -lemma order_tendstoI: - fixes y :: "_ :: order_topology" - assumes "\a. a < y \ eventually (\x. a < f x) F" - assumes "\a. y < a \ eventually (\x. f x < a) F" - shows "(f ---> y) F" -proof (rule topological_tendstoI) - fix S assume "open S" "y \ S" - then show "eventually (\x. f x \ S) F" - unfolding open_generated_order - proof induct - case (UN K) - then obtain k where "y \ k" "k \ K" by auto - with UN(2)[of k] show ?case - by (auto elim: eventually_elim1) - qed (insert assms, auto elim: eventually_elim2) -qed - -lemma order_tendstoD: - fixes y :: "_ :: order_topology" - assumes y: "(f ---> y) F" - shows "a < y \ eventually (\x. a < f x) F" - and "y < a \ eventually (\x. f x < a) F" - using topological_tendstoD[OF y, of "{..< a}"] topological_tendstoD[OF y, of "{a <..}"] by auto - -lemma order_tendsto_iff: - fixes f :: "_ \ 'a :: order_topology" - shows "(f ---> x) F \(\lx. l < f x) F) \ (\u>x. eventually (\x. f x < u) F)" - by (metis order_tendstoI order_tendstoD) - lemma tendsto_Zfun_iff: "(f ---> a) F = Zfun (\x. f x - a) F" by (simp only: tendsto_iff Zfun_def dist_norm) -lemma tendsto_bot [simp]: "(f ---> a) bot" - unfolding tendsto_def by simp - -lemma tendsto_ident_at [tendsto_intros]: "((\x. x) ---> a) (at a)" - unfolding tendsto_def eventually_at_topological by auto - -lemma tendsto_ident_at_within [tendsto_intros]: - "((\x. x) ---> a) (at a within S)" - unfolding tendsto_def eventually_within eventually_at_topological by auto - -lemma tendsto_const [tendsto_intros]: "((\x. k) ---> k) F" - by (simp add: tendsto_def) - -lemma tendsto_unique: - fixes f :: "'a \ 'b::t2_space" - assumes "\ trivial_limit F" and "(f ---> a) F" and "(f ---> b) F" - shows "a = b" -proof (rule ccontr) - assume "a \ b" - obtain U V where "open U" "open V" "a \ U" "b \ V" "U \ V = {}" - using hausdorff [OF `a \ b`] by fast - have "eventually (\x. f x \ U) F" - using `(f ---> a) F` `open U` `a \ U` by (rule topological_tendstoD) - moreover - have "eventually (\x. f x \ V) F" - using `(f ---> b) F` `open V` `b \ V` by (rule topological_tendstoD) - ultimately - have "eventually (\x. False) F" - proof eventually_elim - case (elim x) - hence "f x \ U \ V" by simp - with `U \ V = {}` show ?case by simp - qed - with `\ trivial_limit F` show "False" - by (simp add: trivial_limit_def) -qed - -lemma tendsto_const_iff: - fixes a b :: "'a::t2_space" - assumes "\ trivial_limit F" shows "((\x. a) ---> b) F \ a = b" - by (safe intro!: tendsto_const tendsto_unique [OF assms tendsto_const]) - -lemma tendsto_at_iff_tendsto_nhds: - "(g ---> g l) (at l) \ (g ---> g l) (nhds l)" - unfolding tendsto_def at_def eventually_within - by (intro ext all_cong imp_cong) (auto elim!: eventually_elim1) - -lemma tendsto_compose: - "(g ---> g l) (at l) \ (f ---> l) F \ ((\x. g (f x)) ---> g l) F" - unfolding tendsto_at_iff_tendsto_nhds by (rule filterlim_compose[of g]) - -lemma tendsto_compose_eventually: - "(g ---> m) (at l) \ (f ---> l) F \ eventually (\x. f x \ l) F \ ((\x. g (f x)) ---> m) F" - by (rule filterlim_compose[of g _ "at l"]) (auto simp add: filterlim_at) - -lemma metric_tendsto_imp_tendsto: - assumes f: "(f ---> a) F" - assumes le: "eventually (\x. dist (g x) b \ dist (f x) a) F" - shows "(g ---> b) F" -proof (rule tendstoI) - fix e :: real assume "0 < e" - with f have "eventually (\x. dist (f x) a < e) F" by (rule tendstoD) - with le show "eventually (\x. dist (g x) b < e) F" - using le_less_trans by (rule eventually_elim2) -qed - -lemma increasing_tendsto: - fixes f :: "_ \ 'a::order_topology" - assumes bdd: "eventually (\n. f n \ l) F" - and en: "\x. x < l \ eventually (\n. x < f n) F" - shows "(f ---> l) F" - using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) - -lemma decreasing_tendsto: - fixes f :: "_ \ 'a::order_topology" - assumes bdd: "eventually (\n. l \ f n) F" - and en: "\x. l < x \ eventually (\n. f n < x) F" - shows "(f ---> l) F" - using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) - subsubsection {* Distance and norms *} -lemma tendsto_dist [tendsto_intros]: - assumes f: "(f ---> l) F" and g: "(g ---> m) F" - shows "((\x. dist (f x) (g x)) ---> dist l m) F" -proof (rule tendstoI) - fix e :: real assume "0 < e" - hence e2: "0 < e/2" by simp - from tendstoD [OF f e2] tendstoD [OF g e2] - show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) F" - proof (eventually_elim) - case (elim x) - then show "dist (dist (f x) (g x)) (dist l m) < e" - unfolding dist_real_def - using dist_triangle2 [of "f x" "g x" "l"] - using dist_triangle2 [of "g x" "l" "m"] - using dist_triangle3 [of "l" "m" "f x"] - using dist_triangle [of "f x" "m" "g x"] - by arith - qed -qed - -lemma norm_conv_dist: "norm x = dist x 0" - unfolding dist_norm by simp - lemma tendsto_norm [tendsto_intros]: "(f ---> a) F \ ((\x. norm (f x)) ---> norm a) F" unfolding norm_conv_dist by (intro tendsto_intros) +lemma continuous_norm [continuous_intros]: + "continuous F f \ continuous F (\x. norm (f x))" + unfolding continuous_def by (rule tendsto_norm) + +lemma continuous_on_norm [continuous_on_intros]: + "continuous_on s f \ continuous_on s (\x. norm (f x))" + unfolding continuous_on_def by (auto intro: tendsto_norm) + lemma tendsto_norm_zero: "(f ---> 0) F \ ((\x. norm (f x)) ---> 0) F" by (drule tendsto_norm, simp) @@ -1001,6 +260,14 @@ "(f ---> (l::real)) F \ ((\x. \f x\) ---> \l\) F" by (fold real_norm_def, rule tendsto_norm) +lemma continuous_rabs [continuous_intros]: + "continuous F f \ continuous F (\x. \f x :: real\)" + unfolding real_norm_def[symmetric] by (rule continuous_norm) + +lemma continuous_on_rabs [continuous_on_intros]: + "continuous_on s f \ continuous_on s (\x. \f x :: real\)" + unfolding real_norm_def[symmetric] by (rule continuous_on_norm) + lemma tendsto_rabs_zero: "(f ---> (0::real)) F \ ((\x. \f x\) ---> 0) F" by (fold real_norm_def, rule tendsto_norm_zero) @@ -1020,8 +287,18 @@ shows "\(f ---> a) F; (g ---> b) F\ \ ((\x. f x + g x) ---> a + b) F" by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) +lemma continuous_add [continuous_intros]: + fixes f g :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous F f \ continuous F g \ continuous F (\x. f x + g x)" + unfolding continuous_def by (rule tendsto_add) + +lemma continuous_on_add [continuous_on_intros]: + fixes f g :: "_ \ 'b::real_normed_vector" + shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x + g x)" + unfolding continuous_on_def by (auto intro: tendsto_add) + lemma tendsto_add_zero: - fixes f g :: "'a::type \ 'b::real_normed_vector" + fixes f g :: "_ \ 'b::real_normed_vector" shows "\(f ---> 0) F; (g ---> 0) F\ \ ((\x. f x + g x) ---> 0) F" by (drule (1) tendsto_add, simp) @@ -1030,6 +307,16 @@ shows "(f ---> a) F \ ((\x. - f x) ---> - a) F" by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) +lemma continuous_minus [continuous_intros]: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous F f \ continuous F (\x. - f x)" + unfolding continuous_def by (rule tendsto_minus) + +lemma continuous_on_minus [continuous_on_intros]: + fixes f :: "_ \ 'b::real_normed_vector" + shows "continuous_on s f \ continuous_on s (\x. - f x)" + unfolding continuous_on_def by (auto intro: tendsto_minus) + lemma tendsto_minus_cancel: fixes a :: "'a::real_normed_vector" shows "((\x. - f x) ---> - a) F \ (f ---> a) F" @@ -1045,6 +332,16 @@ shows "\(f ---> a) F; (g ---> b) F\ \ ((\x. f x - g x) ---> a - b) F" by (simp add: diff_minus tendsto_add tendsto_minus) +lemma continuous_diff [continuous_intros]: + fixes f g :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous F f \ continuous F g \ continuous F (\x. f x - g x)" + unfolding continuous_def by (rule tendsto_diff) + +lemma continuous_on_diff [continuous_on_intros]: + fixes f g :: "'a::t2_space \ 'b::real_normed_vector" + shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x - g x)" + unfolding continuous_on_def by (auto intro: tendsto_diff) + lemma tendsto_setsum [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::real_normed_vector" assumes "\i. i \ S \ (f i ---> a i) F" @@ -1057,18 +354,15 @@ by (simp add: tendsto_const) qed -lemma tendsto_sandwich: - fixes f g h :: "'a \ 'b::order_topology" - assumes ev: "eventually (\n. f n \ g n) net" "eventually (\n. g n \ h n) net" - assumes lim: "(f ---> c) net" "(h ---> c) net" - shows "(g ---> c) net" -proof (rule order_tendstoI) - fix a show "a < c \ eventually (\x. a < g x) net" - using order_tendstoD[OF lim(1), of a] ev by (auto elim: eventually_elim2) -next - fix a show "c < a \ eventually (\x. g x < a) net" - using order_tendstoD[OF lim(2), of a] ev by (auto elim: eventually_elim2) -qed +lemma continuous_setsum [continuous_intros]: + fixes f :: "'a \ 'b::t2_space \ 'c::real_normed_vector" + shows "(\i. i \ S \ continuous F (f i)) \ continuous F (\x. \i\S. f i x)" + unfolding continuous_def by (rule tendsto_setsum) + +lemma continuous_on_setsum [continuous_intros]: + fixes f :: "'a \ _ \ 'c::real_normed_vector" + shows "(\i. i \ S \ continuous_on s (f i)) \ continuous_on s (\x. \i\S. f i x)" + unfolding continuous_on_def by (auto intro: tendsto_setsum) lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real] @@ -1078,6 +372,14 @@ "(g ---> a) F \ ((\x. f (g x)) ---> f a) F" by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) +lemma (in bounded_linear) continuous: + "continuous F g \ continuous F (\x. f (g x))" + using tendsto[of g _ F] by (auto simp: continuous_def) + +lemma (in bounded_linear) continuous_on: + "continuous_on s g \ continuous_on s (\x. f (g x))" + using tendsto[of g] by (auto simp: continuous_on_def) + lemma (in bounded_linear) tendsto_zero: "(g ---> 0) F \ ((\x. f (g x)) ---> 0) F" by (drule tendsto, simp only: zero) @@ -1087,6 +389,14 @@ by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right) +lemma (in bounded_bilinear) continuous: + "continuous F f \ continuous F g \ continuous F (\x. f x ** g x)" + using tendsto[of f _ F g] by (auto simp: continuous_def) + +lemma (in bounded_bilinear) continuous_on: + "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x ** g x)" + using tendsto[of f _ _ g] by (auto simp: continuous_on_def) + lemma (in bounded_bilinear) tendsto_zero: assumes f: "(f ---> 0) F" assumes g: "(g ---> 0) F" @@ -1110,6 +420,24 @@ lemmas tendsto_mult [tendsto_intros] = bounded_bilinear.tendsto [OF bounded_bilinear_mult] +lemmas continuous_of_real [continuous_intros] = + bounded_linear.continuous [OF bounded_linear_of_real] + +lemmas continuous_scaleR [continuous_intros] = + bounded_bilinear.continuous [OF bounded_bilinear_scaleR] + +lemmas continuous_mult [continuous_intros] = + bounded_bilinear.continuous [OF bounded_bilinear_mult] + +lemmas continuous_on_of_real [continuous_on_intros] = + bounded_linear.continuous_on [OF bounded_linear_of_real] + +lemmas continuous_on_scaleR [continuous_on_intros] = + bounded_bilinear.continuous_on [OF bounded_bilinear_scaleR] + +lemmas continuous_on_mult [continuous_on_intros] = + bounded_bilinear.continuous_on [OF bounded_bilinear_mult] + lemmas tendsto_mult_zero = bounded_bilinear.tendsto_zero [OF bounded_bilinear_mult] @@ -1124,6 +452,16 @@ shows "(f ---> a) F \ ((\x. f x ^ n) ---> a ^ n) F" by (induct n) (simp_all add: tendsto_const tendsto_mult) +lemma continuous_power [continuous_intros]: + fixes f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" + shows "continuous F f \ continuous F (\x. (f x)^n)" + unfolding continuous_def by (rule tendsto_power) + +lemma continuous_on_power [continuous_on_intros]: + fixes f :: "_ \ 'b::{power,real_normed_algebra}" + shows "continuous_on s f \ continuous_on s (\x. (f x)^n)" + unfolding continuous_on_def by (auto intro: tendsto_power) + lemma tendsto_setprod [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::{real_normed_algebra,comm_ring_1}" assumes "\i. i \ S \ (f i ---> L i) F" @@ -1136,30 +474,15 @@ by (simp add: tendsto_const) qed -lemma tendsto_le: - fixes f g :: "'a \ 'b::linorder_topology" - assumes F: "\ trivial_limit F" - assumes x: "(f ---> x) F" and y: "(g ---> y) F" - assumes ev: "eventually (\x. g x \ f x) F" - shows "y \ x" -proof (rule ccontr) - assume "\ y \ x" - with less_separate[of x y] obtain a b where xy: "x < a" "b < y" "{.. {b<..} = {}" - by (auto simp: not_le) - then have "eventually (\x. f x < a) F" "eventually (\x. b < g x) F" - using x y by (auto intro: order_tendstoD) - with ev have "eventually (\x. False) F" - by eventually_elim (insert xy, fastforce) - with F show False - by (simp add: eventually_False) -qed +lemma continuous_setprod [continuous_intros]: + fixes f :: "'a \ 'b::t2_space \ 'c::{real_normed_algebra,comm_ring_1}" + shows "(\i. i \ S \ continuous F (f i)) \ continuous F (\x. \i\S. f i x)" + unfolding continuous_def by (rule tendsto_setprod) -lemma tendsto_le_const: - fixes f :: "'a \ 'b::linorder_topology" - assumes F: "\ trivial_limit F" - assumes x: "(f ---> x) F" and a: "eventually (\x. a \ f x) F" - shows "a \ x" - using F x tendsto_const a by (rule tendsto_le) +lemma continuous_on_setprod [continuous_intros]: + fixes f :: "'a \ _ \ 'c::{real_normed_algebra,comm_ring_1}" + shows "(\i. i \ S \ continuous_on s (f i)) \ continuous_on s (\x. \i\S. f i x)" + unfolding continuous_on_def by (auto intro: tendsto_setprod) subsubsection {* Inverse and division *} @@ -1270,129 +593,88 @@ unfolding tendsto_Zfun_iff by (rule Zfun_ssubst) qed +lemma continuous_inverse: + fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" + assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" + shows "continuous F (\x. inverse (f x))" + using assms unfolding continuous_def by (rule tendsto_inverse) + +lemma continuous_at_within_inverse[continuous_intros]: + fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" + assumes "continuous (at a within s) f" and "f a \ 0" + shows "continuous (at a within s) (\x. inverse (f x))" + using assms unfolding continuous_within by (rule tendsto_inverse) + +lemma isCont_inverse[continuous_intros, simp]: + fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" + assumes "isCont f a" and "f a \ 0" + shows "isCont (\x. inverse (f x)) a" + using assms unfolding continuous_at by (rule tendsto_inverse) + +lemma continuous_on_inverse[continuous_on_intros]: + fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" + assumes "continuous_on s f" and "\x\s. f x \ 0" + shows "continuous_on s (\x. inverse (f x))" + using assms unfolding continuous_on_def by (fast intro: tendsto_inverse) + lemma tendsto_divide [tendsto_intros]: fixes a b :: "'a::real_normed_field" shows "\(f ---> a) F; (g ---> b) F; b \ 0\ \ ((\x. f x / g x) ---> a / b) F" by (simp add: tendsto_mult tendsto_inverse divide_inverse) +lemma continuous_divide: + fixes f g :: "'a::t2_space \ 'b::real_normed_field" + assumes "continuous F f" and "continuous F g" and "g (Lim F (\x. x)) \ 0" + shows "continuous F (\x. (f x) / (g x))" + using assms unfolding continuous_def by (rule tendsto_divide) + +lemma continuous_at_within_divide[continuous_intros]: + fixes f g :: "'a::t2_space \ 'b::real_normed_field" + assumes "continuous (at a within s) f" "continuous (at a within s) g" and "g a \ 0" + shows "continuous (at a within s) (\x. (f x) / (g x))" + using assms unfolding continuous_within by (rule tendsto_divide) + +lemma isCont_divide[continuous_intros, simp]: + fixes f g :: "'a::t2_space \ 'b::real_normed_field" + assumes "isCont f a" "isCont g a" "g a \ 0" + shows "isCont (\x. (f x) / g x) a" + using assms unfolding continuous_at by (rule tendsto_divide) + +lemma continuous_on_divide[continuous_on_intros]: + fixes f :: "'a::topological_space \ 'b::real_normed_field" + assumes "continuous_on s f" "continuous_on s g" and "\x\s. g x \ 0" + shows "continuous_on s (\x. (f x) / (g x))" + using assms unfolding continuous_on_def by (fast intro: tendsto_divide) + lemma tendsto_sgn [tendsto_intros]: fixes l :: "'a::real_normed_vector" shows "\(f ---> l) F; l \ 0\ \ ((\x. sgn (f x)) ---> sgn l) F" unfolding sgn_div_norm by (simp add: tendsto_intros) -subsection {* Limits to @{const at_top} and @{const at_bot} *} - -lemma filterlim_at_top: - fixes f :: "'a \ ('b::linorder)" - shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z \ f x) F)" - by (auto simp: filterlim_iff eventually_at_top_linorder elim!: eventually_elim1) - -lemma filterlim_at_top_dense: - fixes f :: "'a \ ('b::dense_linorder)" - shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z < f x) F)" - by (metis eventually_elim1[of _ F] eventually_gt_at_top order_less_imp_le - filterlim_at_top[of f F] filterlim_iff[of f at_top F]) - -lemma filterlim_at_top_ge: - fixes f :: "'a \ ('b::linorder)" and c :: "'b" - shows "(LIM x F. f x :> at_top) \ (\Z\c. eventually (\x. Z \ f x) F)" - unfolding filterlim_at_top -proof safe - fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" - with *[THEN spec, of "max Z c"] show "eventually (\x. Z \ f x) F" - by (auto elim!: eventually_elim1) -qed simp +lemma continuous_sgn: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" + shows "continuous F (\x. sgn (f x))" + using assms unfolding continuous_def by (rule tendsto_sgn) -lemma filterlim_at_top_at_top: - fixes f :: "'a::linorder \ 'b::linorder" - assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" - assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" - assumes Q: "eventually Q at_top" - assumes P: "eventually P at_top" - shows "filterlim f at_top at_top" -proof - - from P obtain x where x: "\y. x \ y \ P y" - unfolding eventually_at_top_linorder by auto - show ?thesis - proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) - fix z assume "x \ z" - with x have "P z" by auto - have "eventually (\x. g z \ x) at_top" - by (rule eventually_ge_at_top) - with Q show "eventually (\x. z \ f x) at_top" - by eventually_elim (metis mono bij `P z`) - qed -qed - -lemma filterlim_at_top_gt: - fixes f :: "'a \ ('b::dense_linorder)" and c :: "'b" - shows "(LIM x F. f x :> at_top) \ (\Z>c. eventually (\x. Z \ f x) F)" - by (metis filterlim_at_top order_less_le_trans gt_ex filterlim_at_top_ge) - -lemma filterlim_at_bot: - fixes f :: "'a \ ('b::linorder)" - shows "(LIM x F. f x :> at_bot) \ (\Z. eventually (\x. f x \ Z) F)" - by (auto simp: filterlim_iff eventually_at_bot_linorder elim!: eventually_elim1) +lemma continuous_at_within_sgn[continuous_intros]: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + assumes "continuous (at a within s) f" and "f a \ 0" + shows "continuous (at a within s) (\x. sgn (f x))" + using assms unfolding continuous_within by (rule tendsto_sgn) -lemma filterlim_at_bot_le: - fixes f :: "'a \ ('b::linorder)" and c :: "'b" - shows "(LIM x F. f x :> at_bot) \ (\Z\c. eventually (\x. Z \ f x) F)" - unfolding filterlim_at_bot -proof safe - fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" - with *[THEN spec, of "min Z c"] show "eventually (\x. Z \ f x) F" - by (auto elim!: eventually_elim1) -qed simp - -lemma filterlim_at_bot_lt: - fixes f :: "'a \ ('b::dense_linorder)" and c :: "'b" - shows "(LIM x F. f x :> at_bot) \ (\Zx. Z \ f x) F)" - by (metis filterlim_at_bot filterlim_at_bot_le lt_ex order_le_less_trans) +lemma isCont_sgn[continuous_intros]: + fixes f :: "'a::t2_space \ 'b::real_normed_vector" + assumes "isCont f a" and "f a \ 0" + shows "isCont (\x. sgn (f x)) a" + using assms unfolding continuous_at by (rule tendsto_sgn) -lemma filterlim_at_bot_at_right: - fixes f :: "real \ 'b::linorder" - assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" - assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" - assumes Q: "eventually Q (at_right a)" and bound: "\b. Q b \ a < b" - assumes P: "eventually P at_bot" - shows "filterlim f at_bot (at_right a)" -proof - - from P obtain x where x: "\y. y \ x \ P y" - unfolding eventually_at_bot_linorder by auto - show ?thesis - proof (intro filterlim_at_bot_le[THEN iffD2] allI impI) - fix z assume "z \ x" - with x have "P z" by auto - have "eventually (\x. x \ g z) (at_right a)" - using bound[OF bij(2)[OF `P z`]] - by (auto simp add: eventually_within_less dist_real_def intro!: exI[of _ "g z - a"]) - with Q show "eventually (\x. f x \ z) (at_right a)" - by eventually_elim (metis bij `P z` mono) - qed -qed - -lemma filterlim_at_top_at_left: - fixes f :: "real \ 'b::linorder" - assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" - assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" - assumes Q: "eventually Q (at_left a)" and bound: "\b. Q b \ b < a" - assumes P: "eventually P at_top" - shows "filterlim f at_top (at_left a)" -proof - - from P obtain x where x: "\y. x \ y \ P y" - unfolding eventually_at_top_linorder by auto - show ?thesis - proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) - fix z assume "x \ z" - with x have "P z" by auto - have "eventually (\x. g z \ x) (at_left a)" - using bound[OF bij(2)[OF `P z`]] - by (auto simp add: eventually_within_less dist_real_def intro!: exI[of _ "a - g z"]) - with Q show "eventually (\x. z \ f x) (at_left a)" - by eventually_elim (metis bij `P z` mono) - qed -qed +lemma continuous_on_sgn[continuous_on_intros]: + fixes f :: "'a::topological_space \ 'b::real_normed_vector" + assumes "continuous_on s f" and "\x\s. f x \ 0" + shows "continuous_on s (\x. sgn (f x))" + using assms unfolding continuous_on_def by (fast intro: tendsto_sgn) lemma filterlim_at_infinity: fixes f :: "_ \ 'a\real_normed_vector" @@ -1413,13 +695,6 @@ qed qed force -lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top" - unfolding filterlim_at_top - apply (intro allI) - apply (rule_tac c="natceiling (Z + 1)" in eventually_sequentiallyI) - apply (auto simp: natceiling_le_eq) - done - subsection {* Relate @{const at}, @{const at_left} and @{const at_right} *} text {* @@ -1429,13 +704,7 @@ *} -lemma at_eq_sup_left_right: "at (x::real) = sup (at_left x) (at_right x)" - by (auto simp: eventually_within at_def filter_eq_iff eventually_sup - elim: eventually_elim2 eventually_elim1) - -lemma filterlim_split_at_real: - "filterlim f F (at_left x) \ filterlim f F (at_right x) \ filterlim f F (at (x::real))" - by (subst at_eq_sup_left_right) (rule filterlim_sup) +lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real] lemma filtermap_nhds_shift: "filtermap (\x. x - d) (nhds a) = nhds (a - d::real)" unfolding filter_eq_iff eventually_filtermap eventually_nhds_metric @@ -1486,14 +755,6 @@ "eventually P (at_left (a::real)) \ eventually (\x. P (- x)) (at_right (-a))" unfolding at_left_minus[of a] by (simp add: eventually_filtermap) -lemma filterlim_at_split: - "filterlim f F (at (x::real)) \ filterlim f F (at_left x) \ filterlim f F (at_right x)" - by (subst at_eq_sup_left_right) (simp add: filterlim_def filtermap_sup) - -lemma eventually_at_split: - "eventually P (at (x::real)) \ eventually P (at_left x) \ eventually P (at_right x)" - by (subst at_eq_sup_left_right) (simp add: eventually_sup) - lemma at_top_mirror: "at_top = filtermap uminus (at_bot :: real filter)" unfolding filter_eq_iff eventually_filtermap eventually_at_top_linorder eventually_at_bot_linorder by (metis le_minus_iff minus_minus) @@ -1765,5 +1026,10 @@ by (auto simp: norm_power) qed simp + +(* Unfortunately eventually_within was overwritten by Multivariate_Analysis. + Hence it was references as Limits.within, but now it is Basic_Topology.eventually_within *) +lemmas eventually_within = eventually_within + end diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Log.thy --- a/src/HOL/Log.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Log.thy Sat Mar 23 07:30:53 2013 +0100 @@ -21,6 +21,29 @@ "log a x = ln x / ln a" +lemma tendsto_log [tendsto_intros]: + "\(f ---> a) F; (g ---> b) F; 0 < a; a \ 1; 0 < b\ \ ((\x. log (f x) (g x)) ---> log a b) F" + unfolding log_def by (intro tendsto_intros) auto + +lemma continuous_log: + assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\x. x))" and "f (Lim F (\x. x)) \ 1" and "0 < g (Lim F (\x. x))" + shows "continuous F (\x. log (f x) (g x))" + using assms unfolding continuous_def by (rule tendsto_log) + +lemma continuous_at_within_log[continuous_intros]: + assumes "continuous (at a within s) f" "continuous (at a within s) g" and "0 < f a" and "f a \ 1" and "0 < g a" + shows "continuous (at a within s) (\x. log (f x) (g x))" + using assms unfolding continuous_within by (rule tendsto_log) + +lemma isCont_log[continuous_intros, simp]: + assumes "isCont f a" "isCont g a" "0 < f a" "f a \ 1" "0 < g a" + shows "isCont (\x. log (f x) (g x)) a" + using assms unfolding continuous_at by (rule tendsto_log) + +lemma continuous_on_log[continuous_on_intros]: + assumes "continuous_on s f" "continuous_on s g" and "\x\s. 0 < f x" "\x\s. f x \ 1" "\x\s. 0 < g x" + shows "continuous_on s (\x. log (f x) (g x))" + using assms unfolding continuous_on_def by (fast intro: tendsto_log) lemma powr_one_eq_one [simp]: "1 powr a = 1" by (simp add: powr_def) @@ -221,7 +244,7 @@ lemma root_powr_inverse: "0 < n \ 0 < x \ root n x = x powr (1/n)" -by (auto simp: root_def powr_realpow[symmetric] powr_powr) + by (rule real_root_pos_unique) (auto simp: powr_realpow[symmetric] powr_powr) lemma ln_powr: "0 < x ==> 0 < y ==> ln(x powr y) = y * ln x" by (unfold powr_def, simp) @@ -338,6 +361,26 @@ "\(f ---> a) F; (g ---> b) F; 0 < a\ \ ((\x. f x powr g x) ---> a powr b) F" unfolding powr_def by (intro tendsto_intros) +lemma continuous_powr: + assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\x. x))" + shows "continuous F (\x. (f x) powr (g x))" + using assms unfolding continuous_def by (rule tendsto_powr) + +lemma continuous_at_within_powr[continuous_intros]: + assumes "continuous (at a within s) f" "continuous (at a within s) g" and "0 < f a" + shows "continuous (at a within s) (\x. (f x) powr (g x))" + using assms unfolding continuous_within by (rule tendsto_powr) + +lemma isCont_powr[continuous_intros, simp]: + assumes "isCont f a" "isCont g a" "0 < f a" + shows "isCont (\x. (f x) powr g x) a" + using assms unfolding continuous_at by (rule tendsto_powr) + +lemma continuous_on_powr[continuous_on_intros]: + assumes "continuous_on s f" "continuous_on s g" and "\x\s. 0 < f x" + shows "continuous_on s (\x. (f x) powr (g x))" + using assms unfolding continuous_on_def by (fast intro: tendsto_powr) + (* FIXME: generalize by replacing d by with g x and g ---> d? *) lemma tendsto_zero_powrI: assumes "eventually (\x. 0 < f x ) F" and "(f ---> 0) F" diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Metric_Spaces.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Metric_Spaces.thy Sat Mar 23 07:30:53 2013 +0100 @@ -0,0 +1,629 @@ +(* Title: HOL/Metric_Spaces.thy + Author: Brian Huffman + Author: Johannes Hölzl +*) + +header {* Metric Spaces *} + +theory Metric_Spaces +imports RComplete Topological_Spaces +begin + + +subsection {* Metric spaces *} + +class dist = + fixes dist :: "'a \ 'a \ real" + +class open_dist = "open" + dist + + assumes open_dist: "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" + +class metric_space = open_dist + + assumes dist_eq_0_iff [simp]: "dist x y = 0 \ x = y" + assumes dist_triangle2: "dist x y \ dist x z + dist y z" +begin + +lemma dist_self [simp]: "dist x x = 0" +by simp + +lemma zero_le_dist [simp]: "0 \ dist x y" +using dist_triangle2 [of x x y] by simp + +lemma zero_less_dist_iff: "0 < dist x y \ x \ y" +by (simp add: less_le) + +lemma dist_not_less_zero [simp]: "\ dist x y < 0" +by (simp add: not_less) + +lemma dist_le_zero_iff [simp]: "dist x y \ 0 \ x = y" +by (simp add: le_less) + +lemma dist_commute: "dist x y = dist y x" +proof (rule order_antisym) + show "dist x y \ dist y x" + using dist_triangle2 [of x y x] by simp + show "dist y x \ dist x y" + using dist_triangle2 [of y x y] by simp +qed + +lemma dist_triangle: "dist x z \ dist x y + dist y z" +using dist_triangle2 [of x z y] by (simp add: dist_commute) + +lemma dist_triangle3: "dist x y \ dist a x + dist a y" +using dist_triangle2 [of x y a] by (simp add: dist_commute) + +lemma dist_triangle_alt: + shows "dist y z <= dist x y + dist x z" +by (rule dist_triangle3) + +lemma dist_pos_lt: + shows "x \ y ==> 0 < dist x y" +by (simp add: zero_less_dist_iff) + +lemma dist_nz: + shows "x \ y \ 0 < dist x y" +by (simp add: zero_less_dist_iff) + +lemma dist_triangle_le: + shows "dist x z + dist y z <= e \ dist x y <= e" +by (rule order_trans [OF dist_triangle2]) + +lemma dist_triangle_lt: + shows "dist x z + dist y z < e ==> dist x y < e" +by (rule le_less_trans [OF dist_triangle2]) + +lemma dist_triangle_half_l: + shows "dist x1 y < e / 2 \ dist x2 y < e / 2 \ dist x1 x2 < e" +by (rule dist_triangle_lt [where z=y], simp) + +lemma dist_triangle_half_r: + 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) + +subclass topological_space +proof + have "\e::real. 0 < e" + by (fast intro: zero_less_one) + then show "open UNIV" + unfolding open_dist by simp +next + fix S T assume "open S" "open T" + then show "open (S \ T)" + unfolding open_dist + apply clarify + apply (drule (1) bspec)+ + apply (clarify, rename_tac r s) + apply (rule_tac x="min r s" in exI, simp) + done +next + fix K assume "\S\K. open S" thus "open (\K)" + unfolding open_dist by fast +qed + +lemma open_ball: "open {y. dist x y < d}" +proof (unfold open_dist, intro ballI) + fix y assume *: "y \ {y. dist x y < d}" + then show "\e>0. \z. dist z y < e \ z \ {y. dist x y < d}" + by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt) +qed + +subclass first_countable_topology +proof + fix x + show "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + proof (safe intro!: exI[of _ "\n. {y. dist x y < inverse (Suc n)}"]) + fix S assume "open S" "x \ S" + then obtain e where "0 < e" "{y. dist x y < e} \ S" + by (auto simp: open_dist subset_eq dist_commute) + moreover + then obtain i where "inverse (Suc i) < e" + by (auto dest!: reals_Archimedean) + then have "{y. dist x y < inverse (Suc i)} \ {y. dist x y < e}" + by auto + ultimately show "\i. {y. dist x y < inverse (Suc i)} \ S" + by blast + qed (auto intro: open_ball) +qed + +end + +instance metric_space \ t2_space +proof + fix x y :: "'a::metric_space" + assume xy: "x \ y" + let ?U = "{y'. dist x y' < dist x y / 2}" + let ?V = "{x'. dist y x' < dist x y / 2}" + have th0: "\d x y z. (d x z :: real) \ d x y + d y z \ d y z = d z y + \ \(d x y * 2 < d x z \ d z y * 2 < d x z)" by arith + have "open ?U \ open ?V \ x \ ?U \ y \ ?V \ ?U \ ?V = {}" + using dist_pos_lt[OF xy] th0[of dist, OF dist_triangle dist_commute] + using open_ball[of _ "dist x y / 2"] by auto + then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" + by blast +qed + +lemma eventually_nhds_metric: + fixes a :: "'a :: metric_space" + shows "eventually P (nhds a) \ (\d>0. \x. dist x a < d \ P x)" +unfolding eventually_nhds open_dist +apply safe +apply fast +apply (rule_tac x="{x. dist x a < d}" in exI, simp) +apply clarsimp +apply (rule_tac x="d - dist x a" in exI, clarsimp) +apply (simp only: less_diff_eq) +apply (erule le_less_trans [OF dist_triangle]) +done + +lemma eventually_at: + fixes a :: "'a::metric_space" + shows "eventually P (at a) \ (\d>0. \x. x \ a \ dist x a < d \ P x)" +unfolding at_def eventually_within eventually_nhds_metric by auto + +lemma eventually_within_less: (* COPY FROM Topo/eventually_within *) + fixes a :: "'a :: metric_space" + shows "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" + unfolding eventually_within eventually_at dist_nz by auto + +lemma eventually_within_le: (* COPY FROM Topo/eventually_within_le *) + fixes a :: "'a :: metric_space" + shows "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" + unfolding eventually_within_less by auto (metis dense order_le_less_trans) + +lemma tendstoI: + fixes l :: "'a :: metric_space" + assumes "\e. 0 < e \ eventually (\x. dist (f x) l < e) F" + shows "(f ---> l) F" + apply (rule topological_tendstoI) + apply (simp add: open_dist) + apply (drule (1) bspec, clarify) + apply (drule assms) + apply (erule eventually_elim1, simp) + done + +lemma tendstoD: + fixes l :: "'a :: metric_space" + shows "(f ---> l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" + apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD) + apply (clarsimp simp add: open_dist) + apply (rule_tac x="e - dist x l" in exI, clarsimp) + apply (simp only: less_diff_eq) + apply (erule le_less_trans [OF dist_triangle]) + apply simp + apply simp + done + +lemma tendsto_iff: + fixes l :: "'a :: metric_space" + shows "(f ---> l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" + using tendstoI tendstoD by fast + +lemma metric_tendsto_imp_tendsto: + fixes a :: "'a :: metric_space" and b :: "'b :: metric_space" + assumes f: "(f ---> a) F" + assumes le: "eventually (\x. dist (g x) b \ dist (f x) a) F" + shows "(g ---> b) F" +proof (rule tendstoI) + fix e :: real assume "0 < e" + with f have "eventually (\x. dist (f x) a < e) F" by (rule tendstoD) + with le show "eventually (\x. dist (g x) b < e) F" + using le_less_trans by (rule eventually_elim2) +qed + +lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top" + unfolding filterlim_at_top + apply (intro allI) + apply (rule_tac c="natceiling (Z + 1)" in eventually_sequentiallyI) + apply (auto simp: natceiling_le_eq) + done + +subsubsection {* Limits of Sequences *} + +lemma LIMSEQ_def: "X ----> (L::'a::metric_space) \ (\r>0. \no. \n\no. dist (X n) L < r)" + unfolding tendsto_iff eventually_sequentially .. + +lemma LIMSEQ_iff_nz: "X ----> (L::'a::metric_space) = (\r>0. \no>0. \n\no. dist (X n) L < r)" + unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc) + +lemma metric_LIMSEQ_I: + "(\r. 0 < r \ \no. \n\no. dist (X n) L < r) \ X ----> (L::'a::metric_space)" +by (simp add: LIMSEQ_def) + +lemma metric_LIMSEQ_D: + "\X ----> (L::'a::metric_space); 0 < r\ \ \no. \n\no. dist (X n) L < r" +by (simp add: LIMSEQ_def) + + +subsubsection {* Limits of Functions *} + +lemma LIM_def: "f -- (a::'a::metric_space) --> (L::'b::metric_space) = + (\r > 0. \s > 0. \x. x \ a & dist x a < s + --> dist (f x) L < r)" +unfolding tendsto_iff eventually_at .. + +lemma metric_LIM_I: + "(\r. 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r) + \ f -- (a::'a::metric_space) --> (L::'b::metric_space)" +by (simp add: LIM_def) + +lemma metric_LIM_D: + "\f -- (a::'a::metric_space) --> (L::'b::metric_space); 0 < r\ + \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r" +by (simp add: LIM_def) + +lemma metric_LIM_imp_LIM: + assumes f: "f -- a --> (l::'a::metric_space)" + assumes le: "\x. x \ a \ dist (g x) m \ dist (f x) l" + shows "g -- a --> (m::'b::metric_space)" + by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp add: eventually_at_topological le) + +lemma metric_LIM_equal2: + assumes 1: "0 < R" + assumes 2: "\x. \x \ a; dist x a < R\ \ f x = g x" + shows "g -- a --> l \ f -- (a::'a::metric_space) --> l" +apply (rule topological_tendstoI) +apply (drule (2) topological_tendstoD) +apply (simp add: eventually_at, safe) +apply (rule_tac x="min d R" in exI, safe) +apply (simp add: 1) +apply (simp add: 2) +done + +lemma metric_LIM_compose2: + assumes f: "f -- (a::'a::metric_space) --> b" + assumes g: "g -- b --> c" + assumes inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ b" + shows "(\x. g (f x)) -- a --> c" + using g f inj [folded eventually_at] + by (rule tendsto_compose_eventually) + +lemma metric_isCont_LIM_compose2: + fixes f :: "'a :: metric_space \ _" + assumes f [unfolded isCont_def]: "isCont f a" + assumes g: "g -- f a --> l" + assumes inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ f a" + shows "(\x. g (f x)) -- a --> l" +by (rule metric_LIM_compose2 [OF f g inj]) + +subsubsection {* Boundedness *} + +definition Bfun :: "('a \ 'b::metric_space) \ 'a filter \ bool" where + Bfun_metric_def: "Bfun f F = (\y. \K>0. eventually (\x. dist (f x) y \ K) F)" + +abbreviation Bseq :: "(nat \ 'a::metric_space) \ bool" where + "Bseq X \ Bfun X sequentially" + +lemma Bseq_conv_Bfun: "Bseq X \ Bfun X sequentially" .. + +lemma Bseq_ignore_initial_segment: "Bseq X \ Bseq (\n. X (n + k))" + unfolding Bfun_metric_def by (subst eventually_sequentially_seg) + +lemma Bseq_offset: "Bseq (\n. X (n + k)) \ Bseq X" + unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) + +subsection {* Complete metric spaces *} + +subsection {* Cauchy sequences *} + +definition (in metric_space) Cauchy :: "(nat \ 'a) \ bool" where + "Cauchy X = (\e>0. \M. \m \ M. \n \ M. dist (X m) (X n) < e)" + +subsection {* Cauchy Sequences *} + +lemma metric_CauchyI: + "(\e. 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e) \ Cauchy X" + by (simp add: Cauchy_def) + +lemma metric_CauchyD: + "Cauchy X \ 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e" + by (simp add: Cauchy_def) + +lemma metric_Cauchy_iff2: + "Cauchy X = (\j. (\M. \m \ M. \n \ M. dist (X m) (X n) < inverse(real (Suc j))))" +apply (simp add: Cauchy_def, auto) +apply (drule reals_Archimedean, safe) +apply (drule_tac x = n in spec, auto) +apply (rule_tac x = M in exI, auto) +apply (drule_tac x = m in spec, simp) +apply (drule_tac x = na in spec, auto) +done + +lemma Cauchy_subseq_Cauchy: + "\ Cauchy X; subseq f \ \ Cauchy (X o f)" +apply (auto simp add: Cauchy_def) +apply (drule_tac x=e in spec, clarify) +apply (rule_tac x=M in exI, clarify) +apply (blast intro: le_trans [OF _ seq_suble] dest!: spec) +done + +lemma Cauchy_Bseq: "Cauchy X \ Bseq X" + unfolding Cauchy_def Bfun_metric_def eventually_sequentially + apply (erule_tac x=1 in allE) + apply simp + apply safe + apply (rule_tac x="X M" in exI) + apply (rule_tac x=1 in exI) + apply (erule_tac x=M in allE) + apply simp + apply (rule_tac x=M in exI) + apply (auto simp: dist_commute) + done + +subsubsection {* Cauchy Sequences are Convergent *} + +class complete_space = metric_space + + assumes Cauchy_convergent: "Cauchy X \ convergent X" + +theorem LIMSEQ_imp_Cauchy: + assumes X: "X ----> a" shows "Cauchy X" +proof (rule metric_CauchyI) + fix e::real assume "0 < e" + hence "0 < e/2" by simp + with X have "\N. \n\N. dist (X n) a < e/2" by (rule metric_LIMSEQ_D) + then obtain N where N: "\n\N. dist (X n) a < e/2" .. + show "\N. \m\N. \n\N. dist (X m) (X n) < e" + proof (intro exI allI impI) + fix m assume "N \ m" + hence m: "dist (X m) a < e/2" using N by fast + fix n assume "N \ n" + hence n: "dist (X n) a < e/2" using N by fast + have "dist (X m) (X n) \ dist (X m) a + dist (X n) a" + by (rule dist_triangle2) + also from m n have "\ < e" by simp + finally show "dist (X m) (X n) < e" . + qed +qed + +lemma convergent_Cauchy: "convergent X \ Cauchy X" +unfolding convergent_def +by (erule exE, erule LIMSEQ_imp_Cauchy) + +lemma Cauchy_convergent_iff: + fixes X :: "nat \ 'a::complete_space" + shows "Cauchy X = convergent X" +by (fast intro: Cauchy_convergent convergent_Cauchy) + +subsection {* Uniform Continuity *} + +definition + isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where + "isUCont f = (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" + +lemma isUCont_isCont: "isUCont f ==> isCont f x" +by (simp add: isUCont_def isCont_def LIM_def, force) + +lemma isUCont_Cauchy: + "\isUCont f; Cauchy X\ \ Cauchy (\n. f (X n))" +unfolding isUCont_def +apply (rule metric_CauchyI) +apply (drule_tac x=e in spec, safe) +apply (drule_tac e=s in metric_CauchyD, safe) +apply (rule_tac x=M in exI, simp) +done + +subsection {* The set of real numbers is a complete metric space *} + +instantiation real :: metric_space +begin + +definition dist_real_def: + "dist x y = \x - y\" + +definition open_real_def: + "open (S :: real set) \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" + +instance + by default (auto simp: open_real_def dist_real_def) +end + +instance real :: linorder_topology +proof + show "(open :: real set \ bool) = generate_topology (range lessThan \ range greaterThan)" + proof (rule ext, safe) + fix S :: "real set" assume "open S" + then guess f unfolding open_real_def bchoice_iff .. + then have *: "S = (\x\S. {x - f x <..} \ {..< x + f x})" + by (fastforce simp: dist_real_def) + show "generate_topology (range lessThan \ range greaterThan) S" + apply (subst *) + apply (intro generate_topology_Union generate_topology.Int) + apply (auto intro: generate_topology.Basis) + done + next + fix S :: "real set" assume "generate_topology (range lessThan \ range greaterThan) S" + moreover have "\a::real. open {.. (\y. \y - x\ < a - x \ y \ {..e>0. \y. \y - x\ < e \ y \ {..a::real. open {a <..}" + unfolding open_real_def dist_real_def + proof clarify + fix x a :: real assume "a < x" + hence "0 < x - a \ (\y. \y - x\ < x - a \ y \ {a<..})" by auto + thus "\e>0. \y. \y - x\ < e \ y \ {a<..}" .. + qed + ultimately show "open S" + by induct auto + qed +qed + +lemmas open_real_greaterThan = open_greaterThan[where 'a=real] +lemmas open_real_lessThan = open_lessThan[where 'a=real] +lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real] +lemmas closed_real_atMost = closed_atMost[where 'a=real] +lemmas closed_real_atLeast = closed_atLeast[where 'a=real] +lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real] + +text {* +Proof that Cauchy sequences converge based on the one from +http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html +*} + +text {* + If sequence @{term "X"} is Cauchy, then its limit is the lub of + @{term "{r::real. \N. \n\N. r < X n}"} +*} + +lemma isUb_UNIV_I: "(\y. y \ S \ y \ u) \ isUb UNIV S u" +by (simp add: isUbI setleI) + +lemma increasing_LIMSEQ: + fixes f :: "nat \ real" + assumes inc: "\n. f n \ f (Suc n)" + and bdd: "\n. f n \ l" + and en: "\e. 0 < e \ \n. l \ f n + e" + shows "f ----> l" +proof (rule increasing_tendsto) + fix x assume "x < l" + with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x" + by auto + from en[OF `0 < e`] obtain n where "l - e \ f n" + by (auto simp: field_simps) + with `e < l - x` `0 < e` have "x < f n" by simp + with incseq_SucI[of f, OF inc] show "eventually (\n. x < f n) sequentially" + by (auto simp: eventually_sequentially incseq_def intro: less_le_trans) +qed (insert bdd, auto) + +lemma real_Cauchy_convergent: + fixes X :: "nat \ real" + assumes X: "Cauchy X" + shows "convergent X" +proof - + def S \ "{x::real. \N. \n\N. x < X n}" + then have mem_S: "\N x. \n\N. x < X n \ x \ S" by auto + + { fix N x assume N: "\n\N. X n < x" + have "isUb UNIV S x" + proof (rule isUb_UNIV_I) + fix y::real assume "y \ S" + hence "\M. \n\M. y < X n" + by (simp add: S_def) + then obtain M where "\n\M. y < X n" .. + hence "y < X (max M N)" by simp + also have "\ < x" using N by simp + finally show "y \ x" + by (rule order_less_imp_le) + qed } + note bound_isUb = this + + have "\u. isLub UNIV S u" + proof (rule reals_complete) + obtain N where "\m\N. \n\N. dist (X m) (X n) < 1" + using X[THEN metric_CauchyD, OF zero_less_one] by auto + hence N: "\n\N. dist (X n) (X N) < 1" by simp + show "\x. x \ S" + proof + from N have "\n\N. X N - 1 < X n" + by (simp add: abs_diff_less_iff dist_real_def) + thus "X N - 1 \ S" by (rule mem_S) + qed + show "\u. isUb UNIV S u" + proof + from N have "\n\N. X n < X N + 1" + by (simp add: abs_diff_less_iff dist_real_def) + thus "isUb UNIV S (X N + 1)" + by (rule bound_isUb) + qed + qed + then obtain x where x: "isLub UNIV S x" .. + have "X ----> x" + proof (rule metric_LIMSEQ_I) + fix r::real assume "0 < r" + hence r: "0 < r/2" by simp + obtain N where "\n\N. \m\N. dist (X n) (X m) < r/2" + using metric_CauchyD [OF X r] by auto + hence "\n\N. dist (X n) (X N) < r/2" by simp + hence N: "\n\N. X N - r/2 < X n \ X n < X N + r/2" + by (simp only: dist_real_def abs_diff_less_iff) + + from N have "\n\N. X N - r/2 < X n" by fast + hence "X N - r/2 \ S" by (rule mem_S) + hence 1: "X N - r/2 \ x" using x isLub_isUb isUbD by fast + + from N have "\n\N. X n < X N + r/2" by fast + hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb) + hence 2: "x \ X N + r/2" using x isLub_le_isUb by fast + + show "\N. \n\N. dist (X n) x < r" + proof (intro exI allI impI) + fix n assume n: "N \ n" + from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+ + thus "dist (X n) x < r" using 1 2 + by (simp add: abs_diff_less_iff dist_real_def) + qed + qed + then show ?thesis unfolding convergent_def by auto +qed + +instance real :: complete_space + by intro_classes (rule real_Cauchy_convergent) + +lemma tendsto_dist [tendsto_intros]: + fixes l m :: "'a :: metric_space" + assumes f: "(f ---> l) F" and g: "(g ---> m) F" + shows "((\x. dist (f x) (g x)) ---> dist l m) F" +proof (rule tendstoI) + fix e :: real assume "0 < e" + hence e2: "0 < e/2" by simp + from tendstoD [OF f e2] tendstoD [OF g e2] + show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) F" + proof (eventually_elim) + case (elim x) + then show "dist (dist (f x) (g x)) (dist l m) < e" + unfolding dist_real_def + using dist_triangle2 [of "f x" "g x" "l"] + using dist_triangle2 [of "g x" "l" "m"] + using dist_triangle3 [of "l" "m" "f x"] + using dist_triangle [of "f x" "m" "g x"] + by arith + qed +qed + +lemma continuous_dist[continuous_intros]: + fixes f g :: "_ \ 'a :: metric_space" + shows "continuous F f \ continuous F g \ continuous F (\x. dist (f x) (g x))" + unfolding continuous_def by (rule tendsto_dist) + +lemma continuous_on_dist[continuous_on_intros]: + fixes f g :: "_ \ 'a :: metric_space" + shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. dist (f x) (g x))" + unfolding continuous_on_def by (auto intro: tendsto_dist) + +lemma tendsto_at_topI_sequentially: + fixes f :: "real \ real" + assumes mono: "mono f" + assumes limseq: "(\n. f (real n)) ----> y" + shows "(f ---> y) at_top" +proof (rule tendstoI) + fix e :: real assume "0 < e" + with limseq obtain N :: nat where N: "\n. N \ n \ \f (real n) - y\ < e" + by (auto simp: LIMSEQ_def dist_real_def) + { fix x :: real + from ex_le_of_nat[of x] guess n .. + note monoD[OF mono this] + also have "f (real_of_nat n) \ y" + by (rule LIMSEQ_le_const[OF limseq]) + (auto intro: exI[of _ n] monoD[OF mono] simp: real_eq_of_nat[symmetric]) + finally have "f x \ y" . } + note le = this + have "eventually (\x. real N \ x) at_top" + by (rule eventually_ge_at_top) + then show "eventually (\x. dist (f x) y < e) at_top" + proof eventually_elim + fix x assume N': "real N \ x" + with N[of N] le have "y - f (real N) < e" by auto + moreover note monoD[OF mono N'] + ultimately show "dist (f x) y < e" + using le[of x] by (auto simp: dist_real_def field_simps) + qed +qed + +lemma Cauchy_iff2: + "Cauchy X = (\j. (\M. \m \ M. \n \ M. \X m - X n\ < inverse(real (Suc j))))" + unfolding metric_Cauchy_iff2 dist_real_def .. + +end + diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Sat Mar 23 07:30:53 2013 +0100 @@ -26,14 +26,6 @@ lemma divide_nonneg_nonneg:assumes "a \ 0" "b \ 0" shows "0 \ a / (b::real)" apply(cases "b=0") defer apply(rule divide_nonneg_pos) using assms by auto -lemma continuous_setsum: - fixes f :: "'i \ 'a::t2_space \ 'b::real_normed_vector" - assumes f: "\i. i \ I \ continuous F (f i)" shows "continuous F (\x. \i\I. f i x)" -proof cases - assume "finite I" from this f show ?thesis - by (induct I) (auto intro!: continuous_intros) -qed (auto intro!: continuous_intros) - lemma brouwer_compactness_lemma: fixes f :: "'a::metric_space \ 'b::euclidean_space" assumes "compact s" "continuous_on s f" "\ (\x\s. (f x = 0))" diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Sat Mar 23 07:30:53 2013 +0100 @@ -1155,119 +1155,32 @@ subsection {* Connectedness of convex sets *} -lemma connected_real_lemma: - fixes f :: "real \ 'a::metric_space" - assumes ab: "a \ b" and fa: "f a \ e1" and fb: "f b \ e2" - and dst: "\e x. a <= x \ x <= b \ 0 < e ==> \d > 0. \y. abs(y - x) < d \ dist(f y) (f x) < e" - and e1: "\y \ e1. \e > 0. \y'. dist y' y < e \ y' \ e1" - and e2: "\y \ e2. \e > 0. \y'. dist y' y < e \ y' \ e2" - and e12: "~(\x \ a. x <= b \ f x \ e1 \ f x \ e2)" - shows "\x \ a. x <= b \ f x \ e1 \ f x \ e2" (is "\ x. ?P x") -proof - - let ?S = "{c. \x \ a. x <= c \ f x \ e1}" - have Se: " \x. x \ ?S" - apply (rule exI[where x=a]) - apply (auto simp add: fa) - done - have Sub: "\y. isUb UNIV ?S y" - apply (rule exI[where x= b]) - using ab fb e12 apply (auto simp add: isUb_def setle_def) - done - from reals_complete[OF Se Sub] obtain l where - l: "isLub UNIV ?S l"by blast - have alb: "a \ l" "l \ b" using l ab fa fb e12 - apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) - apply (metis linorder_linear) - done - have ale1: "\z \ a. z < l \ f z \ e1" using l - apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) - apply (metis linorder_linear not_le) - done - have th1: "\z x e d :: real. z <= x + e \ e < d ==> z < x \ abs(z - x) < d" by arith - have th2: "\e x:: real. 0 < e ==> ~(x + e <= x)" by arith - have "\d::real. 0 < d \ 0 < d/2 \ d/2 < d" by simp - then have th3: "\d::real. d > 0 \ \e > 0. e < d" by blast - { assume le2: "f l \ e2" - from le2 fa fb e12 alb have la: "l \ a" by metis - then have lap: "l - a > 0" using alb by arith - from e2[rule_format, OF le2] obtain e where - e: "e > 0" "\y. dist y (f l) < e \ y \ e2" by metis - from dst[OF alb e(1)] obtain d where - d: "d > 0" "\y. \y - l\ < d \ dist (f y) (f l) < e" by metis - let ?d' = "min (d/2) ((l - a)/2)" - have "?d' < d \ 0 < ?d' \ ?d' < l - a" using lap d(1) - by (simp add: min_max.less_infI2) - then have "\d'. d' < d \ d' >0 \ l - d' > a" by auto - then obtain d' where d': "d' > 0" "d' < d" "l - d' > a" by metis - from d e have th0: "\y. \y - l\ < d \ f y \ e2" by metis - from th0[rule_format, of "l - d'"] d' have "f (l - d') \ e2" by auto - moreover - have "f (l - d') \ e1" using ale1[rule_format, of "l -d'"] d' by auto - ultimately have False using e12 alb d' by auto } - moreover - { assume le1: "f l \ e1" - from le1 fa fb e12 alb have lb: "l \ b" by metis - then have blp: "b - l > 0" using alb by arith - from e1[rule_format, OF le1] obtain e where - e: "e > 0" "\y. dist y (f l) < e \ y \ e1" by metis - from dst[OF alb e(1)] obtain d where - d: "d > 0" "\y. \y - l\ < d \ dist (f y) (f l) < e" by metis - have "\d::real. 0 < d \ d/2 < d \ 0 < d/2" by simp - then have "\d'. d' < d \ d' >0" using d(1) by blast - then obtain d' where d': "d' > 0" "d' < d" by metis - from d e have th0: "\y. \y - l\ < d \ f y \ e1" by auto - then have "\y. l \ y \ y \ l + d' \ f y \ e1" using d' by auto - with ale1 have "\y. a \ y \ y \ l + d' \ f y \ e1" by auto - with l d' have False - by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) } - ultimately show ?thesis using alb by metis -qed +lemma connectedD: + "connected S \ open A \ open B \ S \ A \ B \ A \ B \ S = {} \ A \ S = {} \ B \ S = {}" + by (metis connected_def) lemma convex_connected: fixes s :: "'a::real_normed_vector set" assumes "convex s" shows "connected s" -proof - - { fix e1 e2 - assume as:"open e1" "open e2" "e1 \ e2 \ s = {}" "s \ e1 \ e2" - assume "e1 \ s \ {}" "e2 \ s \ {}" - then obtain x1 x2 where x1:"x1\e1" "x1\s" and x2:"x2\e2" "x2\s" by auto - then have n: "norm (x1 - x2) > 0" unfolding zero_less_norm_iff using as(3) by auto - - { fix x e::real assume as:"0 \ x" "x \ 1" "0 < e" - { fix y - have *: "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2) = (y - x) *\<^sub>R x1 - (y - x) *\<^sub>R x2" - by (simp add: algebra_simps) - assume "\y - x\ < e / norm (x1 - x2)" - hence "norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e" - unfolding * and scaleR_right_diff_distrib[symmetric] - unfolding less_divide_eq using n by auto - } - then have "\d>0. \y. \y - x\ < d \ norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e" - apply (rule_tac x="e / norm (x1 - x2)" in exI) - using as - apply auto - unfolding zero_less_divide_iff - using n apply simp - done - } note * = this - - have "\x\0. x \ 1 \ (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \ e1 \ (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \ e2" - apply (rule connected_real_lemma) - apply (simp add: `x1\e1` `x2\e2` dist_commute)+ - using * apply (simp add: dist_norm) - using as(1,2)[unfolded open_dist] apply simp - using as(1,2)[unfolded open_dist] apply simp - using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]] using x1 x2 - using as(3) apply auto - done - then obtain x where "x\0" "x\1" "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \ e1" "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \ e2" - by auto - then have False - using as(4) - using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]] - using x1(2) x2(2) by auto - } - then show ?thesis unfolding connected_def by auto +proof (rule connectedI) + fix A B + assume "open A" "open B" "A \ B \ s = {}" "s \ A \ B" + moreover + assume "A \ s \ {}" "B \ s \ {}" + then obtain a b where a: "a \ A" "a \ s" and b: "b \ B" "b \ s" by auto + def f \ "\u. u *\<^sub>R a + (1 - u) *\<^sub>R b" + then have "continuous_on {0 .. 1} f" + by (auto intro!: continuous_on_intros) + then have "connected (f ` {0 .. 1})" + by (auto intro!: connected_continuous_image) + note connectedD[OF this, of A B] + moreover have "a \ A \ f ` {0 .. 1}" + using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def) + moreover have "b \ B \ f ` {0 .. 1}" + using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def) + moreover have "f ` {0 .. 1} \ s" + using `convex s` a b unfolding convex_def f_def by auto + ultimately show False by auto qed text {* One rather trivial consequence. *} @@ -3491,11 +3404,11 @@ apply(rule,rule) defer apply(rule) unfolding inner_minus_left and neg_less_iff_less proof- from ab have "((\x. inner a x) ` t) *<= (inner a y - b)" apply(erule_tac x=y in ballE) apply(rule setleI) using `y\s` by auto - hence k:"isLub UNIV ((\x. inner a x) ` t) k" unfolding k_def apply(rule_tac Sup) using assms(5) by auto + hence k:"isLub UNIV ((\x. inner a x) ` t) k" unfolding k_def apply(rule_tac isLub_cSup) using assms(5) by auto fix x assume "x\t" thus "inner a x < (k + b / 2)" using `0s" - hence "k \ inner a x - b" unfolding k_def apply(rule_tac Sup_least) using assms(5) + hence "k \ inner a x - b" unfolding k_def apply(rule_tac cSup_least) using assms(5) using ab[THEN bspec[where x=x]] by auto thus "k + b / 2 < inner a x" using `0 < b` by auto qed @@ -3544,9 +3457,9 @@ thus ?thesis apply(rule_tac x=a in exI, rule_tac x="Sup ((\x. inner a x) ` s)" in exI) using `a\0` apply auto - apply (rule Sup[THEN isLubD2]) + apply (rule isLub_cSup[THEN isLubD2]) prefer 4 - apply (rule Sup_least) + apply (rule cSup_least) using assms(3-5) apply (auto simp add: setle_def) apply metis done @@ -6208,7 +6121,7 @@ using interior_subset[of I] `x \ interior I` by auto have "Inf (?F x) \ (f x - f y) / (x - y)" - proof (rule Inf_lower2) + proof (rule cInf_lower2) show "(f x - f t) / (x - t) \ ?F x" using `t \ I` `x < t` by auto show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" @@ -6231,7 +6144,7 @@ with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto have "(f x - f y) / (x - y) \ Inf (?F x)" - proof (rule Inf_greatest) + proof (rule cInf_greatest) have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" using `y < x` by (auto simp: field_simps) also diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Sat Mar 23 07:30:53 2013 +0100 @@ -576,10 +576,6 @@ unfolding FDERIV_conv_has_derivative [symmetric] by (rule FDERIV_unique) -lemma continuous_isCont: "isCont f x = continuous (at x) f" - unfolding isCont_def LIM_def - unfolding continuous_at Lim_at unfolding dist_nz by auto - lemma frechet_derivative_unique_within_closed_interval: fixes f::"'a::ordered_euclidean_space \ 'b::real_normed_vector" assumes "\i\Basis. a\i < b\i" "x \ {a..b}" (is "x\?I") @@ -783,15 +779,12 @@ shows "\x\{a<..x\{a<..xa. f' x xa - (f b - f a) / (b - a) * xa) = (\v. 0)" - apply(rule rolle[OF assms(1), of "\x. f x - (f b - f a) / (b - a) * x"]) - defer - apply(rule continuous_on_intros assms(2))+ - proof + proof (intro rolle[OF assms(1), of "\x. f x - (f b - f a) / (b - a) * x"] ballI) fix x assume x:"x \ {a<..x. f x - (f b - f a) / (b - a) * x) has_derivative (\xa. f' x xa - (f b - f a) / (b - a) * xa)) (at x)" by (intro has_derivative_intros assms(3)[rule_format,OF x] mult_right_has_derivative) - qed(insert assms(1), auto simp add:field_simps) + qed (insert assms(1,2), auto intro!: continuous_on_intros simp: field_simps) then guess x .. thus ?thesis apply(rule_tac x=x in bexI) apply(drule fun_cong[of _ _ "b - a"]) by auto diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy --- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Sat Mar 23 07:30:53 2013 +0100 @@ -470,7 +470,7 @@ fixes f :: "'a::t2_space => real" fixes A assumes "open A" shows "continuous_on A f <-> continuous_on A (ereal o f)" - using continuous_at_iff_ereal assms by (auto simp add: continuous_on_eq_continuous_at) + using continuous_at_iff_ereal assms by (auto simp add: continuous_on_eq_continuous_at cong del: continuous_on_cong) lemma continuous_on_real: "continuous_on (UNIV-{\,(-\::ereal)}) real" @@ -523,7 +523,7 @@ proof - { assume "S ~= {}" { assume ex: "EX B. ALL x:S. B<=x" - then have *: "ALL x:S. Inf S <= x" using Inf_lower_EX[of _ S] ex by metis + then have *: "ALL x:S. Inf S <= x" using cInf_lower_EX[of _ S] ex by metis then have "Inf S : S" apply (subst closed_contains_Inf) using ex `S ~= {}` `closed S` by auto then have "ALL x. (Inf S <= x <-> x:S)" using mono[rule_format, of "Inf S"] * by auto then have "S = {Inf S ..}" by auto diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Fashoda.thy --- a/src/HOL/Multivariate_Analysis/Fashoda.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Fashoda.thy Sat Mar 23 07:30:53 2013 +0100 @@ -18,7 +18,7 @@ subsection {*Fashoda meet theorem. *} lemma infnorm_2: "infnorm (x::real^2) = max (abs(x$1)) (abs(x$2))" - unfolding infnorm_cart UNIV_2 apply(rule Sup_eq) by auto + unfolding infnorm_cart UNIV_2 apply(rule cSup_eq) by auto lemma infnorm_eq_1_2: "infnorm (x::real^2) = 1 \ (abs(x$1) \ 1 \ abs(x$2) \ 1 \ (x$1 = -1 \ x$1 = 1 \ x$2 = -1 \ x$2 = 1))" diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Sat Mar 23 07:30:53 2013 +0100 @@ -8,6 +8,45 @@ "~~/src/HOL/Library/Indicator_Function" begin +lemma cSup_finite_ge_iff: + fixes S :: "real set" + assumes fS: "finite S" and Se: "S \ {}" + shows "a \ Sup S \ (\ x \ S. a \ x)" +by (metis Max_ge Se cSup_eq_Max Max_in fS linorder_not_le less_le_trans) + +lemma cSup_finite_le_iff: + fixes S :: "real set" + assumes fS: "finite S" and Se: "S \ {}" + shows "a \ Sup S \ (\ x \ S. a \ x)" +by (metis Max_ge Se cSup_eq_Max Max_in fS order_trans) + +lemma Inf: (* rename *) + fixes S :: "real set" + shows "S \ {} ==> (\b. b <=* S) ==> isGlb UNIV S (Inf S)" +by (auto simp add: isLb_def setle_def setge_def isGlb_def greatestP_def intro: cInf_lower cInf_greatest) + +lemma real_le_inf_subset: + assumes "t \ {}" "t \ s" "\b. b <=* s" + shows "Inf s <= Inf (t::real set)" + apply (rule isGlb_le_isLb) + apply (rule Inf[OF assms(1)]) + apply (insert assms) + apply (erule exE) + apply (rule_tac x = b in exI) + apply (auto simp: isLb_def setge_def intro: cInf_lower cInf_greatest) + done + +lemma real_ge_sup_subset: + assumes "t \ {}" "t \ s" "\b. s *<= b" + shows "Sup s >= Sup (t::real set)" + apply (rule isLub_le_isUb) + apply (rule isLub_cSup[OF assms(1)]) + apply (insert assms) + apply (erule exE) + apply (rule_tac x = b in exI) + apply (auto simp: isUb_def setle_def intro: cSup_upper cSup_least) + done + (*declare not_less[simp] not_le[simp]*) lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib @@ -49,33 +88,6 @@ shows "bounded_linear f" unfolding bounded_linear_def additive_def bounded_linear_axioms_def using assms by auto -lemma Inf: (* rename *) - fixes S :: "real set" - shows "S \ {} ==> (\b. b <=* S) ==> isGlb UNIV S (Inf S)" -by (auto simp add: isLb_def setle_def setge_def isGlb_def greatestP_def) - -lemma real_le_inf_subset: - assumes "t \ {}" "t \ s" "\b. b <=* s" - shows "Inf s <= Inf (t::real set)" - apply (rule isGlb_le_isLb) - apply (rule Inf[OF assms(1)]) - apply (insert assms) - apply (erule exE) - apply (rule_tac x = b in exI) - apply (auto simp: isLb_def setge_def) - done - -lemma real_ge_sup_subset: - assumes "t \ {}" "t \ s" "\b. s *<= b" - shows "Sup s >= Sup (t::real set)" - apply (rule isLub_le_isUb) - apply (rule Sup[OF assms(1)]) - apply (insert assms) - apply (erule exE) - apply (rule_tac x = b in exI) - apply (auto simp: isUb_def setle_def) - done - lemma bounded_linear_component [intro]: "bounded_linear (\x::'a::euclidean_space. x \ k)" by (rule bounded_linear_inner_left) @@ -387,14 +399,14 @@ interval_upperbound {a..b} = (b::'a::ordered_euclidean_space)" unfolding interval_upperbound_def euclidean_representation_setsum by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setle_def - intro!: Sup_unique) + intro!: cSup_unique) lemma interval_lowerbound[simp]: "\i\Basis. a\i \ b\i \ interval_lowerbound {a..b} = (a::'a::ordered_euclidean_space)" unfolding interval_lowerbound_def euclidean_representation_setsum by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setge_def - intro!: Inf_unique) + intro!: cInf_unique) lemmas interval_bounds = interval_upperbound interval_lowerbound @@ -3201,7 +3213,7 @@ 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 + hence N:"\x\(\(x, k). norm (f x)) ` p. x \ real N" apply(subst(asm) cSup_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]] @@ -5480,7 +5492,7 @@ "\d. d division_of {a..b} \ setsum (\k. norm(integral k f)) d \ B" shows "f absolutely_integrable_on {a..b}" proof- let ?S = "(\d. setsum (\k. norm(integral k f)) d) ` {d. d division_of {a..b} }" def i \ "Sup ?S" - have i:"isLub UNIV ?S i" unfolding i_def apply(rule Sup) + have i:"isLub UNIV ?S i" unfolding i_def apply(rule isLub_cSup) apply(rule elementary_interval) defer apply(rule_tac x=B in exI) apply(rule setleI) using assms(2) by auto show ?thesis apply(rule,rule assms) apply rule apply(subst has_integral[of _ i]) @@ -5693,7 +5705,7 @@ shows "f absolutely_integrable_on UNIV" proof(rule absolutely_integrable_onI,fact,rule) let ?S = "(\d. setsum (\k. norm(integral k f)) d) ` {d. d division_of (\d)}" def i \ "Sup ?S" - have i:"isLub UNIV ?S i" unfolding i_def apply(rule Sup) + have i:"isLub UNIV ?S i" unfolding i_def apply(rule isLub_cSup) apply(rule elementary_interval) defer apply(rule_tac x=B in exI) apply(rule setleI) using assms(2) by auto have f_int:"\a b. f absolutely_integrable_on {a..b}" @@ -6044,7 +6056,7 @@ prefer 5 unfolding real_norm_def apply rule - apply (rule Inf_abs_ge) + apply (rule cInf_abs_ge) prefer 5 apply rule apply (rule_tac g = h in absolutely_integrable_integrable_bound_real) @@ -6065,11 +6077,11 @@ fix x assume x: "x \ s" show "Inf {f j x |j. j \ {m..m + Suc k}} \ Inf {f j x |j. j \ {m..m + k}}" - apply (rule Inf_ge) + apply (rule cInf_ge) unfolding setge_def defer apply rule - apply (subst Inf_finite_le_iff) + apply (subst cInf_finite_le_iff) prefer 3 apply (rule_tac x=xa in bexI) apply auto @@ -6119,7 +6131,7 @@ prefer 3 apply (rule,rule isGlbD1[OF i]) prefer 3 - apply (subst Inf_finite_le_iff) + apply (subst cInf_finite_le_iff) prefer 3 apply (rule_tac x=y in bexI) using N goal1 @@ -6146,7 +6158,7 @@ apply(rule absolutely_integrable_sup_real) prefer 5 unfolding real_norm_def apply rule - apply (rule Sup_abs_le) + apply (rule cSup_abs_le) prefer 5 apply rule apply (rule_tac g=h in absolutely_integrable_integrable_bound_real) @@ -6167,11 +6179,11 @@ fix x assume x: "x\s" show "Sup {f j x |j. j \ {m..m + Suc k}} \ Sup {f j x |j. j \ {m..m + k}}" - apply (rule Sup_le) + apply (rule cSup_le) unfolding setle_def defer apply rule - apply (subst Sup_finite_ge_iff) + apply (subst cSup_finite_ge_iff) prefer 3 apply (rule_tac x=y in bexI) apply auto @@ -6183,7 +6195,7 @@ case goal1 note r=this have i: "isLub UNIV ?S i" unfolding i_def - apply (rule Sup) + apply (rule isLub_cSup) defer apply (rule_tac x="h x" in exI) unfolding setle_def @@ -6221,7 +6233,7 @@ prefer 3 apply (rule, rule isLubD1[OF i]) prefer 3 - apply (subst Sup_finite_ge_iff) + apply (subst cSup_finite_ge_iff) prefer 3 apply (rule_tac x = y in bexI) using N goal1 @@ -6246,7 +6258,7 @@ apply fact+ unfolding real_norm_def apply rule - apply (rule Inf_abs_ge) + apply (rule cInf_abs_ge) using assms(3) apply auto done @@ -6276,7 +6288,7 @@ apply (rule_tac x=N in exI,safe) unfolding real_norm_def apply (rule le_less_trans[of _ "r/2"]) - apply (rule Inf_asclose) + apply (rule cInf_asclose) apply safe defer apply (rule less_imp_le) @@ -6302,7 +6314,7 @@ apply fact+ unfolding real_norm_def apply rule - apply (rule Sup_abs_le) + apply (rule cSup_abs_le) using assms(3) apply auto done @@ -6330,7 +6342,7 @@ apply (rule_tac x=N in exI,safe) unfolding real_norm_def apply (rule le_less_trans[of _ "r/2"]) - apply (rule Sup_asclose) + apply (rule cSup_asclose) apply safe defer apply (rule less_imp_le) @@ -6364,7 +6376,7 @@ assume x: "x \ s" have *: "\x y::real. x \ - y \ - x \ y" by auto show "Inf {f j x |j. n \ j} \ f n x" - apply (rule Inf_lower[where z="- h x"]) + apply (rule cInf_lower[where z="- h x"]) defer apply (rule *) using assms(3)[rule_format,OF x] @@ -6377,7 +6389,7 @@ fix x assume x: "x \ s" show "f n x \ Sup {f j x |j. n \ j}" - apply (rule Sup_upper[where z="h x"]) + apply (rule cSup_upper[where z="h x"]) defer using assms(3)[rule_format,OF x] unfolding real_norm_def abs_le_iff diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Sat Mar 23 07:30:53 2013 +0100 @@ -22,7 +22,7 @@ qed lemma square_continuous: "0 < (e::real) ==> \d. 0 < d \ (\y. abs(y - x) < d \ abs(y * y - x * x) < e)" - using isCont_power[OF isCont_ident, of 2, unfolded isCont_def LIM_eq, rule_format, of e x] + using isCont_power[OF isCont_ident, of x, unfolded isCont_def LIM_eq, rule_format, of e 2] apply (auto simp add: power2_eq_square) apply (rule_tac x="s" in exI) apply auto @@ -2498,6 +2498,9 @@ "{ abs ((x::'a::euclidean_space) \ i) |i. i \ Basis} = (\i. abs(x \ i)) ` Basis" by blast +lemma infnorm_Max: "infnorm (x::'a::euclidean_space) = Max ((\i. abs(x \ i)) ` Basis)" + by (simp add: infnorm_def infnorm_set_image cSup_eq_Max) + lemma infnorm_set_lemma: shows "finite {abs((x::'a::euclidean_space) \ i) |i. i \ Basis}" and "{abs(x \ i) |i. i \ Basis} \ {}" @@ -2505,41 +2508,22 @@ by auto lemma infnorm_pos_le: "0 \ infnorm (x::'a::euclidean_space)" - unfolding infnorm_def - unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma] - unfolding infnorm_set_image - by (auto simp: ex_in_conv) + by (simp add: infnorm_Max Max_ge_iff ex_in_conv) lemma infnorm_triangle: "infnorm ((x::'a::euclidean_space) + y) \ infnorm x + infnorm y" proof - - have th: "\x y (z::real). x - y <= z \ x - z <= y" by arith - have th1: "\S f. f ` S = { f i| i. i \ S}" by blast - have th2: "\x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith + have *: "\a b c d :: real. \a\ \ c \ \b\ \ d \ \a + b\ \ c + d" + by simp show ?thesis - unfolding infnorm_def - unfolding Sup_finite_le_iff[ OF infnorm_set_lemma[where 'a='a]] - apply (subst diff_le_eq[symmetric]) - unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma] - unfolding infnorm_set_image bex_simps - apply (subst th) - unfolding th1 - unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma[where 'a='a]] - unfolding infnorm_set_image ball_simps bex_simps - apply (simp add: inner_add_left) - apply (metis th2) - done + by (auto simp: infnorm_Max inner_add_left intro!: *) qed lemma infnorm_eq_0: "infnorm x = 0 \ (x::_::euclidean_space) = 0" proof - - have "infnorm x <= 0 \ x = 0" - unfolding infnorm_def - unfolding Sup_finite_le_iff[OF infnorm_set_lemma] - unfolding infnorm_set_image ball_simps - apply (subst (1) euclidean_eq_iff) - apply auto - done - then show ?thesis using infnorm_pos_le[of x] by simp + have "infnorm x \ 0 \ x = 0" + unfolding infnorm_Max by (simp add: euclidean_all_zero_iff) + then show ?thesis + using infnorm_pos_le[of x] by simp qed lemma infnorm_0: "infnorm 0 = 0" @@ -2573,40 +2557,24 @@ using infnorm_pos_le[of x] by arith lemma Basis_le_infnorm: - assumes b: "b \ Basis" shows "\x \ b\ \ infnorm (x::'a::euclidean_space)" - unfolding infnorm_def -proof (subst Sup_finite_ge_iff) - let ?S = "{\x \ i\ |i. i \ Basis}" - show "finite ?S" by (rule infnorm_set_lemma) - show "?S \ {}" by auto - show "Bex ?S (op \ \x \ b\)" - using b by (auto intro!: exI[of _ b]) -qed - -lemma infnorm_mul_lemma: "infnorm(a *\<^sub>R x) <= \a\ * infnorm x" - apply (subst infnorm_def) - unfolding Sup_finite_le_iff[OF infnorm_set_lemma] - unfolding infnorm_set_image ball_simps inner_scaleR abs_mult - using Basis_le_infnorm[of _ x] - apply (auto intro: mult_mono) - done + "b \ Basis \ \x \ b\ \ infnorm (x::'a::euclidean_space)" + by (simp add: infnorm_Max) lemma infnorm_mul: "infnorm(a *\<^sub>R x) = abs a * infnorm x" -proof - - { assume a0: "a = 0" then have ?thesis by (simp add: infnorm_0) } - moreover - { assume a0: "a \ 0" - from a0 have th: "(1/a) *\<^sub>R (a *\<^sub>R x) = x" by simp - from a0 have ap: "\a\ > 0" by arith - from infnorm_mul_lemma[of "1/a" "a *\<^sub>R x"] - have "infnorm x \ 1/\a\ * infnorm (a*\<^sub>R x)" - unfolding th by simp - with ap have "\a\ * infnorm x \ \a\ * (1/\a\ * infnorm (a *\<^sub>R x))" by (simp add: field_simps) - then have "\a\ * infnorm x \ infnorm (a*\<^sub>R x)" - using ap by (simp add: field_simps) - with infnorm_mul_lemma[of a x] have ?thesis by arith } - ultimately show ?thesis by blast -qed + unfolding infnorm_Max +proof (safe intro!: Max_eqI) + let ?B = "(\i. \x \ i\) ` Basis" + show "\b :: 'a. b \ Basis \ \a *\<^sub>R x \ b\ \ \a\ * Max ?B" + by (simp add: abs_mult mult_left_mono) + + from Max_in[of ?B] obtain b where "b \ Basis" "Max ?B = \x \ b\" + by (auto simp del: Max_in) + then show "\a\ * Max ((\i. \x \ i\) ` Basis) \ (\i. \a *\<^sub>R x \ i\) ` Basis" + by (intro image_eqI[where x=b]) (auto simp: abs_mult) +qed simp + +lemma infnorm_mul_lemma: "infnorm(a *\<^sub>R x) \ \a\ * infnorm x" + unfolding infnorm_mul .. lemma infnorm_pos_lt: "infnorm x > 0 \ x \ 0" using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith @@ -2614,15 +2582,13 @@ text {* Prove that it differs only up to a bound from Euclidean norm. *} lemma infnorm_le_norm: "infnorm x \ norm x" - unfolding infnorm_def Sup_finite_le_iff[OF infnorm_set_lemma] - unfolding infnorm_set_image ball_simps - by (metis Basis_le_norm) + by (simp add: Basis_le_norm infnorm_Max) lemma euclidean_inner: "inner x y = (\b\Basis. (x \ b) * (y \ b))" by (subst (1 2) euclidean_representation[symmetric, where 'a='a]) (simp add: inner_setsum_left inner_setsum_right setsum_cases inner_Basis ac_simps if_distrib) -lemma norm_le_infnorm: "norm(x) <= sqrt DIM('a) * infnorm(x::'a::euclidean_space)" +lemma norm_le_infnorm: "norm x \ sqrt DIM('a) * infnorm(x::'a::euclidean_space)" proof - let ?d = "DIM('a)" have "real ?d \ 0" by simp @@ -2639,10 +2605,7 @@ apply (auto simp add: power2_eq_square[symmetric]) apply (subst power2_abs[symmetric]) apply (rule power_mono) - unfolding infnorm_def Sup_finite_ge_iff[OF infnorm_set_lemma] - unfolding infnorm_set_image bex_simps - apply (rule_tac x=i in bexI) - apply auto + apply (auto simp: infnorm_Max) done from real_le_lsqrt[OF inner_ge_zero th th1] show ?thesis unfolding norm_eq_sqrt_inner id_def . diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Operator_Norm.thy --- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy Sat Mar 23 07:30:53 2013 +0100 @@ -58,7 +58,7 @@ hence Se: "?S \ {}" by auto from linear_bounded[OF lf] have b: "\ b. ?S *<= b" unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def) - {from Sup[OF Se b, unfolded onorm_def[symmetric]] + { from isLub_cSup[OF Se b, unfolded onorm_def[symmetric]] show "norm (f x) <= onorm f * norm x" apply - apply (rule spec[where x = x]) @@ -66,7 +66,7 @@ by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} { show "\x. norm (f x) <= b * norm x \ onorm f <= b" - using Sup[OF Se b, unfolded onorm_def[symmetric]] + using isLub_cSup[OF Se b, unfolded onorm_def[symmetric]] unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} } @@ -93,7 +93,7 @@ by (auto simp: SOME_Basis intro!: exI[of _ "SOME i. i \ Basis"]) show ?thesis unfolding onorm_def th - apply (rule Sup_unique) by (simp_all add: setle_def) + apply (rule cSup_unique) by (simp_all add: setle_def) qed lemma onorm_pos_lt: assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Path_Connected.thy --- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Sat Mar 23 07:30:53 2013 +0100 @@ -111,106 +111,33 @@ assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" unfolding path_def pathfinish_def pathstart_def - apply rule defer - apply(erule conjE) -proof - - assume as: "continuous_on {0..1} (g1 +++ g2)" - have *: "g1 = (\x. g1 (2 *\<^sub>R x)) \ (\x. (1/2) *\<^sub>R x)" - "g2 = (\x. g2 (2 *\<^sub>R x - 1)) \ (\x. (1/2) *\<^sub>R (x + 1))" - unfolding o_def by (auto simp add: add_divide_distrib) - have "op *\<^sub>R (1 / 2) ` {0::real..1} \ {0..1}" - "(\x. (1 / 2) *\<^sub>R (x + 1)) ` {(0::real)..1} \ {0..1}" +proof safe + assume cont: "continuous_on {0..1} (g1 +++ g2)" + have g1: "continuous_on {0..1} g1 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2))" + by (intro continuous_on_cong refl) (auto simp: joinpaths_def) + have g2: "continuous_on {0..1} g2 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2 + 1/2))" + using assms by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def) + show "continuous_on {0..1} g1" "continuous_on {0..1} g2" + unfolding g1 g2 + by (auto intro!: continuous_on_intros continuous_on_subset[OF cont] simp del: o_apply) +next + assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2" + have 01: "{0 .. 1} = {0..1/2} \ {1/2 .. 1::real}" by auto - then show "continuous_on {0..1} g1 \ continuous_on {0..1} g2" - apply - - apply rule - apply (subst *) defer - apply (subst *) - apply (rule_tac[!] continuous_on_compose) - apply (intro continuous_on_intros) defer - apply (intro continuous_on_intros) - apply (rule_tac[!] continuous_on_eq[of _ "g1 +++ g2"]) defer prefer 3 - apply (rule_tac[1-2] continuous_on_subset[of "{0 .. 1}"]) - apply (rule as, assumption, rule as, assumption) - apply rule defer - apply rule - proof - - fix x - assume "x \ op *\<^sub>R (1 / 2) ` {0::real..1}" - then have "x \ 1 / 2" unfolding image_iff by auto - then show "(g1 +++ g2) x = g1 (2 *\<^sub>R x)" unfolding joinpaths_def by auto - next - fix x - assume "x \ (\x. (1 / 2) *\<^sub>R (x + 1)) ` {0::real..1}" - then have "x \ 1 / 2" unfolding image_iff by auto - then show "(g1 +++ g2) x = g2 (2 *\<^sub>R x - 1)" - proof (cases "x = 1 / 2") - case True - then have "x = (1/2) *\<^sub>R 1" by auto - then show ?thesis - unfolding joinpaths_def - using assms[unfolded pathstart_def pathfinish_def] - by (auto simp add: mult_ac) - qed (auto simp add:le_less joinpaths_def) - qed -next - assume as:"continuous_on {0..1} g1" "continuous_on {0..1} g2" - have *: "{0 .. 1::real} = {0.. (1/2)*\<^sub>R 1} \ {(1/2) *\<^sub>R 1 .. 1}" by auto - have **: "op *\<^sub>R 2 ` {0..(1 / 2) *\<^sub>R 1} = {0..1::real}" - apply (rule set_eqI, rule) - unfolding image_iff - defer - apply (rule_tac x="(1/2)*\<^sub>R x" in bexI) - apply auto - done - have ***: "(\x. 2 *\<^sub>R x - 1) ` {(1 / 2) *\<^sub>R 1..1} = {0..1::real}" - apply (auto simp add: image_def) - apply (rule_tac x="(x + 1) / 2" in bexI) - apply (auto simp add: add_divide_distrib) - done + { fix x :: real assume "0 \ x" "x \ 1" then have "x \ (\x. x * 2) ` {0..1 / 2}" + by (intro image_eqI[where x="x/2"]) auto } + note 1 = this + { fix x :: real assume "0 \ x" "x \ 1" then have "x \ (\x. x * 2 - 1) ` {1 / 2..1}" + by (intro image_eqI[where x="x/2 + 1/2"]) auto } + note 2 = this show "continuous_on {0..1} (g1 +++ g2)" - unfolding * - apply (rule continuous_on_union) - apply (rule closed_real_atLeastAtMost)+ - proof - - show "continuous_on {0..(1 / 2) *\<^sub>R 1} (g1 +++ g2)" - apply (rule continuous_on_eq[of _ "\x. g1 (2 *\<^sub>R x)"]) defer - unfolding o_def[THEN sym] - apply (rule continuous_on_compose) - apply (intro continuous_on_intros) - unfolding ** - apply (rule as(1)) - unfolding joinpaths_def - apply auto - done - next - show "continuous_on {(1/2)*\<^sub>R1..1} (g1 +++ g2)" - apply (rule continuous_on_eq[of _ "g2 \ (\x. 2 *\<^sub>R x - 1)"]) defer - apply (rule continuous_on_compose) - apply (intro continuous_on_intros) - unfolding *** o_def joinpaths_def - apply (rule as(2)) - using assms[unfolded pathstart_def pathfinish_def] - apply (auto simp add: mult_ac) - done - qed + using assms unfolding joinpaths_def 01 + by (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_on_intros) + (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2) qed lemma path_image_join_subset: "path_image(g1 +++ g2) \ (path_image g1 \ path_image g2)" -proof - fix x - assume "x \ path_image (g1 +++ g2)" - then obtain y where y:"y\{0..1}" "x = (if y \ 1 / 2 then g1 (2 *\<^sub>R y) else g2 (2 *\<^sub>R y - 1))" - unfolding path_image_def image_iff joinpaths_def by auto - then show "x \ path_image g1 \ path_image g2" - apply (cases "y \ 1/2") - apply (rule_tac UnI1) defer - apply (rule_tac UnI2) - unfolding y(2) path_image_def - using y(1) - apply (auto intro!: imageI) - done -qed + unfolding path_image_def joinpaths_def by auto lemma subset_path_image_join: assumes "path_image g1 \ s" "path_image g2 \ s" @@ -218,7 +145,7 @@ using path_image_join_subset[of g1 g2] and assms by auto lemma path_image_join: - assumes "path g1" "path g2" "pathfinish g1 = pathstart g2" + assumes "pathfinish g1 = pathstart g2" shows "path_image(g1 +++ g2) = (path_image g1) \ (path_image g2)" apply (rule, rule path_image_join_subset, rule) unfolding Un_iff @@ -240,7 +167,7 @@ then show "x \ path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff apply(rule_tac x="(1/2) *\<^sub>R (y + 1)" in bexI) - using assms(3)[unfolded pathfinish_def pathstart_def] + using assms(1)[unfolded pathfinish_def pathstart_def] apply (auto simp add: add_divide_distrib) done qed @@ -755,8 +682,8 @@ unfolding xy apply (rule_tac x="f \ g" in exI) unfolding path_defs - using assms(1) - apply (auto intro!: continuous_on_compose continuous_on_subset[of _ _ "g ` {0..1}"]) + apply (intro conjI continuous_on_compose continuous_on_subset[OF assms(1)]) + apply auto done qed diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Sat Mar 23 07:30:53 2013 +0100 @@ -34,6 +34,24 @@ "finite I \ (\i. i \ I \ countable (F i)) \ countable (PiE I F)" by (induct I arbitrary: F rule: finite_induct) (auto simp: PiE_insert_eq) +lemma Lim_within_open: + fixes f :: "'a::topological_space \ 'b::topological_space" + shows "a \ S \ open S \ (f ---> l)(at a within S) \ (f ---> l)(at a)" + by (fact tendsto_within_open) + +lemma Lim_within_subset: "(f ---> l) (net within S) \ T \ S \ (f ---> l) (net within T)" + by (fact tendsto_within_subset) + +lemma continuous_on_union: + "closed s \ closed t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" + by (fact continuous_on_closed_Un) + +lemma continuous_on_cases: + "closed s \ closed t \ continuous_on s f \ continuous_on t g \ + \x. (x\s \ \ P x) \ (x \ t \ P x) \ f x = g x \ + continuous_on (s \ t) (\x. if P x then f x else g x)" + by (rule continuous_on_If) auto + subsection {* Topological Basis *} context topological_space @@ -181,44 +199,15 @@ end -class first_countable_topology = topological_space + - assumes first_countable_basis: - "\A. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - -lemma (in first_countable_topology) countable_basis_at_decseq: - obtains A :: "nat \ 'a set" where - "\i. open (A i)" "\i. x \ (A i)" - "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" -proof atomize_elim - from first_countable_basis[of x] obtain A - where "countable A" - and nhds: "\a. a \ A \ open a" "\a. a \ A \ x \ a" - and incl: "\S. open S \ x \ S \ \a\A. a \ S" by auto - then have "A \ {}" by auto - with `countable A` have r: "A = range (from_nat_into A)" by auto - def F \ "\n. \i\n. from_nat_into A i" - show "\A. (\i. open (A i)) \ (\i. x \ A i) \ - (\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially)" - proof (safe intro!: exI[of _ F]) - fix i - show "open (F i)" using nhds(1) r by (auto simp: F_def intro!: open_INT) - show "x \ F i" using nhds(2) r by (auto simp: F_def) - next - fix S assume "open S" "x \ S" - from incl[OF this] obtain i where "F i \ S" - by (subst (asm) r) (auto simp: F_def) - moreover have "\j. i \ j \ F j \ F i" - by (auto simp: F_def) - ultimately show "eventually (\i. F i \ S) sequentially" - by (auto simp: eventually_sequentially) - qed -qed - lemma (in first_countable_topology) first_countable_basisE: obtains A where "countable A" "\a. a \ A \ x \ a" "\a. a \ A \ open a" "\S. open S \ x \ S \ (\a\A. a \ S)" using first_countable_basis[of x] - by atomize_elim auto + apply atomize_elim + apply (elim exE) + apply (rule_tac x="range A" in exI) + apply auto + done lemma (in first_countable_topology) first_countable_basis_Int_stableE: obtains A where "countable A" "\a. a \ A \ x \ a" "\a. a \ A \ open a" @@ -245,77 +234,25 @@ qed qed - -lemma countable_basis: - obtains A :: "nat \ 'a::first_countable_topology set" where - "\i. open (A i)" "\i. x \ A i" - "\F. (\n. F n \ A n) \ F ----> x" -proof atomize_elim - from countable_basis_at_decseq[of x] guess A . note A = this - { fix F S assume "\n. F n \ A n" "open S" "x \ S" - with A(3)[of S] have "eventually (\n. F n \ S) sequentially" - by (auto elim: eventually_elim1 simp: subset_eq) } - with A show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\F. (\n. F n \ A n) \ F ----> x)" - by (intro exI[of _ A]) (auto simp: tendsto_def) -qed - -lemma sequentially_imp_eventually_nhds_within: - fixes a :: "'a::first_countable_topology" - assumes "\f. (\n. f n \ s) \ f ----> a \ eventually (\n. P (f n)) sequentially" - shows "eventually P (nhds a within s)" -proof (rule ccontr) - from countable_basis[of a] guess A . note A = this - assume "\ eventually P (nhds a within s)" - with A have P: "\F. \n. F n \ s \ F n \ A n \ \ P (F n)" - unfolding Limits.eventually_within eventually_nhds by (intro choice) fastforce - then guess F .. - hence F0: "\n. F n \ s" and F2: "\n. F n \ A n" and F3: "\n. \ P (F n)" - by fast+ - with A have "F ----> a" by auto - hence "eventually (\n. P (F n)) sequentially" - using assms F0 by simp - thus "False" by (simp add: F3) -qed - -lemma eventually_nhds_within_iff_sequentially: - fixes a :: "'a::first_countable_topology" - shows "eventually P (nhds a within s) \ - (\f. (\n. f n \ s) \ f ----> a \ eventually (\n. P (f n)) sequentially)" -proof (safe intro!: sequentially_imp_eventually_nhds_within) - assume "eventually P (nhds a within s)" - then obtain S where "open S" "a \ S" "\x\S. x \ s \ P x" - by (auto simp: Limits.eventually_within eventually_nhds) - moreover fix f assume "\n. f n \ s" "f ----> a" - ultimately show "eventually (\n. P (f n)) sequentially" - by (auto dest!: topological_tendstoD elim: eventually_elim1) -qed - -lemma eventually_nhds_iff_sequentially: - fixes a :: "'a::first_countable_topology" - shows "eventually P (nhds a) \ (\f. f ----> a \ eventually (\n. P (f n)) sequentially)" - using eventually_nhds_within_iff_sequentially[of P a UNIV] by simp - -lemma not_eventually_sequentiallyD: - assumes P: "\ eventually P sequentially" - shows "\r. subseq r \ (\n. \ P (r n))" -proof - - from P have "\n. \m\n. \ P m" - unfolding eventually_sequentially by (simp add: not_less) - then obtain r where "\n. r n \ n" "\n. \ P (r n)" - by (auto simp: choice_iff) - then show ?thesis - by (auto intro!: exI[of _ "\n. r (((Suc \ r) ^^ Suc n) 0)"] - simp: less_eq_Suc_le subseq_Suc_iff) -qed - +lemma (in topological_space) first_countableI: + assumes "countable A" and 1: "\a. a \ A \ x \ a" "\a. a \ A \ open a" + and 2: "\S. open S \ x \ S \ \a\A. a \ S" + shows "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" +proof (safe intro!: exI[of _ "from_nat_into A"]) + have "A \ {}" using 2[of UNIV] by auto + { fix i show "x \ from_nat_into A i" "open (from_nat_into A i)" + using range_from_nat_into_subset[OF `A \ {}`] 1 by auto } + { fix S assume "open S" "x\S" from 2[OF this] show "\i. from_nat_into A i \ S" + using subset_range_from_nat_into[OF `countable A`] by auto } +qed instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology proof fix x :: "'a \ 'b" from first_countable_basisE[of "fst x"] guess A :: "'a set set" . note A = this from first_countable_basisE[of "snd x"] guess B :: "'b set set" . note B = this - show "\A::('a\'b) set set. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - proof (intro exI[of _ "(\(a, b). a \ b) ` (A \ B)"], safe) + show "\A::nat \ ('a \ 'b) set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + proof (rule first_countableI[of "(\(a, b). a \ b) ` (A \ B)"], safe) fix a b assume x: "a \ A" "b \ B" with A(2, 3)[of a] B(2, 3)[of b] show "x \ a \ b" "open (a \ b)" unfolding mem_Times_iff by (auto intro: open_Times) @@ -329,23 +266,6 @@ qed (simp add: A B) qed -instance metric_space \ first_countable_topology -proof - fix x :: 'a - show "\A. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - proof (intro exI, safe) - fix S assume "open S" "x \ S" - then obtain r where "0 < r" "{y. dist x y < r} \ S" - by (auto simp: open_dist dist_commute subset_eq) - moreover from reals_Archimedean[OF `0 < r`] guess n .. - moreover - then have "{y. dist x y < inverse (Suc n)} \ {y. dist x y < r}" - by (auto simp: inverse_eq_divide) - ultimately show "\a\range (\n. {y. dist x y < inverse (Suc n)}). a \ S" - by auto - qed (auto intro: open_ball) -qed - class second_countable_topology = topological_space + assumes ex_countable_subbasis: "\B::'a::topological_space set set. countable B \ open = generate_topology B" begin @@ -417,9 +337,9 @@ then have B: "countable B" "topological_basis B" using countable_basis is_basis by (auto simp: countable_basis is_basis) - then show "\A. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - by (intro exI[of _ "{b\B. x \ b}"]) - (fastforce simp: topological_space_class.topological_basis_def) + then show "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + by (intro first_countableI[of "{b\B. x \ b}"]) + (fastforce simp: topological_space_class.topological_basis_def)+ qed subsection {* Polish spaces *} @@ -867,10 +787,6 @@ subsection{* Connectedness *} -definition "connected S \ - ~(\e1 e2. open e1 \ open e2 \ S \ (e1 \ e2) \ (e1 \ e2 \ S = {}) - \ ~(e1 \ S = {}) \ ~(e2 \ S = {}))" - lemma connected_local: "connected S \ ~(\e1 e2. openin (subtopology euclidean S) e1 \ @@ -1340,11 +1256,11 @@ text {* Some property holds "sufficiently close" to the limit point. *} -lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *) +lemma eventually_at: (* FIXME: this replaces Metric_Spaces.eventually_at *) "eventually P (at a) \ (\d>0. \x. 0 < dist x a \ dist x a < d \ P x)" unfolding eventually_at dist_nz by auto -lemma eventually_within: (* FIXME: this replaces Limits.eventually_within *) +lemma eventually_within: (* FIXME: this replaces Topological_Spaces.eventually_within *) "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" by (rule eventually_within_less) @@ -1371,18 +1287,6 @@ subsection {* Limits *} -text{* Notation Lim to avoid collition with lim defined in analysis *} - -definition Lim :: "'a filter \ ('a \ 'b::t2_space) \ 'b" - where "Lim A f = (THE l. (f ---> l) A)" - -text{* Uniqueness of the limit, when nontrivial. *} - -lemma tendsto_Lim: - fixes f :: "'a \ 'b::t2_space" - shows "~(trivial_limit net) \ (f ---> l) net ==> Lim net f = l" - unfolding Lim_def using tendsto_unique[of net f] by auto - lemma Lim: "(f ---> l) net \ trivial_limit net \ @@ -1415,10 +1319,6 @@ lemma Lim_within_empty: "(f ---> l) (net within {})" unfolding tendsto_def Limits.eventually_within by simp -lemma Lim_within_subset: "(f ---> l) (net within S) \ T \ S \ (f ---> l) (net within T)" - unfolding tendsto_def Limits.eventually_within - by (auto elim!: eventually_elim1) - lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)" shows "(f ---> l) (net within (S \ T))" using assms unfolding tendsto_def Limits.eventually_within @@ -1448,12 +1348,12 @@ from assms obtain T where T: "open T" "x \ T" "T \ S" .. { assume "?lhs" then obtain A where "open A" "x \ A" "\y\A. y \ x \ y \ S \ P y" - unfolding Limits.eventually_within Limits.eventually_at_topological + unfolding Limits.eventually_within eventually_at_topological by auto with T have "open (A \ T)" "x \ A \ T" "\y\(A \ T). y \ x \ P y" by auto then have "?rhs" - unfolding Limits.eventually_at_topological by auto + unfolding eventually_at_topological by auto } moreover { assume "?rhs" hence "?lhs" unfolding Limits.eventually_within @@ -1466,16 +1366,6 @@ "x \ interior S \ at x within S = at x" by (simp add: filter_eq_iff eventually_within_interior) -lemma at_within_open: - "\x \ S; open S\ \ at x within S = at x" - by (simp only: at_within_interior interior_open) - -lemma Lim_within_open: - fixes f :: "'a::topological_space \ 'b::topological_space" - assumes"a \ S" "open S" - shows "(f ---> l)(at a within S) \ (f ---> l)(at a)" - using assms by (simp only: at_within_open) - lemma Lim_within_LIMSEQ: fixes a :: "'a::metric_space" assumes "\S. (\n. S n \ a \ S n \ T) \ S ----> a \ (\n. X (S n)) ----> L" @@ -1499,7 +1389,7 @@ show "(\n. f (S n)) ----> Inf (f ` ({x<..} \ I))" proof (rule LIMSEQ_I, rule ccontr) fix r :: real assume "0 < r" - with Inf_close[of "f ` ({x<..} \ I)" r] + with cInf_close[of "f ` ({x<..} \ I)" r] obtain y where y: "x < y" "y \ I" "f y < Inf (f ` ({x <..} \ I)) + r" by auto from `x < y` have "0 < y - x" by auto from S(2)[THEN LIMSEQ_D, OF this] @@ -1507,7 +1397,7 @@ assume "\ (\N. \n\N. norm (f (S n) - Inf (f ` ({x<..} \ I))) < r)" moreover have "\n. Inf (f ` ({x<..} \ I)) \ f (S n)" - using S bnd by (intro Inf_lower[where z=K]) auto + using S bnd by (intro cInf_lower[where z=K]) auto ultimately obtain n where n: "N \ n" "r + Inf (f ` ({x<..} \ I)) \ f (S n)" by (auto simp: not_less field_simps) with N[OF n(1)] mono[OF _ `y \ I`, of "S n"] S(1)[THEN spec, of n] y @@ -1825,11 +1715,11 @@ unfolding closure_approachable proof safe have *: "\x\S. Inf S \ x" - using Inf_lower_EX[of _ S] assms by metis + using cInf_lower_EX[of _ S] assms by metis fix e :: real assume "0 < e" then obtain x where x: "x \ S" "x < Inf S + e" - using Inf_close `S \ {}` by auto + using cInf_close `S \ {}` by auto moreover then have "x > Inf S - e" using * by auto ultimately have "abs (x - Inf S) < e" by (simp add: abs_diff_less_iff) then show "\x\S. dist x (Inf S) < e" @@ -1883,13 +1773,13 @@ lemma infdist_nonneg: shows "0 \ infdist x A" - using assms by (auto simp add: infdist_def) + using assms by (auto simp add: infdist_def intro: cInf_greatest) lemma infdist_le: assumes "a \ A" assumes "d = dist x a" shows "infdist x A \ d" - using assms by (auto intro!: SupInf.Inf_lower[where z=0] simp add: infdist_def) + using assms by (auto intro!: cInf_lower[where z=0] simp add: infdist_def) lemma infdist_zero[simp]: assumes "a \ A" shows "infdist a A = 0" @@ -1905,13 +1795,13 @@ next assume "A \ {}" then obtain a where "a \ A" by auto have "infdist x A \ Inf {dist x y + dist y a |a. a \ A}" - proof + proof (rule cInf_greatest) from `A \ {}` show "{dist x y + dist y a |a. a \ A} \ {}" by simp fix d assume "d \ {dist x y + dist y a |a. a \ A}" then obtain a where d: "d = dist x y + dist y a" "a \ A" by auto show "infdist x A \ d" unfolding infdist_notempty[OF `A \ {}`] - proof (rule Inf_lower2) + proof (rule cInf_lower2) show "dist x a \ {dist x a |a. a \ A}" using `a \ A` by auto show "dist x a \ d" unfolding d by (rule dist_triangle) fix d assume "d \ {dist x a |a. a \ A}" @@ -1920,20 +1810,19 @@ qed qed also have "\ = dist x y + infdist y A" - proof (rule Inf_eq, safe) + proof (rule cInf_eq, safe) fix a assume "a \ A" thus "dist x y + infdist y A \ dist x y + dist y a" by (auto intro: infdist_le) next fix i assume inf: "\d. d \ {dist x y + dist y a |a. a \ A} \ i \ d" hence "i - dist x y \ infdist y A" unfolding infdist_notempty[OF `A \ {}`] using `a \ A` - by (intro Inf_greatest) (auto simp: field_simps) + by (intro cInf_greatest) (auto simp: field_simps) thus "i \ dist x y + infdist y A" by simp qed finally show ?thesis by simp qed -lemma - in_closure_iff_infdist_zero: +lemma in_closure_iff_infdist_zero: assumes "A \ {}" shows "x \ closure A \ infdist x A = 0" proof @@ -1957,13 +1846,12 @@ assume "\ (\y\A. dist y x < e)" hence "infdist x A \ e" using `a \ A` unfolding infdist_def - by (force simp: dist_commute) + by (force simp: dist_commute intro: cInf_greatest) with x `0 < e` show False by auto qed qed -lemma - in_closed_iff_infdist_zero: +lemma in_closed_iff_infdist_zero: assumes "closed A" "A \ {}" shows "x \ A \ infdist x A = 0" proof - @@ -2353,7 +2241,7 @@ apply (subgoal_tac "\x (y::real). 0 < 1 + abs y \ (x <= y \ x <= 1 + abs y)") by metis arith -lemma Bseq_eq_bounded: "Bseq f \ bounded (range f)" +lemma Bseq_eq_bounded: "Bseq f \ bounded (range f::_::real_normed_vector set)" unfolding Bseq_def bounded_pos by auto lemma bounded_Int[intro]: "bounded S \ bounded T \ bounded (S \ T)" @@ -2424,16 +2312,19 @@ proof fix x assume "x\S" thus "x \ Sup S" - by (metis SupInf.Sup_upper abs_le_D1 assms(1) bounded_real) + by (metis cSup_upper abs_le_D1 assms(1) bounded_real) next show "\b. (\x\S. x \ b) \ Sup S \ b" using assms - by (metis SupInf.Sup_least) + by (metis cSup_least) qed lemma Sup_insert: fixes S :: "real set" shows "bounded S ==> Sup(insert x S) = (if S = {} then x else max x (Sup S))" -by auto (metis Int_absorb Sup_insert_nonempty assms bounded_has_Sup(1) disjoint_iff_not_equal) + apply (subst cSup_insert_If) + apply (rule bounded_has_Sup(1)[of S, rule_format]) + apply (auto simp: sup_max) + done lemma Sup_insert_finite: fixes S :: "real set" @@ -2450,16 +2341,19 @@ fix x assume "x\S" from assms(1) obtain a where a:"\x\S. \x\ \ a" unfolding bounded_real by auto thus "x \ Inf S" using `x\S` - by (metis Inf_lower_EX abs_le_D2 minus_le_iff) + by (metis cInf_lower_EX abs_le_D2 minus_le_iff) next show "\b. (\x\S. x >= b) \ Inf S \ b" using assms - by (metis SupInf.Inf_greatest) + by (metis cInf_greatest) qed lemma Inf_insert: fixes S :: "real set" shows "bounded S ==> Inf(insert x S) = (if S = {} then x else min x (Inf S))" -by auto (metis Int_absorb Inf_insert_nonempty bounded_has_Inf(1) disjoint_iff_not_equal) + apply (subst cInf_insert_if) + apply (rule bounded_has_Inf(1)[of S, rule_format]) + apply (auto simp: inf_min) + done lemma Inf_insert_finite: fixes S :: "real set" @@ -2468,28 +2362,6 @@ subsection {* Compactness *} -subsubsection{* Open-cover compactness *} - -definition compact :: "'a::topological_space set \ bool" where - compact_eq_heine_borel: -- "This name is used for backwards compatibility" - "compact S \ (\C. (\c\C. open c) \ S \ \C \ (\D\C. finite D \ S \ \D))" - -lemma compactI: - assumes "\C. \t\C. open t \ s \ \ C \ \C'. C' \ C \ finite C' \ s \ \ C'" - shows "compact s" - unfolding compact_eq_heine_borel using assms by metis - -lemma compactE: - assumes "compact s" and "\t\C. open t" and "s \ \C" - obtains C' where "C' \ C" and "finite C'" and "s \ \C'" - using assms unfolding compact_eq_heine_borel by metis - -lemma compactE_image: - assumes "compact s" and "\t\C. open (f t)" and "s \ (\c\C. f c)" - obtains C' where "C' \ C" and "finite C'" and "s \ (\c\C'. f c)" - using assms unfolding ball_simps[symmetric] SUP_def - by (metis (lifting) finite_subset_image compact_eq_heine_borel[of s]) - subsubsection {* Bolzano-Weierstrass property *} lemma heine_borel_imp_bolzano_weierstrass: @@ -2684,35 +2556,6 @@ thus ?thesis unfolding closed_sequential_limits by fast qed -lemma compact_imp_closed: - fixes s :: "'a::t2_space set" - assumes "compact s" shows "closed s" -unfolding closed_def -proof (rule openI) - fix y assume "y \ - s" - let ?C = "\x\s. {u. open u \ x \ u \ eventually (\y. y \ u) (nhds y)}" - note `compact s` - moreover have "\u\?C. open u" by simp - moreover have "s \ \?C" - proof - fix x assume "x \ s" - with `y \ - s` have "x \ y" by clarsimp - hence "\u v. open u \ open v \ x \ u \ y \ v \ u \ v = {}" - by (rule hausdorff) - with `x \ s` show "x \ \?C" - unfolding eventually_nhds by auto - qed - ultimately obtain D where "D \ ?C" and "finite D" and "s \ \D" - by (rule compactE) - from `D \ ?C` have "\x\D. eventually (\y. y \ x) (nhds y)" by auto - with `finite D` have "eventually (\y. y \ \D) (nhds y)" - by (simp add: eventually_Ball_finite) - with `s \ \D` have "eventually (\y. y \ s) (nhds y)" - by (auto elim!: eventually_mono [rotated]) - thus "\t. open t \ y \ t \ t \ - s" - by (simp add: eventually_nhds subset_eq) -qed - lemma compact_imp_bounded: assumes "compact U" shows "bounded U" proof - @@ -2727,11 +2570,6 @@ text{* In particular, some common special cases. *} -lemma compact_empty[simp]: - "compact {}" - unfolding compact_eq_heine_borel - by auto - lemma compact_union [intro]: assumes "compact s" "compact t" shows " compact (s \ t)" proof (rule compactI) @@ -2751,20 +2589,6 @@ "finite A \ (\x. x \ A \ compact (B x)) \ compact (\x\A. B x)" unfolding SUP_def by (rule compact_Union) auto -lemma compact_inter_closed [intro]: - assumes "compact s" and "closed t" - shows "compact (s \ t)" -proof (rule compactI) - fix C assume C: "\c\C. open c" and cover: "s \ t \ \C" - from C `closed t` have "\c\C \ {-t}. open c" by auto - moreover from cover have "s \ \(C \ {-t})" by auto - ultimately have "\D\C \ {-t}. finite D \ s \ \D" - using `compact s` unfolding compact_eq_heine_borel by auto - then guess D .. - then show "\D\C. finite D \ s \ t \ \D" - by (intro exI[of _ "D - {-t}"]) auto -qed - lemma closed_inter_compact [intro]: assumes "closed s" and "compact t" shows "compact (s \ t)" @@ -3223,7 +3047,7 @@ lemma cauchy_def: "Cauchy s \ (\e>0. \N. \m n. m \ N \ n \ N --> dist(s m)(s n) < e)" -unfolding Cauchy_def by blast +unfolding Cauchy_def by metis fun helper_1::"('a::metric_space set) \ real \ nat \ 'a" where "helper_1 s e n = (SOME y::'a. y \ s \ (\m (dist (helper_1 s e m) y < e)))" @@ -3863,35 +3687,6 @@ subsection {* Continuity *} -text {* Define continuity over a net to take in restrictions of the set. *} - -definition - continuous :: "'a::t2_space filter \ ('a \ 'b::topological_space) \ bool" - where "continuous net f \ (f ---> f(netlimit net)) net" - -lemma continuous_trivial_limit: - "trivial_limit net ==> continuous net f" - unfolding continuous_def tendsto_def trivial_limit_eq by auto - -lemma continuous_within: "continuous (at x within s) f \ (f ---> f(x)) (at x within s)" - unfolding continuous_def - unfolding tendsto_def - using netlimit_within[of x s] - by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually) - -lemma continuous_at: "continuous (at x) f \ (f ---> f(x)) (at x)" - using continuous_within [of x UNIV f] by simp - -lemma continuous_isCont: "isCont f x = continuous (at x) f" - unfolding isCont_def LIM_def - unfolding continuous_at Lim_at unfolding dist_nz by auto - -lemma continuous_at_within: - assumes "continuous (at x) f" shows "continuous (at x within s) f" - using assms unfolding continuous_at continuous_within - by (rule Lim_at_within) - - text{* Derive the epsilon-delta forms, which we often use as "definitions" *} lemma continuous_within_eps_delta: @@ -3937,20 +3732,6 @@ text{* Define setwise continuity in terms of limits within the set. *} -definition - continuous_on :: - "'a set \ ('a::topological_space \ 'b::topological_space) \ bool" -where - "continuous_on s f \ (\x\s. (f ---> f x) (at x within s))" - -lemma continuous_on_topological: - "continuous_on s f \ - (\x\s. \B. open B \ f x \ B \ - (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" -unfolding continuous_on_def tendsto_def -unfolding Limits.eventually_within eventually_at_topological -by (intro ball_cong [OF refl] all_cong imp_cong ex_cong conj_cong refl) auto - lemma continuous_on_iff: "continuous_on s f \ (\x\s. \e>0. \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" @@ -3978,54 +3759,15 @@ unfolding continuous_within continuous_at using Lim_at_within by auto lemma Lim_trivial_limit: "trivial_limit net \ (f ---> l) net" -unfolding tendsto_def by (simp add: trivial_limit_eq) - -lemma continuous_at_imp_continuous_on: - assumes "\x\s. continuous (at x) f" - shows "continuous_on s f" -unfolding continuous_on_def -proof - fix x assume "x \ s" - with assms have *: "(f ---> f (netlimit (at x))) (at x)" - unfolding continuous_def by simp - have "(f ---> f x) (at x)" - proof (cases "trivial_limit (at x)") - case True thus ?thesis - by (rule Lim_trivial_limit) - next - case False - hence 1: "netlimit (at x) = x" - using netlimit_within [of x UNIV] by simp - with * show ?thesis by simp - qed - thus "(f ---> f x) (at x within s)" - by (rule Lim_at_within) -qed - -lemma continuous_on_eq_continuous_within: - "continuous_on s f \ (\x \ s. continuous (at x within s) f)" -unfolding continuous_on_def continuous_def -apply (rule ball_cong [OF refl]) -apply (case_tac "trivial_limit (at x within s)") -apply (simp add: Lim_trivial_limit) -apply (simp add: netlimit_within) -done + by simp lemmas continuous_on = continuous_on_def -- "legacy theorem name" -lemma continuous_on_eq_continuous_at: - shows "open s ==> (continuous_on s f \ (\x \ s. continuous (at x) f))" - by (auto simp add: continuous_on continuous_at Lim_within_open) - lemma continuous_within_subset: "continuous (at x within s) f \ t \ s ==> continuous (at x within t) f" unfolding continuous_within by(metis Lim_within_subset) -lemma continuous_on_subset: - shows "continuous_on s f \ t \ s ==> continuous_on t f" - unfolding continuous_on by (metis subset_eq Lim_within_subset) - lemma continuous_on_interior: shows "continuous_on s f \ x \ interior s \ continuous (at x) f" by (erule interiorE, drule (1) continuous_on_subset, @@ -4150,169 +3892,32 @@ subsubsection {* Structural rules for pointwise continuity *} -ML {* - -structure Continuous_Intros = Named_Thms -( - val name = @{binding continuous_intros} - val description = "Structural introduction rules for pointwise continuity" -) - -*} - -setup Continuous_Intros.setup - -lemma continuous_within_id[continuous_intros]: "continuous (at a within s) (\x. x)" - unfolding continuous_within by (rule tendsto_ident_at_within) - -lemma continuous_at_id[continuous_intros]: "continuous (at a) (\x. x)" - unfolding continuous_at by (rule tendsto_ident_at) - -lemma continuous_const[continuous_intros]: "continuous F (\x. c)" - unfolding continuous_def by (rule tendsto_const) - -lemma continuous_fst[continuous_intros]: "continuous F f \ continuous F (\x. fst (f x))" - unfolding continuous_def by (rule tendsto_fst) - -lemma continuous_snd[continuous_intros]: "continuous F f \ continuous F (\x. snd (f x))" - unfolding continuous_def by (rule tendsto_snd) - -lemma continuous_Pair[continuous_intros]: "continuous F f \ continuous F g \ continuous F (\x. (f x, g x))" - unfolding continuous_def by (rule tendsto_Pair) - -lemma continuous_dist[continuous_intros]: - assumes "continuous F f" and "continuous F g" - shows "continuous F (\x. dist (f x) (g x))" - using assms unfolding continuous_def by (rule tendsto_dist) +lemmas continuous_within_id = continuous_ident + +lemmas continuous_at_id = isCont_ident lemma continuous_infdist[continuous_intros]: assumes "continuous F f" shows "continuous F (\x. infdist (f x) A)" using assms unfolding continuous_def by (rule tendsto_infdist) -lemma continuous_norm[continuous_intros]: - shows "continuous F f \ continuous F (\x. norm (f x))" - unfolding continuous_def by (rule tendsto_norm) - lemma continuous_infnorm[continuous_intros]: shows "continuous F f \ continuous F (\x. infnorm (f x))" unfolding continuous_def by (rule tendsto_infnorm) -lemma continuous_add[continuous_intros]: - fixes f g :: "'a::t2_space \ 'b::real_normed_vector" - shows "\continuous F f; continuous F g\ \ continuous F (\x. f x + g x)" - unfolding continuous_def by (rule tendsto_add) - -lemma continuous_minus[continuous_intros]: - fixes f :: "'a::t2_space \ 'b::real_normed_vector" - shows "continuous F f \ continuous F (\x. - f x)" - unfolding continuous_def by (rule tendsto_minus) - -lemma continuous_diff[continuous_intros]: - fixes f g :: "'a::t2_space \ 'b::real_normed_vector" - shows "\continuous F f; continuous F g\ \ continuous F (\x. f x - g x)" - unfolding continuous_def by (rule tendsto_diff) - -lemma continuous_scaleR[continuous_intros]: - fixes g :: "'a::t2_space \ 'b::real_normed_vector" - shows "\continuous F f; continuous F g\ \ continuous F (\x. f x *\<^sub>R g x)" - unfolding continuous_def by (rule tendsto_scaleR) - -lemma continuous_mult[continuous_intros]: - fixes f g :: "'a::t2_space \ 'b::real_normed_algebra" - shows "\continuous F f; continuous F g\ \ continuous F (\x. f x * g x)" - unfolding continuous_def by (rule tendsto_mult) - lemma continuous_inner[continuous_intros]: assumes "continuous F f" and "continuous F g" shows "continuous F (\x. inner (f x) (g x))" using assms unfolding continuous_def by (rule tendsto_inner) -lemma continuous_inverse[continuous_intros]: - fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" - assumes "continuous F f" and "f (netlimit F) \ 0" - shows "continuous F (\x. inverse (f x))" - using assms unfolding continuous_def by (rule tendsto_inverse) - -lemma continuous_at_within_inverse[continuous_intros]: - fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" - assumes "continuous (at a within s) f" and "f a \ 0" - shows "continuous (at a within s) (\x. inverse (f x))" - using assms unfolding continuous_within by (rule tendsto_inverse) - -lemma continuous_at_inverse[continuous_intros]: - fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" - assumes "continuous (at a) f" and "f a \ 0" - shows "continuous (at a) (\x. inverse (f x))" - using assms unfolding continuous_at by (rule tendsto_inverse) +lemmas continuous_at_inverse = isCont_inverse subsubsection {* Structural rules for setwise continuity *} -ML {* - -structure Continuous_On_Intros = Named_Thms -( - val name = @{binding continuous_on_intros} - val description = "Structural introduction rules for setwise continuity" -) - -*} - -setup Continuous_On_Intros.setup - -lemma continuous_on_id[continuous_on_intros]: "continuous_on s (\x. x)" - unfolding continuous_on_def by (fast intro: tendsto_ident_at_within) - -lemma continuous_on_const[continuous_on_intros]: "continuous_on s (\x. c)" - unfolding continuous_on_def by (auto intro: tendsto_intros) - -lemma continuous_on_norm[continuous_on_intros]: - shows "continuous_on s f \ continuous_on s (\x. norm (f x))" - unfolding continuous_on_def by (fast intro: tendsto_norm) - lemma continuous_on_infnorm[continuous_on_intros]: shows "continuous_on s f \ continuous_on s (\x. infnorm (f x))" unfolding continuous_on by (fast intro: tendsto_infnorm) -lemma continuous_on_minus[continuous_on_intros]: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "continuous_on s f \ continuous_on s (\x. - f x)" - unfolding continuous_on_def by (auto intro: tendsto_intros) - -lemma continuous_on_add[continuous_on_intros]: - fixes f g :: "'a::topological_space \ 'b::real_normed_vector" - shows "continuous_on s f \ continuous_on s g - \ continuous_on s (\x. f x + g x)" - unfolding continuous_on_def by (auto intro: tendsto_intros) - -lemma continuous_on_diff[continuous_on_intros]: - fixes f g :: "'a::topological_space \ 'b::real_normed_vector" - shows "continuous_on s f \ continuous_on s g - \ continuous_on s (\x. f x - g x)" - unfolding continuous_on_def by (auto intro: tendsto_intros) - -lemma (in bounded_linear) continuous_on[continuous_on_intros]: - "continuous_on s g \ continuous_on s (\x. f (g x))" - unfolding continuous_on_def by (fast intro: tendsto) - -lemma (in bounded_bilinear) continuous_on[continuous_on_intros]: - "\continuous_on s f; continuous_on s g\ \ continuous_on s (\x. f x ** g x)" - unfolding continuous_on_def by (fast intro: tendsto) - -lemma continuous_on_scaleR[continuous_on_intros]: - fixes g :: "'a::topological_space \ 'b::real_normed_vector" - assumes "continuous_on s f" and "continuous_on s g" - shows "continuous_on s (\x. f x *\<^sub>R g x)" - using bounded_bilinear_scaleR assms - by (rule bounded_bilinear.continuous_on) - -lemma continuous_on_mult[continuous_on_intros]: - fixes g :: "'a::topological_space \ 'b::real_normed_algebra" - assumes "continuous_on s f" and "continuous_on s g" - shows "continuous_on s (\x. f x * g x)" - using bounded_bilinear_mult assms - by (rule bounded_bilinear.continuous_on) - lemma continuous_on_inner[continuous_on_intros]: fixes g :: "'a::topological_space \ 'b::real_inner" assumes "continuous_on s f" and "continuous_on s g" @@ -4320,12 +3925,6 @@ using bounded_bilinear_inner assms by (rule bounded_bilinear.continuous_on) -lemma continuous_on_inverse[continuous_on_intros]: - fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" - assumes "continuous_on s f" and "\x\s. f x \ 0" - shows "continuous_on s (\x. inverse (f x))" - using assms unfolding continuous_on by (fast intro: tendsto_inverse) - subsubsection {* Structural rules for uniform continuity *} lemma uniformly_continuous_on_id[continuous_on_intros]: @@ -4406,33 +4005,7 @@ text{* Continuity of all kinds is preserved under composition. *} -lemma continuous_within_topological: - "continuous (at x within s) f \ - (\B. open B \ f x \ B \ - (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" -unfolding continuous_within -unfolding tendsto_def Limits.eventually_within eventually_at_topological -by (intro ball_cong [OF refl] all_cong imp_cong ex_cong conj_cong refl) auto - -lemma continuous_within_compose[continuous_intros]: - assumes "continuous (at x within s) f" - assumes "continuous (at (f x) within f ` s) g" - shows "continuous (at x within s) (g o f)" -using assms unfolding continuous_within_topological by simp metis - -lemma continuous_at_compose[continuous_intros]: - assumes "continuous (at x) f" and "continuous (at (f x)) g" - shows "continuous (at x) (g o f)" -proof- - have "continuous (at (f x) within range f) g" using assms(2) - using continuous_within_subset[of "f x" UNIV g "range f"] by simp - thus ?thesis using assms(1) - using continuous_within_compose[of x UNIV f g] by simp -qed - -lemma continuous_on_compose[continuous_on_intros]: - "continuous_on s f \ continuous_on (f ` s) g \ continuous_on s (g o f)" - unfolding continuous_on_topological by simp metis +lemmas continuous_at_compose = isCont_o lemma uniformly_continuous_on_compose[continuous_on_intros]: assumes "uniformly_continuous_on s f" "uniformly_continuous_on (f ` s) g" @@ -4467,70 +4040,18 @@ qed lemma continuous_on_open: - shows "continuous_on s f \ + "continuous_on s f \ (\t. openin (subtopology euclidean (f ` s)) t --> openin (subtopology euclidean s) {x \ s. f x \ t})" (is "?lhs = ?rhs") -proof (safe) - fix t :: "'b set" - assume 1: "continuous_on s f" - assume 2: "openin (subtopology euclidean (f ` s)) t" - from 2 obtain B where B: "open B" and t: "t = f ` s \ B" - unfolding openin_open by auto - def U == "\{A. open A \ (\x\s. x \ A \ f x \ B)}" - have "open U" unfolding U_def by (simp add: open_Union) - moreover have "\x\s. x \ U \ f x \ t" - proof (intro ballI iffI) - fix x assume "x \ s" and "x \ U" thus "f x \ t" - unfolding U_def t by auto - next - fix x assume "x \ s" and "f x \ t" - hence "x \ s" and "f x \ B" - unfolding t by auto - with 1 B obtain A where "open A" "x \ A" "\y\s. y \ A \ f y \ B" - unfolding t continuous_on_topological by metis - then show "x \ U" - unfolding U_def by auto - qed - ultimately have "open U \ {x \ s. f x \ t} = s \ U" by auto - then show "openin (subtopology euclidean s) {x \ s. f x \ t}" - unfolding openin_open by fast -next - assume "?rhs" show "continuous_on s f" - unfolding continuous_on_topological - proof (clarify) - fix x and B assume "x \ s" and "open B" and "f x \ B" - have "openin (subtopology euclidean (f ` s)) (f ` s \ B)" - unfolding openin_open using `open B` by auto - then have "openin (subtopology euclidean s) {x \ s. f x \ f ` s \ B}" - using `?rhs` by fast - then show "\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)" - unfolding openin_open using `x \ s` and `f x \ B` by auto - qed -qed + unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute + by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) text {* Similarly in terms of closed sets. *} lemma continuous_on_closed: shows "continuous_on s f \ (\t. closedin (subtopology euclidean (f ` s)) t --> closedin (subtopology euclidean s) {x \ s. f x \ t})" (is "?lhs = ?rhs") -proof - assume ?lhs - { fix t - have *:"s - {x \ s. f x \ f ` s - t} = {x \ s. f x \ t}" by auto - have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto - assume as:"closedin (subtopology euclidean (f ` s)) t" - hence "closedin (subtopology euclidean (f ` s)) (f ` s - (f ` s - t))" unfolding closedin_def topspace_euclidean_subtopology unfolding ** by auto - hence "closedin (subtopology euclidean s) {x \ s. f x \ t}" using `?lhs`[unfolded continuous_on_open, THEN spec[where x="(f ` s) - t"]] - unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto } - thus ?rhs by auto -next - assume ?rhs - { fix t - have *:"s - {x \ s. f x \ f ` s - t} = {x \ s. f x \ t}" by auto - assume as:"openin (subtopology euclidean (f ` s)) t" - hence "openin (subtopology euclidean s) {x \ s. f x \ t}" using `?rhs`[THEN spec[where x="(f ` s) - t"]] - unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto } - thus ?lhs unfolding continuous_on_open by auto -qed + unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute + by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) text {* Half-global and completely global cases. *} @@ -4667,7 +4188,7 @@ hence "eventually (\y. f y \ a) (at x within s)" using `a \ U` by (fast elim: eventually_mono [rotated]) thus ?thesis - unfolding Limits.eventually_within Limits.eventually_at + unfolding Limits.eventually_within Metric_Spaces.eventually_at by (rule ex_forward, cut_tac `f x \ a`, auto simp: dist_commute) qed @@ -4872,40 +4393,6 @@ qed qed -lemma compact_continuous_image: - assumes "continuous_on s f" and "compact s" - shows "compact (f ` s)" -using assms (* FIXME: long unstructured proof *) -unfolding continuous_on_open -unfolding compact_eq_openin_cover -apply clarify -apply (drule_tac x="image (\t. {x \ s. f x \ t}) C" in spec) -apply (drule mp) -apply (rule conjI) -apply simp -apply clarsimp -apply (drule subsetD) -apply (erule imageI) -apply fast -apply (erule thin_rl) -apply clarify -apply (rule_tac x="image (inv_into C (\t. {x \ s. f x \ t})) D" in exI) -apply (intro conjI) -apply clarify -apply (rule inv_into_into) -apply (erule (1) subsetD) -apply (erule finite_imageI) -apply (clarsimp, rename_tac x) -apply (drule (1) subsetD, clarify) -apply (drule (1) subsetD, clarify) -apply (rule rev_bexI) -apply assumption -apply (subgoal_tac "{x \ s. f x \ t} \ (\t. {x \ s. f x \ t}) ` C") -apply (drule f_inv_into_f) -apply fast -apply (erule imageI) -done - lemma connected_continuous_image: assumes "continuous_on s f" "connected s" shows "connected(f ` s)" @@ -4958,38 +4445,6 @@ qed (insert D, auto) qed auto -text{* Continuity of inverse function on compact domain. *} - -lemma continuous_on_inv: - fixes f :: "'a::topological_space \ 'b::t2_space" - assumes "continuous_on s f" "compact s" "\x \ s. g (f x) = x" - shows "continuous_on (f ` s) g" -unfolding continuous_on_topological -proof (clarsimp simp add: assms(3)) - fix x :: 'a and B :: "'a set" - assume "x \ s" and "open B" and "x \ B" - have 1: "\x\s. f x \ f ` (s - B) \ x \ s - B" - using assms(3) by (auto, metis) - have "continuous_on (s - B) f" - using `continuous_on s f` Diff_subset - by (rule continuous_on_subset) - moreover have "compact (s - B)" - using `open B` and `compact s` - unfolding Diff_eq by (intro compact_inter_closed closed_Compl) - ultimately have "compact (f ` (s - B))" - by (rule compact_continuous_image) - hence "closed (f ` (s - B))" - by (rule compact_imp_closed) - hence "open (- f ` (s - B))" - by (rule open_Compl) - moreover have "f x \ - f ` (s - B)" - using `x \ s` and `x \ B` by (simp add: 1) - moreover have "\y\s. f y \ - f ` (s - B) \ y \ B" - by (simp add: 1) - ultimately show "\A. open A \ f x \ A \ (\y\s. f y \ A \ y \ B)" - by fast -qed - text {* A uniformly convergent limit of continuous functions is continuous. *} lemma continuous_uniform_limit: @@ -5059,56 +4514,6 @@ shows "continuous_on s f \ (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d --> abs(f x' - f x) < e))" unfolding continuous_on_iff dist_norm by simp -lemma compact_attains_sup: - fixes S :: "'a::linorder_topology set" - assumes "compact S" "S \ {}" - shows "\s\S. \t\S. t \ s" -proof (rule classical) - assume "\ (\s\S. \t\S. t \ s)" - then obtain t where t: "\s\S. t s \ S" and "\s\S. s < t s" - by (metis not_le) - then have "\s\S. open {..< t s}" "S \ (\s\S. {..< t s})" - by auto - with `compact S` obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {..< t s})" - by (erule compactE_image) - with `S \ {}` have Max: "Max (t`C) \ t`C" and "\s\t`C. s \ Max (t`C)" - by (auto intro!: Max_in) - with C have "S \ {..< Max (t`C)}" - by (auto intro: less_le_trans simp: subset_eq) - with t Max `C \ S` show ?thesis - by fastforce -qed - -lemma compact_attains_inf: - fixes S :: "'a::linorder_topology set" - assumes "compact S" "S \ {}" - shows "\s\S. \t\S. s \ t" -proof (rule classical) - assume "\ (\s\S. \t\S. s \ t)" - then obtain t where t: "\s\S. t s \ S" and "\s\S. t s < s" - by (metis not_le) - then have "\s\S. open {t s <..}" "S \ (\s\S. {t s <..})" - by auto - with `compact S` obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {t s <..})" - by (erule compactE_image) - with `S \ {}` have Min: "Min (t`C) \ t`C" and "\s\t`C. Min (t`C) \ s" - by (auto intro!: Min_in) - with C have "S \ {Min (t`C) <..}" - by (auto intro: le_less_trans simp: subset_eq) - with t Min `C \ S` show ?thesis - by fastforce -qed - -lemma continuous_attains_sup: - fixes f :: "'a::topological_space \ 'b::linorder_topology" - shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f y \ f x)" - using compact_attains_sup[of "f ` s"] compact_continuous_image[of s f] by auto - -lemma continuous_attains_inf: - fixes f :: "'a::topological_space \ 'b::linorder_topology" - shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f x \ f y)" - using compact_attains_inf[of "f ` s"] compact_continuous_image[of s f] by auto - text {* Hence some handy theorems on distance, diameter etc. of/from a set. *} lemma distance_attains_sup: @@ -5276,8 +4681,7 @@ have "compact (s \ s)" using `compact s` by (intro compact_Times) moreover have "s \ s \ {}" using `s \ {}` by auto moreover have "continuous_on (s \ s) (\x. dist (fst x) (snd x))" - by (intro continuous_at_imp_continuous_on ballI continuous_dist - continuous_isCont[THEN iffD1] isCont_fst isCont_snd isCont_ident) + by (intro continuous_at_imp_continuous_on ballI continuous_intros) ultimately show ?thesis using continuous_attains_sup[of "s \ s" "\x. dist (fst x) (snd x)"] by auto qed @@ -5295,7 +4699,7 @@ from s obtain z d where z: "\x. x \ s \ dist z x \ d" unfolding bounded_def by auto have "dist x y \ Sup ?D" - proof (rule Sup_upper, safe) + proof (rule cSup_upper, safe) fix a b assume "a \ s" "b \ s" with z[of a] z[of b] dist_triangle[of a b z] show "dist a b \ 2 * d" @@ -5317,7 +4721,7 @@ by (auto simp: diameter_def) then have "?D \ {}" by auto ultimately have "Sup ?D \ d" - by (intro Sup_least) (auto simp: not_less) + by (intro cSup_least) (auto simp: not_less) with `d < diameter s` `s \ {}` show False by (auto simp: diameter_def) qed @@ -5337,7 +4741,7 @@ then obtain x y where xys:"x\s" "y\s" and xy:"\u\s. \v\s. dist u v \ dist x y" using compact_sup_maxdistance[OF assms] by auto hence "diameter s \ dist x y" - unfolding diameter_def by clarsimp (rule Sup_least, fast+) + unfolding diameter_def by clarsimp (rule cSup_least, fast+) thus ?thesis by (metis b diameter_bounded_bound order_antisym xys) qed @@ -5967,7 +5371,7 @@ by (rule isCont_open_vimage) lemma open_Collect_less: - fixes f g :: "'a::topological_space \ real" + fixes f g :: "'a::t2_space \ real" assumes f: "\x. isCont f x" assumes g: "\x. isCont g x" shows "open {x. f x < g x}" @@ -5981,7 +5385,7 @@ qed lemma closed_Collect_le: - fixes f g :: "'a::topological_space \ real" + fixes f g :: "'a::t2_space \ real" assumes f: "\x. isCont f x" assumes g: "\x. isCont g x" shows "closed {x. f x \ g x}" @@ -5995,7 +5399,7 @@ qed lemma closed_Collect_eq: - fixes f g :: "'a::topological_space \ 'b::t2_space" + fixes f g :: "'a::t2_space \ 'b::t2_space" assumes f: "\x. isCont f x" assumes g: "\x. isCont g x" shows "closed {x. f x = g x}" @@ -6129,27 +5533,6 @@ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" unfolding tendsto_def trivial_limit_eq by auto -lemma continuous_on_union: - assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f" - shows "continuous_on (s \ t) f" - using assms unfolding continuous_on Lim_within_union - unfolding Lim_topological trivial_limit_within closed_limpt by auto - -lemma continuous_on_cases: - assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g" - "\x. (x\s \ \ P x) \ (x \ t \ P x) \ f x = g x" - shows "continuous_on (s \ t) (\x. if P x then f x else g x)" -proof- - let ?h = "(\x. if P x then f x else g x)" - have "\x\s. f x = (if P x then f x else g x)" using assms(5) by auto - hence "continuous_on s ?h" using continuous_on_eq[of s f ?h] using assms(3) by auto - moreover - have "\x\t. g x = (if P x then f x else g x)" using assms(5) by auto - hence "continuous_on t ?h" using continuous_on_eq[of t g ?h] using assms(4) by auto - ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto -qed - - text{* Some more convenient intermediate-value theorem formulations. *} lemma connected_ivt_hyperplane: @@ -6773,5 +6156,4 @@ declare tendsto_const [intro] (* FIXME: move *) - end diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/NSA/HSEQ.thy --- a/src/HOL/NSA/HSEQ.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/NSA/HSEQ.thy Sat Mar 23 07:30:53 2013 +0100 @@ -343,7 +343,7 @@ text{*Standard Version: easily now proved using equivalence of NS and standard definitions *} -lemma convergent_Bseq: "convergent X ==> Bseq X" +lemma convergent_Bseq: "convergent X ==> Bseq (X::nat \ _::real_normed_vector)" by (simp add: NSconvergent_NSBseq convergent_NSconvergent_iff Bseq_NSBseq_iff) subsubsection{*Upper Bounds and Lubs of Bounded Sequences*} diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/NthRoot.thy --- a/src/HOL/NthRoot.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/NthRoot.thy Sat Mar 23 07:30:53 2013 +0100 @@ -10,6 +10,17 @@ imports Parity Deriv begin +lemma abs_sgn_eq: "abs (sgn x :: real) = (if x = 0 then 0 else 1)" + by (simp add: sgn_real_def) + +lemma inverse_sgn: "sgn (inverse a) = inverse (sgn a :: real)" + by (simp add: sgn_real_def) + +lemma power_eq_iff_eq_base: + fixes a b :: "_ :: linordered_semidom" + shows "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" + using power_eq_imp_eq_base[of a n b] by auto + subsection {* Existence of Nth Root *} text {* Existence follows from the Intermediate Value Theorem *} @@ -43,11 +54,8 @@ text {* Uniqueness of nth positive root *} -lemma realpow_pos_nth_unique: - "\0 < n; 0 < a\ \ \!r. 0 < r \ r ^ n = (a::real)" -apply (auto intro!: realpow_pos_nth) -apply (rule_tac n=n in power_eq_imp_eq_base, simp_all) -done +lemma realpow_pos_nth_unique: "\0 < n; 0 < a\ \ \!r. 0 < r \ r ^ n = (a::real)" + by (auto intro!: realpow_pos_nth simp: power_eq_iff_eq_base) subsection {* Nth Root *} @@ -55,66 +63,86 @@ @{term "root n (- x) = - root n x"}. This allows us to omit side conditions from many theorems. *} -definition - root :: "[nat, real] \ real" where - "root n x = (if 0 < x then (THE u. 0 < u \ u ^ n = x) else - if x < 0 then - (THE u. 0 < u \ u ^ n = - x) else 0)" +lemma inj_sgn_power: assumes "0 < n" shows "inj (\y. sgn y * \y\^n :: real)" (is "inj ?f") +proof (rule injI) + have x: "\a b :: real. (0 < a \ b < 0) \ (a < 0 \ 0 < b) \ a \ b" by auto + fix x y assume "?f x = ?f y" with power_eq_iff_eq_base[of n "\x\" "\y\"] `0a\ ^ n = x \ x = sgn b * \b\ ^ n \ 0 < n \ a = (b::real)" + using inj_sgn_power[THEN injD, of n a b] by simp + +definition root :: "nat \ real \ real" where + "root n x = (if n = 0 then 0 else the_inv (\y. sgn y * \y\^n) x)" + +lemma root_0 [simp]: "root 0 x = 0" + by (simp add: root_def) + +lemma root_sgn_power: "0 < n \ root n (sgn y * \y\^n) = y" + using the_inv_f_f[OF inj_sgn_power] by (simp add: root_def) + +lemma sgn_power_root: + assumes "0 < n" shows "sgn (root n x) * \(root n x)\^n = x" (is "?f (root n x) = x") +proof cases + assume "x \ 0" + with realpow_pos_nth[OF `0 < n`, of "\x\"] obtain r where "0 < r" "r ^ n = \x\" by auto + with `x \ 0` have S: "x \ range ?f" + by (intro image_eqI[of _ _ "sgn x * r"]) + (auto simp: abs_mult sgn_mult power_mult_distrib abs_sgn_eq mult_sgn_abs) + from `0 < n` f_the_inv_into_f[OF inj_sgn_power[OF `0 < n`] this] show ?thesis + by (simp add: root_def) +qed (insert `0 < n` root_sgn_power[of n 0], simp) + +lemma split_root: "P (root n x) \ (n = 0 \ P 0) \ (0 < n \ (\y. sgn y * \y\^n = x \ P y))" + apply (cases "n = 0") + apply simp_all + apply (metis root_sgn_power sgn_power_root) + done lemma real_root_zero [simp]: "root n 0 = 0" -unfolding root_def by simp + by (simp split: split_root add: sgn_zero_iff) + +lemma real_root_minus: "root n (- x) = - root n x" + by (clarsimp split: split_root elim!: sgn_power_injE simp: sgn_minus) -lemma real_root_minus: "0 < n \ root n (- x) = - root n x" -unfolding root_def by simp +lemma real_root_less_mono: "\0 < n; x < y\ \ root n x < root n y" +proof (clarsimp split: split_root) + have x: "\a b :: real. (0 < b \ a < 0) \ \ a > b" by auto + fix a b :: real assume "0 < n" "sgn a * \a\ ^ n < sgn b * \b\ ^ n" then show "a < b" + using power_less_imp_less_base[of a n b] power_less_imp_less_base[of "-b" n "-a"] + by (simp add: sgn_real_def power_less_zero_eq x[of "a ^ n" "- ((- b) ^ n)"] split: split_if_asm) +qed lemma real_root_gt_zero: "\0 < n; 0 < x\ \ 0 < root n x" -apply (simp add: root_def) -apply (drule (1) realpow_pos_nth_unique) -apply (erule theI' [THEN conjunct1]) -done + using real_root_less_mono[of n 0 x] by simp + +lemma real_root_ge_zero: "0 \ x \ 0 \ root n x" + using real_root_gt_zero[of n x] by (cases "n = 0") (auto simp add: le_less) lemma real_root_pow_pos: (* TODO: rename *) "\0 < n; 0 < x\ \ root n x ^ n = x" -apply (simp add: root_def) -apply (drule (1) realpow_pos_nth_unique) -apply (erule theI' [THEN conjunct2]) -done + using sgn_power_root[of n x] real_root_gt_zero[of n x] by simp lemma real_root_pow_pos2 [simp]: (* TODO: rename *) "\0 < n; 0 \ x\ \ root n x ^ n = x" by (auto simp add: order_le_less real_root_pow_pos) +lemma sgn_root: "0 < n \ sgn (root n x) = sgn x" + by (auto split: split_root simp: sgn_real_def power_less_zero_eq) + lemma odd_real_root_pow: "odd n \ root n x ^ n = x" -apply (rule_tac x=0 and y=x in linorder_le_cases) -apply (erule (1) real_root_pow_pos2 [OF odd_pos]) -apply (subgoal_tac "root n (- x) ^ n = - x") -apply (simp add: real_root_minus odd_pos) -apply (simp add: odd_pos) -done - -lemma real_root_ge_zero: "\0 < n; 0 \ x\ \ 0 \ root n x" -by (auto simp add: order_le_less real_root_gt_zero) + using sgn_power_root[of n x] by (simp add: odd_pos sgn_real_def split: split_if_asm) lemma real_root_power_cancel: "\0 < n; 0 \ x\ \ root n (x ^ n) = x" -apply (subgoal_tac "0 \ x ^ n") -apply (subgoal_tac "0 \ root n (x ^ n)") -apply (subgoal_tac "root n (x ^ n) ^ n = x ^ n") -apply (erule (3) power_eq_imp_eq_base) -apply (erule (1) real_root_pow_pos2) -apply (erule (1) real_root_ge_zero) -apply (erule zero_le_power) -done + using root_sgn_power[of n x] by (auto simp add: le_less power_0_left) lemma odd_real_root_power_cancel: "odd n \ root n (x ^ n) = x" -apply (rule_tac x=0 and y=x in linorder_le_cases) -apply (erule (1) real_root_power_cancel [OF odd_pos]) -apply (subgoal_tac "root n ((- x) ^ n) = - x") -apply (simp add: real_root_minus odd_pos) -apply (erule real_root_power_cancel [OF odd_pos], simp) -done + using root_sgn_power[of n x] by (simp add: odd_pos sgn_real_def power_0_left split: split_if_asm) -lemma real_root_pos_unique: - "\0 < n; 0 \ y; y ^ n = x\ \ root n x = y" -by (erule subst, rule real_root_power_cancel) +lemma real_root_pos_unique: "\0 < n; 0 \ y; y ^ n = x\ \ root n x = y" + using root_sgn_power[of n y] by (auto simp add: le_less power_0_left) lemma odd_real_root_unique: "\odd n; y ^ n = x\ \ root n x = y" @@ -125,32 +153,8 @@ text {* Root function is strictly monotonic, hence injective *} -lemma real_root_less_mono_lemma: - "\0 < n; 0 \ x; x < y\ \ root n x < root n y" -apply (subgoal_tac "0 \ y") -apply (subgoal_tac "root n x ^ n < root n y ^ n") -apply (erule power_less_imp_less_base) -apply (erule (1) real_root_ge_zero) -apply simp -apply simp -done - -lemma real_root_less_mono: "\0 < n; x < y\ \ root n x < root n y" -apply (cases "0 \ x") -apply (erule (2) real_root_less_mono_lemma) -apply (cases "0 \ y") -apply (rule_tac y=0 in order_less_le_trans) -apply (subgoal_tac "0 < root n (- x)") -apply (simp add: real_root_minus) -apply (simp add: real_root_gt_zero) -apply (simp add: real_root_ge_zero) -apply (subgoal_tac "root n (- y) < root n (- x)") -apply (simp add: real_root_minus) -apply (simp add: real_root_less_mono_lemma) -done - lemma real_root_le_mono: "\0 < n; x \ y\ \ root n x \ root n y" -by (auto simp add: order_le_less real_root_less_mono) + by (auto simp add: order_le_less real_root_less_mono) lemma real_root_less_iff [simp]: "0 < n \ (root n x < root n y) = (x < y)" @@ -191,26 +195,34 @@ lemma real_root_eq_1_iff [simp]: "0 < n \ (root n x = 1) = (x = 1)" by (insert real_root_eq_iff [where y=1], simp) +text {* Roots of multiplication and division *} + +lemma real_root_mult: "root n (x * y) = root n x * root n y" + by (auto split: split_root elim!: sgn_power_injE simp: sgn_mult abs_mult power_mult_distrib) + +lemma real_root_inverse: "root n (inverse x) = inverse (root n x)" + by (auto split: split_root elim!: sgn_power_injE simp: inverse_sgn power_inverse) + +lemma real_root_divide: "root n (x / y) = root n x / root n y" + by (simp add: divide_inverse real_root_mult real_root_inverse) + +lemma real_root_abs: "0 < n \ root n \x\ = \root n x\" + by (simp add: abs_if real_root_minus) + +lemma real_root_power: "0 < n \ root n (x ^ k) = root n x ^ k" + by (induct k) (simp_all add: real_root_mult) + text {* Roots of roots *} lemma real_root_Suc_0 [simp]: "root (Suc 0) x = x" by (simp add: odd_real_root_unique) -lemma real_root_pos_mult_exp: - "\0 < m; 0 < n; 0 < x\ \ root (m * n) x = root m (root n x)" -by (rule real_root_pos_unique, simp_all add: power_mult) +lemma real_root_mult_exp: "root (m * n) x = root m (root n x)" + by (auto split: split_root elim!: sgn_power_injE + simp: sgn_zero_iff sgn_mult power_mult[symmetric] abs_mult power_mult_distrib abs_sgn_eq) -lemma real_root_mult_exp: - "\0 < m; 0 < n\ \ root (m * n) x = root m (root n x)" -apply (rule linorder_cases [where x=x and y=0]) -apply (subgoal_tac "root (m * n) (- x) = root m (root n (- x))") -apply (simp add: real_root_minus) -apply (simp_all add: real_root_pos_mult_exp) -done - -lemma real_root_commute: - "\0 < m; 0 < n\ \ root m (root n x) = root n (root m x)" -by (simp add: real_root_mult_exp [symmetric] mult_commute) +lemma real_root_commute: "root m (root n x) = root n (root m x)" + by (simp add: real_root_mult_exp [symmetric] mult_commute) text {* Monotonicity in first argument *} @@ -236,107 +248,36 @@ "\0 < n; n < N; 0 \ x; x \ 1\ \ root n x \ root N x" by (auto simp add: order_le_less real_root_strict_increasing) -text {* Roots of multiplication and division *} - -lemma real_root_mult_lemma: - "\0 < n; 0 \ x; 0 \ y\ \ root n (x * y) = root n x * root n y" -by (simp add: real_root_pos_unique mult_nonneg_nonneg power_mult_distrib) - -lemma real_root_inverse_lemma: - "\0 < n; 0 \ x\ \ root n (inverse x) = inverse (root n x)" -by (simp add: real_root_pos_unique power_inverse [symmetric]) - -lemma real_root_mult: - assumes n: "0 < n" - shows "root n (x * y) = root n x * root n y" -proof (rule linorder_le_cases, rule_tac [!] linorder_le_cases) - assume "0 \ x" and "0 \ y" - thus ?thesis by (rule real_root_mult_lemma [OF n]) -next - assume "0 \ x" and "y \ 0" - hence "0 \ x" and "0 \ - y" by simp_all - hence "root n (x * - y) = root n x * root n (- y)" - by (rule real_root_mult_lemma [OF n]) - thus ?thesis by (simp add: real_root_minus [OF n]) -next - assume "x \ 0" and "0 \ y" - hence "0 \ - x" and "0 \ y" by simp_all - hence "root n (- x * y) = root n (- x) * root n y" - by (rule real_root_mult_lemma [OF n]) - thus ?thesis by (simp add: real_root_minus [OF n]) -next - assume "x \ 0" and "y \ 0" - hence "0 \ - x" and "0 \ - y" by simp_all - hence "root n (- x * - y) = root n (- x) * root n (- y)" - by (rule real_root_mult_lemma [OF n]) - thus ?thesis by (simp add: real_root_minus [OF n]) -qed - -lemma real_root_inverse: - assumes n: "0 < n" - shows "root n (inverse x) = inverse (root n x)" -proof (rule linorder_le_cases) - assume "0 \ x" - thus ?thesis by (rule real_root_inverse_lemma [OF n]) -next - assume "x \ 0" - hence "0 \ - x" by simp - hence "root n (inverse (- x)) = inverse (root n (- x))" - by (rule real_root_inverse_lemma [OF n]) - thus ?thesis by (simp add: real_root_minus [OF n]) -qed - -lemma real_root_divide: - "0 < n \ root n (x / y) = root n x / root n y" -by (simp add: divide_inverse real_root_mult real_root_inverse) - -lemma real_root_power: - "0 < n \ root n (x ^ k) = root n x ^ k" -by (induct k, simp_all add: real_root_mult) - -lemma real_root_abs: "0 < n \ root n \x\ = \root n x\" -by (simp add: abs_if real_root_minus) - text {* Continuity and derivatives *} -lemma isCont_root_pos: - assumes n: "0 < n" - assumes x: "0 < x" - shows "isCont (root n) x" -proof - - have "isCont (root n) (root n x ^ n)" - proof (rule isCont_inverse_function [where f="\a. a ^ n"]) - show "0 < root n x" using n x by simp - show "\z. \z - root n x\ \ root n x \ root n (z ^ n) = z" - by (simp add: abs_le_iff real_root_power_cancel n) - show "\z. \z - root n x\ \ root n x \ isCont (\a. a ^ n) z" - by simp - qed - thus ?thesis using n x by simp -qed +lemma isCont_real_root: "isCont (root n) x" +proof cases + assume n: "0 < n" + let ?f = "\y::real. sgn y * \y\^n" + have "continuous_on ({0..} \ {.. 0}) (\x. if 0 < x then x ^ n else - ((-x) ^ n) :: real)" + using n by (intro continuous_on_If continuous_on_intros) auto + then have "continuous_on UNIV ?f" + by (rule continuous_on_cong[THEN iffD1, rotated 2]) (auto simp: not_less real_sgn_neg le_less n) + then have [simp]: "\x. isCont ?f x" + by (simp add: continuous_on_eq_continuous_at) -lemma isCont_root_neg: - "\0 < n; x < 0\ \ isCont (root n) x" -apply (subgoal_tac "isCont (\x. - root n (- x)) x") -apply (simp add: real_root_minus) -apply (rule isCont_o2 [OF isCont_minus [OF isCont_ident]]) -apply (simp add: isCont_root_pos) -done + have "isCont (root n) (?f (root n x))" + by (rule isCont_inverse_function [where f="?f" and d=1]) (auto simp: root_sgn_power n) + then show ?thesis + by (simp add: sgn_power_root n) +qed (simp add: root_def[abs_def]) -lemma isCont_root_zero: - "0 < n \ isCont (root n) 0" -unfolding isCont_def -apply (rule LIM_I) -apply (rule_tac x="r ^ n" in exI, safe) -apply (simp) -apply (simp add: real_root_abs [symmetric]) -apply (rule_tac n="n" in power_less_imp_less_base, simp_all) -done +lemma tendsto_real_root[tendsto_intros]: + "(f ---> x) F \ ((\x. root n (f x)) ---> root n x) F" + using isCont_tendsto_compose[OF isCont_real_root, of f x F] . -lemma isCont_real_root: "0 < n \ isCont (root n) x" -apply (rule_tac x=x and y=0 in linorder_cases) -apply (simp_all add: isCont_root_pos isCont_root_neg isCont_root_zero) -done +lemma continuous_real_root[continuous_intros]: + "continuous F f \ continuous F (\x. root n (f x))" + unfolding continuous_def by (rule tendsto_real_root) + +lemma continuous_on_real_root[continuous_on_intros]: + "continuous_on s f \ continuous_on s (\x. root n (f x))" + unfolding continuous_on_def by (auto intro: tendsto_real_root) lemma DERIV_real_root: assumes n: "0 < n" @@ -351,9 +292,7 @@ by (rule DERIV_pow) show "real n * root n x ^ (n - Suc 0) \ 0" using n x by simp - show "isCont (root n) x" - using n by (rule isCont_real_root) -qed +qed (rule isCont_real_root) lemma DERIV_odd_real_root: assumes n: "odd n" @@ -368,9 +307,7 @@ by (rule DERIV_pow) show "real n * root n x ^ (n - Suc 0) \ 0" using odd_pos [OF n] x by simp - show "isCont (root n) x" - using odd_pos [OF n] by (rule isCont_real_root) -qed +qed (rule isCont_real_root) lemma DERIV_even_real_root: assumes n: "0 < n" and "even n" @@ -384,7 +321,7 @@ proof (rule allI, rule impI, erule conjE) fix y assume "x - 1 < y" and "y < 0" hence "root n (-y) ^ n = -y" using `0 < n` by simp - with real_root_minus[OF `0 < n`] and `even n` + with real_root_minus and `even n` show "- (root n y ^ n) = y" by simp qed next @@ -392,9 +329,7 @@ by (auto intro!: DERIV_intros) show "- real n * root n x ^ (n - Suc 0) \ 0" using n x by simp - show "isCont (root n) x" - using n by (rule isCont_real_root) -qed +qed (rule isCont_real_root) lemma DERIV_real_root_generic: assumes "0 < n" and "x \ 0" @@ -409,8 +344,7 @@ subsection {* Square Root *} -definition - sqrt :: "real \ real" where +definition sqrt :: "real \ real" where "sqrt = root 2" lemma pos2: "0 < (2::nat)" by simp @@ -441,16 +375,16 @@ unfolding sqrt_def by (rule real_root_one [OF pos2]) lemma real_sqrt_minus: "sqrt (- x) = - sqrt x" -unfolding sqrt_def by (rule real_root_minus [OF pos2]) +unfolding sqrt_def by (rule real_root_minus) lemma real_sqrt_mult: "sqrt (x * y) = sqrt x * sqrt y" -unfolding sqrt_def by (rule real_root_mult [OF pos2]) +unfolding sqrt_def by (rule real_root_mult) lemma real_sqrt_inverse: "sqrt (inverse x) = inverse (sqrt x)" -unfolding sqrt_def by (rule real_root_inverse [OF pos2]) +unfolding sqrt_def by (rule real_root_inverse) lemma real_sqrt_divide: "sqrt (x / y) = sqrt x / sqrt y" -unfolding sqrt_def by (rule real_root_divide [OF pos2]) +unfolding sqrt_def by (rule real_root_divide) lemma real_sqrt_power: "sqrt (x ^ k) = sqrt x ^ k" unfolding sqrt_def by (rule real_root_power [OF pos2]) @@ -459,7 +393,7 @@ unfolding sqrt_def by (rule real_root_gt_zero [OF pos2]) lemma real_sqrt_ge_zero: "0 \ x \ 0 \ sqrt x" -unfolding sqrt_def by (rule real_root_ge_zero [OF pos2]) +unfolding sqrt_def by (rule real_root_ge_zero) lemma real_sqrt_less_mono: "x < y \ sqrt x < sqrt y" unfolding sqrt_def by (rule real_root_less_mono [OF pos2]) @@ -489,7 +423,19 @@ lemmas real_sqrt_eq_1_iff [simp] = real_sqrt_eq_iff [where y=1, simplified] lemma isCont_real_sqrt: "isCont sqrt x" -unfolding sqrt_def by (rule isCont_real_root [OF pos2]) +unfolding sqrt_def by (rule isCont_real_root) + +lemma tendsto_real_sqrt[tendsto_intros]: + "(f ---> x) F \ ((\x. sqrt (f x)) ---> sqrt x) F" + unfolding sqrt_def by (rule tendsto_real_root) + +lemma continuous_real_sqrt[continuous_intros]: + "continuous F f \ continuous F (\x. sqrt (f x))" + unfolding sqrt_def by (rule continuous_real_root) + +lemma continuous_on_real_sqrt[continuous_on_intros]: + "continuous_on s f \ 0 < n \ continuous_on s (\x. sqrt (f x))" + unfolding sqrt_def by (rule continuous_on_real_root) lemma DERIV_real_sqrt_generic: assumes "x \ 0" diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Probability/Borel_Space.thy Sat Mar 23 07:30:53 2013 +0100 @@ -675,11 +675,6 @@ by (rule borel_measurable_continuous_Pair) (auto intro: continuous_on_fst continuous_on_snd continuous_on_mult) -lemma continuous_on_dist: - fixes f :: "'a :: t2_space \ 'b :: metric_space" - shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. dist (f x) (g x))" - unfolding continuous_on_eq_continuous_within by (auto simp: continuous_dist) - lemma borel_measurable_dist[measurable (raw)]: fixes g f :: "'a \ 'b::ordered_euclidean_space" assumes f: "f \ borel_measurable M" @@ -794,8 +789,7 @@ have "(\x. if f x \ {0<..} then ln (f x) else ln 0) \ borel_measurable M" proof (rule borel_measurable_continuous_on_open[OF _ _ f]) show "continuous_on {0<..} ln" - by (auto intro!: continuous_at_imp_continuous_on DERIV_ln DERIV_isCont - simp: continuous_isCont[symmetric]) + by (auto intro!: continuous_at_imp_continuous_on DERIV_ln DERIV_isCont) show "open ({0<..}::real set)" by auto qed also have "(\x. if x \ {0<..} then ln x else ln 0) = ln" @@ -808,8 +802,7 @@ unfolding log_def by auto lemma borel_measurable_exp[measurable]: "exp \ borel_measurable borel" - by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI - continuous_isCont[THEN iffD1] isCont_exp) + by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp) lemma measurable_count_space_eq2_countable: fixes f :: "'a => 'c::countable" diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Probability/Fin_Map.thy --- a/src/HOL/Probability/Fin_Map.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Probability/Fin_Map.thy Sat Mar 23 07:30:53 2013 +0100 @@ -222,8 +222,8 @@ hence open_sub: "\i S. i\domain x \ open (S i) \ x i\(S i) \ (\a\A i. a\(S i))" by auto have A_notempty: "\i. i \ domain x \ A i \ {}" using open_sub[of _ "\_. UNIV"] by auto let ?A = "(\f. Pi' (domain x) f) ` (Pi\<^isub>E (domain x) A)" - show "\A. countable A \ (\a\A. x \ a \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S))" - proof (rule exI[where x="?A"], safe) + show "\A::nat \ ('a\\<^isub>F'b) set. (\i. x \ (A i) \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + proof (rule first_countableI[where A="?A"], safe) show "countable ?A" using A by (simp add: countable_PiE) next fix S::"('a \\<^isub>F 'b) set" assume "open S" "x \ S" diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Sat Mar 23 07:30:53 2013 +0100 @@ -1049,7 +1049,7 @@ let ?g = "\x. if x = a then f a else if x = b then f b else if x \ {a <..< b} then f x else 0" from f have "continuous_on {a <..< b} f" - by (subst continuous_on_eq_continuous_at) (auto simp add: continuous_isCont) + by (subst continuous_on_eq_continuous_at) auto then have "?g \ borel_measurable borel" using borel_measurable_continuous_on_open[of "{a <..< b }" f "\x. x" borel 0] by (auto intro!: measurable_If[where P="\x. x = a"] measurable_If[where P="\x. x = b"]) diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Probability/Probability_Measure.thy Sat Mar 23 07:30:53 2013 +0100 @@ -159,11 +159,11 @@ moreover { fix y assume y: "y \ I" with q(2) `open I` have "Sup ((\x. q x + ?F x * (y - x)) ` I) = q y" - by (auto intro!: Sup_eq_maximum convex_le_Inf_differential image_eqI[OF _ y] simp: interior_open) } + by (auto intro!: cSup_eq_maximum convex_le_Inf_differential image_eqI[OF _ y] simp: interior_open) } ultimately have "q (expectation X) = Sup ((\x. q x + ?F x * (expectation X - x)) ` I)" by simp also have "\ \ expectation (\w. q (X w))" - proof (rule Sup_least) + proof (rule cSup_least) show "(\x. q x + ?F x * (expectation X - x)) ` I \ {}" using `I \ {}` by auto next diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/RealVector.thy --- a/src/HOL/RealVector.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/RealVector.thy Sat Mar 23 07:30:53 2013 +0100 @@ -5,7 +5,7 @@ header {* Vector Spaces and Algebras over the Reals *} theory RealVector -imports RComplete +imports RComplete Metric_Spaces SupInf begin subsection {* Locale for additive functions *} @@ -434,231 +434,6 @@ by (rule Reals_cases) auto -subsection {* Topological spaces *} - -class "open" = - fixes "open" :: "'a set \ bool" - -class topological_space = "open" + - assumes open_UNIV [simp, intro]: "open UNIV" - assumes open_Int [intro]: "open S \ open T \ open (S \ T)" - assumes open_Union [intro]: "\S\K. open S \ open (\ K)" -begin - -definition - closed :: "'a set \ bool" where - "closed S \ open (- S)" - -lemma open_empty [intro, simp]: "open {}" - using open_Union [of "{}"] by simp - -lemma open_Un [intro]: "open S \ open T \ open (S \ T)" - using open_Union [of "{S, T}"] by simp - -lemma open_UN [intro]: "\x\A. open (B x) \ open (\x\A. B x)" - unfolding SUP_def by (rule open_Union) auto - -lemma open_Inter [intro]: "finite S \ \T\S. open T \ open (\S)" - by (induct set: finite) auto - -lemma open_INT [intro]: "finite A \ \x\A. open (B x) \ open (\x\A. B x)" - unfolding INF_def by (rule open_Inter) auto - -lemma closed_empty [intro, simp]: "closed {}" - unfolding closed_def by simp - -lemma closed_Un [intro]: "closed S \ closed T \ closed (S \ T)" - unfolding closed_def by auto - -lemma closed_UNIV [intro, simp]: "closed UNIV" - unfolding closed_def by simp - -lemma closed_Int [intro]: "closed S \ closed T \ closed (S \ T)" - unfolding closed_def by auto - -lemma closed_INT [intro]: "\x\A. closed (B x) \ closed (\x\A. B x)" - unfolding closed_def by auto - -lemma closed_Inter [intro]: "\S\K. closed S \ closed (\ K)" - unfolding closed_def uminus_Inf by auto - -lemma closed_Union [intro]: "finite S \ \T\S. closed T \ closed (\S)" - by (induct set: finite) auto - -lemma closed_UN [intro]: "finite A \ \x\A. closed (B x) \ closed (\x\A. B x)" - unfolding SUP_def by (rule closed_Union) auto - -lemma open_closed: "open S \ closed (- S)" - unfolding closed_def by simp - -lemma closed_open: "closed S \ open (- S)" - unfolding closed_def by simp - -lemma open_Diff [intro]: "open S \ closed T \ open (S - T)" - unfolding closed_open Diff_eq by (rule open_Int) - -lemma closed_Diff [intro]: "closed S \ open T \ closed (S - T)" - unfolding open_closed Diff_eq by (rule closed_Int) - -lemma open_Compl [intro]: "closed S \ open (- S)" - unfolding closed_open . - -lemma closed_Compl [intro]: "open S \ closed (- S)" - unfolding open_closed . - -end - -inductive generate_topology for S where - UNIV: "generate_topology S UNIV" -| Int: "generate_topology S a \ generate_topology S b \ generate_topology S (a \ b)" -| UN: "(\k. k \ K \ generate_topology S k) \ generate_topology S (\K)" -| Basis: "s \ S \ generate_topology S s" - -hide_fact (open) UNIV Int UN Basis - -lemma generate_topology_Union: - "(\k. k \ I \ generate_topology S (K k)) \ generate_topology S (\k\I. K k)" - unfolding SUP_def by (intro generate_topology.UN) auto - -lemma topological_space_generate_topology: - "class.topological_space (generate_topology S)" - by default (auto intro: generate_topology.intros) - -class order_topology = order + "open" + - assumes open_generated_order: "open = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" -begin - -subclass topological_space - unfolding open_generated_order - by (rule topological_space_generate_topology) - -lemma open_greaterThan [simp]: "open {a <..}" - unfolding open_generated_order by (auto intro: generate_topology.Basis) - -lemma open_lessThan [simp]: "open {..< a}" - unfolding open_generated_order by (auto intro: generate_topology.Basis) - -lemma open_greaterThanLessThan [simp]: "open {a <..< b}" - unfolding greaterThanLessThan_eq by (simp add: open_Int) - -end - -class linorder_topology = linorder + order_topology - -lemma closed_atMost [simp]: "closed {.. a::'a::linorder_topology}" - by (simp add: closed_open) - -lemma closed_atLeast [simp]: "closed {a::'a::linorder_topology ..}" - by (simp add: closed_open) - -lemma closed_atLeastAtMost [simp]: "closed {a::'a::linorder_topology .. b}" -proof - - have "{a .. b} = {a ..} \ {.. b}" - by auto - then show ?thesis - by (simp add: closed_Int) -qed - -subsection {* Metric spaces *} - -class dist = - fixes dist :: "'a \ 'a \ real" - -class open_dist = "open" + dist + - assumes open_dist: "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" - -class metric_space = open_dist + - assumes dist_eq_0_iff [simp]: "dist x y = 0 \ x = y" - assumes dist_triangle2: "dist x y \ dist x z + dist y z" -begin - -lemma dist_self [simp]: "dist x x = 0" -by simp - -lemma zero_le_dist [simp]: "0 \ dist x y" -using dist_triangle2 [of x x y] by simp - -lemma zero_less_dist_iff: "0 < dist x y \ x \ y" -by (simp add: less_le) - -lemma dist_not_less_zero [simp]: "\ dist x y < 0" -by (simp add: not_less) - -lemma dist_le_zero_iff [simp]: "dist x y \ 0 \ x = y" -by (simp add: le_less) - -lemma dist_commute: "dist x y = dist y x" -proof (rule order_antisym) - show "dist x y \ dist y x" - using dist_triangle2 [of x y x] by simp - show "dist y x \ dist x y" - using dist_triangle2 [of y x y] by simp -qed - -lemma dist_triangle: "dist x z \ dist x y + dist y z" -using dist_triangle2 [of x z y] by (simp add: dist_commute) - -lemma dist_triangle3: "dist x y \ dist a x + dist a y" -using dist_triangle2 [of x y a] by (simp add: dist_commute) - -lemma dist_triangle_alt: - shows "dist y z <= dist x y + dist x z" -by (rule dist_triangle3) - -lemma dist_pos_lt: - shows "x \ y ==> 0 < dist x y" -by (simp add: zero_less_dist_iff) - -lemma dist_nz: - shows "x \ y \ 0 < dist x y" -by (simp add: zero_less_dist_iff) - -lemma dist_triangle_le: - shows "dist x z + dist y z <= e \ dist x y <= e" -by (rule order_trans [OF dist_triangle2]) - -lemma dist_triangle_lt: - shows "dist x z + dist y z < e ==> dist x y < e" -by (rule le_less_trans [OF dist_triangle2]) - -lemma dist_triangle_half_l: - shows "dist x1 y < e / 2 \ dist x2 y < e / 2 \ dist x1 x2 < e" -by (rule dist_triangle_lt [where z=y], simp) - -lemma dist_triangle_half_r: - 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) - -subclass topological_space -proof - have "\e::real. 0 < e" - by (fast intro: zero_less_one) - then show "open UNIV" - unfolding open_dist by simp -next - fix S T assume "open S" "open T" - then show "open (S \ T)" - unfolding open_dist - apply clarify - apply (drule (1) bspec)+ - apply (clarify, rename_tac r s) - apply (rule_tac x="min r s" in exI, simp) - done -next - fix K assume "\S\K. open S" thus "open (\K)" - unfolding open_dist by fast -qed - -lemma (in metric_space) open_ball: "open {y. dist x y < d}" -proof (unfold open_dist, intro ballI) - fix y assume *: "y \ {y. dist x y < d}" - then show "\e>0. \z. dist z y < e \ z \ {y. dist x y < d}" - by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt) -qed - -end - - subsection {* Real normed vector spaces *} class norm = @@ -890,7 +665,6 @@ using norm_triangle_ineq4 [of "x - z" "y - z"] by simp qed - subsection {* Class instances for real numbers *} instantiation real :: real_normed_field @@ -899,16 +673,9 @@ definition real_norm_def [simp]: "norm r = \r\" -definition dist_real_def: - "dist x y = \x - y\" - -definition open_real_def: - "open (S :: real set) \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" - instance apply (intro_classes, unfold real_norm_def real_scaleR_def) apply (rule dist_real_def) -apply (rule open_real_def) apply (simp add: sgn_real_def) apply (rule abs_eq_0) apply (rule abs_triangle_ineq) @@ -918,47 +685,6 @@ end -instance real :: linorder_topology -proof - show "(open :: real set \ bool) = generate_topology (range lessThan \ range greaterThan)" - proof (rule ext, safe) - fix S :: "real set" assume "open S" - then guess f unfolding open_real_def bchoice_iff .. - then have *: "S = (\x\S. {x - f x <..} \ {..< x + f x})" - by (fastforce simp: dist_real_def) - show "generate_topology (range lessThan \ range greaterThan) S" - apply (subst *) - apply (intro generate_topology_Union generate_topology.Int) - apply (auto intro: generate_topology.Basis) - done - next - fix S :: "real set" assume "generate_topology (range lessThan \ range greaterThan) S" - moreover have "\a::real. open {.. (\y. \y - x\ < a - x \ y \ {..e>0. \y. \y - x\ < e \ y \ {..a::real. open {a <..}" - unfolding open_real_def dist_real_def - proof clarify - fix x a :: real assume "a < x" - hence "0 < x - a \ (\y. \y - x\ < x - a \ y \ {a<..})" by auto - thus "\e>0. \y. \y - x\ < e \ y \ {a<..}" .. - qed - ultimately show "open S" - by induct auto - qed -qed - -lemmas open_real_greaterThan = open_greaterThan[where 'a=real] -lemmas open_real_lessThan = open_lessThan[where 'a=real] -lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real] -lemmas closed_real_atMost = closed_atMost[where 'a=real] -lemmas closed_real_atLeast = closed_atLeast[where 'a=real] -lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real] - subsection {* Extra type constraints *} text {* Only allow @{term "open"} in class @{text topological_space}. *} @@ -976,7 +702,6 @@ setup {* Sign.add_const_constraint (@{const_name norm}, SOME @{typ "'a::real_normed_vector \ real"}) *} - subsection {* Sign function *} lemma norm_sgn: @@ -1017,6 +742,8 @@ lemma real_sgn_neg: "(x::real) < 0 \ sgn x = -1" unfolding real_sgn_eq by simp +lemma norm_conv_dist: "norm x = dist x 0" + unfolding dist_norm by simp subsection {* Bounded Linear and Bilinear Operators *} @@ -1182,109 +909,6 @@ lemma bounded_linear_of_real: "bounded_linear (\r. of_real r)" unfolding of_real_def by (rule bounded_linear_scaleR_left) -subsection{* Hausdorff and other separation properties *} - -class t0_space = topological_space + - assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" - -class t1_space = topological_space + - assumes t1_space: "x \ y \ \U. open U \ x \ U \ y \ U" - -instance t1_space \ t0_space -proof qed (fast dest: t1_space) - -lemma separation_t1: - fixes x y :: "'a::t1_space" - shows "x \ y \ (\U. open U \ x \ U \ y \ U)" - using t1_space[of x y] by blast - -lemma closed_singleton: - fixes a :: "'a::t1_space" - shows "closed {a}" -proof - - let ?T = "\{S. open S \ a \ S}" - have "open ?T" by (simp add: open_Union) - also have "?T = - {a}" - by (simp add: set_eq_iff separation_t1, auto) - finally show "closed {a}" unfolding closed_def . -qed - -lemma closed_insert [simp]: - fixes a :: "'a::t1_space" - assumes "closed S" shows "closed (insert a S)" -proof - - from closed_singleton assms - have "closed ({a} \ S)" by (rule closed_Un) - thus "closed (insert a S)" by simp -qed - -lemma finite_imp_closed: - fixes S :: "'a::t1_space set" - shows "finite S \ closed S" -by (induct set: finite, simp_all) - -text {* T2 spaces are also known as Hausdorff spaces. *} - -class t2_space = topological_space + - assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" - -instance t2_space \ t1_space -proof qed (fast dest: hausdorff) - -lemma (in linorder) less_separate: - assumes "x < y" - shows "\a b. x \ {..< a} \ y \ {b <..} \ {..< a} \ {b <..} = {}" -proof cases - assume "\z. x < z \ z < y" - then guess z .. - then have "x \ {..< z} \ y \ {z <..} \ {z <..} \ {..< z} = {}" - by auto - then show ?thesis by blast -next - assume "\ (\z. x < z \ z < y)" - with `x < y` have "x \ {..< y} \ y \ {x <..} \ {x <..} \ {..< y} = {}" - by auto - then show ?thesis by blast -qed - -instance linorder_topology \ t2_space -proof - fix x y :: 'a - from less_separate[of x y] less_separate[of y x] - show "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" - by (elim neqE) (metis open_lessThan open_greaterThan Int_commute)+ -qed - -instance metric_space \ t2_space -proof - fix x y :: "'a::metric_space" - assume xy: "x \ y" - let ?U = "{y'. dist x y' < dist x y / 2}" - let ?V = "{x'. dist y x' < dist x y / 2}" - have th0: "\d x y z. (d x z :: real) \ d x y + d y z \ d y z = d z y - \ \(d x y * 2 < d x z \ d z y * 2 < d x z)" by arith - have "open ?U \ open ?V \ x \ ?U \ y \ ?V \ ?U \ ?V = {}" - using dist_pos_lt[OF xy] th0[of dist, OF dist_triangle dist_commute] - using open_ball[of _ "dist x y / 2"] by auto - then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" - by blast -qed - -lemma separation_t2: - fixes x y :: "'a::t2_space" - shows "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" - using hausdorff[of x y] by blast - -lemma separation_t0: - fixes x y :: "'a::t0_space" - shows "x \ y \ (\U. open U \ ~(x\U \ y\U))" - using t0_space[of x y] by blast - -text {* A perfect space is a topological space with no isolated points. *} - -class perfect_space = topological_space + - assumes not_open_singleton: "\ open {x}" - instance real_normed_algebra_1 \ perfect_space proof fix x::'a @@ -1293,4 +917,180 @@ by (clarsimp, rule_tac x="x + of_real (e/2)" in exI, simp) qed +section {* Connectedness *} + +class linear_continuum_topology = linorder_topology + conditional_complete_linorder + inner_dense_linorder +begin + +lemma Inf_notin_open: + assumes A: "open A" and bnd: "\a\A. x < a" + shows "Inf A \ A" +proof + assume "Inf A \ A" + then obtain b where "b < Inf A" "{b <.. Inf A} \ A" + using open_left[of A "Inf A" x] assms by auto + with dense[of b "Inf A"] obtain c where "c < Inf A" "c \ A" + by (auto simp: subset_eq) + then show False + using cInf_lower[OF `c \ A`, of x] bnd by (metis less_imp_le not_le) +qed + +lemma Sup_notin_open: + assumes A: "open A" and bnd: "\a\A. a < x" + shows "Sup A \ A" +proof + assume "Sup A \ A" + then obtain b where "Sup A < b" "{Sup A ..< b} \ A" + using open_right[of A "Sup A" x] assms by auto + with dense[of "Sup A" b] obtain c where "Sup A < c" "c \ A" + by (auto simp: subset_eq) + then show False + using cSup_upper[OF `c \ A`, of x] bnd by (metis less_imp_le not_le) +qed + end + +instance real :: linear_continuum_topology .. + +lemma connectedI_interval: + fixes U :: "'a :: linear_continuum_topology set" + assumes *: "\x y z. x \ U \ y \ U \ x \ z \ z \ y \ z \ U" + shows "connected U" +proof (rule connectedI) + { fix A B assume "open A" "open B" "A \ B \ U = {}" "U \ A \ B" + fix x y assume "x < y" "x \ A" "y \ B" "x \ U" "y \ U" + + let ?z = "Inf (B \ {x <..})" + + have "x \ ?z" "?z \ y" + using `y \ B` `x < y` by (auto intro: cInf_lower[where z=x] cInf_greatest) + with `x \ U` `y \ U` have "?z \ U" + by (rule *) + moreover have "?z \ B \ {x <..}" + using `open B` by (intro Inf_notin_open) auto + ultimately have "?z \ A" + using `x \ ?z` `A \ B \ U = {}` `x \ A` `U \ A \ B` by auto + + { assume "?z < y" + obtain a where "?z < a" "{?z ..< a} \ A" + using open_right[OF `open A` `?z \ A` `?z < y`] by auto + moreover obtain b where "b \ B" "x < b" "b < min a y" + using cInf_less_iff[of "B \ {x <..}" x "min a y"] `?z < a` `?z < y` `x < y` `y \ B` + by (auto intro: less_imp_le) + moreover then have "?z \ b" + by (intro cInf_lower[where z=x]) auto + moreover have "b \ U" + using `x \ ?z` `?z \ b` `b < min a y` + by (intro *[OF `x \ U` `y \ U`]) (auto simp: less_imp_le) + ultimately have "\b\B. b \ A \ b \ U" + by (intro bexI[of _ b]) auto } + then have False + using `?z \ y` `?z \ A` `y \ B` `y \ U` `A \ B \ U = {}` unfolding le_less by blast } + note not_disjoint = this + + fix A B assume AB: "open A" "open B" "U \ A \ B" "A \ B \ U = {}" + moreover assume "A \ U \ {}" then obtain x where x: "x \ U" "x \ A" by auto + moreover assume "B \ U \ {}" then obtain y where y: "y \ U" "y \ B" by auto + moreover note not_disjoint[of B A y x] not_disjoint[of A B x y] + ultimately show False by (cases x y rule: linorder_cases) auto +qed + +lemma connected_iff_interval: + fixes U :: "'a :: linear_continuum_topology set" + shows "connected U \ (\x\U. \y\U. \z. x \ z \ z \ y \ z \ U)" + by (auto intro: connectedI_interval dest: connectedD_interval) + +lemma connected_UNIV[simp]: "connected (UNIV::'a::linear_continuum_topology set)" + unfolding connected_iff_interval by auto + +lemma connected_Ioi[simp]: "connected {a::'a::linear_continuum_topology <..}" + unfolding connected_iff_interval by auto + +lemma connected_Ici[simp]: "connected {a::'a::linear_continuum_topology ..}" + unfolding connected_iff_interval by auto + +lemma connected_Iio[simp]: "connected {..< a::'a::linear_continuum_topology}" + unfolding connected_iff_interval by auto + +lemma connected_Iic[simp]: "connected {.. a::'a::linear_continuum_topology}" + unfolding connected_iff_interval by auto + +lemma connected_Ioo[simp]: "connected {a <..< b::'a::linear_continuum_topology}" + unfolding connected_iff_interval by auto + +lemma connected_Ioc[simp]: "connected {a <.. b::'a::linear_continuum_topology}" + unfolding connected_iff_interval by auto + +lemma connected_Ico[simp]: "connected {a ..< b::'a::linear_continuum_topology}" + unfolding connected_iff_interval by auto + +lemma connected_Icc[simp]: "connected {a .. b::'a::linear_continuum_topology}" + unfolding connected_iff_interval by auto + +lemma connected_contains_Ioo: + fixes A :: "'a :: linorder_topology set" + assumes A: "connected A" "a \ A" "b \ A" shows "{a <..< b} \ A" + using connectedD_interval[OF A] by (simp add: subset_eq Ball_def less_imp_le) + +subsection {* Intermediate Value Theorem *} + +lemma IVT': + fixes f :: "'a :: linear_continuum_topology \ 'b :: linorder_topology" + assumes y: "f a \ y" "y \ f b" "a \ b" + assumes *: "continuous_on {a .. b} f" + shows "\x. a \ x \ x \ b \ f x = y" +proof - + have "connected {a..b}" + unfolding connected_iff_interval by auto + from connected_continuous_image[OF * this, THEN connectedD_interval, of "f a" "f b" y] y + show ?thesis + by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) +qed + +lemma IVT2': + fixes f :: "'a :: linear_continuum_topology \ 'b :: linorder_topology" + assumes y: "f b \ y" "y \ f a" "a \ b" + assumes *: "continuous_on {a .. b} f" + shows "\x. a \ x \ x \ b \ f x = y" +proof - + have "connected {a..b}" + unfolding connected_iff_interval by auto + from connected_continuous_image[OF * this, THEN connectedD_interval, of "f b" "f a" y] y + show ?thesis + by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) +qed + +lemma IVT: + fixes f :: "'a :: linear_continuum_topology \ 'b :: linorder_topology" + shows "f a \ y \ y \ f b \ a \ b \ (\x. a \ x \ x \ b \ isCont f x) \ \x. a \ x \ x \ b \ f x = y" + by (rule IVT') (auto intro: continuous_at_imp_continuous_on) + +lemma IVT2: + fixes f :: "'a :: linear_continuum_topology \ 'b :: linorder_topology" + shows "f b \ y \ y \ f a \ a \ b \ (\x. a \ x \ x \ b \ isCont f x) \ \x. a \ x \ x \ b \ f x = y" + by (rule IVT2') (auto intro: continuous_at_imp_continuous_on) + +lemma continuous_inj_imp_mono: + fixes f :: "'a::linear_continuum_topology \ 'b :: linorder_topology" + assumes x: "a < x" "x < b" + assumes cont: "continuous_on {a..b} f" + assumes inj: "inj_on f {a..b}" + shows "(f a < f x \ f x < f b) \ (f b < f x \ f x < f a)" +proof - + note I = inj_on_iff[OF inj] + { assume "f x < f a" "f x < f b" + then obtain s t where "x \ s" "s \ b" "a \ t" "t \ x" "f s = f t" "f x < f s" + using IVT'[of f x "min (f a) (f b)" b] IVT2'[of f x "min (f a) (f b)" a] x + by (auto simp: continuous_on_subset[OF cont] less_imp_le) + with x I have False by auto } + moreover + { assume "f a < f x" "f b < f x" + then obtain s t where "x \ s" "s \ b" "a \ t" "t \ x" "f s = f t" "f s < f x" + using IVT'[of f a "max (f a) (f b)" x] IVT2'[of f b "max (f a) (f b)" x] x + by (auto simp: continuous_on_subset[OF cont] less_imp_le) + with x I have False by auto } + ultimately show ?thesis + using I[of a x] I[of x b] x less_trans[OF x] by (auto simp add: le_less less_imp_neq neq_iff) +qed + +end diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/SEQ.thy --- a/src/HOL/SEQ.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/SEQ.thy Sat Mar 23 07:30:53 2013 +0100 @@ -10,303 +10,19 @@ header {* Sequences and Convergence *} theory SEQ -imports Limits RComplete +imports Limits begin -subsection {* Monotone sequences and subsequences *} - -definition - monoseq :: "(nat \ 'a::order) \ bool" where - --{*Definition of monotonicity. - The use of disjunction here complicates proofs considerably. - One alternative is to add a Boolean argument to indicate the direction. - Another is to develop the notions of increasing and decreasing first.*} - "monoseq X = ((\m. \n\m. X m \ X n) | (\m. \n\m. X n \ X m))" - -definition - incseq :: "(nat \ 'a::order) \ bool" where - --{*Increasing sequence*} - "incseq X \ (\m. \n\m. X m \ X n)" - -definition - decseq :: "(nat \ 'a::order) \ bool" where - --{*Decreasing sequence*} - "decseq X \ (\m. \n\m. X n \ X m)" - -definition - subseq :: "(nat \ nat) \ bool" where - --{*Definition of subsequence*} - "subseq f \ (\m. \n>m. f m < f n)" - -lemma incseq_mono: "mono f \ incseq f" - unfolding mono_def incseq_def by auto - -lemma incseq_SucI: - "(\n. X n \ X (Suc n)) \ incseq X" - using lift_Suc_mono_le[of X] - by (auto simp: incseq_def) - -lemma incseqD: "\i j. incseq f \ i \ j \ f i \ f j" - by (auto simp: incseq_def) - -lemma incseq_SucD: "incseq A \ A i \ A (Suc i)" - using incseqD[of A i "Suc i"] by auto - -lemma incseq_Suc_iff: "incseq f \ (\n. f n \ f (Suc n))" - by (auto intro: incseq_SucI dest: incseq_SucD) - -lemma incseq_const[simp, intro]: "incseq (\x. k)" - unfolding incseq_def by auto - -lemma decseq_SucI: - "(\n. X (Suc n) \ X n) \ decseq X" - using order.lift_Suc_mono_le[OF dual_order, of X] - by (auto simp: decseq_def) - -lemma decseqD: "\i j. decseq f \ i \ j \ f j \ f i" - by (auto simp: decseq_def) - -lemma decseq_SucD: "decseq A \ A (Suc i) \ A i" - using decseqD[of A i "Suc i"] by auto - -lemma decseq_Suc_iff: "decseq f \ (\n. f (Suc n) \ f n)" - by (auto intro: decseq_SucI dest: decseq_SucD) - -lemma decseq_const[simp, intro]: "decseq (\x. k)" - unfolding decseq_def by auto - -lemma monoseq_iff: "monoseq X \ incseq X \ decseq X" - unfolding monoseq_def incseq_def decseq_def .. - -lemma monoseq_Suc: - "monoseq X \ (\n. X n \ X (Suc n)) \ (\n. X (Suc n) \ X n)" - unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff .. - -lemma monoI1: "\m. \ n \ m. X m \ X n ==> monoseq X" -by (simp add: monoseq_def) - -lemma monoI2: "\m. \ n \ m. X n \ X m ==> monoseq X" -by (simp add: monoseq_def) - -lemma mono_SucI1: "\n. X n \ X (Suc n) ==> monoseq X" -by (simp add: monoseq_Suc) - -lemma mono_SucI2: "\n. X (Suc n) \ X n ==> monoseq X" -by (simp add: monoseq_Suc) - -lemma monoseq_minus: - fixes a :: "nat \ 'a::ordered_ab_group_add" - assumes "monoseq a" - shows "monoseq (\ n. - a n)" -proof (cases "\ m. \ n \ m. a m \ a n") - case True - hence "\ m. \ n \ m. - a n \ - a m" by auto - thus ?thesis by (rule monoI2) -next - case False - hence "\ m. \ n \ m. - a m \ - a n" using `monoseq a`[unfolded monoseq_def] by auto - thus ?thesis by (rule monoI1) -qed - -text{*Subsequence (alternative definition, (e.g. Hoskins)*} - -lemma subseq_Suc_iff: "subseq f = (\n. (f n) < (f (Suc n)))" -apply (simp add: subseq_def) -apply (auto dest!: less_imp_Suc_add) -apply (induct_tac k) -apply (auto intro: less_trans) -done - -text{* for any sequence, there is a monotonic subsequence *} -lemma seq_monosub: - fixes s :: "nat => 'a::linorder" - shows "\f. subseq f \ monoseq (\ n. (s (f n)))" -proof cases - let "?P p n" = "p > n \ (\m\p. s m \ s p)" - assume *: "\n. \p. ?P p n" - def f \ "nat_rec (SOME p. ?P p 0) (\_ n. SOME p. ?P p n)" - have f_0: "f 0 = (SOME p. ?P p 0)" unfolding f_def by simp - have f_Suc: "\i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc .. - have P_0: "?P (f 0) 0" unfolding f_0 using *[rule_format] by (rule someI2_ex) auto - have P_Suc: "\i. ?P (f (Suc i)) (f i)" unfolding f_Suc using *[rule_format] by (rule someI2_ex) auto - then have "subseq f" unfolding subseq_Suc_iff by auto - moreover have "monoseq (\n. s (f n))" unfolding monoseq_Suc - proof (intro disjI2 allI) - fix n show "s (f (Suc n)) \ s (f n)" - proof (cases n) - case 0 with P_Suc[of 0] P_0 show ?thesis by auto - next - case (Suc m) - from P_Suc[of n] Suc have "f (Suc m) \ f (Suc (Suc m))" by simp - with P_Suc Suc show ?thesis by simp - qed - qed - ultimately show ?thesis by auto -next - let "?P p m" = "m < p \ s m < s p" - assume "\ (\n. \p>n. (\m\p. s m \ s p))" - then obtain N where N: "\p. p > N \ \m>p. s p < s m" by (force simp: not_le le_less) - def f \ "nat_rec (SOME p. ?P p (Suc N)) (\_ n. SOME p. ?P p n)" - have f_0: "f 0 = (SOME p. ?P p (Suc N))" unfolding f_def by simp - have f_Suc: "\i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc .. - have P_0: "?P (f 0) (Suc N)" - unfolding f_0 some_eq_ex[of "\p. ?P p (Suc N)"] using N[of "Suc N"] by auto - { fix i have "N < f i \ ?P (f (Suc i)) (f i)" - unfolding f_Suc some_eq_ex[of "\p. ?P p (f i)"] using N[of "f i"] . } - note P' = this - { fix i have "N < f i \ ?P (f (Suc i)) (f i)" - by (induct i) (insert P_0 P', auto) } - then have "subseq f" "monoseq (\x. s (f x))" - unfolding subseq_Suc_iff monoseq_Suc by (auto simp: not_le intro: less_imp_le) - then show ?thesis by auto -qed - -lemma seq_suble: assumes sf: "subseq f" shows "n \ f n" -proof(induct n) - case 0 thus ?case by simp -next - case (Suc n) - from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps - have "n < f (Suc n)" by arith - thus ?case by arith -qed - -lemma eventually_subseq: - "subseq r \ eventually P sequentially \ eventually (\n. P (r n)) sequentially" - unfolding eventually_sequentially by (metis seq_suble le_trans) - -lemma filterlim_subseq: "subseq f \ filterlim f sequentially sequentially" - unfolding filterlim_iff by (metis eventually_subseq) - -lemma subseq_o: "subseq r \ subseq s \ subseq (r \ s)" - unfolding subseq_def by simp - -lemma subseq_mono: assumes "subseq r" "m < n" shows "r m < r n" - using assms by (auto simp: subseq_def) - -lemma incseq_imp_monoseq: "incseq X \ monoseq X" - by (simp add: incseq_def monoseq_def) - -lemma decseq_imp_monoseq: "decseq X \ monoseq X" - by (simp add: decseq_def monoseq_def) - -lemma decseq_eq_incseq: - fixes X :: "nat \ 'a::ordered_ab_group_add" shows "decseq X = incseq (\n. - X n)" - by (simp add: decseq_def incseq_def) - -lemma INT_decseq_offset: - assumes "decseq F" - shows "(\i. F i) = (\i\{n..}. F i)" -proof safe - fix x i assume x: "x \ (\i\{n..}. F i)" - show "x \ F i" - proof cases - from x have "x \ F n" by auto - also assume "i \ n" with `decseq F` have "F n \ F i" - unfolding decseq_def by simp - finally show ?thesis . - qed (insert x, simp) -qed auto - -subsection {* Defintions of limits *} - -abbreviation (in topological_space) - LIMSEQ :: "[nat \ 'a, 'a] \ bool" - ("((_)/ ----> (_))" [60, 60] 60) where - "X ----> L \ (X ---> L) sequentially" - -definition - lim :: "(nat \ 'a::t2_space) \ 'a" where - --{*Standard definition of limit using choice operator*} - "lim X = (THE L. X ----> L)" - -definition (in topological_space) convergent :: "(nat \ 'a) \ bool" where - "convergent X = (\L. X ----> L)" - -definition - Bseq :: "(nat => 'a::real_normed_vector) => bool" where - --{*Standard definition for bounded sequence*} - "Bseq X = (\K>0.\n. norm (X n) \ K)" - -definition (in metric_space) Cauchy :: "(nat \ 'a) \ bool" where - "Cauchy X = (\e>0. \M. \m \ M. \n \ M. dist (X m) (X n) < e)" - - -subsection {* Bounded Sequences *} - -lemma BseqI': assumes K: "\n. norm (X n) \ K" shows "Bseq X" -unfolding Bseq_def -proof (intro exI conjI allI) - show "0 < max K 1" by simp -next - fix n::nat - have "norm (X n) \ K" by (rule K) - thus "norm (X n) \ max K 1" by simp -qed - -lemma BseqE: "\Bseq X; \K. \0 < K; \n. norm (X n) \ K\ \ Q\ \ Q" -unfolding Bseq_def by auto - -lemma BseqI2': assumes K: "\n\N. norm (X n) \ K" shows "Bseq X" -proof (rule BseqI') - let ?A = "norm ` X ` {..N}" - have 1: "finite ?A" by simp - fix n::nat - show "norm (X n) \ max K (Max ?A)" - proof (cases rule: linorder_le_cases) - assume "n \ N" - hence "norm (X n) \ K" using K by simp - thus "norm (X n) \ max K (Max ?A)" by simp - next - assume "n \ N" - hence "norm (X n) \ ?A" by simp - with 1 have "norm (X n) \ Max ?A" by (rule Max_ge) - thus "norm (X n) \ max K (Max ?A)" by simp - qed -qed - -lemma Bseq_ignore_initial_segment: "Bseq X \ Bseq (\n. X (n + k))" -unfolding Bseq_def by auto - -lemma Bseq_offset: "Bseq (\n. X (n + k)) \ Bseq X" -apply (erule BseqE) -apply (rule_tac N="k" and K="K" in BseqI2') -apply clarify -apply (drule_tac x="n - k" in spec, simp) -done - -lemma Bseq_conv_Bfun: "Bseq X \ Bfun X sequentially" -unfolding Bfun_def eventually_sequentially -apply (rule iffI) -apply (simp add: Bseq_def) -apply (auto intro: BseqI2') -done - - subsection {* Limits of Sequences *} lemma [trans]: "X=Y ==> Y ----> z ==> X ----> z" by simp -lemma LIMSEQ_def: "X ----> L = (\r>0. \no. \n\no. dist (X n) L < r)" -unfolding tendsto_iff eventually_sequentially .. - lemma LIMSEQ_iff: fixes L :: "'a::real_normed_vector" shows "(X ----> L) = (\r>0. \no. \n \ no. norm (X n - L) < r)" unfolding LIMSEQ_def dist_norm .. -lemma LIMSEQ_iff_nz: "X ----> L = (\r>0. \no>0. \n\no. dist (X n) L < r)" - unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc) - -lemma metric_LIMSEQ_I: - "(\r. 0 < r \ \no. \n\no. dist (X n) L < r) \ X ----> L" -by (simp add: LIMSEQ_def) - -lemma metric_LIMSEQ_D: - "\X ----> L; 0 < r\ \ \no. \n\no. dist (X n) L < r" -by (simp add: LIMSEQ_def) - lemma LIMSEQ_I: fixes L :: "'a::real_normed_vector" shows "(\r. 0 < r \ \no. \n\no. norm (X n - L) < r) \ X ----> L" @@ -317,78 +33,10 @@ shows "\X ----> L; 0 < r\ \ \no. \n\no. norm (X n - L) < r" by (simp add: LIMSEQ_iff) -lemma LIMSEQ_const_iff: - fixes k l :: "'a::t2_space" - shows "(\n. k) ----> l \ k = l" - using trivial_limit_sequentially by (rule tendsto_const_iff) - -lemma LIMSEQ_SUP: - "incseq X \ X ----> (SUP i. X i :: 'a :: {complete_linorder, linorder_topology})" - by (intro increasing_tendsto) - (auto simp: SUP_upper less_SUP_iff incseq_def eventually_sequentially intro: less_le_trans) - -lemma LIMSEQ_INF: - "decseq X \ X ----> (INF i. X i :: 'a :: {complete_linorder, linorder_topology})" - by (intro decreasing_tendsto) - (auto simp: INF_lower INF_less_iff decseq_def eventually_sequentially intro: le_less_trans) - -lemma LIMSEQ_ignore_initial_segment: - "f ----> a \ (\n. f (n + k)) ----> a" -apply (rule topological_tendstoI) -apply (drule (2) topological_tendstoD) -apply (simp only: eventually_sequentially) -apply (erule exE, rename_tac N) -apply (rule_tac x=N in exI) -apply simp -done - -lemma LIMSEQ_offset: - "(\n. f (n + k)) ----> a \ f ----> a" -apply (rule topological_tendstoI) -apply (drule (2) topological_tendstoD) -apply (simp only: eventually_sequentially) -apply (erule exE, rename_tac N) -apply (rule_tac x="N + k" in exI) -apply clarify -apply (drule_tac x="n - k" in spec) -apply (simp add: le_diff_conv2) -done - -lemma LIMSEQ_Suc: "f ----> l \ (\n. f (Suc n)) ----> l" -by (drule_tac k="Suc 0" in LIMSEQ_ignore_initial_segment, simp) - -lemma LIMSEQ_imp_Suc: "(\n. f (Suc n)) ----> l \ f ----> l" -by (rule_tac k="Suc 0" in LIMSEQ_offset, simp) - -lemma LIMSEQ_Suc_iff: "(\n. f (Suc n)) ----> l = f ----> l" -by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc) - lemma LIMSEQ_linear: "\ X ----> x ; l > 0 \ \ (\ n. X (n * l)) ----> x" unfolding tendsto_def eventually_sequentially by (metis div_le_dividend div_mult_self1_is_m le_trans nat_mult_commute) -lemma LIMSEQ_unique: - fixes a b :: "'a::t2_space" - shows "\X ----> a; X ----> b\ \ a = b" - using trivial_limit_sequentially by (rule tendsto_unique) - -lemma increasing_LIMSEQ: - fixes f :: "nat \ real" - assumes inc: "\n. f n \ f (Suc n)" - and bdd: "\n. f n \ l" - and en: "\e. 0 < e \ \n. l \ f n + e" - shows "f ----> l" -proof (rule increasing_tendsto) - fix x assume "x < l" - with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x" - by auto - from en[OF `0 < e`] obtain n where "l - e \ f n" - by (auto simp: field_simps) - with `e < l - x` `0 < e` have "x < f n" by simp - with incseq_SucI[of f, OF inc] show "eventually (\n. x < f n) sequentially" - by (auto simp: eventually_sequentially incseq_def intro: less_le_trans) -qed (insert bdd, auto) - lemma Bseq_inverse_lemma: fixes x :: "'a::real_normed_div_algebra" shows "\r \ norm x; 0 < r\ \ norm (inverse x) \ inverse r" @@ -399,7 +47,7 @@ lemma Bseq_inverse: fixes a :: "'a::real_normed_div_algebra" shows "\X ----> a; a \ 0\ \ Bseq (\n. inverse (X n))" -unfolding Bseq_conv_Bfun by (rule Bfun_inverse) + by (rule Bfun_inverse) lemma LIMSEQ_diff_approach_zero: fixes L :: "'a::real_normed_vector" @@ -443,37 +91,8 @@ using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]] by auto -lemma LIMSEQ_le_const: - "\X ----> (x::'a::linorder_topology); \N. \n\N. a \ X n\ \ a \ x" - using tendsto_le_const[of sequentially X x a] by (simp add: eventually_sequentially) - -lemma LIMSEQ_le: - "\X ----> x; Y ----> y; \N. \n\N. X n \ Y n\ \ x \ (y::'a::linorder_topology)" - using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially) - -lemma LIMSEQ_le_const2: - "\X ----> (x::'a::linorder_topology); \N. \n\N. X n \ a\ \ x \ a" - by (rule LIMSEQ_le[of X x "\n. a"]) (auto simp: tendsto_const) - subsection {* Convergence *} -lemma limI: "X ----> L ==> lim X = L" -apply (simp add: lim_def) -apply (blast intro: LIMSEQ_unique) -done - -lemma convergentD: "convergent X ==> \L. (X ----> L)" -by (simp add: convergent_def) - -lemma convergentI: "(X ----> L) ==> convergent X" -by (auto simp add: convergent_def) - -lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)" -by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) - -lemma convergent_const: "convergent (\n. c)" - by (rule convergentI, rule tendsto_const) - lemma convergent_add: fixes X Y :: "nat \ 'a::real_normed_vector" assumes "convergent (\n. X n)" @@ -508,26 +127,28 @@ apply (drule tendsto_minus, auto) done -lemma lim_le: "convergent f \ (\n. f n \ (x::real)) \ lim f \ x" - using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff) - -lemma monoseq_le: - "monoseq a \ a ----> (x::real) \ - ((\ n. a n \ x) \ (\m. \n\m. a m \ a n)) \ ((\ n. x \ a n) \ (\m. \n\m. a n \ a m))" - by (metis LIMSEQ_le_const LIMSEQ_le_const2 decseq_def incseq_def monoseq_iff) - -lemma LIMSEQ_subseq_LIMSEQ: - "\ X ----> L; subseq f \ \ (X o f) ----> L" - unfolding comp_def by (rule filterlim_compose[of X, OF _ filterlim_subseq]) - -lemma convergent_subseq_convergent: - "\convergent X; subseq f\ \ convergent (X o f)" - unfolding convergent_def by (auto intro: LIMSEQ_subseq_LIMSEQ) - subsection {* Bounded Monotonic Sequences *} -text{*Bounded Sequence*} +subsubsection {* Bounded Sequences *} + +lemma BseqI': "(\n. norm (X n) \ K) \ Bseq X" + by (intro BfunI) (auto simp: eventually_sequentially) + +lemma BseqI2': "\n\N. norm (X n) \ K \ Bseq X" + by (intro BfunI) (auto simp: eventually_sequentially) + +lemma Bseq_def: "Bseq X \ (\K>0. \n. norm (X n) \ K)" + unfolding Bfun_def eventually_sequentially +proof safe + fix N K assume "0 < K" "\n\N. norm (X n) \ K" + then show "\K>0. \n. norm (X n) \ K" + by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] min_max.less_supI2) + (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) +qed auto + +lemma BseqE: "\Bseq X; \K. \0 < K; \n. norm (X n) \ K\ \ Q\ \ Q" +unfolding Bseq_def by auto lemma BseqD: "Bseq X ==> \K. 0 < K & (\n. norm (X n) \ K)" by (simp add: Bseq_def) @@ -619,11 +240,11 @@ (* TODO: delete *) (* FIXME: one use in NSA/HSEQ.thy *) lemma Bmonoseq_LIMSEQ: "\n. m \ n --> X n = X m ==> \L. (X ----> L)" -unfolding tendsto_def eventually_sequentially -apply (rule_tac x = "X m" in exI, safe) -apply (rule_tac x = m in exI, safe) -apply (drule spec, erule impE, auto) -done + apply (rule_tac x="X m" in exI) + apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const]) + unfolding eventually_sequentially + apply blast + done text {* A monotone sequence converges to its least upper bound. *} @@ -656,7 +277,7 @@ "Bseq X \ \m. \n \ m. X m \ X n \ convergent (X::nat=>real)" by (metis Bseq_isLub isLub_mono_imp_LIMSEQ convergentI) -lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X" +lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X" by (simp add: Bseq_def) text{*Main monotonicity theorem*} @@ -665,39 +286,11 @@ by (metis monoseq_iff incseq_def decseq_eq_incseq convergent_minus_iff Bseq_minus_iff Bseq_mono_convergent) -subsubsection{*Increasing and Decreasing Series*} - -lemma incseq_le: "incseq X \ X ----> L \ X n \ (L::real)" - by (metis incseq_def LIMSEQ_le_const) - -lemma decseq_le: "decseq X \ X ----> L \ (L::real) \ X n" - by (metis decseq_def LIMSEQ_le_const2) - -subsection {* Cauchy Sequences *} - -lemma metric_CauchyI: - "(\e. 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e) \ Cauchy X" - by (simp add: Cauchy_def) - -lemma metric_CauchyD: - "Cauchy X \ 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e" - by (simp add: Cauchy_def) - lemma Cauchy_iff: fixes X :: "nat \ 'a::real_normed_vector" shows "Cauchy X \ (\e>0. \M. \m\M. \n\M. norm (X m - X n) < e)" unfolding Cauchy_def dist_norm .. -lemma Cauchy_iff2: - "Cauchy X = (\j. (\M. \m \ M. \n \ M. \X m - X n\ < inverse(real (Suc j))))" -apply (simp add: Cauchy_iff, auto) -apply (drule reals_Archimedean, safe) -apply (drule_tac x = n in spec, auto) -apply (rule_tac x = M in exI, auto) -apply (drule_tac x = m in spec, simp) -apply (drule_tac x = na in spec, auto) -done - lemma CauchyI: fixes X :: "nat \ 'a::real_normed_vector" shows "(\e. 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e) \ Cauchy X" @@ -708,13 +301,42 @@ shows "\Cauchy X; 0 < e\ \ \M. \m\M. \n\M. norm (X m - X n) < e" by (simp add: Cauchy_iff) -lemma Cauchy_subseq_Cauchy: - "\ Cauchy X; subseq f \ \ Cauchy (X o f)" -apply (auto simp add: Cauchy_def) -apply (drule_tac x=e in spec, clarify) -apply (rule_tac x=M in exI, clarify) -apply (blast intro: le_trans [OF _ seq_suble] dest!: spec) -done +lemma Bseq_eq_bounded: "range f \ {a .. b::real} \ Bseq f" + apply (simp add: subset_eq) + apply (rule BseqI'[where K="max (norm a) (norm b)"]) + apply (erule_tac x=n in allE) + apply auto + done + +lemma incseq_bounded: "incseq X \ \i. X i \ (B::real) \ Bseq X" + by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) + +lemma decseq_bounded: "decseq X \ \i. (B::real) \ X i \ Bseq X" + by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) + +lemma incseq_convergent: + fixes X :: "nat \ real" + assumes "incseq X" and "\i. X i \ B" + obtains L where "X ----> L" "\i. X i \ L" +proof atomize_elim + from incseq_bounded[OF assms] `incseq X` Bseq_monoseq_convergent[of X] + obtain L where "X ----> L" + by (auto simp: convergent_def monoseq_def incseq_def) + with `incseq X` show "\L. X ----> L \ (\i. X i \ L)" + by (auto intro!: exI[of _ L] incseq_le) +qed + +lemma decseq_convergent: + fixes X :: "nat \ real" + assumes "decseq X" and "\i. B \ X i" + obtains L where "X ----> L" "\i. L \ X i" +proof atomize_elim + from decseq_bounded[OF assms] `decseq X` Bseq_monoseq_convergent[of X] + obtain L where "X ----> L" + by (auto simp: convergent_def monoseq_def decseq_def) + with `decseq X` show "\L. X ----> L \ (\i. L \ X i)" + by (auto intro!: exI[of _ L] decseq_le) +qed subsubsection {* Cauchy Sequences are Bounded *} @@ -729,141 +351,9 @@ apply simp done -lemma Cauchy_Bseq: "Cauchy X ==> Bseq X" -apply (simp add: Cauchy_iff) -apply (drule spec, drule mp, rule zero_less_one, safe) -apply (drule_tac x="M" in spec, simp) -apply (drule lemmaCauchy) -apply (rule_tac k="M" in Bseq_offset) -apply (simp add: Bseq_def) -apply (rule_tac x="1 + norm (X M)" in exI) -apply (rule conjI, rule order_less_le_trans [OF zero_less_one], simp) -apply (simp add: order_less_imp_le) -done - -subsubsection {* Cauchy Sequences are Convergent *} - -class complete_space = metric_space + - assumes Cauchy_convergent: "Cauchy X \ convergent X" - class banach = real_normed_vector + complete_space -theorem LIMSEQ_imp_Cauchy: - assumes X: "X ----> a" shows "Cauchy X" -proof (rule metric_CauchyI) - fix e::real assume "0 < e" - hence "0 < e/2" by simp - with X have "\N. \n\N. dist (X n) a < e/2" by (rule metric_LIMSEQ_D) - then obtain N where N: "\n\N. dist (X n) a < e/2" .. - show "\N. \m\N. \n\N. dist (X m) (X n) < e" - proof (intro exI allI impI) - fix m assume "N \ m" - hence m: "dist (X m) a < e/2" using N by fast - fix n assume "N \ n" - hence n: "dist (X n) a < e/2" using N by fast - have "dist (X m) (X n) \ dist (X m) a + dist (X n) a" - by (rule dist_triangle2) - also from m n have "\ < e" by simp - finally show "dist (X m) (X n) < e" . - qed -qed - -lemma convergent_Cauchy: "convergent X \ Cauchy X" -unfolding convergent_def -by (erule exE, erule LIMSEQ_imp_Cauchy) - -lemma Cauchy_convergent_iff: - fixes X :: "nat \ 'a::complete_space" - shows "Cauchy X = convergent X" -by (fast intro: Cauchy_convergent convergent_Cauchy) - -text {* -Proof that Cauchy sequences converge based on the one from -http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html -*} - -text {* - If sequence @{term "X"} is Cauchy, then its limit is the lub of - @{term "{r::real. \N. \n\N. r < X n}"} -*} - -lemma isUb_UNIV_I: "(\y. y \ S \ y \ u) \ isUb UNIV S u" -by (simp add: isUbI setleI) - -lemma real_Cauchy_convergent: - fixes X :: "nat \ real" - assumes X: "Cauchy X" - shows "convergent X" -proof - - def S \ "{x::real. \N. \n\N. x < X n}" - then have mem_S: "\N x. \n\N. x < X n \ x \ S" by auto - - { fix N x assume N: "\n\N. X n < x" - have "isUb UNIV S x" - proof (rule isUb_UNIV_I) - fix y::real assume "y \ S" - hence "\M. \n\M. y < X n" - by (simp add: S_def) - then obtain M where "\n\M. y < X n" .. - hence "y < X (max M N)" by simp - also have "\ < x" using N by simp - finally show "y \ x" - by (rule order_less_imp_le) - qed } - note bound_isUb = this - - have "\u. isLub UNIV S u" - proof (rule reals_complete) - obtain N where "\m\N. \n\N. norm (X m - X n) < 1" - using CauchyD [OF X zero_less_one] by auto - hence N: "\n\N. norm (X n - X N) < 1" by simp - show "\x. x \ S" - proof - from N have "\n\N. X N - 1 < X n" - by (simp add: abs_diff_less_iff) - thus "X N - 1 \ S" by (rule mem_S) - qed - show "\u. isUb UNIV S u" - proof - from N have "\n\N. X n < X N + 1" - by (simp add: abs_diff_less_iff) - thus "isUb UNIV S (X N + 1)" - by (rule bound_isUb) - qed - qed - then obtain x where x: "isLub UNIV S x" .. - have "X ----> x" - proof (rule LIMSEQ_I) - fix r::real assume "0 < r" - hence r: "0 < r/2" by simp - obtain N where "\n\N. \m\N. norm (X n - X m) < r/2" - using CauchyD [OF X r] by auto - hence "\n\N. norm (X n - X N) < r/2" by simp - hence N: "\n\N. X N - r/2 < X n \ X n < X N + r/2" - by (simp only: real_norm_def abs_diff_less_iff) - - from N have "\n\N. X N - r/2 < X n" by fast - hence "X N - r/2 \ S" by (rule mem_S) - hence 1: "X N - r/2 \ x" using x isLub_isUb isUbD by fast - - from N have "\n\N. X n < X N + r/2" by fast - hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb) - hence 2: "x \ X N + r/2" using x isLub_le_isUb by fast - - show "\N. \n\N. norm (X n - x) < r" - proof (intro exI allI impI) - fix n assume n: "N \ n" - from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+ - thus "norm (X n - x) < r" using 1 2 - by (simp add: abs_diff_less_iff) - qed - qed - then show ?thesis unfolding convergent_def by auto -qed - -instance real :: banach - by intro_classes (rule real_Cauchy_convergent) - +instance real :: banach by default subsection {* Power Sequences *} @@ -922,33 +412,4 @@ lemma LIMSEQ_rabs_realpow_zero2: "\c\ < 1 \ (\n. c ^ n :: real) ----> 0" by (rule LIMSEQ_power_zero) simp -lemma tendsto_at_topI_sequentially: - fixes f :: "real \ real" - assumes mono: "mono f" - assumes limseq: "(\n. f (real n)) ----> y" - shows "(f ---> y) at_top" -proof (rule tendstoI) - fix e :: real assume "0 < e" - with limseq obtain N :: nat where N: "\n. N \ n \ \f (real n) - y\ < e" - by (auto simp: LIMSEQ_def dist_real_def) - { fix x :: real - from ex_le_of_nat[of x] guess n .. - note monoD[OF mono this] - also have "f (real_of_nat n) \ y" - by (rule LIMSEQ_le_const[OF limseq]) - (auto intro: exI[of _ n] monoD[OF mono] simp: real_eq_of_nat[symmetric]) - finally have "f x \ y" . } - note le = this - have "eventually (\x. real N \ x) at_top" - by (rule eventually_ge_at_top) - then show "eventually (\x. dist (f x) y < e) at_top" - proof eventually_elim - fix x assume N': "real N \ x" - with N[of N] le have "y - f (real N) < e" by auto - moreover note monoD[OF mono N'] - ultimately show "dist (f x) y < e" - using le[of x] by (auto simp: dist_real_def field_simps) - qed -qed - end diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Series.thy --- a/src/HOL/Series.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Series.thy Sat Mar 23 07:30:53 2013 +0100 @@ -373,16 +373,13 @@ have "convergent (\n. setsum f {0..n. setsum f {0..n. setsum f {0..n. x"]) - (auto simp add: le pos) + by (intro BseqI'[of _ x]) (auto simp add: setsum_nonneg pos intro: le) next show "\m n. m \ n \ setsum f {0.. setsum f {0.. L" - by (blast dest: convergentD) thus ?thesis - by (force simp add: summable_def sums_def) + by (force simp add: summable_def sums_def convergent_def) qed lemma series_pos_le: diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/SupInf.thy --- a/src/HOL/SupInf.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/SupInf.thy Sat Mar 23 07:30:53 2013 +0100 @@ -6,210 +6,341 @@ imports RComplete begin -instantiation real :: Sup +lemma Sup_fin_eq_Max: "finite X \ X \ {} \ Sup_fin X = Max X" + by (induct X rule: finite_ne_induct) (simp_all add: sup_max) + +lemma Inf_fin_eq_Min: "finite X \ X \ {} \ Inf_fin X = Min X" + by (induct X rule: finite_ne_induct) (simp_all add: inf_min) + +text {* + +To avoid name classes with the @{class complete_lattice}-class we prefix @{const Sup} and +@{const Inf} in theorem names with c. + +*} + +class conditional_complete_lattice = lattice + Sup + Inf + + assumes cInf_lower: "x \ X \ (\a. a \ X \ z \ a) \ Inf X \ x" + and cInf_greatest: "X \ {} \ (\x. x \ X \ z \ x) \ z \ Inf X" + assumes cSup_upper: "x \ X \ (\a. a \ X \ a \ z) \ x \ Sup X" + and cSup_least: "X \ {} \ (\x. x \ X \ x \ z) \ Sup X \ z" begin -definition - Sup_real_def: "Sup X == (LEAST z::real. \x\X. x\z)" + +lemma cSup_eq_maximum: (*REAL_SUP_MAX in HOL4*) + "z \ X \ (\x. x \ X \ x \ z) \ Sup X = z" + by (blast intro: antisym cSup_upper cSup_least) + +lemma cInf_eq_minimum: (*REAL_INF_MIN in HOL4*) + "z \ X \ (\x. x \ X \ z \ x) \ Inf X = z" + by (intro antisym cInf_lower[of z X z] cInf_greatest[of X z]) auto + +lemma cSup_le_iff: "S \ {} \ (\a. a \ S \ a \ z) \ Sup S \ a \ (\x\S. x \ a)" + by (metis order_trans cSup_upper cSup_least) + +lemma le_cInf_iff: "S \ {} \ (\a. a \ S \ z \ a) \ a \ Inf S \ (\x\S. a \ x)" + by (metis order_trans cInf_lower cInf_greatest) + +lemma cSup_singleton [simp]: "Sup {x} = x" + by (intro cSup_eq_maximum) auto + +lemma cInf_singleton [simp]: "Inf {x} = x" + by (intro cInf_eq_minimum) auto + +lemma cSup_upper2: (*REAL_IMP_LE_SUP in HOL4*) + "x \ X \ y \ x \ (\x. x \ X \ x \ z) \ y \ Sup X" + by (metis cSup_upper order_trans) + +lemma cInf_lower2: + "x \ X \ x \ y \ (\x. x \ X \ z \ x) \ Inf X \ y" + by (metis cInf_lower order_trans) + +lemma cSup_upper_EX: "x \ X \ \z. \x. x \ X \ x \ z \ x \ Sup X" + by (blast intro: cSup_upper) + +lemma cInf_lower_EX: "x \ X \ \z. \x. x \ X \ z \ x \ Inf X \ x" + by (blast intro: cInf_lower) + +lemma cSup_eq_non_empty: + assumes 1: "X \ {}" + assumes 2: "\x. x \ X \ x \ a" + assumes 3: "\y. (\x. x \ X \ x \ y) \ a \ y" + shows "Sup X = a" + by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper) + +lemma cInf_eq_non_empty: + assumes 1: "X \ {}" + assumes 2: "\x. x \ X \ a \ x" + assumes 3: "\y. (\x. x \ X \ y \ x) \ y \ a" + shows "Inf X = a" + by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower) + +lemma cSup_insert: + assumes x: "X \ {}" + and z: "\x. x \ X \ x \ z" + shows "Sup (insert a X) = sup a (Sup X)" +proof (intro cSup_eq_non_empty) + fix y assume "\x. x \ insert a X \ x \ y" with x show "sup a (Sup X) \ y" by (auto intro: cSup_least) +qed (auto intro: le_supI2 z cSup_upper) -instance .. +lemma cInf_insert: + assumes x: "X \ {}" + and z: "\x. x \ X \ z \ x" + shows "Inf (insert a X) = inf a (Inf X)" +proof (intro cInf_eq_non_empty) + fix y assume "\x. x \ insert a X \ y \ x" with x show "y \ inf a (Inf X)" by (auto intro: cInf_greatest) +qed (auto intro: le_infI2 z cInf_lower) + +lemma cSup_insert_If: + "(\x. x \ X \ x \ z) \ Sup (insert a X) = (if X = {} then a else sup a (Sup X))" + using cSup_insert[of X z] by simp + +lemma cInf_insert_if: + "(\x. x \ X \ z \ x) \ Inf (insert a X) = (if X = {} then a else inf a (Inf X))" + using cInf_insert[of X z] by simp + +lemma le_cSup_finite: "finite X \ x \ X \ x \ Sup X" +proof (induct X arbitrary: x rule: finite_induct) + case (insert x X y) then show ?case + apply (cases "X = {}") + apply simp + apply (subst cSup_insert[of _ "Sup X"]) + apply (auto intro: le_supI2) + done +qed simp + +lemma cInf_le_finite: "finite X \ x \ X \ Inf X \ x" +proof (induct X arbitrary: x rule: finite_induct) + case (insert x X y) then show ?case + apply (cases "X = {}") + apply simp + apply (subst cInf_insert[of _ "Inf X"]) + apply (auto intro: le_infI2) + done +qed simp + +lemma cSup_eq_Sup_fin: "finite X \ X \ {} \ Sup X = Sup_fin X" +proof (induct X rule: finite_ne_induct) + case (insert x X) then show ?case + using cSup_insert[of X "Sup_fin X" x] le_cSup_finite[of X] by simp +qed simp + +lemma cInf_eq_Inf_fin: "finite X \ X \ {} \ Inf X = Inf_fin X" +proof (induct X rule: finite_ne_induct) + case (insert x X) then show ?case + using cInf_insert[of X "Inf_fin X" x] cInf_le_finite[of X] by simp +qed simp + +lemma cSup_atMost[simp]: "Sup {..x} = x" + by (auto intro!: cSup_eq_maximum) + +lemma cSup_greaterThanAtMost[simp]: "y < x \ Sup {y<..x} = x" + by (auto intro!: cSup_eq_maximum) + +lemma cSup_atLeastAtMost[simp]: "y \ x \ Sup {y..x} = x" + by (auto intro!: cSup_eq_maximum) + +lemma cInf_atLeast[simp]: "Inf {x..} = x" + by (auto intro!: cInf_eq_minimum) + +lemma cInf_atLeastLessThan[simp]: "y < x \ Inf {y.. x \ Inf {y..x} = y" + by (auto intro!: cInf_eq_minimum) + end -instantiation real :: Inf +instance complete_lattice \ conditional_complete_lattice + by default (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest) + +lemma isLub_cSup: + "(S::'a :: conditional_complete_lattice set) \ {} \ (\b. S *<= b) \ isLub UNIV S (Sup S)" + by (auto simp add: isLub_def setle_def leastP_def isUb_def + intro!: setgeI intro: cSup_upper cSup_least) + +lemma cSup_eq: + fixes a :: "'a :: {conditional_complete_lattice, no_bot}" + assumes upper: "\x. x \ X \ x \ a" + assumes least: "\y. (\x. x \ X \ x \ y) \ a \ y" + shows "Sup X = a" +proof cases + assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) +qed (intro cSup_eq_non_empty assms) + +lemma cInf_eq: + fixes a :: "'a :: {conditional_complete_lattice, no_top}" + assumes upper: "\x. x \ X \ a \ x" + assumes least: "\y. (\x. x \ X \ y \ x) \ y \ a" + shows "Inf X = a" +proof cases + assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) +qed (intro cInf_eq_non_empty assms) + +lemma cSup_le: "(S::'a::conditional_complete_lattice set) \ {} \ S *<= b \ Sup S \ b" + by (metis cSup_least setle_def) + +lemma cInf_ge: "(S::'a :: conditional_complete_lattice set) \ {} \ b <=* S \ Inf S \ b" + by (metis cInf_greatest setge_def) + +class conditional_complete_linorder = conditional_complete_lattice + linorder begin + +lemma less_cSup_iff : (*REAL_SUP_LE in HOL4*) + "X \ {} \ (\x. x \ X \ x \ z) \ y < Sup X \ (\x\X. y < x)" + by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans) + +lemma cInf_less_iff: "X \ {} \ (\x. x \ X \ z \ x) \ Inf X < y \ (\x\X. x < y)" + by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans) + +lemma less_cSupE: + assumes "y < Sup X" "X \ {}" obtains x where "x \ X" "y < x" + by (metis cSup_least assms not_le that) + +lemma complete_interval: + assumes "a < b" and "P a" and "\ P b" + shows "\c. a \ c \ c \ b \ (\x. a \ x \ x < c \ P x) \ + (\d. (\x. a \ x \ x < d \ P x) \ d \ c)" +proof (rule exI [where x = "Sup {d. \x. a \ x & x < d --> P x}"], auto) + show "a \ Sup {d. \c. a \ c \ c < d \ P c}" + by (rule cSup_upper [where z=b], auto) + (metis `a < b` `\ P b` linear less_le) +next + show "Sup {d. \c. a \ c \ c < d \ P c} \ b" + apply (rule cSup_least) + apply auto + apply (metis less_le_not_le) + apply (metis `a x" and lt: "x < Sup {d. \c. a \ c \ c < d \ P c}" + show "P x" + apply (rule less_cSupE [OF lt], auto) + apply (metis less_le_not_le) + apply (metis x) + done +next + fix d + assume 0: "\x. a \ x \ x < d \ P x" + thus "d \ Sup {d. \c. a \ c \ c < d \ P c}" + by (rule_tac z="b" in cSup_upper, auto) + (metis `a (\b'x\S. b' < x) \ Sup S = b" + by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def) + +lemma cInf_unique: "b <=* (S::'a :: {conditional_complete_linorder, no_top} set) \ (\b'>b. \x\S. b' > x) \ Inf S = b" + by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def) + +lemma cSup_eq_Max: "finite (X::'a::conditional_complete_linorder set) \ X \ {} \ Sup X = Max X" + using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp + +lemma cInf_eq_Min: "finite (X::'a::conditional_complete_linorder set) \ X \ {} \ Inf X = Min X" + using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp + +lemma cSup_lessThan[simp]: "Sup {.. Sup {y<.. Sup {y.. Inf {y<..x::'a::{conditional_complete_linorder, dense_linorder}} = y" + by (auto intro!: cInf_eq intro: dense_ge_bounded) + +lemma cInf_greaterThanLessThan[simp]: "y < x \ Inf {y<.. LEAST z::real. \x\X. x\z" + +definition + Inf_real_def: "Inf (X::real set) \ - Sup (uminus ` X)" + +instance +proof + { fix z x :: real and X :: "real set" + assume x: "x \ X" and z: "!!x. x \ X \ x \ z" + show "x \ Sup X" + proof (auto simp add: Sup_real_def) + from complete_real[of X] + obtain s where s: "(\y\X. y \ s) & (\z. ((\y\X. y \ z) --> s \ z))" + by (blast intro: x z) + hence "x \ s" + by (blast intro: x z) + also with s have "... = (LEAST z. \x\X. x \ z)" + by (fast intro: Least_equality [symmetric]) + finally show "x \ (LEAST z. \x\X. x \ z)" . + qed } + note Sup_upper = this -instance .. + { fix z :: real and X :: "real set" + assume x: "X \ {}" + and z: "\x. x \ X \ x \ z" + show "Sup X \ z" + proof (auto simp add: Sup_real_def) + from complete_real x + obtain s where s: "(\y\X. y \ s) & (\z. ((\y\X. y \ z) --> s \ z))" + by (blast intro: z) + hence "(LEAST z. \x\X. x \ z) = s" + by (best intro: Least_equality) + also with s z have "... \ z" + by blast + finally show "(LEAST z. \x\X. x \ z) \ z" . + qed } + note Sup_least = this + + { fix x z :: real and X :: "real set" + assume x: "x \ X" and z: "!!x. x \ X \ z \ x" + show "Inf X \ x" + proof - + have "-x \ Sup (uminus ` X)" + by (rule Sup_upper[of _ _ "- z"]) (auto simp add: image_iff x z) + thus ?thesis + by (auto simp add: Inf_real_def) + qed } + + { fix z :: real and X :: "real set" + assume x: "X \ {}" + and z: "\x. x \ X \ z \ x" + show "z \ Inf X" + proof - + have "Sup (uminus ` X) \ -z" + using x z by (force intro: Sup_least) + hence "z \ - Sup (uminus ` X)" + by simp + thus ?thesis + by (auto simp add: Inf_real_def) + qed } +qed end subsection{*Supremum of a set of reals*} -lemma Sup_upper [intro]: (*REAL_SUP_UBOUND in HOL4*) - fixes x :: real - assumes x: "x \ X" - and z: "!!x. x \ X \ x \ z" - shows "x \ Sup X" -proof (auto simp add: Sup_real_def) - from complete_real - obtain s where s: "(\y\X. y \ s) & (\z. ((\y\X. y \ z) --> s \ z))" - by (blast intro: x z) - hence "x \ s" - by (blast intro: x z) - also with s have "... = (LEAST z. \x\X. x \ z)" - by (fast intro: Least_equality [symmetric]) - finally show "x \ (LEAST z. \x\X. x \ z)" . -qed - -lemma Sup_least [intro]: (*REAL_IMP_SUP_LE in HOL4*) - fixes z :: real - assumes x: "X \ {}" - and z: "\x. x \ X \ x \ z" - shows "Sup X \ z" -proof (auto simp add: Sup_real_def) - from complete_real x - obtain s where s: "(\y\X. y \ s) & (\z. ((\y\X. y \ z) --> s \ z))" - by (blast intro: z) - hence "(LEAST z. \x\X. x \ z) = s" - by (best intro: Least_equality) - also with s z have "... \ z" - by blast - finally show "(LEAST z. \x\X. x \ z) \ z" . -qed - -lemma less_SupE: - fixes y :: real - assumes "y < Sup X" "X \ {}" - obtains x where "x\X" "y < x" -by (metis SupInf.Sup_least assms linorder_not_less that) - -lemma Sup_singleton [simp]: "Sup {x::real} = x" - by (force intro: Least_equality simp add: Sup_real_def) - -lemma Sup_eq_maximum: (*REAL_SUP_MAX in HOL4*) - fixes z :: real - assumes X: "z \ X" and z: "!!x. x \ X \ x \ z" - shows "Sup X = z" - by (force intro: Least_equality X z simp add: Sup_real_def) - -lemma Sup_upper2: (*REAL_IMP_LE_SUP in HOL4*) - fixes x :: real - shows "x \ X \ y \ x \ (!!x. x \ X \ x \ z) \ y \ Sup X" - by (metis Sup_upper order_trans) - -lemma Sup_real_iff : (*REAL_SUP_LE in HOL4*) - fixes z :: real - shows "X ~= {} ==> (!!x. x \ X ==> x \ z) ==> (\x\X. y y < Sup X" - by (rule iffI, metis less_le_trans Sup_upper, metis Sup_least not_less) - -lemma Sup_eq: - fixes a :: real - shows "(!!x. x \ X \ x \ a) - \ (!!y. (!!x. x \ X \ x \ y) \ a \ y) \ Sup X = a" - by (metis Sup_least Sup_upper add_le_cancel_left diff_add_cancel insert_absorb - insert_not_empty order_antisym) - -lemma Sup_le: - fixes S :: "real set" - shows "S \ {} \ S *<= b \ Sup S \ b" -by (metis SupInf.Sup_least setle_def) - -lemma Sup_upper_EX: - fixes x :: real - shows "x \ X \ \z. \x. x \ X \ x \ z \ x \ Sup X" - by blast - -lemma Sup_insert_nonempty: - fixes x :: real - assumes x: "x \ X" - and z: "!!x. x \ X \ x \ z" - shows "Sup (insert a X) = max a (Sup X)" -proof (cases "Sup X \ a") - case True - thus ?thesis - apply (simp add: max_def) - apply (rule Sup_eq_maximum) - apply (rule insertI1) - apply (metis Sup_upper [OF _ z] insertE linear order_trans) - done -next - case False - hence 1:"a < Sup X" by simp - have "Sup X \ Sup (insert a X)" - apply (rule Sup_least) - apply (metis ex_in_conv x) - apply (rule Sup_upper_EX) - apply blast - apply (metis insert_iff linear order_trans z) - done - moreover - have "Sup (insert a X) \ Sup X" - apply (rule Sup_least) - apply blast - apply (metis False Sup_upper insertE linear z) - done - ultimately have "Sup (insert a X) = Sup X" - by (blast intro: antisym ) - thus ?thesis - by (metis 1 min_max.le_iff_sup less_le) -qed - -lemma Sup_insert_if: - fixes X :: "real set" - assumes z: "!!x. x \ X \ x \ z" - shows "Sup (insert a X) = (if X={} then a else max a (Sup X))" -by auto (metis Sup_insert_nonempty z) - -lemma Sup: - fixes S :: "real set" - shows "S \ {} \ (\b. S *<= b) \ isLub UNIV S (Sup S)" -by (auto simp add: isLub_def setle_def leastP_def isUb_def intro!: setgeI) - -lemma Sup_finite_Max: - fixes S :: "real set" - assumes fS: "finite S" and Se: "S \ {}" - shows "Sup S = Max S" -using fS Se -proof- - let ?m = "Max S" - from Max_ge[OF fS] have Sm: "\ x\ S. x \ ?m" by metis - with Sup[OF Se] have lub: "isLub UNIV S (Sup S)" by (metis setle_def) - from Max_in[OF fS Se] lub have mrS: "?m \ Sup S" - by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) - moreover - have "Sup S \ ?m" using Sm lub - by (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) - ultimately show ?thesis by arith -qed - -lemma Sup_finite_in: - fixes S :: "real set" - assumes fS: "finite S" and Se: "S \ {}" - shows "Sup S \ S" - using Sup_finite_Max[OF fS Se] Max_in[OF fS Se] by metis - -lemma Sup_finite_ge_iff: - fixes S :: "real set" - assumes fS: "finite S" and Se: "S \ {}" - shows "a \ Sup S \ (\ x \ S. a \ x)" -by (metis Max_ge Se Sup_finite_Max Sup_finite_in fS linorder_not_le less_le_trans) - -lemma Sup_finite_le_iff: - fixes S :: "real set" - assumes fS: "finite S" and Se: "S \ {}" - shows "a \ Sup S \ (\ x \ S. a \ x)" -by (metis Max_ge Se Sup_finite_Max Sup_finite_in fS order_trans) - -lemma Sup_finite_gt_iff: - fixes S :: "real set" - assumes fS: "finite S" and Se: "S \ {}" - shows "a < Sup S \ (\ x \ S. a < x)" -by (metis Se Sup_finite_le_iff fS linorder_not_less) - -lemma Sup_finite_lt_iff: - fixes S :: "real set" - assumes fS: "finite S" and Se: "S \ {}" - shows "a > Sup S \ (\ x \ S. a > x)" -by (metis Se Sup_finite_ge_iff fS linorder_not_less) - -lemma Sup_unique: - fixes S :: "real set" - shows "S *<= b \ (\b' < b. \x \ S. b' < x) \ Sup S = b" -unfolding setle_def -apply (rule Sup_eq, auto) -apply (metis linorder_not_less) -done - -lemma Sup_abs_le: +lemma cSup_abs_le: fixes S :: "real set" shows "S \ {} \ (\x\S. \x\ \ a) \ \Sup S\ \ a" -by (auto simp add: abs_le_interval_iff) (metis Sup_upper2) +by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2) -lemma Sup_bounds: +lemma cSup_bounds: fixes S :: "real set" assumes Se: "S \ {}" and l: "a <=* S" and u: "S *<= b" shows "a \ Sup S \ Sup S \ b" proof- - from Sup[OF Se] u have lub: "isLub UNIV S (Sup S)" by blast + from isLub_cSup[OF Se] u have lub: "isLub UNIV S (Sup S)" by blast hence b: "Sup S \ b" using u by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) from Se obtain y where y: "y \ S" by blast @@ -219,203 +350,66 @@ with b show ?thesis by blast qed -lemma Sup_asclose: +lemma cSup_asclose: fixes S :: "real set" assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\Sup S - l\ \ e" proof- have th: "\(x::real) l e. \x - l\ \ e \ l - e \ x \ x \ l + e" by arith - thus ?thesis using S b Sup_bounds[of S "l - e" "l+e"] unfolding th + thus ?thesis using S b cSup_bounds[of S "l - e" "l+e"] unfolding th by (auto simp add: setge_def setle_def) qed -lemma Sup_lessThan[simp]: "Sup {.. Sup {y<.. Sup {y.. Sup {y<..x} = (x::real)" - by (auto intro!: Sup_eq_maximum) - -lemma Sup_atLeastAtMost[simp]: "y \ x \ Sup {y..x} = (x::real)" - by (auto intro!: Sup_eq_maximum) - - subsection{*Infimum of a set of reals*} -lemma Inf_lower [intro]: - fixes z :: real - assumes x: "x \ X" - and z: "!!x. x \ X \ z \ x" - shows "Inf X \ x" -proof - - have "-x \ Sup (uminus ` X)" - by (rule Sup_upper [where z = "-z"]) (auto simp add: image_iff x z) - thus ?thesis - by (auto simp add: Inf_real_def) -qed - -lemma Inf_greatest [intro]: +lemma cInf_greater: fixes z :: real - assumes x: "X \ {}" - and z: "\x. x \ X \ z \ x" - shows "z \ Inf X" -proof - - have "Sup (uminus ` X) \ -z" using x z by force - hence "z \ - Sup (uminus ` X)" - by simp - thus ?thesis - by (auto simp add: Inf_real_def) -qed - -lemma Inf_singleton [simp]: "Inf {x::real} = x" - by (simp add: Inf_real_def) - -lemma Inf_eq_minimum: (*REAL_INF_MIN in HOL4*) - fixes z :: real - assumes x: "z \ X" and z: "!!x. x \ X \ z \ x" - shows "Inf X = z" -proof - - have "Sup (uminus ` X) = -z" using x z - by (force intro: Sup_eq_maximum x z) - thus ?thesis - by (simp add: Inf_real_def) -qed - -lemma Inf_lower2: - fixes x :: real - shows "x \ X \ x \ y \ (!!x. x \ X \ z \ x) \ Inf X \ y" - by (metis Inf_lower order_trans) - -lemma Inf_real_iff: - fixes z :: real - shows "X \ {} \ (!!x. x \ X \ z \ x) \ (\x\X. x Inf X < y" - by (metis Inf_greatest Inf_lower less_le_not_le linear - order_less_le_trans) + shows "X \ {} \ Inf X < z \ \x\X. x < z" + by (metis cInf_less_iff not_leE) -lemma Inf_eq: - fixes a :: real - shows "(!!x. x \ X \ a \ x) \ (!!y. (!!x. x \ X \ y \ x) \ y \ a) \ Inf X = a" - by (metis Inf_greatest Inf_lower add_le_cancel_left diff_add_cancel - insert_absorb insert_not_empty order_antisym) - -lemma Inf_ge: - fixes S :: "real set" - shows "S \ {} \ b <=* S \ Inf S \ b" -by (metis SupInf.Inf_greatest setge_def) - -lemma Inf_lower_EX: - fixes x :: real - shows "x \ X \ \z. \x. x \ X \ z \ x \ Inf X \ x" - by blast - -lemma Inf_insert_nonempty: - fixes x :: real - assumes x: "x \ X" - and z: "!!x. x \ X \ z \ x" - shows "Inf (insert a X) = min a (Inf X)" -proof (cases "a \ Inf X") - case True - thus ?thesis - by (simp add: min_def) - (blast intro: Inf_eq_minimum order_trans z) -next - case False - hence 1:"Inf X < a" by simp - have "Inf (insert a X) \ Inf X" - apply (rule Inf_greatest) - apply (metis ex_in_conv x) - apply (rule Inf_lower_EX) - apply (erule insertI2) - apply (metis insert_iff linear order_trans z) - done - moreover - have "Inf X \ Inf (insert a X)" - apply (rule Inf_greatest) - apply blast - apply (metis False Inf_lower insertE linear z) - done - ultimately have "Inf (insert a X) = Inf X" - by (blast intro: antisym ) - thus ?thesis - by (metis False min_max.inf_absorb2 linear) -qed - -lemma Inf_insert_if: - fixes X :: "real set" - assumes z: "!!x. x \ X \ z \ x" - shows "Inf (insert a X) = (if X={} then a else min a (Inf X))" -by auto (metis Inf_insert_nonempty z) - -lemma Inf_greater: - fixes z :: real - shows "X \ {} \ Inf X < z \ \x \ X. x < z" - by (metis Inf_real_iff not_leE) - -lemma Inf_close: +lemma cInf_close: fixes e :: real shows "X \ {} \ 0 < e \ \x \ X. x < Inf X + e" - by (metis add_strict_increasing add_commute Inf_greater linorder_not_le pos_add_strict) + by (metis add_strict_increasing add_commute cInf_greater linorder_not_le pos_add_strict) -lemma Inf_finite_Min: - fixes S :: "real set" - shows "finite S \ S \ {} \ Inf S = Min S" -by (simp add: Inf_real_def Sup_finite_Max image_image) - -lemma Inf_finite_in: +lemma cInf_finite_in: fixes S :: "real set" assumes fS: "finite S" and Se: "S \ {}" shows "Inf S \ S" - using Inf_finite_Min[OF fS Se] Min_in[OF fS Se] by metis + using cInf_eq_Min[OF fS Se] Min_in[OF fS Se] by metis -lemma Inf_finite_ge_iff: +lemma cInf_finite_ge_iff: fixes S :: "real set" shows "finite S \ S \ {} \ a \ Inf S \ (\ x \ S. a \ x)" -by (metis Inf_finite_Min Inf_finite_in Min_le order_trans) +by (metis cInf_eq_Min cInf_finite_in Min_le order_trans) -lemma Inf_finite_le_iff: +lemma cInf_finite_le_iff: fixes S :: "real set" shows "finite S \ S \ {} \ a \ Inf S \ (\ x \ S. a \ x)" -by (metis Inf_finite_Min Inf_finite_ge_iff Inf_finite_in Min_le - order_antisym linear) +by (metis cInf_eq_Min cInf_finite_ge_iff cInf_finite_in Min_le order_antisym linear) -lemma Inf_finite_gt_iff: +lemma cInf_finite_gt_iff: fixes S :: "real set" shows "finite S \ S \ {} \ a < Inf S \ (\ x \ S. a < x)" -by (metis Inf_finite_le_iff linorder_not_less) +by (metis cInf_finite_le_iff linorder_not_less) -lemma Inf_finite_lt_iff: +lemma cInf_finite_lt_iff: fixes S :: "real set" shows "finite S \ S \ {} \ a > Inf S \ (\ x \ S. a > x)" -by (metis Inf_finite_ge_iff linorder_not_less) +by (metis cInf_finite_ge_iff linorder_not_less) -lemma Inf_unique: - fixes S :: "real set" - shows "b <=* S \ (\b' > b. \x \ S. b' > x) \ Inf S = b" -unfolding setge_def -apply (rule Inf_eq, auto) -apply (metis less_le_not_le linorder_not_less) -done - -lemma Inf_abs_ge: +lemma cInf_abs_ge: fixes S :: "real set" shows "S \ {} \ (\x\S. \x\ \ a) \ \Inf S\ \ a" -by (simp add: Inf_real_def) (rule Sup_abs_le, auto) +by (simp add: Inf_real_def) (rule cSup_abs_le, auto) -lemma Inf_asclose: +lemma cInf_asclose: fixes S :: "real set" assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\Inf S - l\ \ e" proof - have "\- Sup (uminus ` S) - l\ = \Sup (uminus ` S) - (-l)\" by auto also have "... \ e" - apply (rule Sup_asclose) + apply (rule cSup_asclose) apply (auto simp add: S) apply (metis abs_minus_add_cancel b add_commute diff_minus) done @@ -424,73 +418,16 @@ by (simp add: Inf_real_def) qed -lemma Inf_greaterThanLessThan[simp]: "y < x \ Inf {y<.. Inf {y.. Inf {y<..x} = (y::real)" - by (simp add: Inf_real_def) - -lemma Inf_atLeastAtMost[simp]: "y \ x \ Inf {y..x} = (y::real)" - by (simp add: Inf_real_def) - subsection{*Relate max and min to Sup and Inf.*} -lemma real_max_Sup: +lemma real_max_cSup: fixes x :: real shows "max x y = Sup {x,y}" -proof- - have "Sup {x,y} \ max x y" by (simp add: Sup_finite_le_iff) - moreover - have "max x y \ Sup {x,y}" by (simp add: linorder_linear Sup_finite_ge_iff) - ultimately show ?thesis by arith -qed + by (subst cSup_insert[of _ y]) (simp_all add: sup_max) -lemma real_min_Inf: +lemma real_min_cInf: fixes x :: real shows "min x y = Inf {x,y}" -proof- - have "Inf {x,y} \ min x y" by (simp add: linorder_linear Inf_finite_le_iff) - moreover - have "min x y \ Inf {x,y}" by (simp add: Inf_finite_ge_iff) - ultimately show ?thesis by arith -qed - -lemma reals_complete_interval: - fixes a::real and b::real - assumes "a < b" and "P a" and "~P b" - shows "\c. a \ c & c \ b & (\x. a \ x & x < c --> P x) & - (\d. (\x. a \ x & x < d --> P x) --> d \ c)" -proof (rule exI [where x = "Sup {d. \x. a \ x & x < d --> P x}"], auto) - show "a \ Sup {d. \c. a \ c \ c < d \ P c}" - by (rule SupInf.Sup_upper [where z=b], auto) - (metis `a < b` `\ P b` linear less_le) -next - show "Sup {d. \c. a \ c \ c < d \ P c} \ b" - apply (rule SupInf.Sup_least) - apply (auto simp add: ) - apply (metis less_le_not_le) - apply (metis `a x" and lt: "x < Sup {d. \c. a \ c \ c < d \ P c}" - show "P x" - apply (rule less_SupE [OF lt], auto) - apply (metis less_le_not_le) - apply (metis x) - done -next - fix d - assume 0: "\x. a \ x \ x < d \ P x" - thus "d \ Sup {d. \c. a \ c \ c < d \ P c}" - by (rule_tac z="b" in SupInf.Sup_upper, auto) - (metis `a bool" + +class topological_space = "open" + + assumes open_UNIV [simp, intro]: "open UNIV" + assumes open_Int [intro]: "open S \ open T \ open (S \ T)" + assumes open_Union [intro]: "\S\K. open S \ open (\ K)" +begin + +definition + closed :: "'a set \ bool" where + "closed S \ open (- S)" + +lemma open_empty [intro, simp]: "open {}" + using open_Union [of "{}"] by simp + +lemma open_Un [intro]: "open S \ open T \ open (S \ T)" + using open_Union [of "{S, T}"] by simp + +lemma open_UN [intro]: "\x\A. open (B x) \ open (\x\A. B x)" + unfolding SUP_def by (rule open_Union) auto + +lemma open_Inter [intro]: "finite S \ \T\S. open T \ open (\S)" + by (induct set: finite) auto + +lemma open_INT [intro]: "finite A \ \x\A. open (B x) \ open (\x\A. B x)" + unfolding INF_def by (rule open_Inter) auto + +lemma openI: + assumes "\x. x \ S \ \T. open T \ x \ T \ T \ S" + shows "open S" +proof - + have "open (\{T. open T \ T \ S})" by auto + moreover have "\{T. open T \ T \ S} = S" by (auto dest!: assms) + ultimately show "open S" by simp +qed + +lemma closed_empty [intro, simp]: "closed {}" + unfolding closed_def by simp + +lemma closed_Un [intro]: "closed S \ closed T \ closed (S \ T)" + unfolding closed_def by auto + +lemma closed_UNIV [intro, simp]: "closed UNIV" + unfolding closed_def by simp + +lemma closed_Int [intro]: "closed S \ closed T \ closed (S \ T)" + unfolding closed_def by auto + +lemma closed_INT [intro]: "\x\A. closed (B x) \ closed (\x\A. B x)" + unfolding closed_def by auto + +lemma closed_Inter [intro]: "\S\K. closed S \ closed (\ K)" + unfolding closed_def uminus_Inf by auto + +lemma closed_Union [intro]: "finite S \ \T\S. closed T \ closed (\S)" + by (induct set: finite) auto + +lemma closed_UN [intro]: "finite A \ \x\A. closed (B x) \ closed (\x\A. B x)" + unfolding SUP_def by (rule closed_Union) auto + +lemma open_closed: "open S \ closed (- S)" + unfolding closed_def by simp + +lemma closed_open: "closed S \ open (- S)" + unfolding closed_def by simp + +lemma open_Diff [intro]: "open S \ closed T \ open (S - T)" + unfolding closed_open Diff_eq by (rule open_Int) + +lemma closed_Diff [intro]: "closed S \ open T \ closed (S - T)" + unfolding open_closed Diff_eq by (rule closed_Int) + +lemma open_Compl [intro]: "closed S \ open (- S)" + unfolding closed_open . + +lemma closed_Compl [intro]: "open S \ closed (- S)" + unfolding open_closed . + +end + +subsection{* Hausdorff and other separation properties *} + +class t0_space = topological_space + + assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" + +class t1_space = topological_space + + assumes t1_space: "x \ y \ \U. open U \ x \ U \ y \ U" + +instance t1_space \ t0_space +proof qed (fast dest: t1_space) + +lemma separation_t1: + fixes x y :: "'a::t1_space" + shows "x \ y \ (\U. open U \ x \ U \ y \ U)" + using t1_space[of x y] by blast + +lemma closed_singleton: + fixes a :: "'a::t1_space" + shows "closed {a}" +proof - + let ?T = "\{S. open S \ a \ S}" + have "open ?T" by (simp add: open_Union) + also have "?T = - {a}" + by (simp add: set_eq_iff separation_t1, auto) + finally show "closed {a}" unfolding closed_def . +qed + +lemma closed_insert [simp]: + fixes a :: "'a::t1_space" + assumes "closed S" shows "closed (insert a S)" +proof - + from closed_singleton assms + have "closed ({a} \ S)" by (rule closed_Un) + thus "closed (insert a S)" by simp +qed + +lemma finite_imp_closed: + fixes S :: "'a::t1_space set" + shows "finite S \ closed S" +by (induct set: finite, simp_all) + +text {* T2 spaces are also known as Hausdorff spaces. *} + +class t2_space = topological_space + + assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" + +instance t2_space \ t1_space +proof qed (fast dest: hausdorff) + +lemma separation_t2: + fixes x y :: "'a::t2_space" + shows "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" + using hausdorff[of x y] by blast + +lemma separation_t0: + fixes x y :: "'a::t0_space" + shows "x \ y \ (\U. open U \ ~(x\U \ y\U))" + using t0_space[of x y] by blast + +text {* A perfect space is a topological space with no isolated points. *} + +class perfect_space = topological_space + + assumes not_open_singleton: "\ open {x}" + + +subsection {* Generators for toplogies *} + +inductive generate_topology for S where + UNIV: "generate_topology S UNIV" +| Int: "generate_topology S a \ generate_topology S b \ generate_topology S (a \ b)" +| UN: "(\k. k \ K \ generate_topology S k) \ generate_topology S (\K)" +| Basis: "s \ S \ generate_topology S s" + +hide_fact (open) UNIV Int UN Basis + +lemma generate_topology_Union: + "(\k. k \ I \ generate_topology S (K k)) \ generate_topology S (\k\I. K k)" + unfolding SUP_def by (intro generate_topology.UN) auto + +lemma topological_space_generate_topology: + "class.topological_space (generate_topology S)" + by default (auto intro: generate_topology.intros) + +subsection {* Order topologies *} + +class order_topology = order + "open" + + assumes open_generated_order: "open = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" +begin + +subclass topological_space + unfolding open_generated_order + by (rule topological_space_generate_topology) + +lemma open_greaterThan [simp]: "open {a <..}" + unfolding open_generated_order by (auto intro: generate_topology.Basis) + +lemma open_lessThan [simp]: "open {..< a}" + unfolding open_generated_order by (auto intro: generate_topology.Basis) + +lemma open_greaterThanLessThan [simp]: "open {a <..< b}" + unfolding greaterThanLessThan_eq by (simp add: open_Int) + +end + +class linorder_topology = linorder + order_topology + +lemma closed_atMost [simp]: "closed {.. a::'a::linorder_topology}" + by (simp add: closed_open) + +lemma closed_atLeast [simp]: "closed {a::'a::linorder_topology ..}" + by (simp add: closed_open) + +lemma closed_atLeastAtMost [simp]: "closed {a::'a::linorder_topology .. b}" +proof - + have "{a .. b} = {a ..} \ {.. b}" + by auto + then show ?thesis + by (simp add: closed_Int) +qed + +lemma (in linorder) less_separate: + assumes "x < y" + shows "\a b. x \ {..< a} \ y \ {b <..} \ {..< a} \ {b <..} = {}" +proof cases + assume "\z. x < z \ z < y" + then guess z .. + then have "x \ {..< z} \ y \ {z <..} \ {z <..} \ {..< z} = {}" + by auto + then show ?thesis by blast +next + assume "\ (\z. x < z \ z < y)" + with `x < y` have "x \ {..< y} \ y \ {x <..} \ {x <..} \ {..< y} = {}" + by auto + then show ?thesis by blast +qed + +instance linorder_topology \ t2_space +proof + fix x y :: 'a + from less_separate[of x y] less_separate[of y x] + show "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" + by (elim neqE) (metis open_lessThan open_greaterThan Int_commute)+ +qed + +lemma (in linorder_topology) open_right: + assumes "open S" "x \ S" and gt_ex: "x < y" shows "\b>x. {x ..< b} \ S" + using assms unfolding open_generated_order +proof induction + case (Int A B) + then obtain a b where "a > x" "{x ..< a} \ A" "b > x" "{x ..< b} \ B" by auto + then show ?case by (auto intro!: exI[of _ "min a b"]) +next + case (Basis S) then show ?case by (fastforce intro: exI[of _ y] gt_ex) +qed blast+ + +lemma (in linorder_topology) open_left: + assumes "open S" "x \ S" and lt_ex: "y < x" shows "\b S" + using assms unfolding open_generated_order +proof induction + case (Int A B) + then obtain a b where "a < x" "{a <.. x} \ A" "b < x" "{b <.. x} \ B" by auto + then show ?case by (auto intro!: exI[of _ "max a b"]) +next + case (Basis S) then show ?case by (fastforce intro: exI[of _ y] lt_ex) +qed blast+ + +subsection {* Filters *} + +text {* + This definition also allows non-proper filters. +*} + +locale is_filter = + fixes F :: "('a \ bool) \ bool" + assumes True: "F (\x. True)" + assumes conj: "F (\x. P x) \ F (\x. Q x) \ F (\x. P x \ Q x)" + assumes mono: "\x. P x \ Q x \ F (\x. P x) \ F (\x. Q x)" + +typedef 'a filter = "{F :: ('a \ bool) \ bool. is_filter F}" +proof + show "(\x. True) \ ?filter" by (auto intro: is_filter.intro) +qed + +lemma is_filter_Rep_filter: "is_filter (Rep_filter F)" + using Rep_filter [of F] by simp + +lemma Abs_filter_inverse': + assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F" + using assms by (simp add: Abs_filter_inverse) + + +subsubsection {* Eventually *} + +definition eventually :: "('a \ bool) \ 'a filter \ bool" + where "eventually P F \ Rep_filter F P" + +lemma eventually_Abs_filter: + assumes "is_filter F" shows "eventually P (Abs_filter F) = F P" + unfolding eventually_def using assms by (simp add: Abs_filter_inverse) + +lemma filter_eq_iff: + shows "F = F' \ (\P. eventually P F = eventually P F')" + unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def .. + +lemma eventually_True [simp]: "eventually (\x. True) F" + unfolding eventually_def + by (rule is_filter.True [OF is_filter_Rep_filter]) + +lemma always_eventually: "\x. P x \ eventually P F" +proof - + assume "\x. P x" hence "P = (\x. True)" by (simp add: ext) + thus "eventually P F" by simp +qed + +lemma eventually_mono: + "(\x. P x \ Q x) \ eventually P F \ eventually Q F" + unfolding eventually_def + by (rule is_filter.mono [OF is_filter_Rep_filter]) + +lemma eventually_conj: + assumes P: "eventually (\x. P x) F" + assumes Q: "eventually (\x. Q x) F" + shows "eventually (\x. P x \ Q x) F" + using assms unfolding eventually_def + by (rule is_filter.conj [OF is_filter_Rep_filter]) + +lemma eventually_Ball_finite: + assumes "finite A" and "\y\A. eventually (\x. P x y) net" + shows "eventually (\x. \y\A. P x y) net" +using assms by (induct set: finite, simp, simp add: eventually_conj) + +lemma eventually_all_finite: + fixes P :: "'a \ 'b::finite \ bool" + assumes "\y. eventually (\x. P x y) net" + shows "eventually (\x. \y. P x y) net" +using eventually_Ball_finite [of UNIV P] assms by simp + +lemma eventually_mp: + assumes "eventually (\x. P x \ Q x) F" + assumes "eventually (\x. P x) F" + shows "eventually (\x. Q x) F" +proof (rule eventually_mono) + show "\x. (P x \ Q x) \ P x \ Q x" by simp + show "eventually (\x. (P x \ Q x) \ P x) F" + using assms by (rule eventually_conj) +qed + +lemma eventually_rev_mp: + assumes "eventually (\x. P x) F" + assumes "eventually (\x. P x \ Q x) F" + shows "eventually (\x. Q x) F" +using assms(2) assms(1) by (rule eventually_mp) + +lemma eventually_conj_iff: + "eventually (\x. P x \ Q x) F \ eventually P F \ eventually Q F" + by (auto intro: eventually_conj elim: eventually_rev_mp) + +lemma eventually_elim1: + assumes "eventually (\i. P i) F" + assumes "\i. P i \ Q i" + shows "eventually (\i. Q i) F" + using assms by (auto elim!: eventually_rev_mp) + +lemma eventually_elim2: + assumes "eventually (\i. P i) F" + assumes "eventually (\i. Q i) F" + assumes "\i. P i \ Q i \ R i" + shows "eventually (\i. R i) F" + using assms by (auto elim!: eventually_rev_mp) + +lemma eventually_subst: + assumes "eventually (\n. P n = Q n) F" + shows "eventually P F = eventually Q F" (is "?L = ?R") +proof - + from assms have "eventually (\x. P x \ Q x) F" + and "eventually (\x. Q x \ P x) F" + by (auto elim: eventually_elim1) + then show ?thesis by (auto elim: eventually_elim2) +qed + +ML {* + fun eventually_elim_tac ctxt thms thm = + let + val thy = Proof_Context.theory_of ctxt + val mp_thms = thms RL [@{thm eventually_rev_mp}] + val raw_elim_thm = + (@{thm allI} RS @{thm always_eventually}) + |> fold (fn thm1 => fn thm2 => thm2 RS thm1) mp_thms + |> fold (fn _ => fn thm => @{thm impI} RS thm) thms + val cases_prop = prop_of (raw_elim_thm RS thm) + val cases = (Rule_Cases.make_common (thy, cases_prop) [(("elim", []), [])]) + in + CASES cases (rtac raw_elim_thm 1) thm + end +*} + +method_setup eventually_elim = {* + Scan.succeed (fn ctxt => METHOD_CASES (eventually_elim_tac ctxt)) +*} "elimination of eventually quantifiers" + + +subsubsection {* Finer-than relation *} + +text {* @{term "F \ F'"} means that filter @{term F} is finer than +filter @{term F'}. *} + +instantiation filter :: (type) complete_lattice +begin + +definition le_filter_def: + "F \ F' \ (\P. eventually P F' \ eventually P F)" + +definition + "(F :: 'a filter) < F' \ F \ F' \ \ F' \ F" + +definition + "top = Abs_filter (\P. \x. P x)" + +definition + "bot = Abs_filter (\P. True)" + +definition + "sup F F' = Abs_filter (\P. eventually P F \ eventually P F')" + +definition + "inf F F' = Abs_filter + (\P. \Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" + +definition + "Sup S = Abs_filter (\P. \F\S. eventually P F)" + +definition + "Inf S = Sup {F::'a filter. \F'\S. F \ F'}" + +lemma eventually_top [simp]: "eventually P top \ (\x. P x)" + unfolding top_filter_def + by (rule eventually_Abs_filter, rule is_filter.intro, auto) + +lemma eventually_bot [simp]: "eventually P bot" + unfolding bot_filter_def + by (subst eventually_Abs_filter, rule is_filter.intro, auto) + +lemma eventually_sup: + "eventually P (sup F F') \ eventually P F \ eventually P F'" + unfolding sup_filter_def + by (rule eventually_Abs_filter, rule is_filter.intro) + (auto elim!: eventually_rev_mp) + +lemma eventually_inf: + "eventually P (inf F F') \ + (\Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" + unfolding inf_filter_def + apply (rule eventually_Abs_filter, rule is_filter.intro) + apply (fast intro: eventually_True) + apply clarify + apply (intro exI conjI) + apply (erule (1) eventually_conj) + apply (erule (1) eventually_conj) + apply simp + apply auto + done + +lemma eventually_Sup: + "eventually P (Sup S) \ (\F\S. eventually P F)" + unfolding Sup_filter_def + apply (rule eventually_Abs_filter, rule is_filter.intro) + apply (auto intro: eventually_conj elim!: eventually_rev_mp) + done + +instance proof + fix F F' F'' :: "'a filter" and S :: "'a filter set" + { show "F < F' \ F \ F' \ \ F' \ F" + by (rule less_filter_def) } + { show "F \ F" + unfolding le_filter_def by simp } + { assume "F \ F'" and "F' \ F''" thus "F \ F''" + unfolding le_filter_def by simp } + { assume "F \ F'" and "F' \ F" thus "F = F'" + unfolding le_filter_def filter_eq_iff by fast } + { show "F \ top" + unfolding le_filter_def eventually_top by (simp add: always_eventually) } + { show "bot \ F" + unfolding le_filter_def by simp } + { show "F \ sup F F'" and "F' \ sup F F'" + unfolding le_filter_def eventually_sup by simp_all } + { assume "F \ F''" and "F' \ F''" thus "sup F F' \ F''" + unfolding le_filter_def eventually_sup by simp } + { show "inf F F' \ F" and "inf F F' \ F'" + unfolding le_filter_def eventually_inf by (auto intro: eventually_True) } + { assume "F \ F'" and "F \ F''" thus "F \ inf F' F''" + unfolding le_filter_def eventually_inf + by (auto elim!: eventually_mono intro: eventually_conj) } + { assume "F \ S" thus "F \ Sup S" + unfolding le_filter_def eventually_Sup by simp } + { assume "\F. F \ S \ F \ F'" thus "Sup S \ F'" + unfolding le_filter_def eventually_Sup by simp } + { assume "F'' \ S" thus "Inf S \ F''" + unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } + { assume "\F'. F' \ S \ F \ F'" thus "F \ Inf S" + unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } +qed + +end + +lemma filter_leD: + "F \ F' \ eventually P F' \ eventually P F" + unfolding le_filter_def by simp + +lemma filter_leI: + "(\P. eventually P F' \ eventually P F) \ F \ F'" + unfolding le_filter_def by simp + +lemma eventually_False: + "eventually (\x. False) F \ F = bot" + unfolding filter_eq_iff by (auto elim: eventually_rev_mp) + +abbreviation (input) trivial_limit :: "'a filter \ bool" + where "trivial_limit F \ F = bot" + +lemma trivial_limit_def: "trivial_limit F \ eventually (\x. False) F" + by (rule eventually_False [symmetric]) + +lemma eventually_const: "\ trivial_limit net \ eventually (\x. P) net \ P" + by (cases P) (simp_all add: eventually_False) + + +subsubsection {* Map function for filters *} + +definition filtermap :: "('a \ 'b) \ 'a filter \ 'b filter" + where "filtermap f F = Abs_filter (\P. eventually (\x. P (f x)) F)" + +lemma eventually_filtermap: + "eventually P (filtermap f F) = eventually (\x. P (f x)) F" + unfolding filtermap_def + apply (rule eventually_Abs_filter) + apply (rule is_filter.intro) + apply (auto elim!: eventually_rev_mp) + done + +lemma filtermap_ident: "filtermap (\x. x) F = F" + by (simp add: filter_eq_iff eventually_filtermap) + +lemma filtermap_filtermap: + "filtermap f (filtermap g F) = filtermap (\x. f (g x)) F" + by (simp add: filter_eq_iff eventually_filtermap) + +lemma filtermap_mono: "F \ F' \ filtermap f F \ filtermap f F'" + unfolding le_filter_def eventually_filtermap by simp + +lemma filtermap_bot [simp]: "filtermap f bot = bot" + by (simp add: filter_eq_iff eventually_filtermap) + +lemma filtermap_sup: "filtermap f (sup F1 F2) = sup (filtermap f F1) (filtermap f F2)" + by (auto simp: filter_eq_iff eventually_filtermap eventually_sup) + +subsubsection {* Order filters *} + +definition at_top :: "('a::order) filter" + where "at_top = Abs_filter (\P. \k. \n\k. P n)" + +lemma eventually_at_top_linorder: "eventually P at_top \ (\N::'a::linorder. \n\N. P n)" + unfolding at_top_def +proof (rule eventually_Abs_filter, rule is_filter.intro) + fix P Q :: "'a \ bool" + assume "\i. \n\i. P n" and "\j. \n\j. Q n" + then obtain i j where "\n\i. P n" and "\n\j. Q n" by auto + then have "\n\max i j. P n \ Q n" by simp + then show "\k. \n\k. P n \ Q n" .. +qed auto + +lemma eventually_ge_at_top: + "eventually (\x. (c::_::linorder) \ x) at_top" + unfolding eventually_at_top_linorder by auto + +lemma eventually_at_top_dense: "eventually P at_top \ (\N::'a::dense_linorder. \n>N. P n)" + unfolding eventually_at_top_linorder +proof safe + fix N assume "\n\N. P n" then show "\N. \n>N. P n" by (auto intro!: exI[of _ N]) +next + fix N assume "\n>N. P n" + moreover from gt_ex[of N] guess y .. + ultimately show "\N. \n\N. P n" by (auto intro!: exI[of _ y]) +qed + +lemma eventually_gt_at_top: + "eventually (\x. (c::_::dense_linorder) < x) at_top" + unfolding eventually_at_top_dense by auto + +definition at_bot :: "('a::order) filter" + where "at_bot = Abs_filter (\P. \k. \n\k. P n)" + +lemma eventually_at_bot_linorder: + fixes P :: "'a::linorder \ bool" shows "eventually P at_bot \ (\N. \n\N. P n)" + unfolding at_bot_def +proof (rule eventually_Abs_filter, rule is_filter.intro) + fix P Q :: "'a \ bool" + assume "\i. \n\i. P n" and "\j. \n\j. Q n" + then obtain i j where "\n\i. P n" and "\n\j. Q n" by auto + then have "\n\min i j. P n \ Q n" by simp + then show "\k. \n\k. P n \ Q n" .. +qed auto + +lemma eventually_le_at_bot: + "eventually (\x. x \ (c::_::linorder)) at_bot" + unfolding eventually_at_bot_linorder by auto + +lemma eventually_at_bot_dense: + fixes P :: "'a::dense_linorder \ bool" shows "eventually P at_bot \ (\N. \nn\N. P n" then show "\N. \nnN. \n\N. P n" by (auto intro!: exI[of _ y]) +qed + +lemma eventually_gt_at_bot: + "eventually (\x. x < (c::_::dense_linorder)) at_bot" + unfolding eventually_at_bot_dense by auto + +subsection {* Sequentially *} + +abbreviation sequentially :: "nat filter" + where "sequentially == at_top" + +lemma sequentially_def: "sequentially = Abs_filter (\P. \k. \n\k. P n)" + unfolding at_top_def by simp + +lemma eventually_sequentially: + "eventually P sequentially \ (\N. \n\N. P n)" + by (rule eventually_at_top_linorder) + +lemma sequentially_bot [simp, intro]: "sequentially \ bot" + unfolding filter_eq_iff eventually_sequentially by auto + +lemmas trivial_limit_sequentially = sequentially_bot + +lemma eventually_False_sequentially [simp]: + "\ eventually (\n. False) sequentially" + by (simp add: eventually_False) + +lemma le_sequentially: + "F \ sequentially \ (\N. eventually (\n. N \ n) F)" + unfolding le_filter_def eventually_sequentially + by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp) + +lemma eventually_sequentiallyI: + assumes "\x. c \ x \ P x" + shows "eventually P sequentially" +using assms by (auto simp: eventually_sequentially) + +lemma eventually_sequentially_seg: + "eventually (\n. P (n + k)) sequentially \ eventually P sequentially" + unfolding eventually_sequentially + apply safe + apply (rule_tac x="N + k" in exI) + apply rule + apply (erule_tac x="n - k" in allE) + apply auto [] + apply (rule_tac x=N in exI) + apply auto [] + done + +subsubsection {* Standard filters *} + +definition within :: "'a filter \ 'a set \ 'a filter" (infixr "within" 70) + where "F within S = Abs_filter (\P. eventually (\x. x \ S \ P x) F)" + +lemma eventually_within: + "eventually P (F within S) = eventually (\x. x \ S \ P x) F" + unfolding within_def + by (rule eventually_Abs_filter, rule is_filter.intro) + (auto elim!: eventually_rev_mp) + +lemma within_UNIV [simp]: "F within UNIV = F" + unfolding filter_eq_iff eventually_within by simp + +lemma within_empty [simp]: "F within {} = bot" + unfolding filter_eq_iff eventually_within by simp + +lemma within_within_eq: "(F within S) within T = F within (S \ T)" + by (auto simp: filter_eq_iff eventually_within elim: eventually_elim1) + +lemma within_le: "F within S \ F" + unfolding le_filter_def eventually_within by (auto elim: eventually_elim1) + +lemma le_withinI: "F \ F' \ eventually (\x. x \ S) F \ F \ F' within S" + unfolding le_filter_def eventually_within by (auto elim: eventually_elim2) + +lemma le_within_iff: "eventually (\x. x \ S) F \ F \ F' within S \ F \ F'" + by (blast intro: within_le le_withinI order_trans) + +subsubsection {* Topological filters *} + +definition (in topological_space) nhds :: "'a \ 'a filter" + where "nhds a = Abs_filter (\P. \S. open S \ a \ S \ (\x\S. P x))" + +definition (in topological_space) at :: "'a \ 'a filter" + where "at a = nhds a within - {a}" + +abbreviation (in order_topology) at_right :: "'a \ 'a filter" where + "at_right x \ at x within {x <..}" + +abbreviation (in order_topology) at_left :: "'a \ 'a filter" where + "at_left x \ at x within {..< x}" + +lemma (in topological_space) eventually_nhds: + "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" + unfolding nhds_def +proof (rule eventually_Abs_filter, rule is_filter.intro) + have "open UNIV \ a \ UNIV \ (\x\UNIV. True)" by simp + thus "\S. open S \ a \ S \ (\x\S. True)" .. +next + fix P Q + assume "\S. open S \ a \ S \ (\x\S. P x)" + and "\T. open T \ a \ T \ (\x\T. Q x)" + then obtain S T where + "open S \ a \ S \ (\x\S. P x)" + "open T \ a \ T \ (\x\T. Q x)" by auto + hence "open (S \ T) \ a \ S \ T \ (\x\(S \ T). P x \ Q x)" + by (simp add: open_Int) + thus "\S. open S \ a \ S \ (\x\S. P x \ Q x)" .. +qed auto + +lemma nhds_neq_bot [simp]: "nhds a \ bot" + unfolding trivial_limit_def eventually_nhds by simp + +lemma eventually_at_topological: + "eventually P (at a) \ (\S. open S \ a \ S \ (\x\S. x \ a \ P x))" +unfolding at_def eventually_within eventually_nhds by simp + +lemma at_within_open: "a \ S \ open S \ at a within S = at a" + unfolding filter_eq_iff eventually_within eventually_at_topological by (metis open_Int Int_iff) + +lemma at_eq_bot_iff: "at a = bot \ open {a}" + unfolding trivial_limit_def eventually_at_topological + by (safe, case_tac "S = {a}", simp, fast, fast) + +lemma at_neq_bot [simp]: "at (a::'a::perfect_space) \ bot" + by (simp add: at_eq_bot_iff not_open_singleton) + +lemma eventually_at_right: + fixes x :: "'a :: {no_top, linorder_topology}" + shows "eventually P (at_right x) \ (\b. x < b \ (\z. x < z \ z < b \ P z))" + unfolding eventually_nhds eventually_within at_def +proof safe + from gt_ex[of x] guess y .. + moreover fix S assume "open S" "x \ S" note open_right[OF this, of y] + moreover note gt_ex[of x] + moreover assume "\s\S. s \ - {x} \ s \ {x<..} \ P s" + ultimately show "\b>x. \z. x < z \ z < b \ P z" + by (auto simp: subset_eq Ball_def) +next + fix b assume "x < b" "\z. x < z \ z < b \ P z" + then show "\S. open S \ x \ S \ (\xa\S. xa \ - {x} \ xa \ {x<..} \ P xa)" + by (intro exI[of _ "{..< b}"]) auto +qed + +lemma eventually_at_left: + fixes x :: "'a :: {no_bot, linorder_topology}" + shows "eventually P (at_left x) \ (\b. x > b \ (\z. b < z \ z < x \ P z))" + unfolding eventually_nhds eventually_within at_def +proof safe + from lt_ex[of x] guess y .. + moreover fix S assume "open S" "x \ S" note open_left[OF this, of y] + moreover assume "\s\S. s \ - {x} \ s \ {.. P s" + ultimately show "\bz. b < z \ z < x \ P z" + by (auto simp: subset_eq Ball_def) +next + fix b assume "b < x" "\z. b < z \ z < x \ P z" + then show "\S. open S \ x \ S \ (\xa\S. xa \ - {x} \ xa \ {.. P xa)" + by (intro exI[of _ "{b <..}"]) auto +qed + +lemma trivial_limit_at_left_real [simp]: + "\ trivial_limit (at_left (x::'a::{no_bot, dense_linorder, linorder_topology}))" + unfolding trivial_limit_def eventually_at_left by (auto dest: dense) + +lemma trivial_limit_at_right_real [simp]: + "\ trivial_limit (at_right (x::'a::{no_top, dense_linorder, linorder_topology}))" + unfolding trivial_limit_def eventually_at_right by (auto dest: dense) + +lemma at_within_eq: "at x within T = nhds x within (T - {x})" + unfolding at_def within_within_eq by (simp add: ac_simps Diff_eq) + +lemma at_eq_sup_left_right: "at (x::'a::linorder_topology) = sup (at_left x) (at_right x)" + by (auto simp: eventually_within at_def filter_eq_iff eventually_sup + elim: eventually_elim2 eventually_elim1) + +lemma eventually_at_split: + "eventually P (at (x::'a::linorder_topology)) \ eventually P (at_left x) \ eventually P (at_right x)" + by (subst at_eq_sup_left_right) (simp add: eventually_sup) + +subsection {* Limits *} + +definition filterlim :: "('a \ 'b) \ 'b filter \ 'a filter \ bool" where + "filterlim f F2 F1 \ filtermap f F1 \ F2" + +syntax + "_LIM" :: "pttrns \ 'a \ 'b \ 'a \ bool" ("(3LIM (_)/ (_)./ (_) :> (_))" [1000, 10, 0, 10] 10) + +translations + "LIM x F1. f :> F2" == "CONST filterlim (%x. f) F2 F1" + +lemma filterlim_iff: + "(LIM x F1. f x :> F2) \ (\P. eventually P F2 \ eventually (\x. P (f x)) F1)" + unfolding filterlim_def le_filter_def eventually_filtermap .. + +lemma filterlim_compose: + "filterlim g F3 F2 \ filterlim f F2 F1 \ filterlim (\x. g (f x)) F3 F1" + unfolding filterlim_def filtermap_filtermap[symmetric] by (metis filtermap_mono order_trans) + +lemma filterlim_mono: + "filterlim f F2 F1 \ F2 \ F2' \ F1' \ F1 \ filterlim f F2' F1'" + unfolding filterlim_def by (metis filtermap_mono order_trans) + +lemma filterlim_ident: "LIM x F. x :> F" + by (simp add: filterlim_def filtermap_ident) + +lemma filterlim_cong: + "F1 = F1' \ F2 = F2' \ eventually (\x. f x = g x) F2 \ filterlim f F1 F2 = filterlim g F1' F2'" + by (auto simp: filterlim_def le_filter_def eventually_filtermap elim: eventually_elim2) + +lemma filterlim_within: + "(LIM x F1. f x :> F2 within S) \ (eventually (\x. f x \ S) F1 \ (LIM x F1. f x :> F2))" + unfolding filterlim_def +proof safe + assume "filtermap f F1 \ F2 within S" then show "eventually (\x. f x \ S) F1" + by (auto simp: le_filter_def eventually_filtermap eventually_within elim!: allE[of _ "\x. x \ S"]) +qed (auto intro: within_le order_trans simp: le_within_iff eventually_filtermap) + +lemma filterlim_filtermap: "filterlim f F1 (filtermap g F2) = filterlim (\x. f (g x)) F1 F2" + unfolding filterlim_def filtermap_filtermap .. + +lemma filterlim_sup: + "filterlim f F F1 \ filterlim f F F2 \ filterlim f F (sup F1 F2)" + unfolding filterlim_def filtermap_sup by auto + +lemma filterlim_Suc: "filterlim Suc sequentially sequentially" + by (simp add: filterlim_iff eventually_sequentially) (metis le_Suc_eq) + +subsubsection {* Tendsto *} + +abbreviation (in topological_space) + tendsto :: "('b \ 'a) \ 'a \ 'b filter \ bool" (infixr "--->" 55) where + "(f ---> l) F \ filterlim f (nhds l) F" + +definition (in t2_space) Lim :: "'f filter \ ('f \ 'a) \ 'a" where + "Lim A f = (THE l. (f ---> l) A)" + +lemma tendsto_eq_rhs: "(f ---> x) F \ x = y \ (f ---> y) F" + by simp + +ML {* + +structure Tendsto_Intros = Named_Thms +( + val name = @{binding tendsto_intros} + val description = "introduction rules for tendsto" +) + +*} + +setup {* + Tendsto_Intros.setup #> + Global_Theory.add_thms_dynamic (@{binding tendsto_eq_intros}, + map (fn thm => @{thm tendsto_eq_rhs} OF [thm]) o Tendsto_Intros.get o Context.proof_of); +*} + +lemma (in topological_space) tendsto_def: + "(f ---> l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" + unfolding filterlim_def +proof safe + fix S assume "open S" "l \ S" "filtermap f F \ nhds l" + then show "eventually (\x. f x \ S) F" + unfolding eventually_nhds eventually_filtermap le_filter_def + by (auto elim!: allE[of _ "\x. x \ S"] eventually_rev_mp) +qed (auto elim!: eventually_rev_mp simp: eventually_nhds eventually_filtermap le_filter_def) + +lemma tendsto_within_subset: "(f ---> l) (net within S) \ T \ S \ (f ---> l) (net within T)" + by (auto simp: tendsto_def eventually_within elim!: eventually_elim1) + +lemma filterlim_at: + "(LIM x F. f x :> at b) \ (eventually (\x. f x \ b) F \ (f ---> b) F)" + by (simp add: at_def filterlim_within) + +lemma tendsto_mono: "F \ F' \ (f ---> l) F' \ (f ---> l) F" + unfolding tendsto_def le_filter_def by fast + +lemma (in topological_space) topological_tendstoI: + "(\S. open S \ l \ S \ eventually (\x. f x \ S) F) + \ (f ---> l) F" + unfolding tendsto_def by auto + +lemma (in topological_space) topological_tendstoD: + "(f ---> l) F \ open S \ l \ S \ eventually (\x. f x \ S) F" + unfolding tendsto_def by auto + +lemma order_tendstoI: + fixes y :: "_ :: order_topology" + assumes "\a. a < y \ eventually (\x. a < f x) F" + assumes "\a. y < a \ eventually (\x. f x < a) F" + shows "(f ---> y) F" +proof (rule topological_tendstoI) + fix S assume "open S" "y \ S" + then show "eventually (\x. f x \ S) F" + unfolding open_generated_order + proof induct + case (UN K) + then obtain k where "y \ k" "k \ K" by auto + with UN(2)[of k] show ?case + by (auto elim: eventually_elim1) + qed (insert assms, auto elim: eventually_elim2) +qed + +lemma order_tendstoD: + fixes y :: "_ :: order_topology" + assumes y: "(f ---> y) F" + shows "a < y \ eventually (\x. a < f x) F" + and "y < a \ eventually (\x. f x < a) F" + using topological_tendstoD[OF y, of "{..< a}"] topological_tendstoD[OF y, of "{a <..}"] by auto + +lemma order_tendsto_iff: + fixes f :: "_ \ 'a :: order_topology" + shows "(f ---> x) F \(\lx. l < f x) F) \ (\u>x. eventually (\x. f x < u) F)" + by (metis order_tendstoI order_tendstoD) + +lemma tendsto_bot [simp]: "(f ---> a) bot" + unfolding tendsto_def by simp + +lemma tendsto_ident_at [tendsto_intros]: "((\x. x) ---> a) (at a)" + unfolding tendsto_def eventually_at_topological by auto + +lemma tendsto_ident_at_within [tendsto_intros]: + "((\x. x) ---> a) (at a within S)" + unfolding tendsto_def eventually_within eventually_at_topological by auto + +lemma (in topological_space) tendsto_const [tendsto_intros]: "((\x. k) ---> k) F" + by (simp add: tendsto_def) + +lemma (in t2_space) tendsto_unique: + assumes "\ trivial_limit F" and "(f ---> a) F" and "(f ---> b) F" + shows "a = b" +proof (rule ccontr) + assume "a \ b" + obtain U V where "open U" "open V" "a \ U" "b \ V" "U \ V = {}" + using hausdorff [OF `a \ b`] by fast + have "eventually (\x. f x \ U) F" + using `(f ---> a) F` `open U` `a \ U` by (rule topological_tendstoD) + moreover + have "eventually (\x. f x \ V) F" + using `(f ---> b) F` `open V` `b \ V` by (rule topological_tendstoD) + ultimately + have "eventually (\x. False) F" + proof eventually_elim + case (elim x) + hence "f x \ U \ V" by simp + with `U \ V = {}` show ?case by simp + qed + with `\ trivial_limit F` show "False" + by (simp add: trivial_limit_def) +qed + +lemma (in t2_space) tendsto_const_iff: + assumes "\ trivial_limit F" shows "((\x. a :: 'a) ---> b) F \ a = b" + by (safe intro!: tendsto_const tendsto_unique [OF assms tendsto_const]) + +lemma increasing_tendsto: + fixes f :: "_ \ 'a::order_topology" + assumes bdd: "eventually (\n. f n \ l) F" + and en: "\x. x < l \ eventually (\n. x < f n) F" + shows "(f ---> l) F" + using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) + +lemma decreasing_tendsto: + fixes f :: "_ \ 'a::order_topology" + assumes bdd: "eventually (\n. l \ f n) F" + and en: "\x. l < x \ eventually (\n. f n < x) F" + shows "(f ---> l) F" + using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) + +lemma tendsto_sandwich: + fixes f g h :: "'a \ 'b::order_topology" + assumes ev: "eventually (\n. f n \ g n) net" "eventually (\n. g n \ h n) net" + assumes lim: "(f ---> c) net" "(h ---> c) net" + shows "(g ---> c) net" +proof (rule order_tendstoI) + fix a show "a < c \ eventually (\x. a < g x) net" + using order_tendstoD[OF lim(1), of a] ev by (auto elim: eventually_elim2) +next + fix a show "c < a \ eventually (\x. g x < a) net" + using order_tendstoD[OF lim(2), of a] ev by (auto elim: eventually_elim2) +qed + +lemma tendsto_le: + fixes f g :: "'a \ 'b::linorder_topology" + assumes F: "\ trivial_limit F" + assumes x: "(f ---> x) F" and y: "(g ---> y) F" + assumes ev: "eventually (\x. g x \ f x) F" + shows "y \ x" +proof (rule ccontr) + assume "\ y \ x" + with less_separate[of x y] obtain a b where xy: "x < a" "b < y" "{.. {b<..} = {}" + by (auto simp: not_le) + then have "eventually (\x. f x < a) F" "eventually (\x. b < g x) F" + using x y by (auto intro: order_tendstoD) + with ev have "eventually (\x. False) F" + by eventually_elim (insert xy, fastforce) + with F show False + by (simp add: eventually_False) +qed + +lemma tendsto_le_const: + fixes f :: "'a \ 'b::linorder_topology" + assumes F: "\ trivial_limit F" + assumes x: "(f ---> x) F" and a: "eventually (\x. a \ f x) F" + shows "a \ x" + using F x tendsto_const a by (rule tendsto_le) + +subsubsection {* Rules about @{const Lim} *} + +lemma (in t2_space) tendsto_Lim: + "\(trivial_limit net) \ (f ---> l) net \ Lim net f = l" + unfolding Lim_def using tendsto_unique[of net f] by auto + +lemma Lim_ident_at: "\ trivial_limit (at x) \ Lim (at x) (\x. x) = x" + by (rule tendsto_Lim[OF _ tendsto_ident_at]) auto + +lemma Lim_ident_at_within: "\ trivial_limit (at x within s) \ Lim (at x within s) (\x. x) = x" + by (rule tendsto_Lim[OF _ tendsto_ident_at_within]) auto + +subsection {* Limits to @{const at_top} and @{const at_bot} *} + +lemma filterlim_at_top: + fixes f :: "'a \ ('b::linorder)" + shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z \ f x) F)" + by (auto simp: filterlim_iff eventually_at_top_linorder elim!: eventually_elim1) + +lemma filterlim_at_top_dense: + fixes f :: "'a \ ('b::dense_linorder)" + shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z < f x) F)" + by (metis eventually_elim1[of _ F] eventually_gt_at_top order_less_imp_le + filterlim_at_top[of f F] filterlim_iff[of f at_top F]) + +lemma filterlim_at_top_ge: + fixes f :: "'a \ ('b::linorder)" and c :: "'b" + shows "(LIM x F. f x :> at_top) \ (\Z\c. eventually (\x. Z \ f x) F)" + unfolding filterlim_at_top +proof safe + fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" + with *[THEN spec, of "max Z c"] show "eventually (\x. Z \ f x) F" + by (auto elim!: eventually_elim1) +qed simp + +lemma filterlim_at_top_at_top: + fixes f :: "'a::linorder \ 'b::linorder" + assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" + assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" + assumes Q: "eventually Q at_top" + assumes P: "eventually P at_top" + shows "filterlim f at_top at_top" +proof - + from P obtain x where x: "\y. x \ y \ P y" + unfolding eventually_at_top_linorder by auto + show ?thesis + proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) + fix z assume "x \ z" + with x have "P z" by auto + have "eventually (\x. g z \ x) at_top" + by (rule eventually_ge_at_top) + with Q show "eventually (\x. z \ f x) at_top" + by eventually_elim (metis mono bij `P z`) + qed +qed + +lemma filterlim_at_top_gt: + fixes f :: "'a \ ('b::dense_linorder)" and c :: "'b" + shows "(LIM x F. f x :> at_top) \ (\Z>c. eventually (\x. Z \ f x) F)" + by (metis filterlim_at_top order_less_le_trans gt_ex filterlim_at_top_ge) + +lemma filterlim_at_bot: + fixes f :: "'a \ ('b::linorder)" + shows "(LIM x F. f x :> at_bot) \ (\Z. eventually (\x. f x \ Z) F)" + by (auto simp: filterlim_iff eventually_at_bot_linorder elim!: eventually_elim1) + +lemma filterlim_at_bot_le: + fixes f :: "'a \ ('b::linorder)" and c :: "'b" + shows "(LIM x F. f x :> at_bot) \ (\Z\c. eventually (\x. Z \ f x) F)" + unfolding filterlim_at_bot +proof safe + fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" + with *[THEN spec, of "min Z c"] show "eventually (\x. Z \ f x) F" + by (auto elim!: eventually_elim1) +qed simp + +lemma filterlim_at_bot_lt: + fixes f :: "'a \ ('b::dense_linorder)" and c :: "'b" + shows "(LIM x F. f x :> at_bot) \ (\Zx. Z \ f x) F)" + by (metis filterlim_at_bot filterlim_at_bot_le lt_ex order_le_less_trans) + +lemma filterlim_at_bot_at_right: + fixes f :: "'a::{no_top, linorder_topology} \ 'b::linorder" + assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" + assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" + assumes Q: "eventually Q (at_right a)" and bound: "\b. Q b \ a < b" + assumes P: "eventually P at_bot" + shows "filterlim f at_bot (at_right a)" +proof - + from P obtain x where x: "\y. y \ x \ P y" + unfolding eventually_at_bot_linorder by auto + show ?thesis + proof (intro filterlim_at_bot_le[THEN iffD2] allI impI) + fix z assume "z \ x" + with x have "P z" by auto + have "eventually (\x. x \ g z) (at_right a)" + using bound[OF bij(2)[OF `P z`]] + unfolding eventually_at_right by (auto intro!: exI[of _ "g z"]) + with Q show "eventually (\x. f x \ z) (at_right a)" + by eventually_elim (metis bij `P z` mono) + qed +qed + +lemma filterlim_at_top_at_left: + fixes f :: "'a::{no_bot, linorder_topology} \ 'b::linorder" + assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" + assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" + assumes Q: "eventually Q (at_left a)" and bound: "\b. Q b \ b < a" + assumes P: "eventually P at_top" + shows "filterlim f at_top (at_left a)" +proof - + from P obtain x where x: "\y. x \ y \ P y" + unfolding eventually_at_top_linorder by auto + show ?thesis + proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) + fix z assume "x \ z" + with x have "P z" by auto + have "eventually (\x. g z \ x) (at_left a)" + using bound[OF bij(2)[OF `P z`]] + unfolding eventually_at_left by (auto intro!: exI[of _ "g z"]) + with Q show "eventually (\x. z \ f x) (at_left a)" + by eventually_elim (metis bij `P z` mono) + qed +qed + +lemma filterlim_split_at: + "filterlim f F (at_left x) \ filterlim f F (at_right x) \ filterlim f F (at (x::'a::linorder_topology))" + by (subst at_eq_sup_left_right) (rule filterlim_sup) + +lemma filterlim_at_split: + "filterlim f F (at (x::'a::linorder_topology)) \ filterlim f F (at_left x) \ filterlim f F (at_right x)" + by (subst at_eq_sup_left_right) (simp add: filterlim_def filtermap_sup) + + +subsection {* Limits on sequences *} + +abbreviation (in topological_space) + LIMSEQ :: "[nat \ 'a, 'a] \ bool" + ("((_)/ ----> (_))" [60, 60] 60) where + "X ----> L \ (X ---> L) sequentially" + +abbreviation (in t2_space) lim :: "(nat \ 'a) \ 'a" where + "lim X \ Lim sequentially X" + +definition (in topological_space) convergent :: "(nat \ 'a) \ bool" where + "convergent X = (\L. X ----> L)" + +lemma lim_def: "lim X = (THE L. X ----> L)" + unfolding Lim_def .. + +subsubsection {* Monotone sequences and subsequences *} + +definition + monoseq :: "(nat \ 'a::order) \ bool" where + --{*Definition of monotonicity. + The use of disjunction here complicates proofs considerably. + One alternative is to add a Boolean argument to indicate the direction. + Another is to develop the notions of increasing and decreasing first.*} + "monoseq X = ((\m. \n\m. X m \ X n) | (\m. \n\m. X n \ X m))" + +definition + incseq :: "(nat \ 'a::order) \ bool" where + --{*Increasing sequence*} + "incseq X \ (\m. \n\m. X m \ X n)" + +definition + decseq :: "(nat \ 'a::order) \ bool" where + --{*Decreasing sequence*} + "decseq X \ (\m. \n\m. X n \ X m)" + +definition + subseq :: "(nat \ nat) \ bool" where + --{*Definition of subsequence*} + "subseq f \ (\m. \n>m. f m < f n)" + +lemma incseq_mono: "mono f \ incseq f" + unfolding mono_def incseq_def by auto + +lemma incseq_SucI: + "(\n. X n \ X (Suc n)) \ incseq X" + using lift_Suc_mono_le[of X] + by (auto simp: incseq_def) + +lemma incseqD: "\i j. incseq f \ i \ j \ f i \ f j" + by (auto simp: incseq_def) + +lemma incseq_SucD: "incseq A \ A i \ A (Suc i)" + using incseqD[of A i "Suc i"] by auto + +lemma incseq_Suc_iff: "incseq f \ (\n. f n \ f (Suc n))" + by (auto intro: incseq_SucI dest: incseq_SucD) + +lemma incseq_const[simp, intro]: "incseq (\x. k)" + unfolding incseq_def by auto + +lemma decseq_SucI: + "(\n. X (Suc n) \ X n) \ decseq X" + using order.lift_Suc_mono_le[OF dual_order, of X] + by (auto simp: decseq_def) + +lemma decseqD: "\i j. decseq f \ i \ j \ f j \ f i" + by (auto simp: decseq_def) + +lemma decseq_SucD: "decseq A \ A (Suc i) \ A i" + using decseqD[of A i "Suc i"] by auto + +lemma decseq_Suc_iff: "decseq f \ (\n. f (Suc n) \ f n)" + by (auto intro: decseq_SucI dest: decseq_SucD) + +lemma decseq_const[simp, intro]: "decseq (\x. k)" + unfolding decseq_def by auto + +lemma monoseq_iff: "monoseq X \ incseq X \ decseq X" + unfolding monoseq_def incseq_def decseq_def .. + +lemma monoseq_Suc: + "monoseq X \ (\n. X n \ X (Suc n)) \ (\n. X (Suc n) \ X n)" + unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff .. + +lemma monoI1: "\m. \ n \ m. X m \ X n ==> monoseq X" +by (simp add: monoseq_def) + +lemma monoI2: "\m. \ n \ m. X n \ X m ==> monoseq X" +by (simp add: monoseq_def) + +lemma mono_SucI1: "\n. X n \ X (Suc n) ==> monoseq X" +by (simp add: monoseq_Suc) + +lemma mono_SucI2: "\n. X (Suc n) \ X n ==> monoseq X" +by (simp add: monoseq_Suc) + +lemma monoseq_minus: + fixes a :: "nat \ 'a::ordered_ab_group_add" + assumes "monoseq a" + shows "monoseq (\ n. - a n)" +proof (cases "\ m. \ n \ m. a m \ a n") + case True + hence "\ m. \ n \ m. - a n \ - a m" by auto + thus ?thesis by (rule monoI2) +next + case False + hence "\ m. \ n \ m. - a m \ - a n" using `monoseq a`[unfolded monoseq_def] by auto + thus ?thesis by (rule monoI1) +qed + +text{*Subsequence (alternative definition, (e.g. Hoskins)*} + +lemma subseq_Suc_iff: "subseq f = (\n. (f n) < (f (Suc n)))" +apply (simp add: subseq_def) +apply (auto dest!: less_imp_Suc_add) +apply (induct_tac k) +apply (auto intro: less_trans) +done + +text{* for any sequence, there is a monotonic subsequence *} +lemma seq_monosub: + fixes s :: "nat => 'a::linorder" + shows "\f. subseq f \ monoseq (\ n. (s (f n)))" +proof cases + let "?P p n" = "p > n \ (\m\p. s m \ s p)" + assume *: "\n. \p. ?P p n" + def f \ "nat_rec (SOME p. ?P p 0) (\_ n. SOME p. ?P p n)" + have f_0: "f 0 = (SOME p. ?P p 0)" unfolding f_def by simp + have f_Suc: "\i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc .. + have P_0: "?P (f 0) 0" unfolding f_0 using *[rule_format] by (rule someI2_ex) auto + have P_Suc: "\i. ?P (f (Suc i)) (f i)" unfolding f_Suc using *[rule_format] by (rule someI2_ex) auto + then have "subseq f" unfolding subseq_Suc_iff by auto + moreover have "monoseq (\n. s (f n))" unfolding monoseq_Suc + proof (intro disjI2 allI) + fix n show "s (f (Suc n)) \ s (f n)" + proof (cases n) + case 0 with P_Suc[of 0] P_0 show ?thesis by auto + next + case (Suc m) + from P_Suc[of n] Suc have "f (Suc m) \ f (Suc (Suc m))" by simp + with P_Suc Suc show ?thesis by simp + qed + qed + ultimately show ?thesis by auto +next + let "?P p m" = "m < p \ s m < s p" + assume "\ (\n. \p>n. (\m\p. s m \ s p))" + then obtain N where N: "\p. p > N \ \m>p. s p < s m" by (force simp: not_le le_less) + def f \ "nat_rec (SOME p. ?P p (Suc N)) (\_ n. SOME p. ?P p n)" + have f_0: "f 0 = (SOME p. ?P p (Suc N))" unfolding f_def by simp + have f_Suc: "\i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc .. + have P_0: "?P (f 0) (Suc N)" + unfolding f_0 some_eq_ex[of "\p. ?P p (Suc N)"] using N[of "Suc N"] by auto + { fix i have "N < f i \ ?P (f (Suc i)) (f i)" + unfolding f_Suc some_eq_ex[of "\p. ?P p (f i)"] using N[of "f i"] . } + note P' = this + { fix i have "N < f i \ ?P (f (Suc i)) (f i)" + by (induct i) (insert P_0 P', auto) } + then have "subseq f" "monoseq (\x. s (f x))" + unfolding subseq_Suc_iff monoseq_Suc by (auto simp: not_le intro: less_imp_le) + then show ?thesis by auto +qed + +lemma seq_suble: assumes sf: "subseq f" shows "n \ f n" +proof(induct n) + case 0 thus ?case by simp +next + case (Suc n) + from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps + have "n < f (Suc n)" by arith + thus ?case by arith +qed + +lemma eventually_subseq: + "subseq r \ eventually P sequentially \ eventually (\n. P (r n)) sequentially" + unfolding eventually_sequentially by (metis seq_suble le_trans) + +lemma not_eventually_sequentiallyD: + assumes P: "\ eventually P sequentially" + shows "\r. subseq r \ (\n. \ P (r n))" +proof - + from P have "\n. \m\n. \ P m" + unfolding eventually_sequentially by (simp add: not_less) + then obtain r where "\n. r n \ n" "\n. \ P (r n)" + by (auto simp: choice_iff) + then show ?thesis + by (auto intro!: exI[of _ "\n. r (((Suc \ r) ^^ Suc n) 0)"] + simp: less_eq_Suc_le subseq_Suc_iff) +qed + +lemma filterlim_subseq: "subseq f \ filterlim f sequentially sequentially" + unfolding filterlim_iff by (metis eventually_subseq) + +lemma subseq_o: "subseq r \ subseq s \ subseq (r \ s)" + unfolding subseq_def by simp + +lemma subseq_mono: assumes "subseq r" "m < n" shows "r m < r n" + using assms by (auto simp: subseq_def) + +lemma incseq_imp_monoseq: "incseq X \ monoseq X" + by (simp add: incseq_def monoseq_def) + +lemma decseq_imp_monoseq: "decseq X \ monoseq X" + by (simp add: decseq_def monoseq_def) + +lemma decseq_eq_incseq: + fixes X :: "nat \ 'a::ordered_ab_group_add" shows "decseq X = incseq (\n. - X n)" + by (simp add: decseq_def incseq_def) + +lemma INT_decseq_offset: + assumes "decseq F" + shows "(\i. F i) = (\i\{n..}. F i)" +proof safe + fix x i assume x: "x \ (\i\{n..}. F i)" + show "x \ F i" + proof cases + from x have "x \ F n" by auto + also assume "i \ n" with `decseq F` have "F n \ F i" + unfolding decseq_def by simp + finally show ?thesis . + qed (insert x, simp) +qed auto + +lemma LIMSEQ_const_iff: + fixes k l :: "'a::t2_space" + shows "(\n. k) ----> l \ k = l" + using trivial_limit_sequentially by (rule tendsto_const_iff) + +lemma LIMSEQ_SUP: + "incseq X \ X ----> (SUP i. X i :: 'a :: {complete_linorder, linorder_topology})" + by (intro increasing_tendsto) + (auto simp: SUP_upper less_SUP_iff incseq_def eventually_sequentially intro: less_le_trans) + +lemma LIMSEQ_INF: + "decseq X \ X ----> (INF i. X i :: 'a :: {complete_linorder, linorder_topology})" + by (intro decreasing_tendsto) + (auto simp: INF_lower INF_less_iff decseq_def eventually_sequentially intro: le_less_trans) + +lemma LIMSEQ_ignore_initial_segment: + "f ----> a \ (\n. f (n + k)) ----> a" + unfolding tendsto_def + by (subst eventually_sequentially_seg[where k=k]) + +lemma LIMSEQ_offset: + "(\n. f (n + k)) ----> a \ f ----> a" + unfolding tendsto_def + by (subst (asm) eventually_sequentially_seg[where k=k]) + +lemma LIMSEQ_Suc: "f ----> l \ (\n. f (Suc n)) ----> l" +by (drule_tac k="Suc 0" in LIMSEQ_ignore_initial_segment, simp) + +lemma LIMSEQ_imp_Suc: "(\n. f (Suc n)) ----> l \ f ----> l" +by (rule_tac k="Suc 0" in LIMSEQ_offset, simp) + +lemma LIMSEQ_Suc_iff: "(\n. f (Suc n)) ----> l = f ----> l" +by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc) + +lemma LIMSEQ_unique: + fixes a b :: "'a::t2_space" + shows "\X ----> a; X ----> b\ \ a = b" + using trivial_limit_sequentially by (rule tendsto_unique) + +lemma LIMSEQ_le_const: + "\X ----> (x::'a::linorder_topology); \N. \n\N. a \ X n\ \ a \ x" + using tendsto_le_const[of sequentially X x a] by (simp add: eventually_sequentially) + +lemma LIMSEQ_le: + "\X ----> x; Y ----> y; \N. \n\N. X n \ Y n\ \ x \ (y::'a::linorder_topology)" + using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially) + +lemma LIMSEQ_le_const2: + "\X ----> (x::'a::linorder_topology); \N. \n\N. X n \ a\ \ x \ a" + by (rule LIMSEQ_le[of X x "\n. a"]) (auto simp: tendsto_const) + +lemma convergentD: "convergent X ==> \L. (X ----> L)" +by (simp add: convergent_def) + +lemma convergentI: "(X ----> L) ==> convergent X" +by (auto simp add: convergent_def) + +lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)" +by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) + +lemma convergent_const: "convergent (\n. c)" + by (rule convergentI, rule tendsto_const) + +lemma monoseq_le: + "monoseq a \ a ----> (x::'a::linorder_topology) \ + ((\ n. a n \ x) \ (\m. \n\m. a m \ a n)) \ ((\ n. x \ a n) \ (\m. \n\m. a n \ a m))" + by (metis LIMSEQ_le_const LIMSEQ_le_const2 decseq_def incseq_def monoseq_iff) + +lemma LIMSEQ_subseq_LIMSEQ: + "\ X ----> L; subseq f \ \ (X o f) ----> L" + unfolding comp_def by (rule filterlim_compose[of X, OF _ filterlim_subseq]) + +lemma convergent_subseq_convergent: + "\convergent X; subseq f\ \ convergent (X o f)" + unfolding convergent_def by (auto intro: LIMSEQ_subseq_LIMSEQ) + +lemma limI: "X ----> L ==> lim X = L" +apply (simp add: lim_def) +apply (blast intro: LIMSEQ_unique) +done + +lemma lim_le: "convergent f \ (\n. f n \ (x::'a::linorder_topology)) \ lim f \ x" + using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff) + +subsubsection{*Increasing and Decreasing Series*} + +lemma incseq_le: "incseq X \ X ----> L \ X n \ (L::'a::linorder_topology)" + by (metis incseq_def LIMSEQ_le_const) + +lemma decseq_le: "decseq X \ X ----> L \ (L::'a::linorder_topology) \ X n" + by (metis decseq_def LIMSEQ_le_const2) + +subsection {* First countable topologies *} + +class first_countable_topology = topological_space + + assumes first_countable_basis: + "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" + +lemma (in first_countable_topology) countable_basis_at_decseq: + obtains A :: "nat \ 'a set" where + "\i. open (A i)" "\i. x \ (A i)" + "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" +proof atomize_elim + from first_countable_basis[of x] obtain A :: "nat \ 'a set" where + nhds: "\i. open (A i)" "\i. x \ A i" + and incl: "\S. open S \ x \ S \ \i. A i \ S" by auto + def F \ "\n. \i\n. A i" + show "\A. (\i. open (A i)) \ (\i. x \ A i) \ + (\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially)" + proof (safe intro!: exI[of _ F]) + fix i + show "open (F i)" using nhds(1) by (auto simp: F_def) + show "x \ F i" using nhds(2) by (auto simp: F_def) + next + fix S assume "open S" "x \ S" + from incl[OF this] obtain i where "F i \ S" unfolding F_def by auto + moreover have "\j. i \ j \ F j \ F i" + by (auto simp: F_def) + ultimately show "eventually (\i. F i \ S) sequentially" + by (auto simp: eventually_sequentially) + qed +qed + +lemma (in first_countable_topology) countable_basis: + obtains A :: "nat \ 'a set" where + "\i. open (A i)" "\i. x \ A i" + "\F. (\n. F n \ A n) \ F ----> x" +proof atomize_elim + from countable_basis_at_decseq[of x] guess A . note A = this + { fix F S assume "\n. F n \ A n" "open S" "x \ S" + with A(3)[of S] have "eventually (\n. F n \ S) sequentially" + by (auto elim: eventually_elim1 simp: subset_eq) } + with A show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\F. (\n. F n \ A n) \ F ----> x)" + by (intro exI[of _ A]) (auto simp: tendsto_def) +qed + +lemma (in first_countable_topology) sequentially_imp_eventually_nhds_within: + assumes "\f. (\n. f n \ s) \ f ----> a \ eventually (\n. P (f n)) sequentially" + shows "eventually P (nhds a within s)" +proof (rule ccontr) + from countable_basis[of a] guess A . note A = this + assume "\ eventually P (nhds a within s)" + with A have P: "\F. \n. F n \ s \ F n \ A n \ \ P (F n)" + unfolding eventually_within eventually_nhds by (intro choice) fastforce + then guess F .. + hence F0: "\n. F n \ s" and F2: "\n. F n \ A n" and F3: "\n. \ P (F n)" + by fast+ + with A have "F ----> a" by auto + hence "eventually (\n. P (F n)) sequentially" + using assms F0 by simp + thus "False" by (simp add: F3) +qed + +lemma (in first_countable_topology) eventually_nhds_within_iff_sequentially: + "eventually P (nhds a within s) \ + (\f. (\n. f n \ s) \ f ----> a \ eventually (\n. P (f n)) sequentially)" +proof (safe intro!: sequentially_imp_eventually_nhds_within) + assume "eventually P (nhds a within s)" + then obtain S where "open S" "a \ S" "\x\S. x \ s \ P x" + by (auto simp: eventually_within eventually_nhds) + moreover fix f assume "\n. f n \ s" "f ----> a" + ultimately show "eventually (\n. P (f n)) sequentially" + by (auto dest!: topological_tendstoD elim: eventually_elim1) +qed + +lemma (in first_countable_topology) eventually_nhds_iff_sequentially: + "eventually P (nhds a) \ (\f. f ----> a \ eventually (\n. P (f n)) sequentially)" + using eventually_nhds_within_iff_sequentially[of P a UNIV] by simp + +subsection {* Function limit at a point *} + +abbreviation + LIM :: "('a::topological_space \ 'b::topological_space) \ 'a \ 'b \ bool" + ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where + "f -- a --> L \ (f ---> L) (at a)" + +lemma tendsto_within_open: "a \ S \ open S \ (f ---> l) (at a within S) \ (f -- a --> l)" + unfolding tendsto_def by (simp add: at_within_open) + +lemma LIM_const_not_eq[tendsto_intros]: + fixes a :: "'a::perfect_space" + fixes k L :: "'b::t2_space" + shows "k \ L \ \ (\x. k) -- a --> L" + by (simp add: tendsto_const_iff) + +lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] + +lemma LIM_const_eq: + fixes a :: "'a::perfect_space" + fixes k L :: "'b::t2_space" + shows "(\x. k) -- a --> L \ k = L" + by (simp add: tendsto_const_iff) + +lemma LIM_unique: + fixes a :: "'a::perfect_space" and L M :: "'b::t2_space" + shows "f -- a --> L \ f -- a --> M \ L = M" + using at_neq_bot by (rule tendsto_unique) + +text {* Limits are equal for functions equal except at limit point *} + +lemma LIM_equal: "\x. x \ a --> (f x = g x) \ (f -- a --> l) \ (g -- a --> l)" + unfolding tendsto_def eventually_at_topological by simp + +lemma LIM_cong: "a = b \ (\x. x \ b \ f x = g x) \ l = m \ (f -- a --> l) \ (g -- b --> m)" + by (simp add: LIM_equal) + +lemma LIM_cong_limit: "f -- x --> L \ K = L \ f -- x --> K" + by simp + +lemma tendsto_at_iff_tendsto_nhds: + "g -- l --> g l \ (g ---> g l) (nhds l)" + unfolding tendsto_def at_def eventually_within + by (intro ext all_cong imp_cong) (auto elim!: eventually_elim1) + +lemma tendsto_compose: + "g -- l --> g l \ (f ---> l) F \ ((\x. g (f x)) ---> g l) F" + unfolding tendsto_at_iff_tendsto_nhds by (rule filterlim_compose[of g]) + +lemma LIM_o: "\g -- l --> g l; f -- a --> l\ \ (g \ f) -- a --> g l" + unfolding o_def by (rule tendsto_compose) + +lemma tendsto_compose_eventually: + "g -- l --> m \ (f ---> l) F \ eventually (\x. f x \ l) F \ ((\x. g (f x)) ---> m) F" + by (rule filterlim_compose[of g _ "at l"]) (auto simp add: filterlim_at) + +lemma LIM_compose_eventually: + assumes f: "f -- a --> b" + assumes g: "g -- b --> c" + assumes inj: "eventually (\x. f x \ b) (at a)" + shows "(\x. g (f x)) -- a --> c" + using g f inj by (rule tendsto_compose_eventually) + +subsubsection {* Relation of LIM and LIMSEQ *} + +lemma (in first_countable_topology) sequentially_imp_eventually_within: + "(\f. (\n. f n \ s \ f n \ a) \ f ----> a \ eventually (\n. P (f n)) sequentially) \ + eventually P (at a within s)" + unfolding at_def within_within_eq + by (intro sequentially_imp_eventually_nhds_within) auto + +lemma (in first_countable_topology) sequentially_imp_eventually_at: + "(\f. (\n. f n \ a) \ f ----> a \ eventually (\n. P (f n)) sequentially) \ eventually P (at a)" + using assms sequentially_imp_eventually_within [where s=UNIV] by simp + +lemma LIMSEQ_SEQ_conv1: + fixes f :: "'a::topological_space \ 'b::topological_space" + assumes f: "f -- a --> l" + shows "\S. (\n. S n \ a) \ S ----> a \ (\n. f (S n)) ----> l" + using tendsto_compose_eventually [OF f, where F=sequentially] by simp + +lemma LIMSEQ_SEQ_conv2: + fixes f :: "'a::first_countable_topology \ 'b::topological_space" + assumes "\S. (\n. S n \ a) \ S ----> a \ (\n. f (S n)) ----> l" + shows "f -- a --> l" + using assms unfolding tendsto_def [where l=l] by (simp add: sequentially_imp_eventually_at) + +lemma LIMSEQ_SEQ_conv: + "(\S. (\n. S n \ a) \ S ----> (a::'a::first_countable_topology) \ (\n. X (S n)) ----> L) = + (X -- a --> (L::'b::topological_space))" + using LIMSEQ_SEQ_conv2 LIMSEQ_SEQ_conv1 .. + +subsection {* Continuity *} + +subsubsection {* Continuity on a set *} + +definition continuous_on :: "'a set \ ('a :: topological_space \ 'b :: topological_space) \ bool" where + "continuous_on s f \ (\x\s. (f ---> f x) (at x within s))" + +lemma continuous_on_cong [cong]: + "s = t \ (\x. x \ t \ f x = g x) \ continuous_on s f \ continuous_on t g" + unfolding continuous_on_def by (intro ball_cong filterlim_cong) (auto simp: eventually_within) + +lemma continuous_on_topological: + "continuous_on s f \ + (\x\s. \B. open B \ f x \ B \ (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" + unfolding continuous_on_def tendsto_def eventually_within eventually_at_topological by metis + +lemma continuous_on_open_invariant: + "continuous_on s f \ (\B. open B \ (\A. open A \ A \ s = f -` B \ s))" +proof safe + fix B :: "'b set" assume "continuous_on s f" "open B" + then have "\x\f -` B \ s. (\A. open A \ x \ A \ s \ A \ f -` B)" + by (auto simp: continuous_on_topological subset_eq Ball_def imp_conjL) + then guess A unfolding bchoice_iff .. + then show "\A. open A \ A \ s = f -` B \ s" + by (intro exI[of _ "\x\f -` B \ s. A x"]) auto +next + assume B: "\B. open B \ (\A. open A \ A \ s = f -` B \ s)" + show "continuous_on s f" + unfolding continuous_on_topological + proof safe + fix x B assume "x \ s" "open B" "f x \ B" + with B obtain A where A: "open A" "A \ s = f -` B \ s" by auto + with `x \ s` `f x \ B` show "\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)" + by (intro exI[of _ A]) auto + qed +qed + +lemma continuous_on_open_vimage: + "open s \ continuous_on s f \ (\B. open B \ open (f -` B \ s))" + unfolding continuous_on_open_invariant + by (metis open_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) + +lemma continuous_on_closed_invariant: + "continuous_on s f \ (\B. closed B \ (\A. closed A \ A \ s = f -` B \ s))" +proof - + have *: "\P Q::'b set\bool. (\A. P A \ Q (- A)) \ (\A. P A) \ (\A. Q A)" + by (metis double_compl) + show ?thesis + unfolding continuous_on_open_invariant by (intro *) (auto simp: open_closed[symmetric]) +qed + +lemma continuous_on_closed_vimage: + "closed s \ continuous_on s f \ (\B. closed B \ closed (f -` B \ s))" + unfolding continuous_on_closed_invariant + by (metis closed_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) + +lemma continuous_on_open_Union: + "(\s. s \ S \ open s) \ (\s. s \ S \ continuous_on s f) \ continuous_on (\S) f" + unfolding continuous_on_def by (simp add: tendsto_within_open open_Union) + +lemma continuous_on_open_UN: + "(\s. s \ S \ open (A s)) \ (\s. s \ S \ continuous_on (A s) f) \ continuous_on (\s\S. A s) f" + unfolding Union_image_eq[symmetric] by (rule continuous_on_open_Union) auto + +lemma continuous_on_closed_Un: + "closed s \ closed t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" + by (auto simp add: continuous_on_closed_vimage closed_Un Int_Un_distrib) + +lemma continuous_on_If: + assumes closed: "closed s" "closed t" and cont: "continuous_on s f" "continuous_on t g" + and P: "\x. x \ s \ \ P x \ f x = g x" "\x. x \ t \ P x \ f x = g x" + shows "continuous_on (s \ t) (\x. if P x then f x else g x)" (is "continuous_on _ ?h") +proof- + from P have "\x\s. f x = ?h x" "\x\t. g x = ?h x" + by auto + with cont have "continuous_on s ?h" "continuous_on t ?h" + by simp_all + with closed show ?thesis + by (rule continuous_on_closed_Un) +qed + +ML {* + +structure Continuous_On_Intros = Named_Thms +( + val name = @{binding continuous_on_intros} + val description = "Structural introduction rules for setwise continuity" +) + +*} + +setup Continuous_On_Intros.setup + +lemma continuous_on_id[continuous_on_intros]: "continuous_on s (\x. x)" + unfolding continuous_on_def by (fast intro: tendsto_ident_at_within) + +lemma continuous_on_const[continuous_on_intros]: "continuous_on s (\x. c)" + unfolding continuous_on_def by (auto intro: tendsto_const) + +lemma continuous_on_compose[continuous_on_intros]: + "continuous_on s f \ continuous_on (f ` s) g \ continuous_on s (g o f)" + unfolding continuous_on_topological by simp metis + +lemma continuous_on_compose2: + "continuous_on t g \ continuous_on s f \ t = f ` s \ continuous_on s (\x. g (f x))" + using continuous_on_compose[of s f g] by (simp add: comp_def) + +subsubsection {* Continuity at a point *} + +definition continuous :: "'a::t2_space filter \ ('a \ 'b::topological_space) \ bool" where + "continuous F f \ (f ---> f (Lim F (\x. x))) F" + +ML {* + +structure Continuous_Intros = Named_Thms +( + val name = @{binding continuous_intros} + val description = "Structural introduction rules for pointwise continuity" +) + +*} + +setup Continuous_Intros.setup + +lemma continuous_bot[continuous_intros, simp]: "continuous bot f" + unfolding continuous_def by auto + +lemma continuous_trivial_limit: "trivial_limit net \ continuous net f" + by simp + +lemma continuous_within: "continuous (at x within s) f \ (f ---> f x) (at x within s)" + by (cases "trivial_limit (at x within s)") (auto simp add: Lim_ident_at_within continuous_def) + +lemma continuous_within_topological: + "continuous (at x within s) f \ + (\B. open B \ f x \ B \ (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" + unfolding continuous_within tendsto_def eventually_within eventually_at_topological by metis + +lemma continuous_within_compose[continuous_intros]: + "continuous (at x within s) f \ continuous (at (f x) within f ` s) g \ + continuous (at x within s) (g o f)" + by (simp add: continuous_within_topological) metis + +lemma continuous_within_compose2: + "continuous (at x within s) f \ continuous (at (f x) within f ` s) g \ + continuous (at x within s) (\x. g (f x))" + using continuous_within_compose[of x s f g] by (simp add: comp_def) + +lemma continuous_at: "continuous (at x) f \ f -- x --> f x" + using continuous_within[of x UNIV f] by simp + +lemma continuous_ident[continuous_intros, simp]: "continuous (at x within S) (\x. x)" + unfolding continuous_within by (rule tendsto_ident_at_within) + +lemma continuous_const[continuous_intros, simp]: "continuous F (\x. c)" + unfolding continuous_def by (rule tendsto_const) + +lemma continuous_on_eq_continuous_within: + "continuous_on s f \ (\x\s. continuous (at x within s) f)" + unfolding continuous_on_def continuous_within .. + +abbreviation isCont :: "('a::t2_space \ 'b::topological_space) \ 'a \ bool" where + "isCont f a \ continuous (at a) f" + +lemma isCont_def: "isCont f a \ f -- a --> f a" + by (rule continuous_at) + +lemma continuous_at_within: "isCont f x \ continuous (at x within s) f" + by (auto intro: within_le filterlim_mono simp: continuous_at continuous_within) + +lemma continuous_on_eq_continuous_at: "open s \ continuous_on s f \ (\x\s. isCont f x)" + by (simp add: continuous_on_def continuous_at tendsto_within_open) + +lemma continuous_on_subset: "continuous_on s f \ t \ s \ continuous_on t f" + unfolding continuous_on_def by (metis subset_eq tendsto_within_subset) + +lemma continuous_at_imp_continuous_on: "\x\s. isCont f x \ continuous_on s f" + by (auto intro: continuous_at_within simp: continuous_on_eq_continuous_within) + +lemma isContI_continuous: "continuous (at x within UNIV) f \ isCont f x" + by simp + +lemma isCont_ident[continuous_intros, simp]: "isCont (\x. x) a" + using continuous_ident by (rule isContI_continuous) + +lemmas isCont_const = continuous_const + +lemma isCont_o2: "isCont f a \ isCont g (f a) \ isCont (\x. g (f x)) a" + unfolding isCont_def by (rule tendsto_compose) + +lemma isCont_o[continuous_intros]: "isCont f a \ isCont g (f a) \ isCont (g \ f) a" + unfolding o_def by (rule isCont_o2) + +lemma isCont_tendsto_compose: "isCont g l \ (f ---> l) F \ ((\x. g (f x)) ---> g l) F" + unfolding isCont_def by (rule tendsto_compose) + +lemma continuous_within_compose3: + "isCont g (f x) \ continuous (at x within s) f \ continuous (at x within s) (\x. g (f x))" + using continuous_within_compose2[of x s f g] by (simp add: continuous_at_within) + +subsubsection{* Open-cover compactness *} + +context topological_space +begin + +definition compact :: "'a set \ bool" where + compact_eq_heine_borel: -- "This name is used for backwards compatibility" + "compact S \ (\C. (\c\C. open c) \ S \ \C \ (\D\C. finite D \ S \ \D))" + +lemma compactI: + assumes "\C. \t\C. open t \ s \ \ C \ \C'. C' \ C \ finite C' \ s \ \ C'" + shows "compact s" + unfolding compact_eq_heine_borel using assms by metis + +lemma compact_empty[simp]: "compact {}" + by (auto intro!: compactI) + +lemma compactE: + assumes "compact s" and "\t\C. open t" and "s \ \C" + obtains C' where "C' \ C" and "finite C'" and "s \ \C'" + using assms unfolding compact_eq_heine_borel by metis + +lemma compactE_image: + assumes "compact s" and "\t\C. open (f t)" and "s \ (\c\C. f c)" + obtains C' where "C' \ C" and "finite C'" and "s \ (\c\C'. f c)" + using assms unfolding ball_simps[symmetric] SUP_def + by (metis (lifting) finite_subset_image compact_eq_heine_borel[of s]) + +lemma compact_inter_closed [intro]: + assumes "compact s" and "closed t" + shows "compact (s \ t)" +proof (rule compactI) + fix C assume C: "\c\C. open c" and cover: "s \ t \ \C" + from C `closed t` have "\c\C \ {-t}. open c" by auto + moreover from cover have "s \ \(C \ {-t})" by auto + ultimately have "\D\C \ {-t}. finite D \ s \ \D" + using `compact s` unfolding compact_eq_heine_borel by auto + then guess D .. + then show "\D\C. finite D \ s \ t \ \D" + by (intro exI[of _ "D - {-t}"]) auto +qed + +end + +lemma (in t2_space) compact_imp_closed: + assumes "compact s" shows "closed s" +unfolding closed_def +proof (rule openI) + fix y assume "y \ - s" + let ?C = "\x\s. {u. open u \ x \ u \ eventually (\y. y \ u) (nhds y)}" + note `compact s` + moreover have "\u\?C. open u" by simp + moreover have "s \ \?C" + proof + fix x assume "x \ s" + with `y \ - s` have "x \ y" by clarsimp + hence "\u v. open u \ open v \ x \ u \ y \ v \ u \ v = {}" + by (rule hausdorff) + with `x \ s` show "x \ \?C" + unfolding eventually_nhds by auto + qed + ultimately obtain D where "D \ ?C" and "finite D" and "s \ \D" + by (rule compactE) + from `D \ ?C` have "\x\D. eventually (\y. y \ x) (nhds y)" by auto + with `finite D` have "eventually (\y. y \ \D) (nhds y)" + by (simp add: eventually_Ball_finite) + with `s \ \D` have "eventually (\y. y \ s) (nhds y)" + by (auto elim!: eventually_mono [rotated]) + thus "\t. open t \ y \ t \ t \ - s" + by (simp add: eventually_nhds subset_eq) +qed + +lemma compact_continuous_image: + assumes f: "continuous_on s f" and s: "compact s" + shows "compact (f ` s)" +proof (rule compactI) + fix C assume "\c\C. open c" and cover: "f`s \ \C" + with f have "\c\C. \A. open A \ A \ s = f -` c \ s" + unfolding continuous_on_open_invariant by blast + then guess A unfolding bchoice_iff .. note A = this + with cover have "\c\C. open (A c)" "s \ (\c\C. A c)" + by (fastforce simp add: subset_eq set_eq_iff)+ + from compactE_image[OF s this] obtain D where "D \ C" "finite D" "s \ (\c\D. A c)" . + with A show "\D \ C. finite D \ f`s \ \D" + by (intro exI[of _ D]) (fastforce simp add: subset_eq set_eq_iff)+ +qed + +lemma continuous_on_inv: + fixes f :: "'a::topological_space \ 'b::t2_space" + assumes "continuous_on s f" "compact s" "\x\s. g (f x) = x" + shows "continuous_on (f ` s) g" +unfolding continuous_on_topological +proof (clarsimp simp add: assms(3)) + fix x :: 'a and B :: "'a set" + assume "x \ s" and "open B" and "x \ B" + have 1: "\x\s. f x \ f ` (s - B) \ x \ s - B" + using assms(3) by (auto, metis) + have "continuous_on (s - B) f" + using `continuous_on s f` Diff_subset + by (rule continuous_on_subset) + moreover have "compact (s - B)" + using `open B` and `compact s` + unfolding Diff_eq by (intro compact_inter_closed closed_Compl) + ultimately have "compact (f ` (s - B))" + by (rule compact_continuous_image) + hence "closed (f ` (s - B))" + by (rule compact_imp_closed) + hence "open (- f ` (s - B))" + by (rule open_Compl) + moreover have "f x \ - f ` (s - B)" + using `x \ s` and `x \ B` by (simp add: 1) + moreover have "\y\s. f y \ - f ` (s - B) \ y \ B" + by (simp add: 1) + ultimately show "\A. open A \ f x \ A \ (\y\s. f y \ A \ y \ B)" + by fast +qed + +lemma continuous_on_inv_into: + fixes f :: "'a::topological_space \ 'b::t2_space" + assumes s: "continuous_on s f" "compact s" and f: "inj_on f s" + shows "continuous_on (f ` s) (the_inv_into s f)" + by (rule continuous_on_inv[OF s]) (auto simp: the_inv_into_f_f[OF f]) + +lemma (in linorder_topology) compact_attains_sup: + assumes "compact S" "S \ {}" + shows "\s\S. \t\S. t \ s" +proof (rule classical) + assume "\ (\s\S. \t\S. t \ s)" + then obtain t where t: "\s\S. t s \ S" and "\s\S. s < t s" + by (metis not_le) + then have "\s\S. open {..< t s}" "S \ (\s\S. {..< t s})" + by auto + with `compact S` obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {..< t s})" + by (erule compactE_image) + with `S \ {}` have Max: "Max (t`C) \ t`C" and "\s\t`C. s \ Max (t`C)" + by (auto intro!: Max_in) + with C have "S \ {..< Max (t`C)}" + by (auto intro: less_le_trans simp: subset_eq) + with t Max `C \ S` show ?thesis + by fastforce +qed + +lemma (in linorder_topology) compact_attains_inf: + assumes "compact S" "S \ {}" + shows "\s\S. \t\S. s \ t" +proof (rule classical) + assume "\ (\s\S. \t\S. s \ t)" + then obtain t where t: "\s\S. t s \ S" and "\s\S. t s < s" + by (metis not_le) + then have "\s\S. open {t s <..}" "S \ (\s\S. {t s <..})" + by auto + with `compact S` obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {t s <..})" + by (erule compactE_image) + with `S \ {}` have Min: "Min (t`C) \ t`C" and "\s\t`C. Min (t`C) \ s" + by (auto intro!: Min_in) + with C have "S \ {Min (t`C) <..}" + by (auto intro: le_less_trans simp: subset_eq) + with t Min `C \ S` show ?thesis + by fastforce +qed + +lemma continuous_attains_sup: + fixes f :: "'a::topological_space \ 'b::linorder_topology" + shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f y \ f x)" + using compact_attains_sup[of "f ` s"] compact_continuous_image[of s f] by auto + +lemma continuous_attains_inf: + fixes f :: "'a::topological_space \ 'b::linorder_topology" + shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f x \ f y)" + using compact_attains_inf[of "f ` s"] compact_continuous_image[of s f] by auto + + +subsection {* Connectedness *} + +context topological_space +begin + +definition "connected S \ + \ (\A B. open A \ open B \ S \ A \ B \ A \ B \ S = {} \ A \ S \ {} \ B \ S \ {})" + +lemma connectedI: + "(\A B. open A \ open B \ A \ U \ {} \ B \ U \ {} \ A \ B \ U = {} \ U \ A \ B \ False) + \ connected U" + by (auto simp: connected_def) + +lemma connected_empty[simp]: "connected {}" + by (auto intro!: connectedI) + +end + +lemma (in linorder_topology) connectedD_interval: + assumes "connected U" and xy: "x \ U" "y \ U" and "x \ z" "z \ y" + shows "z \ U" +proof - + have eq: "{.. {z<..} = - {z}" + by auto + { assume "z \ U" "x < z" "z < y" + with xy have "\ connected U" + unfolding connected_def simp_thms + apply (rule_tac exI[of _ "{..< z}"]) + apply (rule_tac exI[of _ "{z <..}"]) + apply (auto simp add: eq) + done } + with assms show "z \ U" + by (metis less_le) +qed + +lemma connected_continuous_image: + assumes *: "continuous_on s f" + assumes "connected s" + shows "connected (f ` s)" +proof (rule connectedI) + fix A B assume A: "open A" "A \ f ` s \ {}" and B: "open B" "B \ f ` s \ {}" and + AB: "A \ B \ f ` s = {}" "f ` s \ A \ B" + obtain A' where A': "open A'" "f -` A \ s = A' \ s" + using * `open A` unfolding continuous_on_open_invariant by metis + obtain B' where B': "open B'" "f -` B \ s = B' \ s" + using * `open B` unfolding continuous_on_open_invariant by metis + + have "\A B. open A \ open B \ s \ A \ B \ A \ B \ s = {} \ A \ s \ {} \ B \ s \ {}" + proof (rule exI[of _ A'], rule exI[of _ B'], intro conjI) + have "s \ (f -` A \ s) \ (f -` B \ s)" using AB by auto + then show "s \ A' \ B'" using A' B' by auto + next + have "(f -` A \ s) \ (f -` B \ s) = {}" using AB by auto + then show "A' \ B' \ s = {}" using A' B' by auto + qed (insert A' B' A B, auto) + with `connected s` show False + unfolding connected_def by blast +qed + +end + diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/Transcendental.thy Sat Mar 23 07:30:53 2013 +0100 @@ -207,7 +207,7 @@ thus "\ N. \ n \ N. norm (- a (2 * n) - 0) < r" by auto qed ultimately - show ?thesis by (rule lemma_nest_unique) + show ?thesis by (rule nested_sequence_unique) qed lemma summable_Leibniz': fixes a :: "nat \ real" @@ -881,6 +881,12 @@ "(f ---> a) F \ ((\x. exp (f x)) ---> exp a) F" by (rule isCont_tendsto_compose [OF isCont_exp]) +lemma continuous_exp [continuous_intros]: "continuous F f \ continuous F (\x. exp (f x))" + unfolding continuous_def by (rule tendsto_exp) + +lemma continuous_on_exp [continuous_on_intros]: "continuous_on s f \ continuous_on s (\x. exp (f x))" + unfolding continuous_on_def by (auto intro: tendsto_exp) + subsubsection {* Properties of the Exponential Function *} lemma powser_zero: @@ -1169,6 +1175,22 @@ "\(f ---> a) F; 0 < a\ \ ((\x. ln (f x)) ---> ln a) F" by (rule isCont_tendsto_compose [OF isCont_ln]) +lemma continuous_ln: + "continuous F f \ 0 < f (Lim F (\x. x)) \ continuous F (\x. ln (f x))" + unfolding continuous_def by (rule tendsto_ln) + +lemma isCont_ln' [continuous_intros]: + "continuous (at x) f \ 0 < f x \ continuous (at x) (\x. ln (f x))" + unfolding continuous_at by (rule tendsto_ln) + +lemma continuous_within_ln [continuous_intros]: + "continuous (at x within s) f \ 0 < f x \ continuous (at x within s) (\x. ln (f x))" + unfolding continuous_within by (rule tendsto_ln) + +lemma continuous_on_ln [continuous_on_intros]: + "continuous_on s f \ (\x\s. 0 < f x) \ continuous_on s (\x. ln (f x))" + unfolding continuous_on_def by (auto intro: tendsto_ln) + lemma DERIV_ln: "0 < x \ DERIV ln x :> inverse x" apply (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"]) apply (erule DERIV_cong [OF DERIV_exp exp_ln]) @@ -1449,6 +1471,22 @@ "(f ---> a) F \ ((\x. cos (f x)) ---> cos a) F" by (rule isCont_tendsto_compose [OF isCont_cos]) +lemma continuous_sin [continuous_intros]: + "continuous F f \ continuous F (\x. sin (f x))" + unfolding continuous_def by (rule tendsto_sin) + +lemma continuous_on_sin [continuous_on_intros]: + "continuous_on s f \ continuous_on s (\x. sin (f x))" + unfolding continuous_on_def by (auto intro: tendsto_sin) + +lemma continuous_cos [continuous_intros]: + "continuous F f \ continuous F (\x. cos (f x))" + unfolding continuous_def by (rule tendsto_cos) + +lemma continuous_on_cos [continuous_on_intros]: + "continuous_on s f \ continuous_on s (\x. cos (f x))" + unfolding continuous_on_def by (auto intro: tendsto_cos) + declare DERIV_exp[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros] DERIV_ln_divide[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros] @@ -2076,6 +2114,22 @@ "\(f ---> a) F; cos a \ 0\ \ ((\x. tan (f x)) ---> tan a) F" by (rule isCont_tendsto_compose [OF isCont_tan]) +lemma continuous_tan: + "continuous F f \ cos (f (Lim F (\x. x))) \ 0 \ continuous F (\x. tan (f x))" + unfolding continuous_def by (rule tendsto_tan) + +lemma isCont_tan'' [continuous_intros]: + "continuous (at x) f \ cos (f x) \ 0 \ continuous (at x) (\x. tan (f x))" + unfolding continuous_at by (rule tendsto_tan) + +lemma continuous_within_tan [continuous_intros]: + "continuous (at x within s) f \ cos (f x) \ 0 \ continuous (at x within s) (\x. tan (f x))" + unfolding continuous_within by (rule tendsto_tan) + +lemma continuous_on_tan [continuous_on_intros]: + "continuous_on s f \ (\x\s. cos (f x) \ 0) \ continuous_on s (\x. tan (f x))" + unfolding continuous_on_def by (auto intro: tendsto_tan) + lemma LIM_cos_div_sin: "(%x. cos(x)/sin(x)) -- pi/2 --> 0" by (rule LIM_cong_limit, (rule tendsto_intros)+, simp_all) @@ -2403,32 +2457,47 @@ lemma arctan_eq_zero_iff [simp]: "arctan x = 0 \ x = 0" using arctan_eq_iff [of x 0] by simp -lemma isCont_inverse_function2: - fixes f g :: "real \ real" shows - "\a < x; x < b; - \z. a \ z \ z \ b \ g (f z) = z; - \z. a \ z \ z \ b \ isCont f z\ - \ isCont g (f x)" -apply (rule isCont_inverse_function - [where f=f and d="min (x - a) (b - x)"]) -apply (simp_all add: abs_le_iff) -done - -lemma isCont_arcsin: "\-1 < x; x < 1\ \ isCont arcsin x" -apply (subgoal_tac "isCont arcsin (sin (arcsin x))", simp) -apply (rule isCont_inverse_function2 [where f=sin]) -apply (erule (1) arcsin_lt_bounded [THEN conjunct1]) -apply (erule (1) arcsin_lt_bounded [THEN conjunct2]) -apply (fast intro: arcsin_sin, simp) -done - -lemma isCont_arccos: "\-1 < x; x < 1\ \ isCont arccos x" -apply (subgoal_tac "isCont arccos (cos (arccos x))", simp) -apply (rule isCont_inverse_function2 [where f=cos]) -apply (erule (1) arccos_lt_bounded [THEN conjunct1]) -apply (erule (1) arccos_lt_bounded [THEN conjunct2]) -apply (fast intro: arccos_cos, simp) -done +lemma continuous_on_arcsin': "continuous_on {-1 .. 1} arcsin" +proof - + have "continuous_on (sin ` {- pi / 2 .. pi / 2}) arcsin" + by (rule continuous_on_inv) (auto intro: continuous_on_intros simp: arcsin_sin) + also have "sin ` {- pi / 2 .. pi / 2} = {-1 .. 1}" + proof safe + fix x :: real assume "x \ {-1..1}" then show "x \ sin ` {- pi / 2..pi / 2}" + using arcsin_lbound arcsin_ubound by (intro image_eqI[where x="arcsin x"]) auto + qed simp + finally show ?thesis . +qed + +lemma continuous_on_arcsin [continuous_on_intros]: + "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arcsin (f x))" + using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arcsin']] + by (auto simp: comp_def subset_eq) + +lemma isCont_arcsin: "-1 < x \ x < 1 \ isCont arcsin x" + using continuous_on_arcsin'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] + by (auto simp: continuous_on_eq_continuous_at subset_eq) + +lemma continuous_on_arccos': "continuous_on {-1 .. 1} arccos" +proof - + have "continuous_on (cos ` {0 .. pi}) arccos" + by (rule continuous_on_inv) (auto intro: continuous_on_intros simp: arccos_cos) + also have "cos ` {0 .. pi} = {-1 .. 1}" + proof safe + fix x :: real assume "x \ {-1..1}" then show "x \ cos ` {0..pi}" + using arccos_lbound arccos_ubound by (intro image_eqI[where x="arccos x"]) auto + qed simp + finally show ?thesis . +qed + +lemma continuous_on_arccos [continuous_on_intros]: + "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arccos (f x))" + using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arccos']] + by (auto simp: comp_def subset_eq) + +lemma isCont_arccos: "-1 < x \ x < 1 \ isCont arccos x" + using continuous_on_arccos'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] + by (auto simp: continuous_on_eq_continuous_at subset_eq) lemma isCont_arctan: "isCont arctan x" apply (rule arctan_lbound [of x, THEN dense, THEN exE], clarify) @@ -2439,6 +2508,15 @@ apply (metis cos_gt_zero_pi isCont_tan order_less_le_trans less_le) done +lemma tendsto_arctan [tendsto_intros]: "(f ---> x) F \ ((\x. arctan (f x)) ---> arctan x) F" + by (rule isCont_tendsto_compose [OF isCont_arctan]) + +lemma continuous_arctan [continuous_intros]: "continuous F f \ continuous F (\x. arctan (f x))" + unfolding continuous_def by (rule tendsto_arctan) + +lemma continuous_on_arctan [continuous_on_intros]: "continuous_on s f \ continuous_on s (\x. arctan (f x))" + unfolding continuous_on_def by (auto intro: tendsto_arctan) + lemma DERIV_arcsin: "\-1 < x; x < 1\ \ DERIV arcsin x :> inverse (sqrt (1 - x\))" apply (rule DERIV_inverse_function [where f=sin and a="-1" and b="1"]) diff -r 637aa1649ac7 -r 0a0c9a45d294 src/HOL/ex/Gauge_Integration.thy --- a/src/HOL/ex/Gauge_Integration.thy Fri Mar 22 00:39:16 2013 +0100 +++ b/src/HOL/ex/Gauge_Integration.thy Sat Mar 23 07:30:53 2013 +0100 @@ -97,11 +97,7 @@ assumes 2: "\a b c. \P a b; P b c; a \ b; b \ c\ \ P a c" assumes 3: "\x. \d>0. \a b. a \ x & x \ b & (b-a) < d \ P a b" shows "P a b" -apply (subgoal_tac "split P (a,b)", simp) -apply (rule lemma_BOLZANO [OF _ _ 1]) -apply (clarify, erule (3) 2) -apply (clarify, rule 3) -done + using 1 2 3 by (rule Bolzano) text{*We can always find a division that is fine wrt any gauge*}