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