| author | wenzelm | 
| Sat, 14 Jan 2012 13:22:39 +0100 | |
| changeset 46206 | d3d62b528487 | 
| parent 45892 | 8dcf6692433f | 
| child 46886 | 4cd29473c65d | 
| permissions | -rw-r--r-- | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 1 | (* Title : Limits.thy | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 2 | Author : Brian Huffman | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 3 | *) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 4 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 5 | header {* Filters and Limits *}
 | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 6 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 7 | theory Limits | 
| 36822 | 8 | imports RealVector | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 9 | begin | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 10 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 11 | subsection {* Filters *}
 | 
| 31392 | 12 | |
| 13 | text {*
 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 14 | This definition also allows non-proper filters. | 
| 31392 | 15 | *} | 
| 16 | ||
| 36358 
246493d61204
define nets directly as filters, instead of as filter bases
 huffman parents: 
31902diff
changeset | 17 | locale is_filter = | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 18 |   fixes F :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 19 | assumes True: "F (\<lambda>x. True)" | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 20 | assumes conj: "F (\<lambda>x. P x) \<Longrightarrow> F (\<lambda>x. Q x) \<Longrightarrow> F (\<lambda>x. P x \<and> Q x)" | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 21 | assumes mono: "\<forall>x. P x \<longrightarrow> Q x \<Longrightarrow> F (\<lambda>x. P x) \<Longrightarrow> F (\<lambda>x. Q x)" | 
| 36358 
246493d61204
define nets directly as filters, instead of as filter bases
 huffman parents: 
31902diff
changeset | 22 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 23 | typedef (open) 'a filter = "{F :: ('a \<Rightarrow> bool) \<Rightarrow> bool. is_filter F}"
 | 
| 31392 | 24 | proof | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 25 | show "(\<lambda>x. True) \<in> ?filter" by (auto intro: is_filter.intro) | 
| 31392 | 26 | qed | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 27 | |
| 44195 | 28 | lemma is_filter_Rep_filter: "is_filter (Rep_filter F)" | 
| 29 | using Rep_filter [of F] by simp | |
| 31392 | 30 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 31 | lemma Abs_filter_inverse': | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 32 | assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F" | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 33 | using assms by (simp add: Abs_filter_inverse) | 
| 31392 | 34 | |
| 35 | ||
| 36 | subsection {* Eventually *}
 | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 37 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 38 | definition eventually :: "('a \<Rightarrow> bool) \<Rightarrow> 'a filter \<Rightarrow> bool"
 | 
| 44195 | 39 | where "eventually P F \<longleftrightarrow> Rep_filter F P" | 
| 36358 
246493d61204
define nets directly as filters, instead of as filter bases
 huffman parents: 
31902diff
changeset | 40 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 41 | lemma eventually_Abs_filter: | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 42 | assumes "is_filter F" shows "eventually P (Abs_filter F) = F P" | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 43 | unfolding eventually_def using assms by (simp add: Abs_filter_inverse) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 44 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 45 | lemma filter_eq_iff: | 
| 44195 | 46 | shows "F = F' \<longleftrightarrow> (\<forall>P. eventually P F = eventually P F')" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 47 | unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def .. | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 48 | |
| 44195 | 49 | lemma eventually_True [simp]: "eventually (\<lambda>x. True) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 50 | unfolding eventually_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 51 | by (rule is_filter.True [OF is_filter_Rep_filter]) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 52 | |
| 44195 | 53 | lemma always_eventually: "\<forall>x. P x \<Longrightarrow> eventually P F" | 
| 36630 | 54 | proof - | 
| 55 | assume "\<forall>x. P x" hence "P = (\<lambda>x. True)" by (simp add: ext) | |
| 44195 | 56 | thus "eventually P F" by simp | 
| 36630 | 57 | qed | 
| 58 | ||
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 59 | lemma eventually_mono: | 
| 44195 | 60 | "(\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually P F \<Longrightarrow> eventually Q F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 61 | unfolding eventually_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 62 | by (rule is_filter.mono [OF is_filter_Rep_filter]) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 63 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 64 | lemma eventually_conj: | 
| 44195 | 65 | assumes P: "eventually (\<lambda>x. P x) F" | 
| 66 | assumes Q: "eventually (\<lambda>x. Q x) F" | |
| 67 | shows "eventually (\<lambda>x. P x \<and> Q x) F" | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 68 | using assms unfolding eventually_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 69 | by (rule is_filter.conj [OF is_filter_Rep_filter]) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 70 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 71 | lemma eventually_mp: | 
| 44195 | 72 | assumes "eventually (\<lambda>x. P x \<longrightarrow> Q x) F" | 
| 73 | assumes "eventually (\<lambda>x. P x) F" | |
| 74 | shows "eventually (\<lambda>x. Q x) F" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 75 | proof (rule eventually_mono) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 76 | show "\<forall>x. (P x \<longrightarrow> Q x) \<and> P x \<longrightarrow> Q x" by simp | 
| 44195 | 77 | show "eventually (\<lambda>x. (P x \<longrightarrow> Q x) \<and> P x) F" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 78 | using assms by (rule eventually_conj) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 79 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 80 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 81 | lemma eventually_rev_mp: | 
| 44195 | 82 | assumes "eventually (\<lambda>x. P x) F" | 
| 83 | assumes "eventually (\<lambda>x. P x \<longrightarrow> Q x) F" | |
| 84 | shows "eventually (\<lambda>x. Q x) F" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 85 | using assms(2) assms(1) by (rule eventually_mp) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 86 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 87 | lemma eventually_conj_iff: | 
| 44195 | 88 | "eventually (\<lambda>x. P x \<and> Q x) F \<longleftrightarrow> eventually P F \<and> eventually Q F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 89 | by (auto intro: eventually_conj elim: eventually_rev_mp) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 90 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 91 | lemma eventually_elim1: | 
| 44195 | 92 | assumes "eventually (\<lambda>i. P i) F" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 93 | assumes "\<And>i. P i \<Longrightarrow> Q i" | 
| 44195 | 94 | shows "eventually (\<lambda>i. Q i) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 95 | using assms by (auto elim!: eventually_rev_mp) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 96 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 97 | lemma eventually_elim2: | 
| 44195 | 98 | assumes "eventually (\<lambda>i. P i) F" | 
| 99 | assumes "eventually (\<lambda>i. Q i) F" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 100 | assumes "\<And>i. P i \<Longrightarrow> Q i \<Longrightarrow> R i" | 
| 44195 | 101 | shows "eventually (\<lambda>i. R i) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 102 | using assms by (auto elim!: eventually_rev_mp) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 103 | |
| 45892 | 104 | lemma eventually_subst: | 
| 105 | assumes "eventually (\<lambda>n. P n = Q n) F" | |
| 106 | shows "eventually P F = eventually Q F" (is "?L = ?R") | |
| 107 | proof - | |
| 108 | from assms have "eventually (\<lambda>x. P x \<longrightarrow> Q x) F" | |
| 109 | and "eventually (\<lambda>x. Q x \<longrightarrow> P x) F" | |
| 110 | by (auto elim: eventually_elim1) | |
| 111 | then show ?thesis by (auto elim: eventually_elim2) | |
| 112 | qed | |
| 113 | ||
| 114 | ||
| 115 | ||
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 116 | subsection {* Finer-than relation *}
 | 
