| author | desharna | 
| Wed, 12 Feb 2025 08:48:15 +0100 | |
| changeset 82136 | b97cea26d3a3 | 
| parent 78750 | f229090cb8a3 | 
| child 82520 | 1b17f0a41aa3 | 
| permissions | -rw-r--r-- | 
| 69874 | 1 | theory Abstract_Limits | 
| 2 | imports | |
| 3 | Abstract_Topology | |
| 4 | begin | |
| 5 | ||
| 6 | subsection\<open>nhdsin and atin\<close> | |
| 7 | ||
| 8 | definition nhdsin :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a filter" | |
| 9 | where "nhdsin X a = | |
| 70337 
48609a6af1a0
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
70065diff
changeset | 10 |            (if a \<in> topspace X then (INF S\<in>{S. openin X S \<and> a \<in> S}. principal S) else bot)"
 | 
| 69874 | 11 | |
| 78750 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 12 | definition atin_within :: "['a topology, 'a, 'a set] \<Rightarrow> 'a filter" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 13 |   where "atin_within X a S = inf (nhdsin X a) (principal (topspace X \<inter> S - {a}))"
 | 
| 69874 | 14 | |
| 78750 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 15 | abbreviation atin :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a filter" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 16 | where "atin X a \<equiv> atin_within X a UNIV" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 17 | |
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 18 | lemma atin_def: "atin X a = inf (nhdsin X a) (principal (topspace X - {a}))"
 | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 19 | by (simp add: atin_within_def) | 
| 69874 | 20 | |
| 21 | lemma nhdsin_degenerate [simp]: "a \<notin> topspace X \<Longrightarrow> nhdsin X a = bot" | |
| 22 | and atin_degenerate [simp]: "a \<notin> topspace X \<Longrightarrow> atin X a = bot" | |
| 23 | by (simp_all add: nhdsin_def atin_def) | |
| 24 | ||
| 25 | lemma eventually_nhdsin: | |
| 26 | "eventually P (nhdsin X a) \<longleftrightarrow> a \<notin> topspace X \<or> (\<exists>S. openin X S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))" | |
| 27 | proof (cases "a \<in> topspace X") | |
| 28 | case True | |
| 70337 
48609a6af1a0
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
70065diff
changeset | 29 |   hence "nhdsin X a = (INF S\<in>{S. openin X S \<and> a \<in> S}. principal S)"
 | 
| 69874 | 30 | by (simp add: nhdsin_def) | 
| 31 | also have "eventually P \<dots> \<longleftrightarrow> (\<exists>S. openin X S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))" | |
| 32 | using True by (subst eventually_INF_base) (auto simp: eventually_principal) | |
| 33 | finally show ?thesis using True by simp | |
| 34 | qed auto | |
| 35 | ||
| 78750 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 36 | lemma eventually_atin_within: | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 37 | "eventually P (atin_within X a S) \<longleftrightarrow> a \<notin> topspace X \<or> (\<exists>T. openin X T \<and> a \<in> T \<and> (\<forall>x\<in>T. x \<in> S \<and> x \<noteq> a \<longrightarrow> P x))" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 38 | proof (cases "a \<in> topspace X") | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 39 | case True | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 40 | hence "eventually P (atin_within X a S) \<longleftrightarrow> | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 41 | (\<exists>T. openin X T \<and> a \<in> T \<and> | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 42 | (\<forall>x\<in>T. x \<in> topspace X \<and> x \<in> S \<and> x \<noteq> a \<longrightarrow> P x))" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 43 | by (simp add: atin_within_def eventually_inf_principal eventually_nhdsin) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 44 | also have "\<dots> \<longleftrightarrow> (\<exists>T. openin X T \<and> a \<in> T \<and> (\<forall>x\<in>T. x \<in> S \<and> x \<noteq> a \<longrightarrow> P x))" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 45 | using openin_subset by (intro ex_cong) auto | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 46 | finally show ?thesis by (simp add: True) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 47 | qed (simp add: atin_within_def) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 48 | |
| 69874 | 49 | lemma eventually_atin: | 
| 50 | "eventually P (atin X a) \<longleftrightarrow> a \<notin> topspace X \<or> | |
| 51 |              (\<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x \<in> U - {a}. P x))"
 | |
