# HG changeset patch # User huffman # Date 1312856813 25200 # Node ID 730f7cced3a6ab5b0270f29f6f2419c3c3fdcd8c # Parent 53d95b52954c5725827f9bdbf56cbb09c99d5cf0 rename type 'a net to 'a filter, following standard mathematical terminology diff -r 53d95b52954c -r 730f7cced3a6 NEWS --- a/NEWS Mon Aug 08 18:36:32 2011 -0700 +++ b/NEWS Mon Aug 08 19:26:53 2011 -0700 @@ -170,6 +170,9 @@ Every theorem name containing "inat", "Fin", "Infty", or "iSuc" has been renamed accordingly. +* Limits.thy: Type "'a net" has been renamed to "'a filter", in +accordance with standard mathematical terminology. INCOMPATIBILITY. + *** Document preparation *** diff -r 53d95b52954c -r 730f7cced3a6 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Mon Aug 08 18:36:32 2011 -0700 +++ b/src/HOL/Limits.thy Mon Aug 08 19:26:53 2011 -0700 @@ -8,263 +8,262 @@ imports RealVector begin -subsection {* Nets *} +subsection {* Filters *} text {* - A net is now defined simply as a filter on a set. - The definition also allows non-proper filters. + This definition also allows non-proper filters. *} locale is_filter = - fixes net :: "('a \ bool) \ bool" - assumes True: "net (\x. True)" - assumes conj: "net (\x. P x) \ net (\x. Q x) \ net (\x. P x \ Q x)" - assumes mono: "\x. P x \ Q x \ net (\x. P x) \ net (\x. Q x)" + 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 (open) 'a net = - "{net :: ('a \ bool) \ bool. is_filter net}" +typedef (open) 'a filter = "{F :: ('a \ bool) \ bool. is_filter F}" proof - show "(\x. True) \ ?net" by (auto intro: is_filter.intro) + show "(\x. True) \ ?filter" by (auto intro: is_filter.intro) qed -lemma is_filter_Rep_net: "is_filter (Rep_net net)" -using Rep_net [of net] by simp +lemma is_filter_Rep_filter: "is_filter (Rep_filter A)" + using Rep_filter [of A] by simp -lemma Abs_net_inverse': - assumes "is_filter net" shows "Rep_net (Abs_net net) = net" -using assms by (simp add: Abs_net_inverse) +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 net \ bool" where - "eventually P net \ Rep_net net P" +definition eventually :: "('a \ bool) \ 'a filter \ bool" + where "eventually P A \ Rep_filter A P" -lemma eventually_Abs_net: - assumes "is_filter net" shows "eventually P (Abs_net net) = net P" -unfolding eventually_def using assms by (simp add: Abs_net_inverse) +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 expand_net_eq: - shows "net = net' \ (\P. eventually P net = eventually P net')" -unfolding Rep_net_inject [symmetric] fun_eq_iff eventually_def .. +lemma filter_eq_iff: + shows "A = B \ (\P. eventually P A = eventually P B)" + unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def .. -lemma eventually_True [simp]: "eventually (\x. True) net" -unfolding eventually_def -by (rule is_filter.True [OF is_filter_Rep_net]) +lemma eventually_True [simp]: "eventually (\x. True) A" + unfolding eventually_def + by (rule is_filter.True [OF is_filter_Rep_filter]) -lemma always_eventually: "\x. P x \ eventually P net" +lemma always_eventually: "\x. P x \ eventually P A" proof - assume "\x. P x" hence "P = (\x. True)" by (simp add: ext) - thus "eventually P net" by simp + thus "eventually P A" by simp qed lemma eventually_mono: - "(\x. P x \ Q x) \ eventually P net \ eventually Q net" -unfolding eventually_def -by (rule is_filter.mono [OF is_filter_Rep_net]) + "(\x. P x \ Q x) \ eventually P A \ eventually Q A" + unfolding eventually_def + by (rule is_filter.mono [OF is_filter_Rep_filter]) lemma eventually_conj: - assumes P: "eventually (\x. P x) net" - assumes Q: "eventually (\x. Q x) net" - shows "eventually (\x. P x \ Q x) net" -using assms unfolding eventually_def -by (rule is_filter.conj [OF is_filter_Rep_net]) + assumes P: "eventually (\x. P x) A" + assumes Q: "eventually (\x. Q x) A" + shows "eventually (\x. P x \ Q x) A" + using assms unfolding eventually_def + by (rule is_filter.conj [OF is_filter_Rep_filter]) lemma eventually_mp: - assumes "eventually (\x. P x \ Q x) net" - assumes "eventually (\x. P x) net" - shows "eventually (\x. Q x) net" + assumes "eventually (\x. P x \ Q x) A" + assumes "eventually (\x. P x) A" + shows "eventually (\x. Q x) A" 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) net" + show "eventually (\x. (P x \ Q x) \ P x) A" using assms by (rule eventually_conj) qed lemma eventually_rev_mp: - assumes "eventually (\x. P x) net" - assumes "eventually (\x. P x \ Q x) net" - shows "eventually (\x. Q x) net" + assumes "eventually (\x. P x) A" + assumes "eventually (\x. P x \ Q x) A" + shows "eventually (\x. Q x) A" using assms(2) assms(1) by (rule eventually_mp) lemma eventually_conj_iff: - "eventually (\x. P x \ Q x) net \ eventually P net \ eventually Q net" -by (auto intro: eventually_conj elim: eventually_rev_mp) + "eventually (\x. P x \ Q x) A \ eventually P A \ eventually Q A" + by (auto intro: eventually_conj elim: eventually_rev_mp) lemma eventually_elim1: - assumes "eventually (\i. P i) net" + assumes "eventually (\i. P i) A" assumes "\i. P i \ Q i" - shows "eventually (\i. Q i) net" -using assms by (auto elim!: eventually_rev_mp) + shows "eventually (\i. Q i) A" + using assms by (auto elim!: eventually_rev_mp) lemma eventually_elim2: - assumes "eventually (\i. P i) net" - assumes "eventually (\i. Q i) net" + assumes "eventually (\i. P i) A" + assumes "eventually (\i. Q i) A" assumes "\i. P i \ Q i \ R i" - shows "eventually (\i. R i) net" -using assms by (auto elim!: eventually_rev_mp) + shows "eventually (\i. R i) A" + using assms by (auto elim!: eventually_rev_mp) subsection {* Finer-than relation *} -text {* @{term "net \ net'"} means that @{term net} is finer than -@{term net'}. *} +text {* @{term "A \ B"} means that filter @{term A} is finer than +filter @{term B}. *} -instantiation net :: (type) complete_lattice +instantiation filter :: (type) complete_lattice begin -definition - le_net_def: "net \ net' \ (\P. eventually P net' \ eventually P net)" +definition le_filter_def: + "A \ B \ (\P. eventually P B \ eventually P A)" definition - less_net_def: "(net :: 'a net) < net' \ net \ net' \ \ net' \ net" + "(A :: 'a filter) < B \ A \ B \ \ B \ A" definition - top_net_def: "top = Abs_net (\P. \x. P x)" + "top = Abs_filter (\P. \x. P x)" definition - bot_net_def: "bot = Abs_net (\P. True)" + "bot = Abs_filter (\P. True)" definition - sup_net_def: "sup net net' = Abs_net (\P. eventually P net \ eventually P net')" + "sup A B = Abs_filter (\P. eventually P A \ eventually P B)" definition - inf_net_def: "inf a b = Abs_net - (\P. \Q R. eventually Q a \ eventually R b \ (\x. Q x \ R x \ P x))" + "inf A B = Abs_filter + (\P. \Q R. eventually Q A \ eventually R B \ (\x. Q x \ R x \ P x))" definition - Sup_net_def: "Sup A = Abs_net (\P. \net\A. eventually P net)" + "Sup S = Abs_filter (\P. \A\S. eventually P A)" definition - Inf_net_def: "Inf A = Sup {x::'a net. \y\A. x \ y}" + "Inf S = Sup {A::'a filter. \B\S. A \ B}" lemma eventually_top [simp]: "eventually P top \ (\x. P x)" -unfolding top_net_def -by (rule eventually_Abs_net, rule is_filter.intro, auto) + unfolding top_filter_def + by (rule eventually_Abs_filter, rule is_filter.intro, auto) lemma eventually_bot [simp]: "eventually P bot" -unfolding bot_net_def -by (subst eventually_Abs_net, rule is_filter.intro, auto) + unfolding bot_filter_def + by (subst eventually_Abs_filter, rule is_filter.intro, auto) lemma eventually_sup: - "eventually P (sup net net') \ eventually P net \ eventually P net'" -unfolding sup_net_def -by (rule eventually_Abs_net, rule is_filter.intro) - (auto elim!: eventually_rev_mp) + "eventually P (sup A B) \ eventually P A \ eventually P B" + unfolding sup_filter_def + by (rule eventually_Abs_filter, rule is_filter.intro) + (auto elim!: eventually_rev_mp) lemma eventually_inf: - "eventually P (inf a b) \ - (\Q R. eventually Q a \ eventually R b \ (\x. Q x \ R x \ P x))" -unfolding inf_net_def -apply (rule eventually_Abs_net, 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 + "eventually P (inf A B) \ + (\Q R. eventually Q A \ eventually R B \ (\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 A) \ (\net\A. eventually P net)" -unfolding Sup_net_def -apply (rule eventually_Abs_net, rule is_filter.intro) -apply (auto intro: eventually_conj elim!: eventually_rev_mp) -done + "eventually P (Sup S) \ (\A\S. eventually P A)" + 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 x y :: "'a net" show "x < y \ x \ y \ \ y \ x" - by (rule less_net_def) + fix A B :: "'a filter" show "A < B \ A \ B \ \ B \ A" + by (rule less_filter_def) next - fix x :: "'a net" show "x \ x" - unfolding le_net_def by simp + fix A :: "'a filter" show "A \ A" + unfolding le_filter_def by simp next - fix x y z :: "'a net" assume "x \ y" and "y \ z" thus "x \ z" - unfolding le_net_def by simp + fix A B C :: "'a filter" assume "A \ B" and "B \ C" thus "A \ C" + unfolding le_filter_def by simp next - fix x y :: "'a net" assume "x \ y" and "y \ x" thus "x = y" - unfolding le_net_def expand_net_eq by fast + fix A B :: "'a filter" assume "A \ B" and "B \ A" thus "A = B" + unfolding le_filter_def filter_eq_iff by fast next - fix x :: "'a net" show "x \ top" - unfolding le_net_def eventually_top by (simp add: always_eventually) + fix A :: "'a filter" show "A \ top" + unfolding le_filter_def eventually_top by (simp add: always_eventually) next - fix x :: "'a net" show "bot \ x" - unfolding le_net_def by simp + fix A :: "'a filter" show "bot \ A" + unfolding le_filter_def by simp next - fix x y :: "'a net" show "x \ sup x y" and "y \ sup x y" - unfolding le_net_def eventually_sup by simp_all + fix A B :: "'a filter" show "A \ sup A B" and "B \ sup A B" + unfolding le_filter_def eventually_sup by simp_all next - fix x y z :: "'a net" assume "x \ z" and "y \ z" thus "sup x y \ z" - unfolding le_net_def eventually_sup by simp + fix A B C :: "'a filter" assume "A \ C" and "B \ C" thus "sup A B \ C" + unfolding le_filter_def eventually_sup by simp next - fix x y :: "'a net" show "inf x y \ x" and "inf x y \ y" - unfolding le_net_def eventually_inf by (auto intro: eventually_True) + fix A B :: "'a filter" show "inf A B \ A" and "inf A B \ B" + unfolding le_filter_def eventually_inf by (auto intro: eventually_True) next - fix x y z :: "'a net" assume "x \ y" and "x \ z" thus "x \ inf y z" - unfolding le_net_def eventually_inf + fix A B C :: "'a filter" assume "A \ B" and "A \ C" thus "A \ inf B C" + unfolding le_filter_def eventually_inf by (auto elim!: eventually_mono intro: eventually_conj) next - fix x :: "'a net" and A assume "x \ A" thus "x \ Sup A" - unfolding le_net_def eventually_Sup by simp + fix A :: "'a filter" and S assume "A \ S" thus "A \ Sup S" + unfolding le_filter_def eventually_Sup by simp next - fix A and y :: "'a net" assume "\x. x \ A \ x \ y" thus "Sup A \ y" - unfolding le_net_def eventually_Sup by simp + fix S and B :: "'a filter" assume "\A. A \ S \ A \ B" thus "Sup S \ B" + unfolding le_filter_def eventually_Sup by simp next - fix z :: "'a net" and A assume "z \ A" thus "Inf A \ z" - unfolding le_net_def Inf_net_def eventually_Sup Ball_def by simp + fix C :: "'a filter" and S assume "C \ S" thus "Inf S \ C" + unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp next - fix A and x :: "'a net" assume "\y. y \ A \ x \ y" thus "x \ Inf A" - unfolding le_net_def Inf_net_def eventually_Sup Ball_def by simp + fix S and A :: "'a filter" assume "\B. B \ S \ A \ B" thus "A \ Inf S" + unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp qed end -lemma net_leD: - "net \ net' \ eventually P net' \ eventually P net" -unfolding le_net_def by simp +lemma filter_leD: + "A \ B \ eventually P B \ eventually P A" + unfolding le_filter_def by simp -lemma net_leI: - "(\P. eventually P net' \ eventually P net) \ net \ net'" -unfolding le_net_def by simp +lemma filter_leI: + "(\P. eventually P B \ eventually P A) \ A \ B" + unfolding le_filter_def by simp lemma eventually_False: - "eventually (\x. False) net \ net = bot" -unfolding expand_net_eq by (auto elim: eventually_rev_mp) + "eventually (\x. False) A \ A = bot" + unfolding filter_eq_iff by (auto elim: eventually_rev_mp) -subsection {* Map function for nets *} +subsection {* Map function for filters *} -definition netmap :: "('a \ 'b) \ 'a net \ 'b net" where - "netmap f net = Abs_net (\P. eventually (\x. P (f x)) net)" +definition filtermap :: "('a \ 'b) \ 'a filter \ 'b filter" + where "filtermap f A = Abs_filter (\P. eventually (\x. P (f x)) A)" -lemma eventually_netmap: - "eventually P (netmap f net) = eventually (\x. P (f x)) net" -unfolding netmap_def -apply (rule eventually_Abs_net) -apply (rule is_filter.intro) -apply (auto elim!: eventually_rev_mp) -done +lemma eventually_filtermap: + "eventually P (filtermap f A) = eventually (\x. P (f x)) A" + unfolding filtermap_def + apply (rule eventually_Abs_filter) + apply (rule is_filter.intro) + apply (auto elim!: eventually_rev_mp) + done -lemma netmap_ident: "netmap (\x. x) net = net" -by (simp add: expand_net_eq eventually_netmap) - -lemma netmap_netmap: "netmap f (netmap g net) = netmap (\x. f (g x)) net" -by (simp add: expand_net_eq eventually_netmap) +lemma filtermap_ident: "filtermap (\x. x) A = A" + by (simp add: filter_eq_iff eventually_filtermap) -lemma netmap_mono: "net \ net' \ netmap f net \ netmap f net'" -unfolding le_net_def eventually_netmap by simp +lemma filtermap_filtermap: + "filtermap f (filtermap g A) = filtermap (\x. f (g x)) A" + by (simp add: filter_eq_iff eventually_filtermap) -lemma netmap_bot [simp]: "netmap f bot = bot" -by (simp add: expand_net_eq eventually_netmap) +lemma filtermap_mono: "A \ B \ filtermap f A \ filtermap f B" + unfolding le_filter_def eventually_filtermap by simp + +lemma filtermap_bot [simp]: "filtermap f bot = bot" + by (simp add: filter_eq_iff eventually_filtermap) subsection {* Sequentially *} -definition sequentially :: "nat net" where - "sequentially = Abs_net (\P. \k. \n\k. P n)" +definition sequentially :: "nat filter" + where "sequentially = Abs_filter (\P. \k. \n\k. P n)" lemma eventually_sequentially: "eventually P sequentially \ (\N. \n\N. P n)" unfolding sequentially_def -proof (rule eventually_Abs_net, rule is_filter.intro) +proof (rule eventually_Abs_filter, rule is_filter.intro) fix P Q :: "nat \ 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 @@ -273,49 +272,48 @@ qed auto lemma sequentially_bot [simp]: "sequentially \ bot" -unfolding expand_net_eq eventually_sequentially by auto + unfolding filter_eq_iff eventually_sequentially by auto lemma eventually_False_sequentially [simp]: "\ eventually (\n. False) sequentially" -by (simp add: eventually_False) + by (simp add: eventually_False) lemma le_sequentially: - "net \ sequentially \ (\N. eventually (\n. N \ n) net)" -unfolding le_net_def eventually_sequentially -by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp) + "A \ sequentially \ (\N. eventually (\n. N \ n) A)" + unfolding le_filter_def eventually_sequentially + by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp) -definition - trivial_limit :: "'a net \ bool" where - "trivial_limit net \ eventually (\x. False) net" +definition trivial_limit :: "'a filter \ bool" + where "trivial_limit A \ eventually (\x. False) A" -lemma trivial_limit_sequentially[intro]: "\ trivial_limit sequentially" +lemma trivial_limit_sequentially [intro]: "\ trivial_limit sequentially" by (auto simp add: trivial_limit_def eventually_sequentially) -subsection {* Standard Nets *} +subsection {* Standard filters *} -definition within :: "'a net \ 'a set \ 'a net" (infixr "within" 70) where - "net within S = Abs_net (\P. eventually (\x. x \ S \ P x) net)" +definition within :: "'a filter \ 'a set \ 'a filter" (infixr "within" 70) + where "A within S = Abs_filter (\P. eventually (\x. x \ S \ P x) A)" -definition nhds :: "'a::topological_space \ 'a net" where - "nhds a = Abs_net (\P. \S. open S \ a \ S \ (\x\S. P x))" +definition nhds :: "'a::topological_space \ 'a filter" + where "nhds a = Abs_filter (\P. \S. open S \ a \ S \ (\x\S. P x))" -definition at :: "'a::topological_space \ 'a net" where - "at a = nhds a within - {a}" +definition at :: "'a::topological_space \ 'a filter" + where "at a = nhds a within - {a}" lemma eventually_within: - "eventually P (net within S) = eventually (\x. x \ S \ P x) net" -unfolding within_def -by (rule eventually_Abs_net, rule is_filter.intro) - (auto elim!: eventually_rev_mp) + "eventually P (A within S) = eventually (\x. x \ S \ P x) A" + unfolding within_def + by (rule eventually_Abs_filter, rule is_filter.intro) + (auto elim!: eventually_rev_mp) -lemma within_UNIV: "net within UNIV = net" - unfolding expand_net_eq eventually_within by simp +lemma within_UNIV: "A within UNIV = A" + unfolding filter_eq_iff eventually_within by simp lemma eventually_nhds: "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" unfolding nhds_def -proof (rule eventually_Abs_net, rule is_filter.intro) +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)" by - rule next @@ -354,52 +352,52 @@ subsection {* Boundedness *} -definition Bfun :: "('a \ 'b::real_normed_vector) \ 'a net \ bool" where - "Bfun f net = (\K>0. eventually (\x. norm (f x) \ K) net)" +definition Bfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" + where "Bfun f A = (\K>0. eventually (\x. norm (f x) \ K) A)" lemma BfunI: - assumes K: "eventually (\x. norm (f x) \ K) net" shows "Bfun f net" + assumes K: "eventually (\x. norm (f x) \ K) A" shows "Bfun f A" unfolding Bfun_def proof (intro exI conjI allI) show "0 < max K 1" by simp next - show "eventually (\x. norm (f x) \ max K 1) net" + show "eventually (\x. norm (f x) \ max K 1) A" using K by (rule eventually_elim1, simp) qed lemma BfunE: - assumes "Bfun f net" - obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) net" + assumes "Bfun f A" + obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) A" using assms unfolding Bfun_def by fast subsection {* Convergence to Zero *} -definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a net \ bool" where - "Zfun f net = (\r>0. eventually (\x. norm (f x) < r) net)" +definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" + where "Zfun f A = (\r>0. eventually (\x. norm (f x) < r) A)" lemma ZfunI: - "(\r. 0 < r \ eventually (\x. norm (f x) < r) net) \ Zfun f net" -unfolding Zfun_def by simp + "(\r. 0 < r \ eventually (\x. norm (f x) < r) A) \ Zfun f A" + unfolding Zfun_def by simp lemma ZfunD: - "\Zfun f net; 0 < r\ \ eventually (\x. norm (f x) < r) net" -unfolding Zfun_def by simp + "\Zfun f A; 0 < r\ \ eventually (\x. norm (f x) < r) A" + unfolding Zfun_def by simp lemma Zfun_ssubst: - "eventually (\x. f x = g x) net \ Zfun g net \ Zfun f net" -unfolding Zfun_def by (auto elim!: eventually_rev_mp) + "eventually (\x. f x = g x) A \ Zfun g A \ Zfun f A" + unfolding Zfun_def by (auto elim!: eventually_rev_mp) -lemma Zfun_zero: "Zfun (\x. 0) net" -unfolding Zfun_def by simp +lemma Zfun_zero: "Zfun (\x. 0) A" + unfolding Zfun_def by simp -lemma Zfun_norm_iff: "Zfun (\x. norm (f x)) net = Zfun (\x. f x) net" -unfolding Zfun_def by simp +lemma Zfun_norm_iff: "Zfun (\x. norm (f x)) A = Zfun (\x. f x) A" + unfolding Zfun_def by simp lemma Zfun_imp_Zfun: - assumes f: "Zfun f net" - assumes g: "eventually (\x. norm (g x) \ norm (f x) * K) net" - shows "Zfun (\x. g x) net" + assumes f: "Zfun f A" + assumes g: "eventually (\x. norm (g x) \ norm (f x) * K) A" + shows "Zfun (\x. g x) A" proof (cases) assume K: "0 < K" show ?thesis @@ -407,9 +405,9 @@ fix r::real assume "0 < r" hence "0 < r / K" using K by (rule divide_pos_pos) - then have "eventually (\x. norm (f x) < r / K) net" + then have "eventually (\x. norm (f x) < r / K) A" using ZfunD [OF f] by fast - with g show "eventually (\x. norm (g x) < r) net" + with g show "eventually (\x. norm (g x) < r) A" proof (rule eventually_elim2) fix x assume *: "norm (g x) \ norm (f x) * K" @@ -427,7 +425,7 @@ proof (rule ZfunI) fix r :: real assume "0 < r" - from g show "eventually (\x. norm (g x) < r) net" + from g show "eventually (\x. norm (g x) < r) A" proof (rule eventually_elim1) fix x assume "norm (g x) \ norm (f x) * K" @@ -439,22 +437,22 @@ qed qed -lemma Zfun_le: "\Zfun g net; \x. norm (f x) \ norm (g x)\ \ Zfun f net" -by (erule_tac K="1" in Zfun_imp_Zfun, simp) +lemma Zfun_le: "\Zfun g A; \x. norm (f x) \ norm (g x)\ \ Zfun f A" + by (erule_tac K="1" in Zfun_imp_Zfun, simp) lemma Zfun_add: - assumes f: "Zfun f net" and g: "Zfun g net" - shows "Zfun (\x. f x + g x) net" + assumes f: "Zfun f A" and g: "Zfun g A" + shows "Zfun (\x. f x + g x) A" proof (rule ZfunI) fix r::real assume "0 < r" hence r: "0 < r / 2" by simp - have "eventually (\x. norm (f x) < r/2) net" + have "eventually (\x. norm (f x) < r/2) A" using f r by (rule ZfunD) moreover - have "eventually (\x. norm (g x) < r/2) net" + have "eventually (\x. norm (g x) < r/2) A" using g r by (rule ZfunD) ultimately - show "eventually (\x. norm (f x + g x) < r) net" + show "eventually (\x. norm (f x + g x) < r) A" proof (rule eventually_elim2) fix x assume *: "norm (f x) < r/2" "norm (g x) < r/2" @@ -467,28 +465,28 @@ qed qed -lemma Zfun_minus: "Zfun f net \ Zfun (\x. - f x) net" -unfolding Zfun_def by simp +lemma Zfun_minus: "Zfun f A \ Zfun (\x. - f x) A" + unfolding Zfun_def by simp -lemma Zfun_diff: "\Zfun f net; Zfun g net\ \ Zfun (\x. f x - g x) net" -by (simp only: diff_minus Zfun_add Zfun_minus) +lemma Zfun_diff: "\Zfun f A; Zfun g A\ \ Zfun (\x. f x - g x) A" + by (simp only: diff_minus Zfun_add Zfun_minus) lemma (in bounded_linear) Zfun: - assumes g: "Zfun g net" - shows "Zfun (\x. f (g x)) net" + assumes g: "Zfun g A" + shows "Zfun (\x. f (g x)) A" proof - obtain K where "\x. norm (f x) \ norm x * K" using bounded by fast - then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) net" + then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) A" by simp with g show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Zfun: - assumes f: "Zfun f net" - assumes g: "Zfun g net" - shows "Zfun (\x. f x ** g x) net" + assumes f: "Zfun f A" + assumes g: "Zfun g A" + shows "Zfun (\x. f x ** g x) A" proof (rule ZfunI) fix r::real assume r: "0 < r" obtain K where K: "0 < K" @@ -496,13 +494,13 @@ using pos_bounded by fast from K have K': "0 < inverse K" by (rule positive_imp_inverse_positive) - have "eventually (\x. norm (f x) < r) net" + have "eventually (\x. norm (f x) < r) A" using f r by (rule ZfunD) moreover - have "eventually (\x. norm (g x) < inverse K) net" + have "eventually (\x. norm (g x) < inverse K) A" using g K' by (rule ZfunD) ultimately - show "eventually (\x. norm (f x ** g x) < r) net" + show "eventually (\x. norm (f x ** g x) < r) A" proof (rule eventually_elim2) fix x assume *: "norm (f x) < r" "norm (g x) < inverse K" @@ -517,12 +515,12 @@ qed lemma (in bounded_bilinear) Zfun_left: - "Zfun f net \ Zfun (\x. f x ** a) net" -by (rule bounded_linear_left [THEN bounded_linear.Zfun]) + "Zfun f A \ Zfun (\x. f x ** a) A" + by (rule bounded_linear_left [THEN bounded_linear.Zfun]) lemma (in bounded_bilinear) Zfun_right: - "Zfun f net \ Zfun (\x. a ** f x) net" -by (rule bounded_linear_right [THEN bounded_linear.Zfun]) + "Zfun f A \ Zfun (\x. a ** f x) A" + by (rule bounded_linear_right [THEN bounded_linear.Zfun]) lemmas Zfun_mult = mult.Zfun lemmas Zfun_mult_right = mult.Zfun_right @@ -531,9 +529,9 @@ subsection {* Limits *} -definition tendsto :: "('a \ 'b::topological_space) \ 'b \ 'a net \ bool" +definition tendsto :: "('a \ 'b::topological_space) \ 'b \ 'a filter \ bool" (infixr "--->" 55) where - "(f ---> l) net \ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" + "(f ---> l) A \ (\S. open S \ l \ S \ eventually (\x. f x \ S) A)" ML {* structure Tendsto_Intros = Named_Thms @@ -545,74 +543,74 @@ setup Tendsto_Intros.setup -lemma tendsto_mono: "net \ net' \ (f ---> l) net' \ (f ---> l) net" -unfolding tendsto_def le_net_def by fast +lemma tendsto_mono: "A \ A' \ (f ---> l) A' \ (f ---> l) A" + unfolding tendsto_def le_filter_def by fast lemma topological_tendstoI: - "(\S. open S \ l \ S \ eventually (\x. f x \ S) net) - \ (f ---> l) net" + "(\S. open S \ l \ S \ eventually (\x. f x \ S) A) + \ (f ---> l) A" unfolding tendsto_def by auto lemma topological_tendstoD: - "(f ---> l) net \ open S \ l \ S \ eventually (\x. f x \ S) net" + "(f ---> l) A \ open S \ l \ S \ eventually (\x. f x \ S) A" unfolding tendsto_def by auto lemma tendstoI: - assumes "\e. 0 < e \ eventually (\x. dist (f x) l < e) net" - shows "(f ---> l) net" -apply (rule topological_tendstoI) -apply (simp add: open_dist) -apply (drule (1) bspec, clarify) -apply (drule assms) -apply (erule eventually_elim1, simp) -done + assumes "\e. 0 < e \ eventually (\x. dist (f x) l < e) A" + shows "(f ---> l) A" + 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) net \ 0 < e \ eventually (\x. dist (f x) l < e) net" -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 + "(f ---> l) A \ 0 < e \ eventually (\x. dist (f x) l < e) A" + 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) net \ (\e>0. eventually (\x. dist (f x) l < e) net)" -using tendstoI tendstoD by fast + "(f ---> l) A \ (\e>0. eventually (\x. dist (f x) l < e) A)" + using tendstoI tendstoD by fast -lemma tendsto_Zfun_iff: "(f ---> a) net = Zfun (\x. f x - a) net" -by (simp only: tendsto_iff Zfun_def dist_norm) +lemma tendsto_Zfun_iff: "(f ---> a) A = Zfun (\x. f x - a) A" + by (simp only: tendsto_iff Zfun_def dist_norm) lemma tendsto_ident_at [tendsto_intros]: "((\x. x) ---> a) (at a)" -unfolding tendsto_def eventually_at_topological by auto + 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 + unfolding tendsto_def eventually_within eventually_at_topological by auto -lemma tendsto_const [tendsto_intros]: "((\x. k) ---> k) net" -by (simp add: tendsto_def) +lemma tendsto_const [tendsto_intros]: "((\x. k) ---> k) A" + by (simp add: tendsto_def) lemma tendsto_const_iff: fixes k l :: "'a::metric_space" - assumes "net \ bot" shows "((\n. k) ---> l) net \ k = l" -apply (safe intro!: tendsto_const) -apply (rule ccontr) -apply (drule_tac e="dist k l" in tendstoD) -apply (simp add: zero_less_dist_iff) -apply (simp add: eventually_False assms) -done + assumes "A \ bot" shows "((\n. k) ---> l) A \ k = l" + apply (safe intro!: tendsto_const) + apply (rule ccontr) + apply (drule_tac e="dist k l" in tendstoD) + apply (simp add: zero_less_dist_iff) + apply (simp add: eventually_False assms) + done lemma tendsto_dist [tendsto_intros]: - assumes f: "(f ---> l) net" and g: "(g ---> m) net" - shows "((\x. dist (f x) (g x)) ---> dist l m) net" + assumes f: "(f ---> l) A" and g: "(g ---> m) A" + shows "((\x. dist (f x) (g x)) ---> dist l m) A" 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) net" + show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) A" proof (rule eventually_elim2) fix x assume "dist (f x) l < e/2" "dist (g x) m < e/2" then show "dist (dist (f x) (g x)) (dist l m) < e" @@ -626,48 +624,48 @@ qed lemma norm_conv_dist: "norm x = dist x 0" -unfolding dist_norm by simp + unfolding dist_norm by simp lemma tendsto_norm [tendsto_intros]: - "(f ---> a) net \ ((\x. norm (f x)) ---> norm a) net" -unfolding norm_conv_dist by (intro tendsto_intros) + "(f ---> a) A \ ((\x. norm (f x)) ---> norm a) A" + unfolding norm_conv_dist by (intro tendsto_intros) lemma tendsto_norm_zero: - "(f ---> 0) net \ ((\x. norm (f x)) ---> 0) net" -by (drule tendsto_norm, simp) + "(f ---> 0) A \ ((\x. norm (f x)) ---> 0) A" + by (drule tendsto_norm, simp) lemma tendsto_norm_zero_cancel: - "((\x. norm (f x)) ---> 0) net \ (f ---> 0) net" -unfolding tendsto_iff dist_norm by simp + "((\x. norm (f x)) ---> 0) A \ (f ---> 0) A" + unfolding tendsto_iff dist_norm by simp lemma tendsto_norm_zero_iff: - "((\x. norm (f x)) ---> 0) net \ (f ---> 0) net" -unfolding tendsto_iff dist_norm by simp + "((\x. norm (f x)) ---> 0) A \ (f ---> 0) A" + unfolding tendsto_iff dist_norm by simp lemma tendsto_add [tendsto_intros]: fixes a b :: "'a::real_normed_vector" - shows "\(f ---> a) net; (g ---> b) net\ \ ((\x. f x + g x) ---> a + b) net" -by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) + shows "\(f ---> a) A; (g ---> b) A\ \ ((\x. f x + g x) ---> a + b) A" + by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) lemma tendsto_minus [tendsto_intros]: fixes a :: "'a::real_normed_vector" - shows "(f ---> a) net \ ((\x. - f x) ---> - a) net" -by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) + shows "(f ---> a) A \ ((\x. - f x) ---> - a) A" + by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) lemma tendsto_minus_cancel: fixes a :: "'a::real_normed_vector" - shows "((\x. - f x) ---> - a) net \ (f ---> a) net" -by (drule tendsto_minus, simp) + shows "((\x. - f x) ---> - a) A \ (f ---> a) A" + by (drule tendsto_minus, simp) lemma tendsto_diff [tendsto_intros]: fixes a b :: "'a::real_normed_vector" - shows "\(f ---> a) net; (g ---> b) net\ \ ((\x. f x - g x) ---> a - b) net" -by (simp add: diff_minus tendsto_add tendsto_minus) + shows "\(f ---> a) A; (g ---> b) A\ \ ((\x. f x - g x) ---> a - b) A" + by (simp add: diff_minus tendsto_add tendsto_minus) lemma tendsto_setsum [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::real_normed_vector" - assumes "\i. i \ S \ (f i ---> a i) net" - shows "((\x. \i\S. f i x) ---> (\i\S. a i)) net" + assumes "\i. i \ S \ (f i ---> a i) A" + shows "((\x. \i\S. f i x) ---> (\i\S. a i)) A" proof (cases "finite S") assume "finite S" thus ?thesis using assms proof (induct set: finite) @@ -683,29 +681,29 @@ qed lemma (in bounded_linear) tendsto [tendsto_intros]: - "(g ---> a) net \ ((\x. f (g x)) ---> f a) net" -by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) + "(g ---> a) A \ ((\x. f (g x)) ---> f a) A" + by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) lemma (in bounded_bilinear) tendsto [tendsto_intros]: - "\(f ---> a) net; (g ---> b) net\ \ ((\x. f x ** g x) ---> a ** b) net" -by (simp only: tendsto_Zfun_iff prod_diff_prod - Zfun_add Zfun Zfun_left Zfun_right) + "\(f ---> a) A; (g ---> b) A\ \ ((\x. f x ** g x) ---> a ** b) A" + by (simp only: tendsto_Zfun_iff prod_diff_prod + Zfun_add Zfun Zfun_left Zfun_right) subsection {* Continuity of Inverse *} lemma (in bounded_bilinear) Zfun_prod_Bfun: - assumes f: "Zfun f net" - assumes g: "Bfun g net" - shows "Zfun (\x. f x ** g x) net" + assumes f: "Zfun f A" + assumes g: "Bfun g A" + shows "Zfun (\x. f x ** g x) A" proof - obtain K where K: "0 \ K" and norm_le: "\x y. norm (x ** y) \ norm x * norm y * K" using nonneg_bounded by fast obtain B where B: "0 < B" - and norm_g: "eventually (\x. norm (g x) \ B) net" + and norm_g: "eventually (\x. norm (g x) \ B) A" using g by (rule BfunE) - have "eventually (\x. norm (f x ** g x) \ norm (f x) * (B * K)) net" + have "eventually (\x. norm (f x ** g x) \ norm (f x) * (B * K)) A" using norm_g proof (rule eventually_elim1) fix x assume *: "norm (g x) \ B" @@ -724,39 +722,39 @@ lemma (in bounded_bilinear) flip: "bounded_bilinear (\x y. y ** x)" -apply default -apply (rule add_right) -apply (rule add_left) -apply (rule scaleR_right) -apply (rule scaleR_left) -apply (subst mult_commute) -using bounded by fast + apply default + apply (rule add_right) + apply (rule add_left) + apply (rule scaleR_right) + apply (rule scaleR_left) + apply (subst mult_commute) + using bounded by fast lemma (in bounded_bilinear) Bfun_prod_Zfun: - assumes f: "Bfun f net" - assumes g: "Zfun g net" - shows "Zfun (\x. f x ** g x) net" -using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) + assumes f: "Bfun f A" + assumes g: "Zfun g A" + shows "Zfun (\x. f x ** g x) A" + using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) lemma Bfun_inverse_lemma: fixes x :: "'a::real_normed_div_algebra" shows "\r \ norm x; 0 < r\ \ norm (inverse x) \ inverse r" -apply (subst nonzero_norm_inverse, clarsimp) -apply (erule (1) le_imp_inverse_le) -done + apply (subst nonzero_norm_inverse, clarsimp) + apply (erule (1) le_imp_inverse_le) + done lemma Bfun_inverse: fixes a :: "'a::real_normed_div_algebra" - assumes f: "(f ---> a) net" + assumes f: "(f ---> a) A" assumes a: "a \ 0" - shows "Bfun (\x. inverse (f x)) net" + shows "Bfun (\x. inverse (f x)) A" proof - from a have "0 < norm a" by simp hence "\r>0. r < norm a" by (rule dense) then obtain r where r1: "0 < r" and r2: "r < norm a" by fast - have "eventually (\x. dist (f x) a < r) net" + have "eventually (\x. dist (f x) a < r) A" using tendstoD [OF f r1] by fast - hence "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) net" + hence "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) A" proof (rule eventually_elim1) fix x assume "dist (f x) a < r" @@ -783,29 +781,29 @@ lemma tendsto_inverse_lemma: fixes a :: "'a::real_normed_div_algebra" - shows "\(f ---> a) net; a \ 0; eventually (\x. f x \ 0) net\ - \ ((\x. inverse (f x)) ---> inverse a) net" -apply (subst tendsto_Zfun_iff) -apply (rule Zfun_ssubst) -apply (erule eventually_elim1) -apply (erule (1) inverse_diff_inverse) -apply (rule Zfun_minus) -apply (rule Zfun_mult_left) -apply (rule mult.Bfun_prod_Zfun) -apply (erule (1) Bfun_inverse) -apply (simp add: tendsto_Zfun_iff) -done + shows "\(f ---> a) A; a \ 0; eventually (\x. f x \ 0) A\ + \ ((\x. inverse (f x)) ---> inverse a) A" + apply (subst tendsto_Zfun_iff) + apply (rule Zfun_ssubst) + apply (erule eventually_elim1) + apply (erule (1) inverse_diff_inverse) + apply (rule Zfun_minus) + apply (rule Zfun_mult_left) + apply (rule mult.Bfun_prod_Zfun) + apply (erule (1) Bfun_inverse) + apply (simp add: tendsto_Zfun_iff) + done lemma tendsto_inverse [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" - assumes f: "(f ---> a) net" + assumes f: "(f ---> a) A" assumes a: "a \ 0" - shows "((\x. inverse (f x)) ---> inverse a) net" + shows "((\x. inverse (f x)) ---> inverse a) A" proof - from a have "0 < norm a" by simp - with f have "eventually (\x. dist (f x) a < norm a) net" + with f have "eventually (\x. dist (f x) a < norm a) A" by (rule tendstoD) - then have "eventually (\x. f x \ 0) net" + then have "eventually (\x. f x \ 0) A" unfolding dist_norm by (auto elim!: eventually_elim1) with f a show ?thesis by (rule tendsto_inverse_lemma) @@ -813,32 +811,32 @@ lemma tendsto_divide [tendsto_intros]: fixes a b :: "'a::real_normed_field" - shows "\(f ---> a) net; (g ---> b) net; b \ 0\ - \ ((\x. f x / g x) ---> a / b) net" -by (simp add: mult.tendsto tendsto_inverse divide_inverse) + shows "\(f ---> a) A; (g ---> b) A; b \ 0\ + \ ((\x. f x / g x) ---> a / b) A" + by (simp add: mult.tendsto tendsto_inverse divide_inverse) lemma tendsto_unique: fixes f :: "'a \ 'b::t2_space" - assumes "\ trivial_limit net" "(f ---> l) net" "(f ---> l') net" + assumes "\ trivial_limit A" "(f ---> l) A" "(f ---> l') A" shows "l = l'" proof (rule ccontr) assume "l \ l'" obtain U V where "open U" "open V" "l \ U" "l' \ V" "U \ V = {}" using hausdorff [OF `l \ l'`] by fast - have "eventually (\x. f x \ U) net" - using `(f ---> l) net` `open U` `l \ U` by (rule topological_tendstoD) + have "eventually (\x. f x \ U) A" + using `(f ---> l) A` `open U` `l \ U` by (rule topological_tendstoD) moreover - have "eventually (\x. f x \ V) net" - using `(f ---> l') net` `open V` `l' \ V` by (rule topological_tendstoD) + have "eventually (\x. f x \ V) A" + using `(f ---> l') A` `open V` `l' \ V` by (rule topological_tendstoD) ultimately - have "eventually (\x. False) net" + have "eventually (\x. False) A" proof (rule eventually_elim2) fix x assume "f x \ U" "f x \ V" hence "f x \ U \ V" by simp with `U \ V = {}` show "False" by simp qed - with `\ trivial_limit net` show "False" + with `\ trivial_limit A` show "False" by (simp add: trivial_limit_def) qed diff -r 53d95b52954c -r 730f7cced3a6 src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Mon Aug 08 18:36:32 2011 -0700 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Mon Aug 08 19:26:53 2011 -0700 @@ -18,7 +18,7 @@ nets of a particular form. This lets us prove theorems generally and use "at a" or "at a within s" for restriction to a set (1-sided on R etc.) *} -definition has_derivative :: "('a::real_normed_vector \ 'b::real_normed_vector) \ ('a \ 'b) \ ('a net \ bool)" +definition has_derivative :: "('a::real_normed_vector \ 'b::real_normed_vector) \ ('a \ 'b) \ ('a filter \ bool)" (infixl "(has'_derivative)" 12) where "(f has_derivative f') net \ bounded_linear f' \ ((\y. (1 / (norm (y - netlimit net))) *\<^sub>R (f y - (f (netlimit net) + f'(y - netlimit net)))) ---> 0) net" @@ -291,7 +291,7 @@ no_notation Deriv.differentiable (infixl "differentiable" 60) -definition differentiable :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a net \ bool" (infixr "differentiable" 30) where +definition differentiable :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a filter \ bool" (infixr "differentiable" 30) where "f differentiable net \ (\f'. (f has_derivative f') net)" definition differentiable_on :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a set \ bool" (infixr "differentiable'_on" 30) where @@ -469,25 +469,25 @@ subsection {* Composition rules stated just for differentiability. *} -lemma differentiable_const[intro]: "(\z. c) differentiable (net::'a::real_normed_vector net)" +lemma differentiable_const[intro]: "(\z. c) differentiable (net::'a::real_normed_vector filter)" unfolding differentiable_def using has_derivative_const by auto -lemma differentiable_id[intro]: "(\z. z) differentiable (net::'a::real_normed_vector net)" +lemma differentiable_id[intro]: "(\z. z) differentiable (net::'a::real_normed_vector filter)" unfolding differentiable_def using has_derivative_id by auto -lemma differentiable_cmul[intro]: "f differentiable net \ (\x. c *\<^sub>R f(x)) differentiable (net::'a::real_normed_vector net)" +lemma differentiable_cmul[intro]: "f differentiable net \ (\x. c *\<^sub>R f(x)) differentiable (net::'a::real_normed_vector filter)" unfolding differentiable_def apply(erule exE, drule has_derivative_cmul) by auto -lemma differentiable_neg[intro]: "f differentiable net \ (\z. -(f z)) differentiable (net::'a::real_normed_vector net)" +lemma differentiable_neg[intro]: "f differentiable net \ (\z. -(f z)) differentiable (net::'a::real_normed_vector filter)" unfolding differentiable_def apply(erule exE, drule has_derivative_neg) by auto lemma differentiable_add: "f differentiable net \ g differentiable net - \ (\z. f z + g z) differentiable (net::'a::real_normed_vector net)" + \ (\z. f z + g z) differentiable (net::'a::real_normed_vector filter)" unfolding differentiable_def apply(erule exE)+ apply(rule_tac x="\z. f' z + f'a z" in exI) apply(rule has_derivative_add) by auto lemma differentiable_sub: "f differentiable net \ g differentiable net - \ (\z. f z - g z) differentiable (net::'a::real_normed_vector net)" + \ (\z. f z - g z) differentiable (net::'a::real_normed_vector filter)" unfolding differentiable_def apply(erule exE)+ apply(rule_tac x="\z. f' z - f'a z" in exI) apply(rule has_derivative_sub) by auto @@ -1259,7 +1259,7 @@ subsection {* Considering derivative @{typ "real \ 'b\real_normed_vector"} as a vector. *} -definition has_vector_derivative :: "(real \ 'b::real_normed_vector) \ 'b \ (real net \ bool)" +definition has_vector_derivative :: "(real \ 'b::real_normed_vector) \ 'b \ (real filter \ bool)" (infixl "has'_vector'_derivative" 12) where "(f has_vector_derivative f') net \ (f has_derivative (\x. x *\<^sub>R f')) net" diff -r 53d95b52954c -r 730f7cced3a6 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Aug 08 18:36:32 2011 -0700 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Aug 08 19:26:53 2011 -0700 @@ -882,24 +882,25 @@ using frontier_complement frontier_subset_eq[of "- S"] unfolding open_closed by auto -subsection {* Nets and the ``eventually true'' quantifier *} - -text {* Common nets and The "within" modifier for nets. *} +subsection {* Filters and the ``eventually true'' quantifier *} + +text {* Common filters and The "within" modifier for filters. *} definition - at_infinity :: "'a::real_normed_vector net" where - "at_infinity = Abs_net (\P. \r. \x. r \ norm x \ P x)" + at_infinity :: "'a::real_normed_vector filter" where + "at_infinity = Abs_filter (\P. \r. \x. r \ norm x \ P x)" definition - indirection :: "'a::real_normed_vector \ 'a \ 'a net" (infixr "indirection" 70) where + indirection :: "'a::real_normed_vector \ 'a \ 'a filter" + (infixr "indirection" 70) where "a indirection v = (at a) within {b. \c\0. b - a = scaleR c v}" -text{* Prove That They are all nets. *} +text{* Prove That They are all filters. *} lemma eventually_at_infinity: "eventually P at_infinity \ (\b. \x. norm x >= b \ P x)" unfolding at_infinity_def -proof (rule eventually_Abs_net, rule is_filter.intro) +proof (rule eventually_Abs_filter, rule is_filter.intro) fix P Q :: "'a \ bool" assume "\r. \x. r \ norm x \ P x" and "\s. \x. s \ norm x \ Q x" then obtain r s where @@ -944,7 +945,7 @@ by (simp add: trivial_limit_at_iff) lemma trivial_limit_at_infinity: - "\ trivial_limit (at_infinity :: ('a::{real_normed_vector,perfect_space}) net)" + "\ trivial_limit (at_infinity :: ('a::{real_normed_vector,perfect_space}) filter)" unfolding trivial_limit_def eventually_at_infinity apply clarsimp apply (subgoal_tac "\x::'a. x \ 0", clarify) @@ -972,12 +973,6 @@ unfolding trivial_limit_def by (auto elim: eventually_rev_mp) -lemma always_eventually: "(\x. P x) ==> eventually P net" -proof - - assume "\x. P x" hence "P = (\x. True)" by (simp add: ext) - thus "eventually P net" by simp -qed - lemma trivial_limit_eventually: "trivial_limit net \ eventually P net" unfolding trivial_limit_def by (auto elim: eventually_rev_mp) @@ -1012,10 +1007,10 @@ subsection {* Limits *} - text{* Notation Lim to avoid collition with lim defined in analysis *} -definition - Lim :: "'a net \ ('a \ 'b::t2_space) \ 'b" where - "Lim net f = (THE l. (f ---> l) net)" +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)" lemma Lim: "(f ---> l) net \ @@ -1281,9 +1276,9 @@ using assms by (rule scaleR.tendsto) lemma Lim_inv: - fixes f :: "'a \ real" - assumes "(f ---> l) (net::'a net)" "l \ 0" - shows "((inverse o f) ---> inverse l) net" + fixes f :: "'a \ real" and A :: "'a filter" + assumes "(f ---> l) A" and "l \ 0" + shows "((inverse o f) ---> inverse l) A" unfolding o_def using assms by (rule tendsto_inverse) lemma Lim_vmul: @@ -1485,10 +1480,10 @@ thus "?lhs" by (rule topological_tendstoI) qed -text{* It's also sometimes useful to extract the limit point from the net. *} +text{* It's also sometimes useful to extract the limit point from the filter. *} definition - netlimit :: "'a::t2_space net \ 'a" where + netlimit :: "'a::t2_space filter \ 'a" where "netlimit net = (SOME a. ((\x. x) ---> a) net)" lemma netlimit_within: @@ -1943,7 +1938,7 @@ lemma at_within_interior: "x \ interior S \ at x within S = at x" - by (simp add: expand_net_eq eventually_within_interior) + by (simp add: filter_eq_iff eventually_within_interior) lemma lim_within_interior: "x \ interior S \ (f ---> l) (at x within S) \ (f ---> l) (at x)" @@ -3338,8 +3333,8 @@ text {* Define continuity over a net to take in restrictions of the set. *} definition - continuous :: "'a::t2_space net \ ('a \ 'b::topological_space) \ bool" where - "continuous net f \ (f ---> f(netlimit net)) net" + 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"