| 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 117 | |
| 44195 | 118 | text {* @{term "F \<le> F'"} means that filter @{term F} is finer than
 | 
| 119 | filter @{term F'}. *}
 | |
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 120 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 121 | instantiation filter :: (type) complete_lattice | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 122 | begin | 
| 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 123 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 124 | definition le_filter_def: | 
| 44195 | 125 | "F \<le> F' \<longleftrightarrow> (\<forall>P. eventually P F' \<longrightarrow> eventually P F)" | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 126 | |
| 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 127 | definition | 
| 44195 | 128 | "(F :: 'a filter) < F' \<longleftrightarrow> F \<le> F' \<and> \<not> F' \<le> F" | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 129 | |
| 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 130 | definition | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 131 | "top = Abs_filter (\<lambda>P. \<forall>x. P x)" | 
| 36630 | 132 | |
| 133 | definition | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 134 | "bot = Abs_filter (\<lambda>P. True)" | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 135 | |
| 36630 | 136 | definition | 
| 44195 | 137 | "sup F F' = Abs_filter (\<lambda>P. eventually P F \<and> eventually P F')" | 
| 36630 | 138 | |
| 139 | definition | |
| 44195 | 140 | "inf F F' = Abs_filter | 
| 141 | (\<lambda>P. \<exists>Q R. eventually Q F \<and> eventually R F' \<and> (\<forall>x. Q x \<and> R x \<longrightarrow> P x))" | |
| 36630 | 142 | |
| 143 | definition | |
| 44195 | 144 | "Sup S = Abs_filter (\<lambda>P. \<forall>F\<in>S. eventually P F)" | 
| 36630 | 145 | |
| 146 | definition | |
| 44195 | 147 |   "Inf S = Sup {F::'a filter. \<forall>F'\<in>S. F \<le> F'}"
 | 
| 36630 | 148 | |
| 149 | lemma eventually_top [simp]: "eventually P top \<longleftrightarrow> (\<forall>x. P x)" | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 150 | unfolding top_filter_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 151 | by (rule eventually_Abs_filter, rule is_filter.intro, auto) | 
| 36630 | 152 | |
| 36629 
de62713aec6e
swap ordering on nets, so x <= y means 'x is finer than y'
 huffman parents: 
36360diff
changeset | 153 | lemma eventually_bot [simp]: "eventually P bot" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 154 | unfolding bot_filter_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 155 | by (subst eventually_Abs_filter, rule is_filter.intro, auto) | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 156 | |
| 36630 | 157 | lemma eventually_sup: | 
| 44195 | 158 | "eventually P (sup F F') \<longleftrightarrow> eventually P F \<and> eventually P F'" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 159 | unfolding sup_filter_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 160 | by (rule eventually_Abs_filter, rule is_filter.intro) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 161 | (auto elim!: eventually_rev_mp) | 
| 36630 | 162 | |
| 163 | lemma eventually_inf: | |
| 44195 | 164 | "eventually P (inf F F') \<longleftrightarrow> | 
| 165 | (\<exists>Q R. eventually Q F \<and> eventually R F' \<and> (\<forall>x. Q x \<and> R x \<longrightarrow> P x))" | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 166 | unfolding inf_filter_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 167 | apply (rule eventually_Abs_filter, rule is_filter.intro) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 168 | apply (fast intro: eventually_True) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 169 | apply clarify | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 170 | apply (intro exI conjI) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 171 | apply (erule (1) eventually_conj) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 172 | apply (erule (1) eventually_conj) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 173 | apply simp | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 174 | apply auto | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 175 | done | 
| 36630 | 176 | |
| 177 | lemma eventually_Sup: | |
| 44195 | 178 | "eventually P (Sup S) \<longleftrightarrow> (\<forall>F\<in>S. eventually P F)" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 179 | unfolding Sup_filter_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 180 | apply (rule eventually_Abs_filter, rule is_filter.intro) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 181 | apply (auto intro: eventually_conj elim!: eventually_rev_mp) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 182 | done | 
| 36630 | 183 | |
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 184 | instance proof | 
| 44195 | 185 | fix F F' F'' :: "'a filter" and S :: "'a filter set" | 
| 186 |   { show "F < F' \<longleftrightarrow> F \<le> F' \<and> \<not> F' \<le> F"
 | |
| 187 | by (rule less_filter_def) } | |
| 188 |   { show "F \<le> F"
 | |
| 189 | unfolding le_filter_def by simp } | |
| 190 |   { assume "F \<le> F'" and "F' \<le> F''" thus "F \<le> F''"
 | |
| 191 | unfolding le_filter_def by simp } | |
| 192 |   { assume "F \<le> F'" and "F' \<le> F" thus "F = F'"
 | |
| 193 | unfolding le_filter_def filter_eq_iff by fast } | |
| 194 |   { show "F \<le> top"
 | |
| 195 | unfolding le_filter_def eventually_top by (simp add: always_eventually) } | |
| 196 |   { show "bot \<le> F"
 | |
| 197 | unfolding le_filter_def by simp } | |
| 198 |   { show "F \<le> sup F F'" and "F' \<le> sup F F'"
 | |
| 199 | unfolding le_filter_def eventually_sup by simp_all } | |
| 200 |   { assume "F \<le> F''" and "F' \<le> F''" thus "sup F F' \<le> F''"
 | |
| 201 | unfolding le_filter_def eventually_sup by simp } | |
| 202 |   { show "inf F F' \<le> F" and "inf F F' \<le> F'"
 | |
| 203 | unfolding le_filter_def eventually_inf by (auto intro: eventually_True) } | |
| 204 |   { assume "F \<le> F'" and "F \<le> F''" thus "F \<le> inf F' F''"
 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 205 | unfolding le_filter_def eventually_inf | 
| 44195 | 206 | by (auto elim!: eventually_mono intro: eventually_conj) } | 
| 207 |   { assume "F \<in> S" thus "F \<le> Sup S"
 | |
| 208 | unfolding le_filter_def eventually_Sup by simp } | |
| 209 |   { assume "\<And>F. F \<in> S \<Longrightarrow> F \<le> F'" thus "Sup S \<le> F'"
 | |
| 210 | unfolding le_filter_def eventually_Sup by simp } | |
| 211 |   { assume "F'' \<in> S" thus "Inf S \<le> F''"
 | |
| 212 | unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } | |
| 213 |   { assume "\<And>F'. F' \<in> S \<Longrightarrow> F \<le> F'" thus "F \<le> Inf S"
 | |
| 214 | unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } | |
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 215 | qed | 
| 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 216 | |
| 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 217 | end | 
| 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 218 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 219 | lemma filter_leD: | 
| 44195 | 220 | "F \<le> F' \<Longrightarrow> eventually P F' \<Longrightarrow> eventually P F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 221 | unfolding le_filter_def by simp | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 222 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 223 | lemma filter_leI: | 
| 44195 | 224 | "(\<And>P. eventually P F' \<Longrightarrow> eventually P F) \<Longrightarrow> F \<le> F'" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 225 | unfolding le_filter_def by simp | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 226 | |
| 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 227 | lemma eventually_False: | 
| 44195 | 228 | "eventually (\<lambda>x. False) F \<longleftrightarrow> F = bot" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 229 | unfolding filter_eq_iff by (auto elim: eventually_rev_mp) | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 230 | |
| 44342 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 231 | abbreviation (input) trivial_limit :: "'a filter \<Rightarrow> bool" | 
| 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 232 | where "trivial_limit F \<equiv> F = bot" | 
| 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 233 | |
| 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 234 | lemma trivial_limit_def: "trivial_limit F \<longleftrightarrow> eventually (\<lambda>x. False) F" | 
| 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 235 | by (rule eventually_False [symmetric]) | 
| 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 236 | |
| 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 237 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 238 | subsection {* Map function for filters *}
 | 
| 36654 | 239 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 240 | definition filtermap :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a filter \<Rightarrow> 'b filter"
 | 
| 44195 | 241 | where "filtermap f F = Abs_filter (\<lambda>P. eventually (\<lambda>x. P (f x)) F)" | 
| 36654 | 242 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 243 | lemma eventually_filtermap: | 
| 44195 | 244 | "eventually P (filtermap f F) = eventually (\<lambda>x. P (f x)) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 245 | unfolding filtermap_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 246 | apply (rule eventually_Abs_filter) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 247 | apply (rule is_filter.intro) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 248 | apply (auto elim!: eventually_rev_mp) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 249 | done | 
| 36654 | 250 | |
| 44195 | 251 | lemma filtermap_ident: "filtermap (\<lambda>x. x) F = F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 252 | by (simp add: filter_eq_iff eventually_filtermap) | 
| 36654 | 253 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 254 | lemma filtermap_filtermap: | 
| 44195 | 255 | "filtermap f (filtermap g F) = filtermap (\<lambda>x. f (g x)) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 256 | by (simp add: filter_eq_iff eventually_filtermap) | 
| 36654 | 257 | |
| 44195 | 258 | lemma filtermap_mono: "F \<le> F' \<Longrightarrow> filtermap f F \<le> filtermap f F'" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 259 | unfolding le_filter_def eventually_filtermap by simp | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 260 | |
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 261 | lemma filtermap_bot [simp]: "filtermap f bot = bot" | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 262 | by (simp add: filter_eq_iff eventually_filtermap) | 
| 36654 | 263 | |
| 264 | ||
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 265 | subsection {* Sequentially *}
 | 
| 31392 | 266 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 267 | definition sequentially :: "nat filter" | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 268 | where "sequentially = Abs_filter (\<lambda>P. \<exists>k. \<forall>n\<ge>k. P n)" | 
| 31392 | 269 | |
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 270 | lemma eventually_sequentially: | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 271 | "eventually P sequentially \<longleftrightarrow> (\<exists>N. \<forall>n\<ge>N. P n)" | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 272 | unfolding sequentially_def | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 273 | proof (rule eventually_Abs_filter, rule is_filter.intro) | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 274 | fix P Q :: "nat \<Rightarrow> bool" | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 275 | assume "\<exists>i. \<forall>n\<ge>i. P n" and "\<exists>j. \<forall>n\<ge>j. Q n" | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 276 | then obtain i j where "\<forall>n\<ge>i. P n" and "\<forall>n\<ge>j. Q n" by auto | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 277 | then have "\<forall>n\<ge>max i j. P n \<and> Q n" by simp | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 278 | then show "\<exists>k. \<forall>n\<ge>k. P n \<and> Q n" .. | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 279 | qed auto | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 280 | |
| 44342 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 281 | lemma sequentially_bot [simp, intro]: "sequentially \<noteq> bot" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 282 | unfolding filter_eq_iff eventually_sequentially by auto | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 283 | |
| 44342 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 284 | lemmas trivial_limit_sequentially = sequentially_bot | 
| 
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
 huffman parents: 
44282diff
changeset | 285 | |
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 286 | lemma eventually_False_sequentially [simp]: | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 287 | "\<not> eventually (\<lambda>n. False) sequentially" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 288 | by (simp add: eventually_False) | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 289 | |
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 290 | lemma le_sequentially: | 
| 44195 | 291 | "F \<le> sequentially \<longleftrightarrow> (\<forall>N. eventually (\<lambda>n. N \<le> n) F)" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 292 | unfolding le_filter_def eventually_sequentially | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 293 | by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp) | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 294 | |
| 45892 | 295 | lemma eventually_sequentiallyI: | 
| 296 | assumes "\<And>x. c \<le> x \<Longrightarrow> P x" | |
| 297 | shows "eventually P sequentially" | |
| 298 | using assms by (auto simp: eventually_sequentially) | |
| 299 | ||
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 300 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 301 | subsection {* Standard filters *}
 | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 302 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 303 | definition within :: "'a filter \<Rightarrow> 'a set \<Rightarrow> 'a filter" (infixr "within" 70) | 
| 44195 | 304 | where "F within S = Abs_filter (\<lambda>P. eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) F)" | 
| 31392 | 305 | |
| 44206 
5e4a1664106e
locale-ize some constant definitions, so complete_space can inherit from metric_space
 huffman parents: 
44205diff
changeset | 306 | definition (in topological_space) nhds :: "'a \<Rightarrow> 'a filter" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 307 | where "nhds a = Abs_filter (\<lambda>P. \<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))" | 
| 36654 | 308 | |
| 44206 
5e4a1664106e
locale-ize some constant definitions, so complete_space can inherit from metric_space
 huffman parents: 
44205diff
changeset | 309 | definition (in topological_space) at :: "'a \<Rightarrow> 'a filter" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 310 |   where "at a = nhds a within - {a}"
 | 
| 31447 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 311 | |
| 31392 | 312 | lemma eventually_within: | 
| 44195 | 313 | "eventually P (F within S) = eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 314 | unfolding within_def | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 315 | by (rule eventually_Abs_filter, rule is_filter.intro) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 316 | (auto elim!: eventually_rev_mp) | 
| 31392 | 317 | |
| 45031 | 318 | lemma within_UNIV [simp]: "F within UNIV = F" | 
| 319 | unfolding filter_eq_iff eventually_within by simp | |
| 320 | ||
| 321 | lemma within_empty [simp]: "F within {} = bot"
 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 322 | unfolding filter_eq_iff eventually_within by simp | 
| 36360 
9d8f7efd9289
define finer-than ordering on net type; move some theorems into Limits.thy
 huffman parents: 
36358diff
changeset | 323 | |
| 36654 | 324 | lemma eventually_nhds: | 
| 325 | "eventually P (nhds a) \<longleftrightarrow> (\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))" | |
| 326 | unfolding nhds_def | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 327 | proof (rule eventually_Abs_filter, rule is_filter.intro) | 
| 36654 | 328 | have "open UNIV \<and> a \<in> UNIV \<and> (\<forall>x\<in>UNIV. True)" by simp | 
| 329 | thus "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. True)" by - rule | |
| 36358 
246493d61204
define nets directly as filters, instead of as filter bases
 huffman parents: 
31902diff
changeset | 330 | next | 
| 
246493d61204
define nets directly as filters, instead of as filter bases
 huffman parents: 
31902diff
changeset | 331 | fix P Q | 
| 36654 | 332 | assume "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x)" | 
| 333 | and "\<exists>T. open T \<and> a \<in> T \<and> (\<forall>x\<in>T. Q x)" | |
| 36358 
246493d61204
define nets directly as filters, instead of as filter bases
 huffman parents: 
31902diff
changeset | 334 | then obtain S T where | 
| 36654 | 335 | "open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x)" | 
| 336 | "open T \<and> a \<in> T \<and> (\<forall>x\<in>T. Q x)" by auto | |
| 337 | hence "open (S \<inter> T) \<and> a \<in> S \<inter> T \<and> (\<forall>x\<in>(S \<inter> T). P x \<and> Q x)" | |
| 36358 
246493d61204
define nets directly as filters, instead of as filter bases
 huffman parents: 
31902diff
changeset | 338 | by (simp add: open_Int) | 
| 36654 | 339 | thus "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x \<and> Q x)" by - rule | 
| 36358 
246493d61204
define nets directly as filters, instead of as filter bases
 huffman parents: 
31902diff
changeset | 340 | qed auto | 
| 31447 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 341 | |
| 36656 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 342 | lemma eventually_nhds_metric: | 
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 343 | "eventually P (nhds a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. dist x a < d \<longrightarrow> P x)" | 
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 344 | unfolding eventually_nhds open_dist | 
| 31447 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 345 | apply safe | 
| 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 346 | apply fast | 
| 31492 
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
 huffman parents: 
31488diff
changeset | 347 | apply (rule_tac x="{x. dist x a < d}" in exI, simp)
 | 
| 31447 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 348 | apply clarsimp | 
| 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 349 | apply (rule_tac x="d - dist x a" in exI, clarsimp) | 
| 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 350 | apply (simp only: less_diff_eq) | 
| 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 351 | apply (erule le_less_trans [OF dist_triangle]) | 
| 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 352 | done | 
| 
97bab1ac463e
generalize type of 'at' to topological_space; generalize some lemmas
 huffman parents: 
31392diff
changeset | 353 | |
| 44571 | 354 | lemma nhds_neq_bot [simp]: "nhds a \<noteq> bot" | 
| 355 | unfolding trivial_limit_def eventually_nhds by simp | |
| 356 | ||
| 36656 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 357 | lemma eventually_at_topological: | 
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 358 | "eventually P (at a) \<longleftrightarrow> (\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. x \<noteq> a \<longrightarrow> P x))" | 
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 359 | unfolding at_def eventually_within eventually_nhds by simp | 
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 360 | |
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 361 | lemma eventually_at: | 
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 362 | fixes a :: "'a::metric_space" | 
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 363 | shows "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> P x)" | 
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 364 | unfolding at_def eventually_within eventually_nhds_metric by auto | 
| 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 365 | |
| 44571 | 366 | lemma at_eq_bot_iff: "at a = bot \<longleftrightarrow> open {a}"
 | 
| 367 | unfolding trivial_limit_def eventually_at_topological | |
| 368 |   by (safe, case_tac "S = {a}", simp, fast, fast)
 | |
| 369 | ||
| 370 | lemma at_neq_bot [simp]: "at (a::'a::perfect_space) \<noteq> bot" | |
| 371 | by (simp add: at_eq_bot_iff not_open_singleton) | |
| 372 | ||
| 31392 | 373 | |
| 31355 | 374 | subsection {* Boundedness *}
 | 
| 375 | ||
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 376 | definition Bfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
 | 
| 44195 | 377 | where "Bfun f F = (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)" | 
| 31355 | 378 | |
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 379 | lemma BfunI: | 
| 44195 | 380 | assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" shows "Bfun f F" | 
| 31355 | 381 | unfolding Bfun_def | 
| 382 | proof (intro exI conjI allI) | |
| 383 | show "0 < max K 1" by simp | |
| 384 | next | |
| 44195 | 385 | show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F" | 
| 31355 | 386 | using K by (rule eventually_elim1, simp) | 
| 387 | qed | |
| 388 | ||
| 389 | lemma BfunE: | |
| 44195 | 390 | assumes "Bfun f F" | 
| 391 | obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F" | |
| 31355 | 392 | using assms unfolding Bfun_def by fast | 
| 393 | ||
| 394 | ||
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 395 | subsection {* Convergence to Zero *}
 | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 396 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 397 | definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
 | 
| 44195 | 398 | where "Zfun f F = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) F)" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 399 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 400 | lemma ZfunI: | 
| 44195 | 401 | "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 402 | unfolding Zfun_def by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 403 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 404 | lemma ZfunD: | 
| 44195 | 405 | "\<lbrakk>Zfun f F; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 406 | unfolding Zfun_def by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 407 | |
| 31355 | 408 | lemma Zfun_ssubst: | 
| 44195 | 409 | "eventually (\<lambda>x. f x = g x) F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun f F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 410 | unfolding Zfun_def by (auto elim!: eventually_rev_mp) | 
| 31355 | 411 | |
| 44195 | 412 | lemma Zfun_zero: "Zfun (\<lambda>x. 0) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 413 | unfolding Zfun_def by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 414 | |
| 44195 | 415 | lemma Zfun_norm_iff: "Zfun (\<lambda>x. norm (f x)) F = Zfun (\<lambda>x. f x) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 416 | unfolding Zfun_def by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 417 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 418 | lemma Zfun_imp_Zfun: | 
| 44195 | 419 | assumes f: "Zfun f F" | 
| 420 | assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F" | |
| 421 | shows "Zfun (\<lambda>x. g x) F" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 422 | proof (cases) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 423 | assume K: "0 < K" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 424 | show ?thesis | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 425 | proof (rule ZfunI) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 426 | fix r::real assume "0 < r" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 427 | hence "0 < r / K" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 428 | using K by (rule divide_pos_pos) | 
| 44195 | 429 | then have "eventually (\<lambda>x. norm (f x) < r / K) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 430 | using ZfunD [OF f] by fast | 
| 44195 | 431 | with g show "eventually (\<lambda>x. norm (g x) < r) F" | 
| 31355 | 432 | proof (rule eventually_elim2) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 433 | fix x | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 434 | assume *: "norm (g x) \<le> norm (f x) * K" | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 435 | assume "norm (f x) < r / K" | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 436 | hence "norm (f x) * K < r" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 437 | by (simp add: pos_less_divide_eq K) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 438 | thus "norm (g x) < r" | 
| 31355 | 439 | by (simp add: order_le_less_trans [OF *]) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 440 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 441 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 442 | next | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 443 | assume "\<not> 0 < K" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 444 | hence K: "K \<le> 0" by (simp only: not_less) | 
| 31355 | 445 | show ?thesis | 
| 446 | proof (rule ZfunI) | |
| 447 | fix r :: real | |
| 448 | assume "0 < r" | |
| 44195 | 449 | from g show "eventually (\<lambda>x. norm (g x) < r) F" | 
| 31355 | 450 | proof (rule eventually_elim1) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 451 | fix x | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 452 | assume "norm (g x) \<le> norm (f x) * K" | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 453 | also have "\<dots> \<le> norm (f x) * 0" | 
| 31355 | 454 | using K norm_ge_zero by (rule mult_left_mono) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 455 | finally show "norm (g x) < r" | 
| 31355 | 456 | using `0 < r` by simp | 
| 457 | qed | |
| 458 | qed | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 459 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 460 | |
| 44195 | 461 | lemma Zfun_le: "\<lbrakk>Zfun g F; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 462 | by (erule_tac K="1" in Zfun_imp_Zfun, simp) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 463 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 464 | lemma Zfun_add: | 
| 44195 | 465 | assumes f: "Zfun f F" and g: "Zfun g F" | 
| 466 | shows "Zfun (\<lambda>x. f x + g x) F" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 467 | proof (rule ZfunI) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 468 | fix r::real assume "0 < r" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 469 | hence r: "0 < r / 2" by simp | 
| 44195 | 470 | have "eventually (\<lambda>x. norm (f x) < r/2) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 471 | using f r by (rule ZfunD) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 472 | moreover | 
| 44195 | 473 | have "eventually (\<lambda>x. norm (g x) < r/2) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 474 | using g r by (rule ZfunD) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 475 | ultimately | 
| 44195 | 476 | show "eventually (\<lambda>x. norm (f x + g x) < r) F" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 477 | proof (rule eventually_elim2) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 478 | fix x | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 479 | assume *: "norm (f x) < r/2" "norm (g x) < r/2" | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 480 | have "norm (f x + g x) \<le> norm (f x) + norm (g x)" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 481 | by (rule norm_triangle_ineq) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 482 | also have "\<dots> < r/2 + r/2" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 483 | using * by (rule add_strict_mono) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 484 | finally show "norm (f x + g x) < r" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 485 | by simp | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 486 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 487 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 488 | |
| 44195 | 489 | lemma Zfun_minus: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. - f x) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 490 | unfolding Zfun_def by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 491 | |
| 44195 | 492 | lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 493 | by (simp only: diff_minus Zfun_add Zfun_minus) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 494 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 495 | lemma (in bounded_linear) Zfun: | 
| 44195 | 496 | assumes g: "Zfun g F" | 
| 497 | shows "Zfun (\<lambda>x. f (g x)) F" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 498 | proof - | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 499 | obtain K where "\<And>x. norm (f x) \<le> norm x * K" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 500 | using bounded by fast | 
| 44195 | 501 | then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) F" | 
| 31355 | 502 | by simp | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 503 | with g show ?thesis | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 504 | by (rule Zfun_imp_Zfun) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 505 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 506 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 507 | lemma (in bounded_bilinear) Zfun: | 
| 44195 | 508 | assumes f: "Zfun f F" | 
| 509 | assumes g: "Zfun g F" | |
| 510 | shows "Zfun (\<lambda>x. f x ** g x) F" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 511 | proof (rule ZfunI) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 512 | fix r::real assume r: "0 < r" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 513 | obtain K where K: "0 < K" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 514 | and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 515 | using pos_bounded by fast | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 516 | from K have K': "0 < inverse K" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 517 | by (rule positive_imp_inverse_positive) | 
| 44195 | 518 | have "eventually (\<lambda>x. norm (f x) < r) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 519 | using f r by (rule ZfunD) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 520 | moreover | 
| 44195 | 521 | have "eventually (\<lambda>x. norm (g x) < inverse K) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 522 | using g K' by (rule ZfunD) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 523 | ultimately | 
| 44195 | 524 | show "eventually (\<lambda>x. norm (f x ** g x) < r) F" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 525 | proof (rule eventually_elim2) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 526 | fix x | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 527 | assume *: "norm (f x) < r" "norm (g x) < inverse K" | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 528 | have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 529 | by (rule norm_le) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 530 | also have "norm (f x) * norm (g x) * K < r * inverse K * K" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 531 | by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero * K) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 532 | also from K have "r * inverse K * K = r" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 533 | by simp | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 534 | finally show "norm (f x ** g x) < r" . | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 535 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 536 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 537 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 538 | lemma (in bounded_bilinear) Zfun_left: | 
| 44195 | 539 | "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. f x ** a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 540 | by (rule bounded_linear_left [THEN bounded_linear.Zfun]) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 541 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 542 | lemma (in bounded_bilinear) Zfun_right: | 
| 44195 | 543 | "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. a ** f x) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 544 | by (rule bounded_linear_right [THEN bounded_linear.Zfun]) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 545 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 546 | lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 547 | lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 548 | lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 549 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 550 | |
| 31902 | 551 | subsection {* Limits *}
 | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 552 | |
