consistently use variable name 'F' for filters
authorhuffman
Sun, 14 Aug 2011 08:45:38 -0700
changeset 45066f5363511b212
parent 45065 0639898074ae
child 45067 3588f71abb50
child 45076 18da2a87421c
consistently use variable name 'F' for filters
src/HOL/Limits.thy
     1.1 --- a/src/HOL/Limits.thy	Sun Aug 14 07:54:24 2011 -0700
     1.2 +++ b/src/HOL/Limits.thy	Sun Aug 14 08:45:38 2011 -0700
     1.3 @@ -25,8 +25,8 @@
     1.4    show "(\<lambda>x. True) \<in> ?filter" by (auto intro: is_filter.intro)
     1.5  qed
     1.6  
     1.7 -lemma is_filter_Rep_filter: "is_filter (Rep_filter A)"
     1.8 -  using Rep_filter [of A] by simp
     1.9 +lemma is_filter_Rep_filter: "is_filter (Rep_filter F)"
    1.10 +  using Rep_filter [of F] by simp
    1.11  
    1.12  lemma Abs_filter_inverse':
    1.13    assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F"
    1.14 @@ -36,84 +36,84 @@
    1.15  subsection {* Eventually *}
    1.16  
    1.17  definition eventually :: "('a \<Rightarrow> bool) \<Rightarrow> 'a filter \<Rightarrow> bool"
    1.18 -  where "eventually P A \<longleftrightarrow> Rep_filter A P"
    1.19 +  where "eventually P F \<longleftrightarrow> Rep_filter F P"
    1.20  
    1.21  lemma eventually_Abs_filter:
    1.22    assumes "is_filter F" shows "eventually P (Abs_filter F) = F P"
    1.23    unfolding eventually_def using assms by (simp add: Abs_filter_inverse)
    1.24  
    1.25  lemma filter_eq_iff:
    1.26 -  shows "A = B \<longleftrightarrow> (\<forall>P. eventually P A = eventually P B)"
    1.27 +  shows "F = F' \<longleftrightarrow> (\<forall>P. eventually P F = eventually P F')"
    1.28    unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def ..
    1.29  
    1.30 -lemma eventually_True [simp]: "eventually (\<lambda>x. True) A"
    1.31 +lemma eventually_True [simp]: "eventually (\<lambda>x. True) F"
    1.32    unfolding eventually_def
    1.33    by (rule is_filter.True [OF is_filter_Rep_filter])
    1.34  
    1.35 -lemma always_eventually: "\<forall>x. P x \<Longrightarrow> eventually P A"
    1.36 +lemma always_eventually: "\<forall>x. P x \<Longrightarrow> eventually P F"
    1.37  proof -
    1.38    assume "\<forall>x. P x" hence "P = (\<lambda>x. True)" by (simp add: ext)
    1.39 -  thus "eventually P A" by simp
    1.40 +  thus "eventually P F" by simp
    1.41  qed
    1.42  
    1.43  lemma eventually_mono:
    1.44 -  "(\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually P A \<Longrightarrow> eventually Q A"
    1.45 +  "(\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually P F \<Longrightarrow> eventually Q F"
    1.46    unfolding eventually_def
    1.47    by (rule is_filter.mono [OF is_filter_Rep_filter])
    1.48  
    1.49  lemma eventually_conj:
    1.50 -  assumes P: "eventually (\<lambda>x. P x) A"
    1.51 -  assumes Q: "eventually (\<lambda>x. Q x) A"
    1.52 -  shows "eventually (\<lambda>x. P x \<and> Q x) A"
    1.53 +  assumes P: "eventually (\<lambda>x. P x) F"
    1.54 +  assumes Q: "eventually (\<lambda>x. Q x) F"
    1.55 +  shows "eventually (\<lambda>x. P x \<and> Q x) F"
    1.56    using assms unfolding eventually_def
    1.57    by (rule is_filter.conj [OF is_filter_Rep_filter])
    1.58  
    1.59  lemma eventually_mp:
    1.60 -  assumes "eventually (\<lambda>x. P x \<longrightarrow> Q x) A"
    1.61 -  assumes "eventually (\<lambda>x. P x) A"
    1.62 -  shows "eventually (\<lambda>x. Q x) A"
    1.63 +  assumes "eventually (\<lambda>x. P x \<longrightarrow> Q x) F"
    1.64 +  assumes "eventually (\<lambda>x. P x) F"
    1.65 +  shows "eventually (\<lambda>x. Q x) F"
    1.66  proof (rule eventually_mono)
    1.67    show "\<forall>x. (P x \<longrightarrow> Q x) \<and> P x \<longrightarrow> Q x" by simp
    1.68 -  show "eventually (\<lambda>x. (P x \<longrightarrow> Q x) \<and> P x) A"
    1.69 +  show "eventually (\<lambda>x. (P x \<longrightarrow> Q x) \<and> P x) F"
    1.70      using assms by (rule eventually_conj)
    1.71  qed
    1.72  
    1.73  lemma eventually_rev_mp:
    1.74 -  assumes "eventually (\<lambda>x. P x) A"
    1.75 -  assumes "eventually (\<lambda>x. P x \<longrightarrow> Q x) A"
    1.76 -  shows "eventually (\<lambda>x. Q x) A"
    1.77 +  assumes "eventually (\<lambda>x. P x) F"
    1.78 +  assumes "eventually (\<lambda>x. P x \<longrightarrow> Q x) F"
    1.79 +  shows "eventually (\<lambda>x. Q x) F"
    1.80  using assms(2) assms(1) by (rule eventually_mp)
    1.81  
    1.82  lemma eventually_conj_iff:
    1.83 -  "eventually (\<lambda>x. P x \<and> Q x) A \<longleftrightarrow> eventually P A \<and> eventually Q A"
    1.84 +  "eventually (\<lambda>x. P x \<and> Q x) F \<longleftrightarrow> eventually P F \<and> eventually Q F"
    1.85    by (auto intro: eventually_conj elim: eventually_rev_mp)
    1.86  
    1.87  lemma eventually_elim1:
    1.88 -  assumes "eventually (\<lambda>i. P i) A"
    1.89 +  assumes "eventually (\<lambda>i. P i) F"
    1.90    assumes "\<And>i. P i \<Longrightarrow> Q i"
    1.91 -  shows "eventually (\<lambda>i. Q i) A"
    1.92 +  shows "eventually (\<lambda>i. Q i) F"
    1.93    using assms by (auto elim!: eventually_rev_mp)
    1.94  
    1.95  lemma eventually_elim2:
    1.96 -  assumes "eventually (\<lambda>i. P i) A"
    1.97 -  assumes "eventually (\<lambda>i. Q i) A"
    1.98 +  assumes "eventually (\<lambda>i. P i) F"
    1.99 +  assumes "eventually (\<lambda>i. Q i) F"
   1.100    assumes "\<And>i. P i \<Longrightarrow> Q i \<Longrightarrow> R i"
   1.101 -  shows "eventually (\<lambda>i. R i) A"
   1.102 +  shows "eventually (\<lambda>i. R i) F"
   1.103    using assms by (auto elim!: eventually_rev_mp)
   1.104  
   1.105  subsection {* Finer-than relation *}
   1.106  
   1.107 -text {* @{term "A \<le> B"} means that filter @{term A} is finer than
   1.108 -filter @{term B}. *}
   1.109 +text {* @{term "F \<le> F'"} means that filter @{term F} is finer than
   1.110 +filter @{term F'}. *}
   1.111  
   1.112  instantiation filter :: (type) complete_lattice
   1.113  begin
   1.114  
   1.115  definition le_filter_def:
   1.116 -  "A \<le> B \<longleftrightarrow> (\<forall>P. eventually P B \<longrightarrow> eventually P A)"
   1.117 +  "F \<le> F' \<longleftrightarrow> (\<forall>P. eventually P F' \<longrightarrow> eventually P F)"
   1.118  
   1.119  definition
   1.120 -  "(A :: 'a filter) < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> A"
   1.121 +  "(F :: 'a filter) < F' \<longleftrightarrow> F \<le> F' \<and> \<not> F' \<le> F"
   1.122  
   1.123  definition
   1.124    "top = Abs_filter (\<lambda>P. \<forall>x. P x)"
   1.125 @@ -122,17 +122,17 @@
   1.126    "bot = Abs_filter (\<lambda>P. True)"
   1.127  
   1.128  definition
   1.129 -  "sup A B = Abs_filter (\<lambda>P. eventually P A \<and> eventually P B)"
   1.130 +  "sup F F' = Abs_filter (\<lambda>P. eventually P F \<and> eventually P F')"
   1.131  
   1.132  definition
   1.133 -  "inf A B = Abs_filter
   1.134 -      (\<lambda>P. \<exists>Q R. eventually Q A \<and> eventually R B \<and> (\<forall>x. Q x \<and> R x \<longrightarrow> P x))"
   1.135 +  "inf F F' = Abs_filter
   1.136 +      (\<lambda>P. \<exists>Q R. eventually Q F \<and> eventually R F' \<and> (\<forall>x. Q x \<and> R x \<longrightarrow> P x))"
   1.137  
   1.138  definition
   1.139 -  "Sup S = Abs_filter (\<lambda>P. \<forall>A\<in>S. eventually P A)"
   1.140 +  "Sup S = Abs_filter (\<lambda>P. \<forall>F\<in>S. eventually P F)"
   1.141  
   1.142  definition
   1.143 -  "Inf S = Sup {A::'a filter. \<forall>B\<in>S. A \<le> B}"
   1.144 +  "Inf S = Sup {F::'a filter. \<forall>F'\<in>S. F \<le> F'}"
   1.145  
   1.146  lemma eventually_top [simp]: "eventually P top \<longleftrightarrow> (\<forall>x. P x)"
   1.147    unfolding top_filter_def
   1.148 @@ -143,14 +143,14 @@
   1.149    by (subst eventually_Abs_filter, rule is_filter.intro, auto)
   1.150  
   1.151  lemma eventually_sup:
   1.152 -  "eventually P (sup A B) \<longleftrightarrow> eventually P A \<and> eventually P B"
   1.153 +  "eventually P (sup F F') \<longleftrightarrow> eventually P F \<and> eventually P F'"
   1.154    unfolding sup_filter_def
   1.155    by (rule eventually_Abs_filter, rule is_filter.intro)
   1.156       (auto elim!: eventually_rev_mp)
   1.157  
   1.158  lemma eventually_inf:
   1.159 -  "eventually P (inf A B) \<longleftrightarrow>
   1.160 -   (\<exists>Q R. eventually Q A \<and> eventually R B \<and> (\<forall>x. Q x \<and> R x \<longrightarrow> P x))"
   1.161 +  "eventually P (inf F F') \<longleftrightarrow>
   1.162 +   (\<exists>Q R. eventually Q F \<and> eventually R F' \<and> (\<forall>x. Q x \<and> R x \<longrightarrow> P x))"
   1.163    unfolding inf_filter_def
   1.164    apply (rule eventually_Abs_filter, rule is_filter.intro)
   1.165    apply (fast intro: eventually_True)
   1.166 @@ -163,92 +163,80 @@
   1.167    done
   1.168  
   1.169  lemma eventually_Sup:
   1.170 -  "eventually P (Sup S) \<longleftrightarrow> (\<forall>A\<in>S. eventually P A)"
   1.171 +  "eventually P (Sup S) \<longleftrightarrow> (\<forall>F\<in>S. eventually P F)"
   1.172    unfolding Sup_filter_def
   1.173    apply (rule eventually_Abs_filter, rule is_filter.intro)
   1.174    apply (auto intro: eventually_conj elim!: eventually_rev_mp)
   1.175    done
   1.176  
   1.177  instance proof
   1.178 -  fix A B :: "'a filter" show "A < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> A"
   1.179 -    by (rule less_filter_def)
   1.180 -next
   1.181 -  fix A :: "'a filter" show "A \<le> A"
   1.182 -    unfolding le_filter_def by simp
   1.183 -next
   1.184 -  fix A B C :: "'a filter" assume "A \<le> B" and "B \<le> C" thus "A \<le> C"
   1.185 -    unfolding le_filter_def by simp
   1.186 -next
   1.187 -  fix A B :: "'a filter" assume "A \<le> B" and "B \<le> A" thus "A = B"
   1.188 -    unfolding le_filter_def filter_eq_iff by fast
   1.189 -next
   1.190 -  fix A :: "'a filter" show "A \<le> top"
   1.191 -    unfolding le_filter_def eventually_top by (simp add: always_eventually)
   1.192 -next
   1.193 -  fix A :: "'a filter" show "bot \<le> A"
   1.194 -    unfolding le_filter_def by simp
   1.195 -next
   1.196 -  fix A B :: "'a filter" show "A \<le> sup A B" and "B \<le> sup A B"
   1.197 -    unfolding le_filter_def eventually_sup by simp_all
   1.198 -next
   1.199 -  fix A B C :: "'a filter" assume "A \<le> C" and "B \<le> C" thus "sup A B \<le> C"
   1.200 -    unfolding le_filter_def eventually_sup by simp
   1.201 -next
   1.202 -  fix A B :: "'a filter" show "inf A B \<le> A" and "inf A B \<le> B"
   1.203 -    unfolding le_filter_def eventually_inf by (auto intro: eventually_True)
   1.204 -next
   1.205 -  fix A B C :: "'a filter" assume "A \<le> B" and "A \<le> C" thus "A \<le> inf B C"
   1.206 +  fix F F' F'' :: "'a filter" and S :: "'a filter set"
   1.207 +  { show "F < F' \<longleftrightarrow> F \<le> F' \<and> \<not> F' \<le> F"
   1.208 +    by (rule less_filter_def) }
   1.209 +  { show "F \<le> F"
   1.210 +    unfolding le_filter_def by simp }
   1.211 +  { assume "F \<le> F'" and "F' \<le> F''" thus "F \<le> F''"
   1.212 +    unfolding le_filter_def by simp }
   1.213 +  { assume "F \<le> F'" and "F' \<le> F" thus "F = F'"
   1.214 +    unfolding le_filter_def filter_eq_iff by fast }
   1.215 +  { show "F \<le> top"
   1.216 +    unfolding le_filter_def eventually_top by (simp add: always_eventually) }
   1.217 +  { show "bot \<le> F"
   1.218 +    unfolding le_filter_def by simp }
   1.219 +  { show "F \<le> sup F F'" and "F' \<le> sup F F'"
   1.220 +    unfolding le_filter_def eventually_sup by simp_all }
   1.221 +  { assume "F \<le> F''" and "F' \<le> F''" thus "sup F F' \<le> F''"
   1.222 +    unfolding le_filter_def eventually_sup by simp }
   1.223 +  { show "inf F F' \<le> F" and "inf F F' \<le> F'"
   1.224 +    unfolding le_filter_def eventually_inf by (auto intro: eventually_True) }
   1.225 +  { assume "F \<le> F'" and "F \<le> F''" thus "F \<le> inf F' F''"
   1.226      unfolding le_filter_def eventually_inf
   1.227 -    by (auto elim!: eventually_mono intro: eventually_conj)
   1.228 -next
   1.229 -  fix A :: "'a filter" and S assume "A \<in> S" thus "A \<le> Sup S"
   1.230 -    unfolding le_filter_def eventually_Sup by simp
   1.231 -next
   1.232 -  fix S and B :: "'a filter" assume "\<And>A. A \<in> S \<Longrightarrow> A \<le> B" thus "Sup S \<le> B"
   1.233 -    unfolding le_filter_def eventually_Sup by simp
   1.234 -next
   1.235 -  fix C :: "'a filter" and S assume "C \<in> S" thus "Inf S \<le> C"
   1.236 -    unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp
   1.237 -next
   1.238 -  fix S and A :: "'a filter" assume "\<And>B. B \<in> S \<Longrightarrow> A \<le> B" thus "A \<le> Inf S"
   1.239 -    unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp
   1.240 +    by (auto elim!: eventually_mono intro: eventually_conj) }
   1.241 +  { assume "F \<in> S" thus "F \<le> Sup S"
   1.242 +    unfolding le_filter_def eventually_Sup by simp }
   1.243 +  { assume "\<And>F. F \<in> S \<Longrightarrow> F \<le> F'" thus "Sup S \<le> F'"
   1.244 +    unfolding le_filter_def eventually_Sup by simp }
   1.245 +  { assume "F'' \<in> S" thus "Inf S \<le> F''"
   1.246 +    unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp }
   1.247 +  { assume "\<And>F'. F' \<in> S \<Longrightarrow> F \<le> F'" thus "F \<le> Inf S"
   1.248 +    unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp }
   1.249  qed
   1.250  
   1.251  end
   1.252  
   1.253  lemma filter_leD:
   1.254 -  "A \<le> B \<Longrightarrow> eventually P B \<Longrightarrow> eventually P A"
   1.255 +  "F \<le> F' \<Longrightarrow> eventually P F' \<Longrightarrow> eventually P F"
   1.256    unfolding le_filter_def by simp
   1.257  
   1.258  lemma filter_leI:
   1.259 -  "(\<And>P. eventually P B \<Longrightarrow> eventually P A) \<Longrightarrow> A \<le> B"
   1.260 +  "(\<And>P. eventually P F' \<Longrightarrow> eventually P F) \<Longrightarrow> F \<le> F'"
   1.261    unfolding le_filter_def by simp
   1.262  
   1.263  lemma eventually_False:
   1.264 -  "eventually (\<lambda>x. False) A \<longleftrightarrow> A = bot"
   1.265 +  "eventually (\<lambda>x. False) F \<longleftrightarrow> F = bot"
   1.266    unfolding filter_eq_iff by (auto elim: eventually_rev_mp)
   1.267  
   1.268  subsection {* Map function for filters *}
   1.269  
   1.270  definition filtermap :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a filter \<Rightarrow> 'b filter"
   1.271 -  where "filtermap f A = Abs_filter (\<lambda>P. eventually (\<lambda>x. P (f x)) A)"
   1.272 +  where "filtermap f F = Abs_filter (\<lambda>P. eventually (\<lambda>x. P (f x)) F)"
   1.273  
   1.274  lemma eventually_filtermap:
   1.275 -  "eventually P (filtermap f A) = eventually (\<lambda>x. P (f x)) A"
   1.276 +  "eventually P (filtermap f F) = eventually (\<lambda>x. P (f x)) F"
   1.277    unfolding filtermap_def
   1.278    apply (rule eventually_Abs_filter)
   1.279    apply (rule is_filter.intro)
   1.280    apply (auto elim!: eventually_rev_mp)
   1.281    done
   1.282  
   1.283 -lemma filtermap_ident: "filtermap (\<lambda>x. x) A = A"
   1.284 +lemma filtermap_ident: "filtermap (\<lambda>x. x) F = F"
   1.285    by (simp add: filter_eq_iff eventually_filtermap)
   1.286  
   1.287  lemma filtermap_filtermap:
   1.288 -  "filtermap f (filtermap g A) = filtermap (\<lambda>x. f (g x)) A"
   1.289 +  "filtermap f (filtermap g F) = filtermap (\<lambda>x. f (g x)) F"
   1.290    by (simp add: filter_eq_iff eventually_filtermap)
   1.291  
   1.292 -lemma filtermap_mono: "A \<le> B \<Longrightarrow> filtermap f A \<le> filtermap f B"
   1.293 +lemma filtermap_mono: "F \<le> F' \<Longrightarrow> filtermap f F \<le> filtermap f F'"
   1.294    unfolding le_filter_def eventually_filtermap by simp
   1.295  
   1.296  lemma filtermap_bot [simp]: "filtermap f bot = bot"
   1.297 @@ -279,13 +267,13 @@
   1.298    by (simp add: eventually_False)
   1.299  
   1.300  lemma le_sequentially:
   1.301 -  "A \<le> sequentially \<longleftrightarrow> (\<forall>N. eventually (\<lambda>n. N \<le> n) A)"
   1.302 +  "F \<le> sequentially \<longleftrightarrow> (\<forall>N. eventually (\<lambda>n. N \<le> n) F)"
   1.303    unfolding le_filter_def eventually_sequentially
   1.304    by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp)
   1.305  
   1.306  
   1.307  definition trivial_limit :: "'a filter \<Rightarrow> bool"
   1.308 -  where "trivial_limit A \<longleftrightarrow> eventually (\<lambda>x. False) A"
   1.309 +  where "trivial_limit F \<longleftrightarrow> eventually (\<lambda>x. False) F"
   1.310  
   1.311  lemma trivial_limit_sequentially [intro]: "\<not> trivial_limit sequentially"
   1.312    by (auto simp add: trivial_limit_def eventually_sequentially)
   1.313 @@ -293,7 +281,7 @@
   1.314  subsection {* Standard filters *}
   1.315  
   1.316  definition within :: "'a filter \<Rightarrow> 'a set \<Rightarrow> 'a filter" (infixr "within" 70)
   1.317 -  where "A within S = Abs_filter (\<lambda>P. eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) A)"
   1.318 +  where "F within S = Abs_filter (\<lambda>P. eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) F)"
   1.319  
   1.320  definition nhds :: "'a::topological_space \<Rightarrow> 'a filter"
   1.321    where "nhds a = Abs_filter (\<lambda>P. \<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))"
   1.322 @@ -302,12 +290,12 @@
   1.323    where "at a = nhds a within - {a}"
   1.324  
   1.325  lemma eventually_within:
   1.326 -  "eventually P (A within S) = eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) A"
   1.327 +  "eventually P (F within S) = eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) F"
   1.328    unfolding within_def
   1.329    by (rule eventually_Abs_filter, rule is_filter.intro)
   1.330       (auto elim!: eventually_rev_mp)
   1.331  
   1.332 -lemma within_UNIV: "A within UNIV = A"
   1.333 +lemma within_UNIV: "F within UNIV = F"
   1.334    unfolding filter_eq_iff eventually_within by simp
   1.335  
   1.336  lemma eventually_nhds:
   1.337 @@ -353,51 +341,51 @@
   1.338  subsection {* Boundedness *}
   1.339  
   1.340  definition Bfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
   1.341 -  where "Bfun f A = (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) A)"
   1.342 +  where "Bfun f F = (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)"
   1.343  
   1.344  lemma BfunI:
   1.345 -  assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) A" shows "Bfun f A"
   1.346 +  assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" shows "Bfun f F"
   1.347  unfolding Bfun_def
   1.348  proof (intro exI conjI allI)
   1.349    show "0 < max K 1" by simp
   1.350  next
   1.351 -  show "eventually (\<lambda>x. norm (f x) \<le> max K 1) A"
   1.352 +  show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F"
   1.353      using K by (rule eventually_elim1, simp)
   1.354  qed
   1.355  
   1.356  lemma BfunE:
   1.357 -  assumes "Bfun f A"
   1.358 -  obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) A"
   1.359 +  assumes "Bfun f F"
   1.360 +  obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F"
   1.361  using assms unfolding Bfun_def by fast
   1.362  
   1.363  
   1.364  subsection {* Convergence to Zero *}
   1.365  
   1.366  definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
   1.367 -  where "Zfun f A = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) A)"
   1.368 +  where "Zfun f F = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) F)"
   1.369  
   1.370  lemma ZfunI:
   1.371 -  "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) A) \<Longrightarrow> Zfun f A"
   1.372 +  "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F"
   1.373    unfolding Zfun_def by simp
   1.374  
   1.375  lemma ZfunD:
   1.376 -  "\<lbrakk>Zfun f A; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) A"
   1.377 +  "\<lbrakk>Zfun f F; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F"
   1.378    unfolding Zfun_def by simp
   1.379  
   1.380  lemma Zfun_ssubst:
   1.381 -  "eventually (\<lambda>x. f x = g x) A \<Longrightarrow> Zfun g A \<Longrightarrow> Zfun f A"
   1.382 +  "eventually (\<lambda>x. f x = g x) F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun f F"
   1.383    unfolding Zfun_def by (auto elim!: eventually_rev_mp)
   1.384  
   1.385 -lemma Zfun_zero: "Zfun (\<lambda>x. 0) A"
   1.386 +lemma Zfun_zero: "Zfun (\<lambda>x. 0) F"
   1.387    unfolding Zfun_def by simp
   1.388  
   1.389 -lemma Zfun_norm_iff: "Zfun (\<lambda>x. norm (f x)) A = Zfun (\<lambda>x. f x) A"
   1.390 +lemma Zfun_norm_iff: "Zfun (\<lambda>x. norm (f x)) F = Zfun (\<lambda>x. f x) F"
   1.391    unfolding Zfun_def by simp
   1.392  
   1.393  lemma Zfun_imp_Zfun:
   1.394 -  assumes f: "Zfun f A"
   1.395 -  assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) A"
   1.396 -  shows "Zfun (\<lambda>x. g x) A"
   1.397 +  assumes f: "Zfun f F"
   1.398 +  assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F"
   1.399 +  shows "Zfun (\<lambda>x. g x) F"
   1.400  proof (cases)
   1.401    assume K: "0 < K"
   1.402    show ?thesis
   1.403 @@ -405,9 +393,9 @@
   1.404      fix r::real assume "0 < r"
   1.405      hence "0 < r / K"
   1.406        using K by (rule divide_pos_pos)
   1.407 -    then have "eventually (\<lambda>x. norm (f x) < r / K) A"
   1.408 +    then have "eventually (\<lambda>x. norm (f x) < r / K) F"
   1.409        using ZfunD [OF f] by fast
   1.410 -    with g show "eventually (\<lambda>x. norm (g x) < r) A"
   1.411 +    with g show "eventually (\<lambda>x. norm (g x) < r) F"
   1.412      proof (rule eventually_elim2)
   1.413        fix x
   1.414        assume *: "norm (g x) \<le> norm (f x) * K"
   1.415 @@ -425,7 +413,7 @@
   1.416    proof (rule ZfunI)
   1.417      fix r :: real
   1.418      assume "0 < r"
   1.419 -    from g show "eventually (\<lambda>x. norm (g x) < r) A"
   1.420 +    from g show "eventually (\<lambda>x. norm (g x) < r) F"
   1.421      proof (rule eventually_elim1)
   1.422        fix x
   1.423        assume "norm (g x) \<le> norm (f x) * K"
   1.424 @@ -437,22 +425,22 @@
   1.425    qed
   1.426  qed
   1.427  
   1.428 -lemma Zfun_le: "\<lbrakk>Zfun g A; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f A"
   1.429 +lemma Zfun_le: "\<lbrakk>Zfun g F; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f F"
   1.430    by (erule_tac K="1" in Zfun_imp_Zfun, simp)
   1.431  
   1.432  lemma Zfun_add:
   1.433 -  assumes f: "Zfun f A" and g: "Zfun g A"
   1.434 -  shows "Zfun (\<lambda>x. f x + g x) A"
   1.435 +  assumes f: "Zfun f F" and g: "Zfun g F"
   1.436 +  shows "Zfun (\<lambda>x. f x + g x) F"
   1.437  proof (rule ZfunI)
   1.438    fix r::real assume "0 < r"
   1.439    hence r: "0 < r / 2" by simp
   1.440 -  have "eventually (\<lambda>x. norm (f x) < r/2) A"
   1.441 +  have "eventually (\<lambda>x. norm (f x) < r/2) F"
   1.442      using f r by (rule ZfunD)
   1.443    moreover
   1.444 -  have "eventually (\<lambda>x. norm (g x) < r/2) A"
   1.445 +  have "eventually (\<lambda>x. norm (g x) < r/2) F"
   1.446      using g r by (rule ZfunD)
   1.447    ultimately
   1.448 -  show "eventually (\<lambda>x. norm (f x + g x) < r) A"
   1.449 +  show "eventually (\<lambda>x. norm (f x + g x) < r) F"
   1.450    proof (rule eventually_elim2)
   1.451      fix x
   1.452      assume *: "norm (f x) < r/2" "norm (g x) < r/2"
   1.453 @@ -465,28 +453,28 @@
   1.454    qed
   1.455  qed
   1.456  
   1.457 -lemma Zfun_minus: "Zfun f A \<Longrightarrow> Zfun (\<lambda>x. - f x) A"
   1.458 +lemma Zfun_minus: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. - f x) F"
   1.459    unfolding Zfun_def by simp
   1.460  
   1.461 -lemma Zfun_diff: "\<lbrakk>Zfun f A; Zfun g A\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) A"
   1.462 +lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F"
   1.463    by (simp only: diff_minus Zfun_add Zfun_minus)
   1.464  
   1.465  lemma (in bounded_linear) Zfun:
   1.466 -  assumes g: "Zfun g A"
   1.467 -  shows "Zfun (\<lambda>x. f (g x)) A"
   1.468 +  assumes g: "Zfun g F"
   1.469 +  shows "Zfun (\<lambda>x. f (g x)) F"
   1.470  proof -
   1.471    obtain K where "\<And>x. norm (f x) \<le> norm x * K"
   1.472      using bounded by fast
   1.473 -  then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) A"
   1.474 +  then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) F"
   1.475      by simp
   1.476    with g show ?thesis
   1.477      by (rule Zfun_imp_Zfun)
   1.478  qed
   1.479  
   1.480  lemma (in bounded_bilinear) Zfun:
   1.481 -  assumes f: "Zfun f A"
   1.482 -  assumes g: "Zfun g A"
   1.483 -  shows "Zfun (\<lambda>x. f x ** g x) A"
   1.484 +  assumes f: "Zfun f F"
   1.485 +  assumes g: "Zfun g F"
   1.486 +  shows "Zfun (\<lambda>x. f x ** g x) F"
   1.487  proof (rule ZfunI)
   1.488    fix r::real assume r: "0 < r"
   1.489    obtain K where K: "0 < K"
   1.490 @@ -494,13 +482,13 @@
   1.491      using pos_bounded by fast
   1.492    from K have K': "0 < inverse K"
   1.493      by (rule positive_imp_inverse_positive)
   1.494 -  have "eventually (\<lambda>x. norm (f x) < r) A"
   1.495 +  have "eventually (\<lambda>x. norm (f x) < r) F"
   1.496      using f r by (rule ZfunD)
   1.497    moreover
   1.498 -  have "eventually (\<lambda>x. norm (g x) < inverse K) A"
   1.499 +  have "eventually (\<lambda>x. norm (g x) < inverse K) F"
   1.500      using g K' by (rule ZfunD)
   1.501    ultimately
   1.502 -  show "eventually (\<lambda>x. norm (f x ** g x) < r) A"
   1.503 +  show "eventually (\<lambda>x. norm (f x ** g x) < r) F"
   1.504    proof (rule eventually_elim2)
   1.505      fix x
   1.506      assume *: "norm (f x) < r" "norm (g x) < inverse K"
   1.507 @@ -515,11 +503,11 @@
   1.508  qed
   1.509  
   1.510  lemma (in bounded_bilinear) Zfun_left:
   1.511 -  "Zfun f A \<Longrightarrow> Zfun (\<lambda>x. f x ** a) A"
   1.512 +  "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. f x ** a) F"
   1.513    by (rule bounded_linear_left [THEN bounded_linear.Zfun])
   1.514  
   1.515  lemma (in bounded_bilinear) Zfun_right:
   1.516 -  "Zfun f A \<Longrightarrow> Zfun (\<lambda>x. a ** f x) A"
   1.517 +  "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. a ** f x) F"
   1.518    by (rule bounded_linear_right [THEN bounded_linear.Zfun])
   1.519  
   1.520  lemmas Zfun_mult = mult.Zfun
   1.521 @@ -531,7 +519,7 @@
   1.522  
   1.523  definition tendsto :: "('a \<Rightarrow> 'b::topological_space) \<Rightarrow> 'b \<Rightarrow> 'a filter \<Rightarrow> bool"
   1.524      (infixr "--->" 55) where
   1.525 -  "(f ---> l) A \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) A)"
   1.526 +  "(f ---> l) F \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) F)"
   1.527  
   1.528  ML {*
   1.529  structure Tendsto_Intros = Named_Thms
   1.530 @@ -543,21 +531,21 @@
   1.531  
   1.532  setup Tendsto_Intros.setup
   1.533  
   1.534 -lemma tendsto_mono: "A \<le> A' \<Longrightarrow> (f ---> l) A' \<Longrightarrow> (f ---> l) A"
   1.535 +lemma tendsto_mono: "F \<le> F' \<Longrightarrow> (f ---> l) F' \<Longrightarrow> (f ---> l) F"
   1.536    unfolding tendsto_def le_filter_def by fast
   1.537  
   1.538  lemma topological_tendstoI:
   1.539 -  "(\<And>S. open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) A)
   1.540 -    \<Longrightarrow> (f ---> l) A"
   1.541 +  "(\<And>S. open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) F)
   1.542 +    \<Longrightarrow> (f ---> l) F"
   1.543    unfolding tendsto_def by auto
   1.544  
   1.545  lemma topological_tendstoD:
   1.546 -  "(f ---> l) A \<Longrightarrow> open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) A"
   1.547 +  "(f ---> l) F \<Longrightarrow> open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) F"
   1.548    unfolding tendsto_def by auto
   1.549  
   1.550  lemma tendstoI:
   1.551 -  assumes "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) A"
   1.552 -  shows "(f ---> l) A"
   1.553 +  assumes "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F"
   1.554 +  shows "(f ---> l) F"
   1.555    apply (rule topological_tendstoI)
   1.556    apply (simp add: open_dist)
   1.557    apply (drule (1) bspec, clarify)
   1.558 @@ -566,7 +554,7 @@
   1.559    done
   1.560  
   1.561  lemma tendstoD:
   1.562 -  "(f ---> l) A \<Longrightarrow> 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) A"
   1.563 +  "(f ---> l) F \<Longrightarrow> 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F"
   1.564    apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD)
   1.565    apply (clarsimp simp add: open_dist)
   1.566    apply (rule_tac x="e - dist x l" in exI, clarsimp)
   1.567 @@ -577,10 +565,10 @@
   1.568    done
   1.569  
   1.570  lemma tendsto_iff:
   1.571 -  "(f ---> l) A \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) A)"
   1.572 +  "(f ---> l) F \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) F)"
   1.573    using tendstoI tendstoD by fast
   1.574  
   1.575 -lemma tendsto_Zfun_iff: "(f ---> a) A = Zfun (\<lambda>x. f x - a) A"
   1.576 +lemma tendsto_Zfun_iff: "(f ---> a) F = Zfun (\<lambda>x. f x - a) F"
   1.577    by (simp only: tendsto_iff Zfun_def dist_norm)
   1.578  
   1.579  lemma tendsto_ident_at [tendsto_intros]: "((\<lambda>x. x) ---> a) (at a)"
   1.580 @@ -590,12 +578,12 @@
   1.581    "((\<lambda>x. x) ---> a) (at a within S)"
   1.582    unfolding tendsto_def eventually_within eventually_at_topological by auto
   1.583  
   1.584 -lemma tendsto_const [tendsto_intros]: "((\<lambda>x. k) ---> k) A"
   1.585 +lemma tendsto_const [tendsto_intros]: "((\<lambda>x. k) ---> k) F"
   1.586    by (simp add: tendsto_def)
   1.587  
   1.588  lemma tendsto_const_iff:
   1.589    fixes k l :: "'a::metric_space"
   1.590 -  assumes "A \<noteq> bot" shows "((\<lambda>n. k) ---> l) A \<longleftrightarrow> k = l"
   1.591 +  assumes "F \<noteq> bot" shows "((\<lambda>n. k) ---> l) F \<longleftrightarrow> k = l"
   1.592    apply (safe intro!: tendsto_const)
   1.593    apply (rule ccontr)
   1.594    apply (drule_tac e="dist k l" in tendstoD)
   1.595 @@ -604,13 +592,13 @@
   1.596    done
   1.597  
   1.598  lemma tendsto_dist [tendsto_intros]:
   1.599 -  assumes f: "(f ---> l) A" and g: "(g ---> m) A"
   1.600 -  shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) A"
   1.601 +  assumes f: "(f ---> l) F" and g: "(g ---> m) F"
   1.602 +  shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) F"
   1.603  proof (rule tendstoI)
   1.604    fix e :: real assume "0 < e"
   1.605    hence e2: "0 < e/2" by simp
   1.606    from tendstoD [OF f e2] tendstoD [OF g e2]
   1.607 -  show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) A"
   1.608 +  show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F"
   1.609    proof (rule eventually_elim2)
   1.610      fix x assume "dist (f x) l < e/2" "dist (g x) m < e/2"
   1.611      then show "dist (dist (f x) (g x)) (dist l m) < e"
   1.612 @@ -629,68 +617,68 @@
   1.613    unfolding dist_norm by simp
   1.614  
   1.615  lemma tendsto_norm [tendsto_intros]:
   1.616 -  "(f ---> a) A \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> norm a) A"
   1.617 +  "(f ---> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> norm a) F"
   1.618    unfolding norm_conv_dist by (intro tendsto_intros)
   1.619  
   1.620  lemma tendsto_norm_zero:
   1.621 -  "(f ---> 0) A \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> 0) A"
   1.622 +  "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> 0) F"
   1.623    by (drule tendsto_norm, simp)
   1.624  
   1.625  lemma tendsto_norm_zero_cancel:
   1.626 -  "((\<lambda>x. norm (f x)) ---> 0) A \<Longrightarrow> (f ---> 0) A"
   1.627 +  "((\<lambda>x. norm (f x)) ---> 0) F \<Longrightarrow> (f ---> 0) F"
   1.628    unfolding tendsto_iff dist_norm by simp
   1.629  
   1.630  lemma tendsto_norm_zero_iff:
   1.631 -  "((\<lambda>x. norm (f x)) ---> 0) A \<longleftrightarrow> (f ---> 0) A"
   1.632 +  "((\<lambda>x. norm (f x)) ---> 0) F \<longleftrightarrow> (f ---> 0) F"
   1.633    unfolding tendsto_iff dist_norm by simp
   1.634  
   1.635  lemma tendsto_rabs [tendsto_intros]:
   1.636 -  "(f ---> (l::real)) A \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> \<bar>l\<bar>) A"
   1.637 +  "(f ---> (l::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> \<bar>l\<bar>) F"
   1.638    by (fold real_norm_def, rule tendsto_norm)
   1.639  
   1.640  lemma tendsto_rabs_zero:
   1.641 -  "(f ---> (0::real)) A \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> 0) A"
   1.642 +  "(f ---> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> 0) F"
   1.643    by (fold real_norm_def, rule tendsto_norm_zero)
   1.644  
   1.645  lemma tendsto_rabs_zero_cancel:
   1.646 -  "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) A \<Longrightarrow> (f ---> 0) A"
   1.647 +  "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<Longrightarrow> (f ---> 0) F"
   1.648    by (fold real_norm_def, rule tendsto_norm_zero_cancel)
   1.649  
   1.650  lemma tendsto_rabs_zero_iff:
   1.651 -  "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) A \<longleftrightarrow> (f ---> 0) A"
   1.652 +  "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<longleftrightarrow> (f ---> 0) F"
   1.653    by (fold real_norm_def, rule tendsto_norm_zero_iff)
   1.654  
   1.655  subsubsection {* Addition and subtraction *}
   1.656  
   1.657  lemma tendsto_add [tendsto_intros]:
   1.658    fixes a b :: "'a::real_normed_vector"
   1.659 -  shows "\<lbrakk>(f ---> a) A; (g ---> b) A\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> a + b) A"
   1.660 +  shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> a + b) F"
   1.661    by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add)
   1.662  
   1.663  lemma tendsto_add_zero:
   1.664    fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"
   1.665 -  shows "\<lbrakk>(f ---> 0) A; (g ---> 0) A\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> 0) A"
   1.666 +  shows "\<lbrakk>(f ---> 0) F; (g ---> 0) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> 0) F"
   1.667    by (drule (1) tendsto_add, simp)
   1.668  
   1.669  lemma tendsto_minus [tendsto_intros]:
   1.670    fixes a :: "'a::real_normed_vector"
   1.671 -  shows "(f ---> a) A \<Longrightarrow> ((\<lambda>x. - f x) ---> - a) A"
   1.672 +  shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. - f x) ---> - a) F"
   1.673    by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus)
   1.674  
   1.675  lemma tendsto_minus_cancel:
   1.676    fixes a :: "'a::real_normed_vector"
   1.677 -  shows "((\<lambda>x. - f x) ---> - a) A \<Longrightarrow> (f ---> a) A"
   1.678 +  shows "((\<lambda>x. - f x) ---> - a) F \<Longrightarrow> (f ---> a) F"
   1.679    by (drule tendsto_minus, simp)
   1.680  
   1.681  lemma tendsto_diff [tendsto_intros]:
   1.682    fixes a b :: "'a::real_normed_vector"
   1.683 -  shows "\<lbrakk>(f ---> a) A; (g ---> b) A\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) A"
   1.684 +  shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) F"
   1.685    by (simp add: diff_minus tendsto_add tendsto_minus)
   1.686  
   1.687  lemma tendsto_setsum [tendsto_intros]:
   1.688    fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::real_normed_vector"
   1.689 -  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> a i) A"
   1.690 -  shows "((\<lambda>x. \<Sum>i\<in>S. f i x) ---> (\<Sum>i\<in>S. a i)) A"
   1.691 +  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> a i) F"
   1.692 +  shows "((\<lambda>x. \<Sum>i\<in>S. f i x) ---> (\<Sum>i\<in>S. a i)) F"
   1.693  proof (cases "finite S")
   1.694    assume "finite S" thus ?thesis using assms
   1.695      by (induct, simp add: tendsto_const, simp add: tendsto_add)
   1.696 @@ -702,43 +690,43 @@
   1.697  subsubsection {* Linear operators and multiplication *}
   1.698  
   1.699  lemma (in bounded_linear) tendsto [tendsto_intros]:
   1.700 -  "(g ---> a) A \<Longrightarrow> ((\<lambda>x. f (g x)) ---> f a) A"
   1.701 +  "(g ---> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> f a) F"
   1.702    by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun)
   1.703  
   1.704  lemma (in bounded_linear) tendsto_zero:
   1.705 -  "(g ---> 0) A \<Longrightarrow> ((\<lambda>x. f (g x)) ---> 0) A"
   1.706 +  "(g ---> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> 0) F"
   1.707    by (drule tendsto, simp only: zero)
   1.708  
   1.709  lemma (in bounded_bilinear) tendsto [tendsto_intros]:
   1.710 -  "\<lbrakk>(f ---> a) A; (g ---> b) A\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) ---> a ** b) A"
   1.711 +  "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) ---> a ** b) F"
   1.712    by (simp only: tendsto_Zfun_iff prod_diff_prod
   1.713                   Zfun_add Zfun Zfun_left Zfun_right)
   1.714  
   1.715  lemma (in bounded_bilinear) tendsto_zero:
   1.716 -  assumes f: "(f ---> 0) A"
   1.717 -  assumes g: "(g ---> 0) A"
   1.718 -  shows "((\<lambda>x. f x ** g x) ---> 0) A"
   1.719 +  assumes f: "(f ---> 0) F"
   1.720 +  assumes g: "(g ---> 0) F"
   1.721 +  shows "((\<lambda>x. f x ** g x) ---> 0) F"
   1.722    using tendsto [OF f g] by (simp add: zero_left)
   1.723  
   1.724  lemma (in bounded_bilinear) tendsto_left_zero:
   1.725 -  "(f ---> 0) A \<Longrightarrow> ((\<lambda>x. f x ** c) ---> 0) A"
   1.726 +  "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. f x ** c) ---> 0) F"
   1.727    by (rule bounded_linear.tendsto_zero [OF bounded_linear_left])
   1.728  
   1.729  lemma (in bounded_bilinear) tendsto_right_zero:
   1.730 -  "(f ---> 0) A \<Longrightarrow> ((\<lambda>x. c ** f x) ---> 0) A"
   1.731 +  "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. c ** f x) ---> 0) F"
   1.732    by (rule bounded_linear.tendsto_zero [OF bounded_linear_right])
   1.733  
   1.734  lemmas tendsto_mult = mult.tendsto
   1.735  
   1.736  lemma tendsto_power [tendsto_intros]:
   1.737    fixes f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
   1.738 -  shows "(f ---> a) A \<Longrightarrow> ((\<lambda>x. f x ^ n) ---> a ^ n) A"
   1.739 +  shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) ---> a ^ n) F"
   1.740    by (induct n) (simp_all add: tendsto_const tendsto_mult)
   1.741  
   1.742  lemma tendsto_setprod [tendsto_intros]:
   1.743    fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
   1.744 -  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> L i) A"
   1.745 -  shows "((\<lambda>x. \<Prod>i\<in>S. f i x) ---> (\<Prod>i\<in>S. L i)) A"
   1.746 +  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> L i) F"
   1.747 +  shows "((\<lambda>x. \<Prod>i\<in>S. f i x) ---> (\<Prod>i\<in>S. L i)) F"
   1.748  proof (cases "finite S")
   1.749    assume "finite S" thus ?thesis using assms
   1.750      by (induct, simp add: tendsto_const, simp add: tendsto_mult)
   1.751 @@ -750,17 +738,17 @@
   1.752  subsubsection {* Inverse and division *}
   1.753  
   1.754  lemma (in bounded_bilinear) Zfun_prod_Bfun:
   1.755 -  assumes f: "Zfun f A"
   1.756 -  assumes g: "Bfun g A"
   1.757 -  shows "Zfun (\<lambda>x. f x ** g x) A"
   1.758 +  assumes f: "Zfun f F"
   1.759 +  assumes g: "Bfun g F"
   1.760 +  shows "Zfun (\<lambda>x. f x ** g x) F"
   1.761  proof -
   1.762    obtain K where K: "0 \<le> K"
   1.763      and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
   1.764      using nonneg_bounded by fast
   1.765    obtain B where B: "0 < B"
   1.766 -    and norm_g: "eventually (\<lambda>x. norm (g x) \<le> B) A"
   1.767 +    and norm_g: "eventually (\<lambda>x. norm (g x) \<le> B) F"
   1.768      using g by (rule BfunE)
   1.769 -  have "eventually (\<lambda>x. norm (f x ** g x) \<le> norm (f x) * (B * K)) A"
   1.770 +  have "eventually (\<lambda>x. norm (f x ** g x) \<le> norm (f x) * (B * K)) F"
   1.771    using norm_g proof (rule eventually_elim1)
   1.772      fix x
   1.773      assume *: "norm (g x) \<le> B"
   1.774 @@ -788,9 +776,9 @@
   1.775    using bounded by fast
   1.776  
   1.777  lemma (in bounded_bilinear) Bfun_prod_Zfun:
   1.778 -  assumes f: "Bfun f A"
   1.779 -  assumes g: "Zfun g A"
   1.780 -  shows "Zfun (\<lambda>x. f x ** g x) A"
   1.781 +  assumes f: "Bfun f F"
   1.782 +  assumes g: "Zfun g F"
   1.783 +  shows "Zfun (\<lambda>x. f x ** g x) F"
   1.784    using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun)
   1.785  
   1.786  lemma Bfun_inverse_lemma:
   1.787 @@ -802,16 +790,16 @@
   1.788  
   1.789  lemma Bfun_inverse:
   1.790    fixes a :: "'a::real_normed_div_algebra"
   1.791 -  assumes f: "(f ---> a) A"
   1.792 +  assumes f: "(f ---> a) F"
   1.793    assumes a: "a \<noteq> 0"
   1.794 -  shows "Bfun (\<lambda>x. inverse (f x)) A"
   1.795 +  shows "Bfun (\<lambda>x. inverse (f x)) F"
   1.796  proof -
   1.797    from a have "0 < norm a" by simp
   1.798    hence "\<exists>r>0. r < norm a" by (rule dense)
   1.799    then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
   1.800 -  have "eventually (\<lambda>x. dist (f x) a < r) A"
   1.801 +  have "eventually (\<lambda>x. dist (f x) a < r) F"
   1.802      using tendstoD [OF f r1] by fast
   1.803 -  hence "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) A"
   1.804 +  hence "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F"
   1.805    proof (rule eventually_elim1)
   1.806      fix x
   1.807      assume "dist (f x) a < r"
   1.808 @@ -838,8 +826,8 @@
   1.809  
   1.810  lemma tendsto_inverse_lemma:
   1.811    fixes a :: "'a::real_normed_div_algebra"
   1.812 -  shows "\<lbrakk>(f ---> a) A; a \<noteq> 0; eventually (\<lambda>x. f x \<noteq> 0) A\<rbrakk>
   1.813 -         \<Longrightarrow> ((\<lambda>x. inverse (f x)) ---> inverse a) A"
   1.814 +  shows "\<lbrakk>(f ---> a) F; a \<noteq> 0; eventually (\<lambda>x. f x \<noteq> 0) F\<rbrakk>
   1.815 +         \<Longrightarrow> ((\<lambda>x. inverse (f x)) ---> inverse a) F"
   1.816    apply (subst tendsto_Zfun_iff)
   1.817    apply (rule Zfun_ssubst)
   1.818    apply (erule eventually_elim1)
   1.819 @@ -853,14 +841,14 @@
   1.820  
   1.821  lemma tendsto_inverse [tendsto_intros]:
   1.822    fixes a :: "'a::real_normed_div_algebra"
   1.823 -  assumes f: "(f ---> a) A"
   1.824 +  assumes f: "(f ---> a) F"
   1.825    assumes a: "a \<noteq> 0"
   1.826 -  shows "((\<lambda>x. inverse (f x)) ---> inverse a) A"
   1.827 +  shows "((\<lambda>x. inverse (f x)) ---> inverse a) F"
   1.828  proof -
   1.829    from a have "0 < norm a" by simp
   1.830 -  with f have "eventually (\<lambda>x. dist (f x) a < norm a) A"
   1.831 +  with f have "eventually (\<lambda>x. dist (f x) a < norm a) F"
   1.832      by (rule tendstoD)
   1.833 -  then have "eventually (\<lambda>x. f x \<noteq> 0) A"
   1.834 +  then have "eventually (\<lambda>x. f x \<noteq> 0) F"
   1.835      unfolding dist_norm by (auto elim!: eventually_elim1)
   1.836    with f a show ?thesis
   1.837      by (rule tendsto_inverse_lemma)
   1.838 @@ -868,39 +856,39 @@
   1.839  
   1.840  lemma tendsto_divide [tendsto_intros]:
   1.841    fixes a b :: "'a::real_normed_field"
   1.842 -  shows "\<lbrakk>(f ---> a) A; (g ---> b) A; b \<noteq> 0\<rbrakk>
   1.843 -    \<Longrightarrow> ((\<lambda>x. f x / g x) ---> a / b) A"
   1.844 +  shows "\<lbrakk>(f ---> a) F; (g ---> b) F; b \<noteq> 0\<rbrakk>
   1.845 +    \<Longrightarrow> ((\<lambda>x. f x / g x) ---> a / b) F"
   1.846    by (simp add: mult.tendsto tendsto_inverse divide_inverse)
   1.847  
   1.848  lemma tendsto_sgn [tendsto_intros]:
   1.849    fixes l :: "'a::real_normed_vector"
   1.850 -  shows "\<lbrakk>(f ---> l) A; l \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. sgn (f x)) ---> sgn l) A"
   1.851 +  shows "\<lbrakk>(f ---> l) F; l \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. sgn (f x)) ---> sgn l) F"
   1.852    unfolding sgn_div_norm by (simp add: tendsto_intros)
   1.853  
   1.854  subsubsection {* Uniqueness *}
   1.855  
   1.856  lemma tendsto_unique:
   1.857    fixes f :: "'a \<Rightarrow> 'b::t2_space"
   1.858 -  assumes "\<not> trivial_limit A"  "(f ---> l) A"  "(f ---> l') A"
   1.859 +  assumes "\<not> trivial_limit F"  "(f ---> l) F"  "(f ---> l') F"
   1.860    shows "l = l'"
   1.861  proof (rule ccontr)
   1.862    assume "l \<noteq> l'"
   1.863    obtain U V where "open U" "open V" "l \<in> U" "l' \<in> V" "U \<inter> V = {}"
   1.864      using hausdorff [OF `l \<noteq> l'`] by fast
   1.865 -  have "eventually (\<lambda>x. f x \<in> U) A"
   1.866 -    using `(f ---> l) A` `open U` `l \<in> U` by (rule topological_tendstoD)
   1.867 +  have "eventually (\<lambda>x. f x \<in> U) F"
   1.868 +    using `(f ---> l) F` `open U` `l \<in> U` by (rule topological_tendstoD)
   1.869    moreover
   1.870 -  have "eventually (\<lambda>x. f x \<in> V) A"
   1.871 -    using `(f ---> l') A` `open V` `l' \<in> V` by (rule topological_tendstoD)
   1.872 +  have "eventually (\<lambda>x. f x \<in> V) F"
   1.873 +    using `(f ---> l') F` `open V` `l' \<in> V` by (rule topological_tendstoD)
   1.874    ultimately
   1.875 -  have "eventually (\<lambda>x. False) A"
   1.876 +  have "eventually (\<lambda>x. False) F"
   1.877    proof (rule eventually_elim2)
   1.878      fix x
   1.879      assume "f x \<in> U" "f x \<in> V"
   1.880      hence "f x \<in> U \<inter> V" by simp
   1.881      with `U \<inter> V = {}` show "False" by simp
   1.882    qed
   1.883 -  with `\<not> trivial_limit A` show "False"
   1.884 +  with `\<not> trivial_limit F` show "False"
   1.885      by (simp add: trivial_limit_def)
   1.886  qed
   1.887