# HG changeset patch # User huffman # Date 1272221919 25200 # Node ID 246493d61204a415134994bc49561e95a8fe2889 # Parent 2623a1987e1d92deafeb5c3c11f60696a03d6734 define nets directly as filters, instead of as filter bases diff -r 2623a1987e1d -r 246493d61204 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Sun Apr 25 10:23:03 2010 -0700 +++ b/src/HOL/Limits.thy Sun Apr 25 11:58:39 2010 -0700 @@ -11,62 +11,55 @@ subsection {* Nets *} text {* - A net is now defined as a filter base. - The definition also allows non-proper filter bases. + A net is now defined simply as a filter. + The 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)" + typedef (open) 'a net = - "{net :: 'a set set. (\A. A \ net) - \ (\A\net. \B\net. \C\net. C \ A \ C \ B)}" + "{net :: ('a \ bool) \ bool. is_filter net}" proof - show "UNIV \ ?net" by auto + show "(\x. True) \ ?net" by (auto intro: is_filter.intro) qed -lemma Rep_net_nonempty: "\A. A \ Rep_net net" -using Rep_net [of net] by simp - -lemma Rep_net_directed: - "A \ Rep_net net \ B \ Rep_net net \ \C\Rep_net net. C \ A \ C \ B" +lemma is_filter_Rep_net: "is_filter (Rep_net net)" using Rep_net [of net] by simp lemma Abs_net_inverse': - assumes "\A. A \ net" - assumes "\A B. A \ net \ B \ net \ \C\net. C \ A \ C \ B" - shows "Rep_net (Abs_net net) = net" + assumes "is_filter net" shows "Rep_net (Abs_net net) = net" using assms by (simp add: Abs_net_inverse) -lemma image_nonempty: "\x. x \ A \ \x. x \ f ` A" -by auto - subsection {* Eventually *} definition eventually :: "('a \ bool) \ 'a net \ bool" where - [code del]: "eventually P net \ (\A\Rep_net net. \x\A. P x)" + [code del]: "eventually P net \ Rep_net net 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_True [simp]: "eventually (\x. True) net" -unfolding eventually_def using Rep_net_nonempty [of net] by fast +unfolding eventually_def +by (rule is_filter.True [OF is_filter_Rep_net]) lemma eventually_mono: "(\x. P x \ Q x) \ eventually P net \ eventually Q net" -unfolding eventually_def by blast +unfolding eventually_def +by (rule is_filter.mono [OF is_filter_Rep_net]) 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" -proof - - obtain A where A: "A \ Rep_net net" "\x\A. P x" - using P unfolding eventually_def by fast - obtain B where B: "B \ Rep_net net" "\x\B. Q x" - using Q unfolding eventually_def by fast - obtain C where C: "C \ Rep_net net" "C \ A" "C \ B" - using Rep_net_directed [OF A(1) B(1)] by fast - then have "\x\C. P x \ Q x" "C \ Rep_net net" - using A(2) B(2) by auto - then show ?thesis unfolding eventually_def .. -qed +using assms unfolding eventually_def +by (rule is_filter.conj [OF is_filter_Rep_net]) lemma eventually_mp: assumes "eventually (\x. P x \ Q x) net" @@ -105,57 +98,54 @@ subsection {* Standard Nets *} definition - sequentially :: "nat net" where - [code del]: "sequentially = Abs_net (range (\n. {n..}))" - -definition - within :: "'a net \ 'a set \ 'a net" (infixr "within" 70) where - [code del]: "net within S = Abs_net ((\A. A \ S) ` Rep_net net)" + sequentially :: "nat net" +where [code del]: + "sequentially = Abs_net (\P. \k. \n\k. P n)" definition - at :: "'a::topological_space \ 'a net" where - [code del]: "at a = Abs_net ((\S. S - {a}) ` {S. open S \ a \ S})" - -lemma Rep_net_sequentially: - "Rep_net sequentially = range (\n. {n..})" -unfolding sequentially_def -apply (rule Abs_net_inverse') -apply (rule image_nonempty, simp) -apply (clarsimp, rename_tac m n) -apply (rule_tac x="max m n" in exI, auto) -done + within :: "'a net \ 'a set \ 'a net" (infixr "within" 70) +where [code del]: + "net within S = Abs_net (\P. eventually (\x. x \ S \ P x) net)" -lemma Rep_net_within: - "Rep_net (net within S) = (\A. A \ S) ` Rep_net net" -unfolding within_def -apply (rule Abs_net_inverse') -apply (rule image_nonempty, rule Rep_net_nonempty) -apply (clarsimp, rename_tac A B) -apply (drule (1) Rep_net_directed) -apply (clarify, rule_tac x=C in bexI, auto) -done - -lemma Rep_net_at: - "Rep_net (at a) = ((\S. S - {a}) ` {S. open S \ a \ S})" -unfolding at_def -apply (rule Abs_net_inverse') -apply (rule image_nonempty) -apply (rule_tac x="UNIV" in exI, simp) -apply (clarsimp, rename_tac S T) -apply (rule_tac x="S \ T" in exI, auto simp add: open_Int) -done +definition + at :: "'a::topological_space \ 'a net" +where [code del]: + "at a = Abs_net (\P. \S. open S \ a \ S \ (\x\S. x \ a \ P x))" lemma eventually_sequentially: "eventually P sequentially \ (\N. \n\N. P n)" -unfolding eventually_def Rep_net_sequentially by auto +unfolding sequentially_def +proof (rule eventually_Abs_net, 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 + 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_within: "eventually P (net within S) = eventually (\x. x \ S \ P x) net" -unfolding eventually_def Rep_net_within by auto +unfolding within_def +by (rule eventually_Abs_net, rule is_filter.intro) + (auto elim!: eventually_rev_mp) lemma eventually_at_topological: "eventually P (at a) \ (\S. open S \ a \ S \ (\x\S. x \ a \ P x))" -unfolding eventually_def Rep_net_at by auto +unfolding at_def +proof (rule eventually_Abs_net, rule is_filter.intro) + have "open UNIV \ a \ UNIV \ (\x\UNIV. x \ a \ True)" by simp + thus "\S. open S \ a \ S \ (\x\S. x \ a \ True)" by - rule +next + fix P Q + assume "\S. open S \ a \ S \ (\x\S. x \ a \ P x)" + and "\T. open T \ a \ T \ (\x\T. x \ a \ Q x)" + then obtain S T where + "open S \ a \ S \ (\x\S. x \ a \ P x)" + "open T \ a \ T \ (\x\T. x \ a \ Q x)" by auto + hence "open (S \ T) \ a \ S \ T \ (\x\(S \ T). x \ a \ P x \ Q x)" + by (simp add: open_Int) + thus "\S. open S \ a \ S \ (\x\S. x \ a \ P x \ Q x)" by - rule +qed auto lemma eventually_at: fixes a :: "'a::metric_space" diff -r 2623a1987e1d -r 246493d61204 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Sun Apr 25 10:23:03 2010 -0700 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Sun Apr 25 11:58:39 2010 -0700 @@ -968,7 +968,7 @@ definition at_infinity :: "'a::real_normed_vector net" where - "at_infinity = Abs_net (range (\r. {x. r \ norm x}))" + "at_infinity = Abs_net (\P. \r. \x. r \ norm x \ P x)" definition indirection :: "'a::real_normed_vector \ 'a \ 'a net" (infixr "indirection" 70) where @@ -976,23 +976,32 @@ text{* Prove That They are all nets. *} -lemma Rep_net_at_infinity: - "Rep_net at_infinity = range (\r. {x. r \ norm x})" +(* TODO: move to HOL/Limits.thy *) +lemma expand_net_eq: + "net = net' \ (\P. eventually P net = eventually P net')" + unfolding Rep_net_inject [symmetric] expand_fun_eq eventually_def .. + +(* TODO: move to HOL/Limits.thy *) +lemma within_UNIV: "net within UNIV = net" + unfolding expand_net_eq eventually_within by simp + +lemma eventually_at_infinity: + "eventually P at_infinity \ (\b. \x. norm x >= b \ P x)" unfolding at_infinity_def -apply (rule Abs_net_inverse') -apply (rule image_nonempty, simp) -apply (clarsimp, rename_tac r s) -apply (rule_tac x="max r s" in exI, auto) -done - -lemma within_UNIV: "net within UNIV = net" - by (simp add: Rep_net_inject [symmetric] Rep_net_within) +proof (rule eventually_Abs_net, 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 + "\x. r \ norm x \ P x" and "\x. s \ norm x \ Q x" by auto + then have "\x. max r s \ norm x \ P x \ Q x" by simp + then show "\r. \x. r \ norm x \ P x \ Q x" .. +qed auto subsection{* Identify Trivial limits, where we can't approach arbitrarily closely. *} definition trivial_limit :: "'a net \ bool" where - "trivial_limit net \ {} \ Rep_net net" + "trivial_limit net \ eventually (\x. False) net" lemma trivial_limit_within: shows "trivial_limit (at a within S) \ \ a islimpt S" @@ -1000,21 +1009,21 @@ assume "trivial_limit (at a within S)" thus "\ a islimpt S" unfolding trivial_limit_def - unfolding Rep_net_within Rep_net_at + unfolding eventually_within eventually_at_topological unfolding islimpt_def apply (clarsimp simp add: expand_set_eq) apply (rename_tac T, rule_tac x=T in exI) - apply (clarsimp, drule_tac x=y in spec, simp) + apply (clarsimp, drule_tac x=y in bspec, simp_all) done next assume "\ a islimpt S" thus "trivial_limit (at a within S)" unfolding trivial_limit_def - unfolding Rep_net_within Rep_net_at + unfolding eventually_within eventually_at_topological unfolding islimpt_def - apply (clarsimp simp add: image_image) - apply (rule_tac x=T in image_eqI) - apply (auto simp add: expand_set_eq) + apply clarsimp + apply (rule_tac x=T in exI) + apply auto done qed @@ -1030,14 +1039,14 @@ lemma trivial_limit_at_infinity: "\ trivial_limit (at_infinity :: ('a::{real_normed_vector,zero_neq_one}) net)" (* FIXME: find a more appropriate type class *) - unfolding trivial_limit_def Rep_net_at_infinity - apply (clarsimp simp add: expand_set_eq) - apply (drule_tac x="scaleR r (sgn 1)" in spec) + unfolding trivial_limit_def eventually_at_infinity + apply clarsimp + apply (rule_tac x="scaleR b (sgn 1)" in exI) apply (simp add: norm_sgn) done lemma trivial_limit_sequentially[intro]: "\ trivial_limit sequentially" - by (auto simp add: trivial_limit_def Rep_net_sequentially) + by (auto simp add: trivial_limit_def eventually_sequentially) subsection{* Some property holds "sufficiently close" to the limit point. *} @@ -1045,10 +1054,6 @@ "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_at_infinity: - "eventually P at_infinity \ (\b. \x. norm x >= b \ P x)" -unfolding eventually_def Rep_net_at_infinity by auto - lemma 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 @@ -1059,18 +1064,20 @@ by auto (metis Rats_dense_in_nn_real order_le_less_trans order_refl) lemma eventually_happens: "eventually P net ==> trivial_limit net \ (\x. P x)" - unfolding eventually_def trivial_limit_def - using Rep_net_nonempty [of net] by auto + unfolding trivial_limit_def + by (auto elim: eventually_rev_mp) lemma always_eventually: "(\x. P x) ==> eventually P net" - unfolding eventually_def trivial_limit_def - using Rep_net_nonempty [of net] by auto +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 eventually_def by auto + unfolding trivial_limit_def by (auto elim: eventually_rev_mp) lemma eventually_False: "eventually (\x. False) net \ trivial_limit net" - unfolding trivial_limit_def eventually_def by auto + unfolding trivial_limit_def .. lemma trivial_limit_eq: "trivial_limit net \ (\P. eventually P net)" apply (safe elim!: trivial_limit_eventually)