| 44206 
5e4a1664106e
locale-ize some constant definitions, so complete_space can inherit from metric_space
 huffman parents: 
44205diff
changeset | 553 | definition (in topological_space) | 
| 
5e4a1664106e
locale-ize some constant definitions, so complete_space can inherit from metric_space
 huffman parents: 
44205diff
changeset | 554 |   tendsto :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b filter \<Rightarrow> bool" (infixr "--->" 55) where
 | 
| 44195 | 555 | "(f ---> l) F \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) F)" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 556 | |
| 45892 | 557 | definition real_tendsto_inf :: "('a \<Rightarrow> real) \<Rightarrow> 'a filter \<Rightarrow> bool" where
 | 
| 558 | "real_tendsto_inf f F \<equiv> \<forall>x. eventually (\<lambda>y. x < f y) F" | |
| 559 | ||
| 31902 | 560 | ML {*
 | 
| 561 | structure Tendsto_Intros = Named_Thms | |
| 562 | ( | |
| 45294 | 563 |   val name = @{binding tendsto_intros}
 | 
| 31902 | 564 | val description = "introduction rules for tendsto" | 
| 565 | ) | |
| 31565 | 566 | *} | 
| 567 | ||
| 31902 | 568 | setup Tendsto_Intros.setup | 
| 31565 | 569 | |
| 44195 | 570 | lemma tendsto_mono: "F \<le> F' \<Longrightarrow> (f ---> l) F' \<Longrightarrow> (f ---> l) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 571 | unfolding tendsto_def le_filter_def by fast | 
| 36656 
fec55067ae9b
add lemmas eventually_nhds_metric and tendsto_mono
 huffman parents: 