| 78750 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 52 | by (auto simp add: eventually_atin_within) | 
| 69874 | 53 | |
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 54 | lemma nontrivial_limit_atin: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 55 | "atin X a \<noteq> bot \<longleftrightarrow> a \<in> X derived_set_of topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 56 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 57 | assume L: "atin X a \<noteq> bot" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 58 | then have "a \<in> topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 59 | by (meson atin_degenerate) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 60 |   moreover have "\<not> openin X {a}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 61 | using L by (auto simp: eventually_atin trivial_limit_eq) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 62 | ultimately | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 63 | show "a \<in> X derived_set_of topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 64 | by (auto simp: derived_set_of_topspace) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 65 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 66 | assume a: "a \<in> X derived_set_of topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 67 | show "atin X a \<noteq> bot" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 68 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 69 | assume "atin X a = bot" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 70 | then have "eventually (\<lambda>_. False) (atin X a)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 71 | by simp | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 72 | then show False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 73 | by (smt (verit, best) a eventually_atin in_derived_set_of insertE insert_Diff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 74 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 75 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 76 | |
| 78750 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 77 | lemma eventually_atin_subtopology: | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 78 | assumes "a \<in> topspace X" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 79 | shows "eventually P (atin (subtopology X S) a) \<longleftrightarrow> | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 80 |     (a \<in> S \<longrightarrow> (\<exists>U. openin (subtopology X S) U \<and> a \<in> U \<and> (\<forall>x\<in>U - {a}. P x)))"
 | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 81 | using assms by (simp add: eventually_atin) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 82 | |
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 83 | lemma eventually_within_imp: | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 84 | "eventually P (atin_within X a S) \<longleftrightarrow> eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) (atin X a)" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 85 | by (auto simp add: eventually_atin_within eventually_atin) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 86 | |
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 87 | lemma atin_subtopology_within: | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 88 | assumes "a \<in> S" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 89 | shows "atin (subtopology X S) a = atin_within X a S" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 90 | proof - | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 91 | have "eventually P (atin (subtopology X S) a) \<longleftrightarrow> eventually P (atin_within X a S)" for P | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 92 | unfolding eventually_atin eventually_atin_within openin_subtopology | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 93 | using assms by auto | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 94 | then show ?thesis | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 95 | by (meson le_filter_def order.eq_iff) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 96 | qed | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 97 | |
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 98 | lemma atin_subtopology_within_if: | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 99 | shows "atin (subtopology X S) a = (if a \<in> S then atin_within X a S else bot)" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 100 | by (simp add: atin_subtopology_within) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 101 | |
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 102 | lemma trivial_limit_atpointof_within: | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 103 | "trivial_limit(atin_within X a S) \<longleftrightarrow> | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 104 | (a \<notin> X derived_set_of S)" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 105 | by (auto simp: trivial_limit_def eventually_atin_within in_derived_set_of) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 106 | |
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 107 | lemma derived_set_of_trivial_limit: | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 108 | "a \<in> X derived_set_of S \<longleftrightarrow> \<not> trivial_limit(atin_within X a S)" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 109 | by (simp add: trivial_limit_atpointof_within) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 110 | |
| 69874 | 111 | |
| 112 | subsection\<open>Limits in a topological space\<close> | |
| 113 | ||
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 114 | definition limitin :: "'a topology \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b filter \<Rightarrow> bool" where
 | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 115 | "limitin X f l F \<equiv> l \<in> topspace X \<and> (\<forall>U. openin X U \<and> l \<in> U \<longrightarrow> eventually (\<lambda>x. f x \<in> U) F)" | 
| 69874 | 116 | |
| 78750 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 117 | lemma limit_within_subset: | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 118 | "\<lbrakk>limitin X f l (atin_within Y a S); T \<subseteq> S\<rbrakk> \<Longrightarrow> limitin X f l (atin_within Y a T)" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 119 | by (smt (verit) eventually_atin_within limitin_def subset_eq) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 120 | |
| 77934 | 121 | lemma limitinD: "\<lbrakk>limitin X f l F; openin X U; l \<in> U\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. f x \<in> U) F" | 
| 122 | by (simp add: limitin_def) | |
| 123 | ||
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 124 | lemma limitin_canonical_iff [simp]: "limitin euclidean f l F \<longleftrightarrow> (f \<longlongrightarrow> l) F" | 
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 125 | by (auto simp: limitin_def tendsto_def) | 
| 69874 | 126 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 127 | lemma limitin_topspace: "limitin X f l F \<Longrightarrow> l \<in> topspace X" | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 128 | by (simp add: limitin_def) | 
| 69874 | 129 | |
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 130 | lemma limitin_const_iff [simp]: "limitin X (\<lambda>a. l) l F \<longleftrightarrow> l \<in> topspace X" | 
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 131 | by (simp add: limitin_def) | 
| 69874 | 132 | |
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 133 | lemma limitin_const: "limitin euclidean (\<lambda>a. l) l F" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 134 | by simp | 
| 69874 | 135 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 136 | lemma limitin_eventually: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 137 | "\<lbrakk>l \<in> topspace X; eventually (\<lambda>x. f x = l) F\<rbrakk> \<Longrightarrow> limitin X f l F" | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 138 | by (auto simp: limitin_def eventually_mono) | 
| 69874 | 139 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 140 | lemma limitin_subsequence: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 141 | "\<lbrakk>strict_mono r; limitin X f l sequentially\<rbrakk> \<Longrightarrow> limitin X (f \<circ> r) l sequentially" | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 142 | unfolding limitin_def using eventually_subseq by fastforce | 
| 69874 | 143 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 144 | lemma limitin_subtopology: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 145 | "limitin (subtopology X S) f l F | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 146 | \<longleftrightarrow> l \<in> S \<and> eventually (\<lambda>a. f a \<in> S) F \<and> limitin X f l F" (is "?lhs = ?rhs") | 
| 69874 | 147 | proof (cases "l \<in> S \<inter> topspace X") | 
| 148 | case True | |
| 149 | show ?thesis | |
| 150 | proof | |
| 151 | assume L: ?lhs | |
| 152 | with True | |
| 153 | have "\<forall>\<^sub>F b in F. f b \<in> topspace X \<inter> S" | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 154 | by (metis (no_types) limitin_def openin_topspace topspace_subtopology) | 
| 69874 | 155 | with L show ?rhs | 
| 71172 | 156 | apply (clarsimp simp add: limitin_def eventually_mono openin_subtopology_alt) | 
| 69874 | 157 | apply (drule_tac x="S \<inter> U" in spec, force simp: elim: eventually_mono) | 
| 158 | done | |
| 159 | next | |
| 160 | assume ?rhs | |
| 161 | then show ?lhs | |
| 162 | using eventually_elim2 | |
| 71172 | 163 | by (fastforce simp add: limitin_def openin_subtopology_alt) | 
| 69874 | 164 | qed | 
| 71172 | 165 | qed (auto simp: limitin_def) | 
| 69874 | 166 | |
| 167 | ||
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 168 | lemma limitin_canonical_iff_gen [simp]: | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 169 | assumes "open S" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 170 | shows "limitin (top_of_set S) f l F \<longleftrightarrow> (f \<longlongrightarrow> l) F \<and> l \<in> S" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 171 | using assms by (auto simp: limitin_subtopology tendsto_def) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 172 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 173 | lemma limitin_sequentially: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 174 | "limitin X S l sequentially \<longleftrightarrow> | 
| 69874 | 175 | l \<in> topspace X \<and> (\<forall>U. openin X U \<and> l \<in> U \<longrightarrow> (\<exists>N. \<forall>n. N \<le> n \<longrightarrow> S n \<in> U))" | 
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 176 | by (simp add: limitin_def eventually_sequentially) | 
| 69874 | 177 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 178 | lemma limitin_sequentially_offset: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 179 | "limitin X f l sequentially \<Longrightarrow> limitin X (\<lambda>i. f (i + k)) l sequentially" | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 180 | unfolding limitin_sequentially | 
| 69874 | 181 | by (metis add.commute le_add2 order_trans) | 
| 182 | ||
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 183 | lemma limitin_sequentially_offset_rev: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 184 | assumes "limitin X (\<lambda>i. f (i + k)) l sequentially" | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 185 | shows "limitin X f l sequentially" | 
| 69874 | 186 | proof - | 
| 187 | have "\<exists>N. \<forall>n\<ge>N. f n \<in> U" if U: "openin X U" "l \<in> U" for U | |
| 188 | proof - | |
| 189 | obtain N where "\<And>n. n\<ge>N \<Longrightarrow> f (n + k) \<in> U" | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 190 | using assms U unfolding limitin_sequentially by blast | 
| 69874 | 191 | then have "\<forall>n\<ge>N+k. f n \<in> U" | 
| 192 | by (metis add_leD2 le_add_diff_inverse ordered_cancel_comm_monoid_diff_class.le_diff_conv2 add.commute) | |
| 193 | then show ?thesis .. | |
| 194 | qed | |
| 195 | with assms show ?thesis | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 196 | unfolding limitin_sequentially | 
| 69874 | 197 | by simp | 
| 198 | qed | |
| 199 | ||
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 200 | lemma limitin_atin: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 201 | "limitin Y f y (atin X x) \<longleftrightarrow> | 
| 69874 | 202 | y \<in> topspace Y \<and> | 
| 203 | (x \<in> topspace X | |
| 204 | \<longrightarrow> (\<forall>V. openin Y V \<and> y \<in> V | |
| 205 |                  \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> f ` (U - {x}) \<subseteq> V)))"
 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 206 | by (auto simp: limitin_def eventually_atin image_subset_iff) | 
| 69874 | 207 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 208 | lemma limitin_atin_self: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 209 | "limitin Y f (f a) (atin X a) \<longleftrightarrow> | 
| 69874 | 210 | f a \<in> topspace Y \<and> | 
| 211 | (a \<in> topspace X | |
| 212 | \<longrightarrow> (\<forall>V. openin Y V \<and> f a \<in> V | |
| 213 | \<longrightarrow> (\<exists>U. openin X U \<and> a \<in> U \<and> f ` U \<subseteq> V)))" | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 214 | unfolding limitin_atin by fastforce | 
| 69874 | 215 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 216 | lemma limitin_trivial: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 217 | "\<lbrakk>trivial_limit F; y \<in> topspace X\<rbrakk> \<Longrightarrow> limitin X f y F" | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 218 | by (simp add: limitin_def) | 
| 69874 | 219 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 220 | lemma limitin_transform_eventually: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 221 | "\<lbrakk>eventually (\<lambda>x. f x = g x) F; limitin X f l F\<rbrakk> \<Longrightarrow> limitin X g l F" | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 222 | unfolding limitin_def using eventually_elim2 by fastforce | 
| 69874 | 223 | |
| 224 | lemma continuous_map_limit: | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 225 | assumes "continuous_map X Y g" and f: "limitin X f l F" | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 226 | shows "limitin Y (g \<circ> f) (g l) F" | 
| 69874 | 227 | proof - | 
| 228 | have "g l \<in> topspace Y" | |
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 229 | by (meson assms continuous_map f image_eqI in_mono limitin_def) | 
| 69874 | 230 | moreover | 
| 231 | have "\<And>U. \<lbrakk>\<forall>V. openin X V \<and> l \<in> V \<longrightarrow> (\<forall>\<^sub>F x in F. f x \<in> V); openin Y U; g l \<in> U\<rbrakk> | |
| 232 | \<Longrightarrow> \<forall>\<^sub>F x in F. g (f x) \<in> U" | |
| 233 | using assms eventually_mono | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 234 | by (fastforce simp: limitin_def dest!: openin_continuous_map_preimage) | 
| 69874 | 235 | ultimately show ?thesis | 
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 236 | using f by (fastforce simp add: limitin_def) | 
| 69874 | 237 | qed | 
| 238 | ||
| 239 | ||
| 240 | subsection\<open>Pointwise continuity in topological spaces\<close> | |
| 241 | ||
| 242 | definition topcontinuous_at where | |
| 243 | "topcontinuous_at X Y f x \<longleftrightarrow> | |
| 244 | x \<in> topspace X \<and> | |
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 245 | f \<in> topspace X \<rightarrow> topspace Y \<and> | 
| 69874 | 246 | (\<forall>V. openin Y V \<and> f x \<in> V | 
| 247 | \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>y \<in> U. f y \<in> V)))" | |
| 248 | ||
| 249 | lemma topcontinuous_at_atin: | |
| 250 | "topcontinuous_at X Y f x \<longleftrightarrow> | |
| 251 | x \<in> topspace X \<and> | |
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 252 | f \<in> topspace X \<rightarrow> topspace Y \<and> | 
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 253 | limitin Y f (f x) (atin X x)" | 
| 69874 | 254 | unfolding topcontinuous_at_def | 
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 255 | by (fastforce simp add: limitin_atin)+ | 
| 69874 | 256 | |
| 257 | lemma continuous_map_eq_topcontinuous_at: | |
| 258 | "continuous_map X Y f \<longleftrightarrow> (\<forall>x \<in> topspace X. topcontinuous_at X Y f x)" | |
| 259 | (is "?lhs = ?rhs") | |
| 260 | proof | |
| 261 | assume ?lhs | |
| 262 | then show ?rhs | |
| 263 | by (auto simp: continuous_map_def topcontinuous_at_def) | |
| 264 | next | |
| 265 | assume R: ?rhs | |
| 266 | then show ?lhs | |
| 267 | apply (auto simp: continuous_map_def topcontinuous_at_def) | |
| 78750 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 268 | by (smt (verit, del_insts) mem_Collect_eq openin_subopen openin_subset subset_eq) | 
| 69874 | 269 | qed | 
| 270 | ||
| 271 | lemma continuous_map_atin: | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 272 | "continuous_map X Y f \<longleftrightarrow> (\<forall>x \<in> topspace X. limitin Y f (f x) (atin X x))" | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 273 | by (auto simp: limitin_def topcontinuous_at_atin continuous_map_eq_topcontinuous_at) | 
| 69874 | 274 | |
| 69875 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 275 | lemma limitin_continuous_map: | 
| 
03bc14eab432
renamed the constant "limit" as it is too "generic"
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 276 | "\<lbrakk>continuous_map X Y f; a \<in> topspace X; f a = b\<rbrakk> \<Longrightarrow> limitin Y f b (atin X a)" | 
| 69874 | 277 | by (auto simp: continuous_map_atin) | 
| 278 | ||
| 78750 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 279 | lemma limit_continuous_map_within: | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 280 | "\<lbrakk>continuous_map (subtopology X S) Y f; a \<in> S; a \<in> topspace X\<rbrakk> | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 281 | \<Longrightarrow> limitin Y f (f a) (atin_within X a S)" | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 282 | by (metis Int_iff atin_subtopology_within continuous_map_atin topspace_subtopology) | 
| 
f229090cb8a3
atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
 paulson <lp15@cam.ac.uk> parents: 