36655diff
changeset | 572 | |
| 31488 | 573 | lemma topological_tendstoI: | 
| 44195 | 574 | "(\<And>S. open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) F) | 
| 575 | \<Longrightarrow> (f ---> l) F" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 576 | unfolding tendsto_def by auto | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 577 | |
| 31488 | 578 | lemma topological_tendstoD: | 
| 44195 | 579 | "(f ---> l) F \<Longrightarrow> open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) F" | 
| 31488 | 580 | unfolding tendsto_def by auto | 
| 581 | ||
| 582 | lemma tendstoI: | |
| 44195 | 583 | assumes "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F" | 
| 584 | shows "(f ---> l) F" | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 585 | apply (rule topological_tendstoI) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 586 | apply (simp add: open_dist) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 587 | apply (drule (1) bspec, clarify) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 588 | apply (drule assms) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 589 | apply (erule eventually_elim1, simp) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 590 | done | 
| 31488 | 591 | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 592 | lemma tendstoD: | 
| 44195 | 593 | "(f ---> l) F \<Longrightarrow> 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 594 |   apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD)
 | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 595 | apply (clarsimp simp add: open_dist) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 596 | apply (rule_tac x="e - dist x l" in exI, clarsimp) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 597 | apply (simp only: less_diff_eq) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 598 | apply (erule le_less_trans [OF dist_triangle]) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 599 | apply simp | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 600 | apply simp | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 601 | done | 
| 31488 | 602 | |
| 603 | lemma tendsto_iff: | |
| 44195 | 604 | "(f ---> l) F \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) F)" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 605 | using tendstoI tendstoD by fast | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 606 | |
| 44195 | 607 | lemma tendsto_Zfun_iff: "(f ---> a) F = Zfun (\<lambda>x. f x - a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 608 | by (simp only: tendsto_iff Zfun_def dist_norm) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 609 | |
| 45031 | 610 | lemma tendsto_bot [simp]: "(f ---> a) bot" | 
| 611 | unfolding tendsto_def by simp | |
| 612 | ||
| 31565 | 613 | lemma tendsto_ident_at [tendsto_intros]: "((\<lambda>x. x) ---> a) (at a)" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 614 | unfolding tendsto_def eventually_at_topological by auto | 
| 31565 | 615 | |
| 616 | lemma tendsto_ident_at_within [tendsto_intros]: | |
| 36655 | 617 | "((\<lambda>x. x) ---> a) (at a within S)" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 618 | unfolding tendsto_def eventually_within eventually_at_topological by auto | 
| 31565 | 619 | |
| 44195 | 620 | lemma tendsto_const [tendsto_intros]: "((\<lambda>x. k) ---> k) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 621 | by (simp add: tendsto_def) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 622 | |
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 623 | lemma tendsto_unique: | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 624 | fixes f :: "'a \<Rightarrow> 'b::t2_space" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 625 | assumes "\<not> trivial_limit F" and "(f ---> a) F" and "(f ---> b) F" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 626 | shows "a = b" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 627 | proof (rule ccontr) | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 628 | assume "a \<noteq> b" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 629 |   obtain U V where "open U" "open V" "a \<in> U" "b \<in> V" "U \<inter> V = {}"
 | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 630 | using hausdorff [OF `a \<noteq> b`] by fast | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 631 | have "eventually (\<lambda>x. f x \<in> U) F" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 632 | using `(f ---> a) F` `open U` `a \<in> U` by (rule topological_tendstoD) | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 633 | moreover | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 634 | have "eventually (\<lambda>x. f x \<in> V) F" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 635 | using `(f ---> b) F` `open V` `b \<in> V` by (rule topological_tendstoD) | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 636 | ultimately | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 637 | have "eventually (\<lambda>x. False) F" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 638 | proof (rule eventually_elim2) | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 639 | fix x | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 640 | assume "f x \<in> U" "f x \<in> V" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 641 | hence "f x \<in> U \<inter> V" by simp | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 642 |     with `U \<inter> V = {}` show "False" by simp
 | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 643 | qed | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 644 | with `\<not> trivial_limit F` show "False" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 645 | by (simp add: trivial_limit_def) | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 646 | qed | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 647 | |
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 648 | lemma tendsto_const_iff: | 
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 649 | fixes a b :: "'a::t2_space" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 650 | assumes "\<not> trivial_limit F" shows "((\<lambda>x. a) ---> b) F \<longleftrightarrow> a = b" | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 651 | by (safe intro!: tendsto_const tendsto_unique [OF assms tendsto_const]) | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 652 | |
| 44218 | 653 | lemma tendsto_compose: | 
| 654 | assumes g: "(g ---> g l) (at l)" | |
| 655 | assumes f: "(f ---> l) F" | |
| 656 | shows "((\<lambda>x. g (f x)) ---> g l) F" | |
| 657 | proof (rule topological_tendstoI) | |
| 658 | fix B assume B: "open B" "g l \<in> B" | |
| 659 | obtain A where A: "open A" "l \<in> A" | |
| 660 | and gB: "\<forall>y. y \<in> A \<longrightarrow> g y \<in> B" | |
| 661 | using topological_tendstoD [OF g B] B(2) | |
| 662 | unfolding eventually_at_topological by fast | |
| 663 | hence "\<forall>x. f x \<in> A \<longrightarrow> g (f x) \<in> B" by simp | |
| 664 | from this topological_tendstoD [OF f A] | |
| 665 | show "eventually (\<lambda>x. g (f x) \<in> B) F" | |
| 666 | by (rule eventually_mono) | |
| 667 | qed | |
| 668 | ||
| 44253 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 669 | lemma tendsto_compose_eventually: | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 670 | assumes g: "(g ---> m) (at l)" | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 671 | assumes f: "(f ---> l) F" | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 672 | assumes inj: "eventually (\<lambda>x. f x \<noteq> l) F" | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 673 | shows "((\<lambda>x. g (f x)) ---> m) F" | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 674 | proof (rule topological_tendstoI) | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 675 | fix B assume B: "open B" "m \<in> B" | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 676 | obtain A where A: "open A" "l \<in> A" | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 677 | and gB: "\<And>y. y \<in> A \<Longrightarrow> y \<noteq> l \<Longrightarrow> g y \<in> B" | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 678 | using topological_tendstoD [OF g B] | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 679 | unfolding eventually_at_topological by fast | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 680 | show "eventually (\<lambda>x. g (f x) \<in> B) F" | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 681 | using topological_tendstoD [OF f A] inj | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 682 | by (rule eventually_elim2) (simp add: gB) | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 683 | qed | 
| 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 684 | |
| 44251 | 685 | lemma metric_tendsto_imp_tendsto: | 
| 686 | assumes f: "(f ---> a) F" | |
| 687 | assumes le: "eventually (\<lambda>x. dist (g x) b \<le> dist (f x) a) F" | |
| 688 | shows "(g ---> b) F" | |
| 689 | proof (rule tendstoI) | |
| 690 | fix e :: real assume "0 < e" | |
| 691 | with f have "eventually (\<lambda>x. dist (f x) a < e) F" by (rule tendstoD) | |
| 692 | with le show "eventually (\<lambda>x. dist (g x) b < e) F" | |
| 693 | using le_less_trans by (rule eventually_elim2) | |
| 694 | qed | |
| 695 | ||
| 45892 | 696 | lemma real_tendsto_inf_real: "real_tendsto_inf real sequentially" | 
| 697 | proof (unfold real_tendsto_inf_def, rule allI) | |
| 698 | fix x show "eventually (\<lambda>y. x < real y) sequentially" | |
| 699 | by (rule eventually_sequentiallyI[of "natceiling (x + 1)"]) | |
| 700 | (simp add: natceiling_le_eq) | |
| 701 | qed | |
| 702 | ||
| 703 | ||
| 704 | ||
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44195diff
changeset | 705 | subsubsection {* Distance and norms *}
 | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 706 | |