78320diff
changeset | 283 | |
| 69874 | 284 | |
| 285 | subsection\<open>Combining theorems for continuous functions into the reals\<close> | |
| 286 | ||
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 287 | lemma continuous_map_canonical_const [continuous_intros]: | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 288 | "continuous_map X euclidean (\<lambda>x. c)" | 
| 69874 | 289 | by simp | 
| 290 | ||
| 291 | lemma continuous_map_real_mult [continuous_intros]: | |
| 292 | "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk> | |
| 293 | \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x * g x)" | |
| 294 | by (simp add: continuous_map_atin tendsto_mult) | |
| 295 | ||
| 296 | lemma continuous_map_real_pow [continuous_intros]: | |
| 297 | "continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x ^ n)" | |
| 298 | by (induction n) (auto simp: continuous_map_real_mult) | |
| 299 | ||
| 300 | lemma continuous_map_real_mult_left: | |
| 301 | "continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. c * f x)" | |
| 302 | by (simp add: continuous_map_atin tendsto_mult) | |
| 303 | ||
| 304 | lemma continuous_map_real_mult_left_eq: | |
| 305 | "continuous_map X euclideanreal (\<lambda>x. c * f x) \<longleftrightarrow> c = 0 \<or> continuous_map X euclideanreal f" | |
| 306 | proof (cases "c = 0") | |
| 307 | case False | |
| 308 | have "continuous_map X euclideanreal (\<lambda>x. c * f x) \<Longrightarrow> continuous_map X euclideanreal f" | |
| 309 | apply (frule continuous_map_real_mult_left [where c="inverse c"]) | |
| 310 | apply (simp add: field_simps False) | |
| 311 | done | |
| 312 | with False show ?thesis | |
| 313 | using continuous_map_real_mult_left by blast | |
| 314 | qed simp | |
| 315 | ||
| 316 | lemma continuous_map_real_mult_right: | |
| 317 | "continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x * c)" | |
| 318 | by (simp add: continuous_map_atin tendsto_mult) | |
| 319 | ||
| 320 | lemma continuous_map_real_mult_right_eq: | |
| 321 | "continuous_map X euclideanreal (\<lambda>x. f x * c) \<longleftrightarrow> c = 0 \<or> continuous_map X euclideanreal f" | |
| 322 | by (simp add: mult.commute flip: continuous_map_real_mult_left_eq) | |
| 323 | ||
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 324 | lemma continuous_map_minus [continuous_intros]: | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 325 | fixes f :: "'a\<Rightarrow>'b::real_normed_vector" | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 326 | shows "continuous_map X euclidean f \<Longrightarrow> continuous_map X euclidean (\<lambda>x. - f x)" | 
| 69874 | 327 | by (simp add: continuous_map_atin tendsto_minus) | 
| 328 | ||
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 329 | lemma continuous_map_minus_eq [simp]: | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 330 | fixes f :: "'a\<Rightarrow>'b::real_normed_vector" | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 331 | shows "continuous_map X euclidean (\<lambda>x. - f x) \<longleftrightarrow> continuous_map X euclidean f" | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 332 | using continuous_map_minus add.inverse_inverse continuous_map_eq by fastforce | 
| 69874 | 333 | |
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 334 | lemma continuous_map_add [continuous_intros]: | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 335 | fixes f :: "'a\<Rightarrow>'b::real_normed_vector" | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 336 | shows "\<lbrakk>continuous_map X euclidean f; continuous_map X euclidean g\<rbrakk> \<Longrightarrow> continuous_map X euclidean (\<lambda>x. f x + g x)" | 
| 69874 | 337 | by (simp add: continuous_map_atin tendsto_add) | 
| 338 | ||
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 339 | lemma continuous_map_diff [continuous_intros]: | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 340 | fixes f :: "'a\<Rightarrow>'b::real_normed_vector" | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 341 | shows "\<lbrakk>continuous_map X euclidean f; continuous_map X euclidean g\<rbrakk> \<Longrightarrow> continuous_map X euclidean (\<lambda>x. f x - g x)" | 
| 69874 | 342 | by (simp add: continuous_map_atin tendsto_diff) | 
| 343 | ||
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 344 | lemma continuous_map_norm [continuous_intros]: | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 345 | fixes f :: "'a\<Rightarrow>'b::real_normed_vector" | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 346 | shows "continuous_map X euclidean f \<Longrightarrow> continuous_map X euclidean (\<lambda>x. norm(f x))" | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 347 | by (simp add: continuous_map_atin tendsto_norm) | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 348 | |
| 69874 | 349 | lemma continuous_map_real_abs [continuous_intros]: | 
| 350 | "continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. abs(f x))" | |
| 351 | by (simp add: continuous_map_atin tendsto_rabs) | |
| 352 | ||
| 353 | lemma continuous_map_real_max [continuous_intros]: | |
| 354 | "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk> | |
| 355 | \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. max (f x) (g x))" | |
| 356 | by (simp add: continuous_map_atin tendsto_max) | |
| 357 | ||
| 358 | lemma continuous_map_real_min [continuous_intros]: | |
| 359 | "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk> | |
| 360 | \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. min (f x) (g x))" | |
| 361 | by (simp add: continuous_map_atin tendsto_min) | |
| 362 | ||
| 363 | lemma continuous_map_sum [continuous_intros]: | |
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 364 | fixes f :: "'a\<Rightarrow>'b\<Rightarrow>'c::real_normed_vector" | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 365 | shows "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> continuous_map X euclidean (\<lambda>x. f x i)\<rbrakk> | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69875diff
changeset | 366 | \<Longrightarrow> continuous_map X euclidean (\<lambda>x. sum (f x) I)" | 
| 69874 | 367 | by (simp add: continuous_map_atin tendsto_sum) | 
| 368 | ||
| 369 | lemma continuous_map_prod [continuous_intros]: | |
| 370 | "\<lbrakk>finite I; | |
| 371 | \<And>i. i \<in> I \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x i)\<rbrakk> | |
| 372 | \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. prod (f x) I)" | |
| 373 | by (simp add: continuous_map_atin tendsto_prod) | |
| 374 | ||
| 375 | lemma continuous_map_real_inverse [continuous_intros]: | |
| 376 | "\<lbrakk>continuous_map X euclideanreal f; \<And>x. x \<in> topspace X \<Longrightarrow> f x \<noteq> 0\<rbrakk> | |
| 377 | \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. inverse(f x))" | |
| 378 | by (simp add: continuous_map_atin tendsto_inverse) | |
| 379 | ||
| 380 | lemma continuous_map_real_divide [continuous_intros]: | |
| 381 | "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g; \<And>x. x \<in> topspace X \<Longrightarrow> g x \<noteq> 0\<rbrakk> | |
| 382 | \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x / g x)" | |
| 383 | by (simp add: continuous_map_atin tendsto_divide) | |
| 384 | ||
| 385 | end |