| 31565 | 707 | lemma tendsto_dist [tendsto_intros]: | 
| 44195 | 708 | assumes f: "(f ---> l) F" and g: "(g ---> m) F" | 
| 709 | shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) F" | |
| 31565 | 710 | proof (rule tendstoI) | 
| 711 | fix e :: real assume "0 < e" | |
| 712 | hence e2: "0 < e/2" by simp | |
| 713 | from tendstoD [OF f e2] tendstoD [OF g e2] | |
| 44195 | 714 | show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F" | 
| 31565 | 715 | proof (rule eventually_elim2) | 
| 716 | fix x assume "dist (f x) l < e/2" "dist (g x) m < e/2" | |
| 717 | then show "dist (dist (f x) (g x)) (dist l m) < e" | |
| 718 | unfolding dist_real_def | |
| 719 | using dist_triangle2 [of "f x" "g x" "l"] | |
| 720 | using dist_triangle2 [of "g x" "l" "m"] | |
| 721 | using dist_triangle3 [of "l" "m" "f x"] | |
| 722 | using dist_triangle [of "f x" "m" "g x"] | |
| 723 | by arith | |
| 724 | qed | |
| 725 | qed | |
| 726 | ||
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 727 | lemma norm_conv_dist: "norm x = dist x 0" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 728 | unfolding dist_norm by simp | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 729 | |
| 31565 | 730 | lemma tendsto_norm [tendsto_intros]: | 
| 44195 | 731 | "(f ---> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> norm a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 732 | unfolding norm_conv_dist by (intro tendsto_intros) | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 733 | |
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 734 | lemma tendsto_norm_zero: | 
| 44195 | 735 | "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> 0) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 736 | by (drule tendsto_norm, simp) | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 737 | |
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 738 | lemma tendsto_norm_zero_cancel: | 
| 44195 | 739 | "((\<lambda>x. norm (f x)) ---> 0) F \<Longrightarrow> (f ---> 0) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 740 | unfolding tendsto_iff dist_norm by simp | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 741 | |
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 742 | lemma tendsto_norm_zero_iff: | 
| 44195 | 743 | "((\<lambda>x. norm (f x)) ---> 0) F \<longleftrightarrow> (f ---> 0) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 744 | unfolding tendsto_iff dist_norm by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 745 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 746 | lemma tendsto_rabs [tendsto_intros]: | 
| 44195 | 747 | "(f ---> (l::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> \<bar>l\<bar>) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 748 | by (fold real_norm_def, rule tendsto_norm) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 749 | |
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 750 | lemma tendsto_rabs_zero: | 
| 44195 | 751 | "(f ---> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 752 | by (fold real_norm_def, rule tendsto_norm_zero) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 753 | |
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 754 | lemma tendsto_rabs_zero_cancel: | 
| 44195 | 755 | "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<Longrightarrow> (f ---> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 756 | by (fold real_norm_def, rule tendsto_norm_zero_cancel) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 757 | |
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 758 | lemma tendsto_rabs_zero_iff: | 
| 44195 | 759 | "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<longleftrightarrow> (f ---> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 760 | by (fold real_norm_def, rule tendsto_norm_zero_iff) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 761 | |
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 762 | subsubsection {* Addition and subtraction *}
 | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 763 | |
| 31565 | 764 | lemma tendsto_add [tendsto_intros]: | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 765 | fixes a b :: "'a::real_normed_vector" | 
| 44195 | 766 | shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> a + b) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 767 | by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 768 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 769 | lemma tendsto_add_zero: | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 770 | fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector" | 
| 44195 | 771 | shows "\<lbrakk>(f ---> 0) F; (g ---> 0) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 772 | by (drule (1) tendsto_add, simp) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 773 | |
| 31565 | 774 | lemma tendsto_minus [tendsto_intros]: | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 775 | fixes a :: "'a::real_normed_vector" | 
| 44195 | 776 | shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. - f x) ---> - a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 777 | by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 778 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 779 | lemma tendsto_minus_cancel: | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 780 | fixes a :: "'a::real_normed_vector" | 
| 44195 | 781 | shows "((\<lambda>x. - f x) ---> - a) F \<Longrightarrow> (f ---> a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 782 | by (drule tendsto_minus, simp) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 783 | |
| 31565 | 784 | lemma tendsto_diff [tendsto_intros]: | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 785 | fixes a b :: "'a::real_normed_vector" | 
| 44195 | 786 | shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 787 | by (simp add: diff_minus tendsto_add tendsto_minus) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 788 | |
| 31588 | 789 | lemma tendsto_setsum [tendsto_intros]: | 
| 790 | fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::real_normed_vector" | |
| 44195 | 791 | assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> a i) F" | 
| 792 | shows "((\<lambda>x. \<Sum>i\<in>S. f i x) ---> (\<Sum>i\<in>S. a i)) F" | |
| 31588 | 793 | proof (cases "finite S") | 
| 794 | assume "finite S" thus ?thesis using assms | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 795 | by (induct, simp add: tendsto_const, simp add: tendsto_add) | 
| 31588 | 796 | next | 
| 797 | assume "\<not> finite S" thus ?thesis | |
| 798 | by (simp add: tendsto_const) | |
| 799 | qed | |
| 800 | ||
| 45892 | 801 | lemma real_tendsto_sandwich: | 
| 802 | fixes f g h :: "'a \<Rightarrow> real" | |
| 803 | assumes ev: "eventually (\<lambda>n. f n \<le> g n) net" "eventually (\<lambda>n. g n \<le> h n) net" | |
| 804 | assumes lim: "(f ---> c) net" "(h ---> c) net" | |
| 805 | shows "(g ---> c) net" | |
| 806 | proof - | |
| 807 | have "((\<lambda>n. g n - f n) ---> 0) net" | |
| 808 | proof (rule metric_tendsto_imp_tendsto) | |
| 809 | show "eventually (\<lambda>n. dist (g n - f n) 0 \<le> dist (h n - f n) 0) net" | |
| 810 | using ev by (rule eventually_elim2) (simp add: dist_real_def) | |
| 811 | show "((\<lambda>n. h n - f n) ---> 0) net" | |
| 812 | using tendsto_diff[OF lim(2,1)] by simp | |
| 813 | qed | |
| 814 | from tendsto_add[OF this lim(1)] show ?thesis by simp | |
| 815 | qed | |
| 816 | ||
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 817 | subsubsection {* Linear operators and multiplication *}
 | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 818 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 819 | lemma (in bounded_linear) tendsto: | 
| 44195 | 820 | "(g ---> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> f a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 821 | by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 822 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 823 | lemma (in bounded_linear) tendsto_zero: | 
| 44195 | 824 | "(g ---> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 825 | by (drule tendsto, simp only: zero) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 826 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 827 | lemma (in bounded_bilinear) tendsto: | 
| 44195 | 828 | "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) ---> a ** b) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 829 | by (simp only: tendsto_Zfun_iff prod_diff_prod | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 830 | Zfun_add Zfun Zfun_left Zfun_right) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 831 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 832 | lemma (in bounded_bilinear) tendsto_zero: | 
| 44195 | 833 | assumes f: "(f ---> 0) F" | 
| 834 | assumes g: "(g ---> 0) F" | |
| 835 | shows "((\<lambda>x. f x ** g x) ---> 0) F" | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 836 | using tendsto [OF f g] by (simp add: zero_left) | 
| 31355 | 837 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 838 | lemma (in bounded_bilinear) tendsto_left_zero: | 
| 44195 | 839 | "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. f x ** c) ---> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 840 | by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 841 | |
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 842 | lemma (in bounded_bilinear) tendsto_right_zero: | 
| 44195 | 843 | "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. c ** f x) ---> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 844 | by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 845 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 846 | lemmas tendsto_of_real [tendsto_intros] = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 847 | bounded_linear.tendsto [OF bounded_linear_of_real] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 848 | |
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 849 | lemmas tendsto_scaleR [tendsto_intros] = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 850 | bounded_bilinear.tendsto [OF bounded_bilinear_scaleR] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 851 | |
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 852 | lemmas tendsto_mult [tendsto_intros] = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 853 | bounded_bilinear.tendsto [OF bounded_bilinear_mult] | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 854 | |
| 44568 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 855 | lemmas tendsto_mult_zero = | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 856 | bounded_bilinear.tendsto_zero [OF bounded_bilinear_mult] | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 857 | |
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 858 | lemmas tendsto_mult_left_zero = | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 859 | bounded_bilinear.tendsto_left_zero [OF bounded_bilinear_mult] | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 860 | |
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 861 | lemmas tendsto_mult_right_zero = | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 862 | bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult] | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 863 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 864 | lemma tendsto_power [tendsto_intros]: | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 865 |   fixes f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
 | 
| 44195 | 866 | shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) ---> a ^ n) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 867 | by (induct n) (simp_all add: tendsto_const tendsto_mult) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 868 | |
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 869 | lemma tendsto_setprod [tendsto_intros]: | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 870 |   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
 | 
| 44195 | 871 | assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> L i) F" | 
| 872 | shows "((\<lambda>x. \<Prod>i\<in>S. f i x) ---> (\<Prod>i\<in>S. L i)) F" | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 873 | proof (cases "finite S") | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 874 | assume "finite S" thus ?thesis using assms | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 875 | by (induct, simp add: tendsto_const, simp add: tendsto_mult) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 876 | next | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 877 | assume "\<not> finite S" thus ?thesis | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 878 | by (simp add: tendsto_const) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 879 | qed | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 880 | |
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 881 | subsubsection {* Inverse and division *}
 | 
| 31355 | 882 | |
| 883 | lemma (in bounded_bilinear) Zfun_prod_Bfun: | |
| 44195 | 884 | assumes f: "Zfun f F" | 
| 885 | assumes g: "Bfun g F" | |
| 886 | shows "Zfun (\<lambda>x. f x ** g x) F" | |
| 31355 | 887 | proof - | 
| 888 | obtain K where K: "0 \<le> K" | |
| 889 | and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K" | |
| 890 | using nonneg_bounded by fast | |
| 891 | obtain B where B: "0 < B" | |
| 44195 | 892 | and norm_g: "eventually (\<lambda>x. norm (g x) \<le> B) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 893 | using g by (rule BfunE) | 
| 44195 | 894 | have "eventually (\<lambda>x. norm (f x ** g x) \<le> norm (f x) * (B * K)) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 895 | using norm_g proof (rule eventually_elim1) | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 896 | fix x | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 897 | assume *: "norm (g x) \<le> B" | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 898 | have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K" | 
| 31355 | 899 | by (rule norm_le) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 900 | also have "\<dots> \<le> norm (f x) * B * K" | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 901 | by (intro mult_mono' order_refl norm_g norm_ge_zero | 
| 31355 | 902 | mult_nonneg_nonneg K *) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 903 | also have "\<dots> = norm (f x) * (B * K)" | 
| 31355 | 904 | by (rule mult_assoc) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 905 | finally show "norm (f x ** g x) \<le> norm (f x) * (B * K)" . | 
| 31355 | 906 | qed | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 907 | with f show ?thesis | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 908 | by (rule Zfun_imp_Zfun) | 
| 31355 | 909 | qed | 
| 910 | ||
| 911 | lemma (in bounded_bilinear) flip: | |
| 912 | "bounded_bilinear (\<lambda>x y. y ** x)" | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 913 | apply default | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 914 | apply (rule add_right) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 915 | apply (rule add_left) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 916 | apply (rule scaleR_right) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 917 | apply (rule scaleR_left) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 918 | apply (subst mult_commute) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 919 | using bounded by fast | 
| 31355 | 920 | |
| 921 | lemma (in bounded_bilinear) Bfun_prod_Zfun: | |
| 44195 | 922 | assumes f: "Bfun f F" | 
| 923 | assumes g: "Zfun g F" | |
| 924 | shows "Zfun (\<lambda>x. f x ** g x) F" | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 925 | using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) | 
| 31355 | 926 | |
| 927 | lemma Bfun_inverse_lemma: | |
| 928 | fixes x :: "'a::real_normed_div_algebra" | |
| 929 | shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r" | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 930 | apply (subst nonzero_norm_inverse, clarsimp) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 931 | apply (erule (1) le_imp_inverse_le) | 
| 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 932 | done | 
| 31355 | 933 | |
| 934 | lemma Bfun_inverse: | |
| 935 | fixes a :: "'a::real_normed_div_algebra" | |
| 44195 | 936 | assumes f: "(f ---> a) F" | 
| 31355 | 937 | assumes a: "a \<noteq> 0" | 
| 44195 | 938 | shows "Bfun (\<lambda>x. inverse (f x)) F" | 
| 31355 | 939 | proof - | 
| 940 | from a have "0 < norm a" by simp | |
| 941 | hence "\<exists>r>0. r < norm a" by (rule dense) | |
| 942 | then obtain r where r1: "0 < r" and r2: "r < norm a" by fast | |
| 44195 | 943 | have "eventually (\<lambda>x. dist (f x) a < r) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 944 | using tendstoD [OF f r1] by fast | 
| 44195 | 945 | hence "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F" | 
| 31355 | 946 | proof (rule eventually_elim1) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 947 | fix x | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 948 | assume "dist (f x) a < r" | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 949 | hence 1: "norm (f x - a) < r" | 
| 31355 | 950 | by (simp add: dist_norm) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 951 | hence 2: "f x \<noteq> 0" using r2 by auto | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 952 | hence "norm (inverse (f x)) = inverse (norm (f x))" | 
| 31355 | 953 | by (rule nonzero_norm_inverse) | 
| 954 | also have "\<dots> \<le> inverse (norm a - r)" | |
| 955 | proof (rule le_imp_inverse_le) | |
| 956 | show "0 < norm a - r" using r2 by simp | |
| 957 | next | |
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 958 | have "norm a - norm (f x) \<le> norm (a - f x)" | 
| 31355 | 959 | by (rule norm_triangle_ineq2) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 960 | also have "\<dots> = norm (f x - a)" | 
| 31355 | 961 | by (rule norm_minus_commute) | 
| 962 | also have "\<dots> < r" using 1 . | |
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 963 | finally show "norm a - r \<le> norm (f x)" by simp | 
| 31355 | 964 | qed | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 965 | finally show "norm (inverse (f x)) \<le> inverse (norm a - r)" . | 
| 31355 | 966 | qed | 
| 967 | thus ?thesis by (rule BfunI) | |
| 968 | qed | |
| 969 | ||
| 31565 | 970 | lemma tendsto_inverse [tendsto_intros]: | 
| 31355 | 971 | fixes a :: "'a::real_normed_div_algebra" | 
| 44195 | 972 | assumes f: "(f ---> a) F" | 
| 31355 | 973 | assumes a: "a \<noteq> 0" | 
| 44195 | 974 | shows "((\<lambda>x. inverse (f x)) ---> inverse a) F" | 
| 31355 | 975 | proof - | 
| 976 | from a have "0 < norm a" by simp | |
| 44195 | 977 | with f have "eventually (\<lambda>x. dist (f x) a < norm a) F" | 
| 31355 | 978 | by (rule tendstoD) | 
| 44195 | 979 | then have "eventually (\<lambda>x. f x \<noteq> 0) F" | 
| 31355 | 980 | unfolding dist_norm by (auto elim!: eventually_elim1) | 
| 44627 | 981 | with a have "eventually (\<lambda>x. inverse (f x) - inverse a = | 
| 982 | - (inverse (f x) * (f x - a) * inverse a)) F" | |
| 983 | by (auto elim!: eventually_elim1 simp: inverse_diff_inverse) | |
| 984 | moreover have "Zfun (\<lambda>x. - (inverse (f x) * (f x - a) * inverse a)) F" | |
| 985 | by (intro Zfun_minus Zfun_mult_left | |
| 986 | bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult] | |
| 987 | Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff]) | |
| 988 | ultimately show ?thesis | |
| 989 | unfolding tendsto_Zfun_iff by (rule Zfun_ssubst) | |
| 31355 | 990 | qed | 
| 991 | ||
| 31565 | 992 | lemma tendsto_divide [tendsto_intros]: | 
| 31355 | 993 | fixes a b :: "'a::real_normed_field" | 
| 44195 | 994 | shows "\<lbrakk>(f ---> a) F; (g ---> b) F; b \<noteq> 0\<rbrakk> | 
| 995 | \<Longrightarrow> ((\<lambda>x. f x / g x) ---> a / b) F" | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 996 | by (simp add: tendsto_mult tendsto_inverse divide_inverse) | 
| 31355 | 997 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 998 | lemma tendsto_sgn [tendsto_intros]: | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 999 | fixes l :: "'a::real_normed_vector" | 
| 44195 | 1000 | shows "\<lbrakk>(f ---> l) F; l \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. sgn (f x)) ---> sgn l) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 1001 | unfolding sgn_div_norm by (simp add: tendsto_intros) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 1002 | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 1003 | end |