| author | wenzelm | 
| Thu, 15 Oct 2015 22:25:57 +0200 | |
| changeset 61456 | b521b8b400f7 | 
| parent 61426 | d53db136e8fd | 
| child 61520 | 8f85bb443d33 | 
| permissions | -rw-r--r-- | 
| 52265 | 1 | (* Title: HOL/Topological_Spaces.thy | 
| 51471 | 2 | Author: Brian Huffman | 
| 3 | Author: Johannes Hölzl | |
| 4 | *) | |
| 5 | ||
| 60758 | 6 | section \<open>Topological Spaces\<close> | 
| 51471 | 7 | |
| 8 | theory Topological_Spaces | |
| 51773 | 9 | imports Main Conditionally_Complete_Lattices | 
| 51471 | 10 | begin | 
| 11 | ||
| 57953 | 12 | named_theorems continuous_intros "structural introduction rules for continuity" | 
| 13 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 14 | |
| 60758 | 15 | subsection \<open>Topological space\<close> | 
| 51471 | 16 | |
| 17 | class "open" = | |
| 18 | fixes "open" :: "'a set \<Rightarrow> bool" | |
| 19 | ||
| 20 | class topological_space = "open" + | |
| 21 | assumes open_UNIV [simp, intro]: "open UNIV" | |
| 22 | assumes open_Int [intro]: "open S \<Longrightarrow> open T \<Longrightarrow> open (S \<inter> T)" | |
| 60585 | 23 | assumes open_Union [intro]: "\<forall>S\<in>K. open S \<Longrightarrow> open (\<Union>K)" | 
| 51471 | 24 | begin | 
| 25 | ||
| 26 | definition | |
| 27 | closed :: "'a set \<Rightarrow> bool" where | |
| 28 | "closed S \<longleftrightarrow> open (- S)" | |
| 29 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 30 | lemma open_empty [continuous_intros, intro, simp]: "open {}"
 | 
| 51471 | 31 |   using open_Union [of "{}"] by simp
 | 
| 32 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 33 | lemma open_Un [continuous_intros, intro]: "open S \<Longrightarrow> open T \<Longrightarrow> open (S \<union> T)" | 
| 51471 | 34 |   using open_Union [of "{S, T}"] by simp
 | 
| 35 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 36 | lemma open_UN [continuous_intros, intro]: "\<forall>x\<in>A. open (B x) \<Longrightarrow> open (\<Union>x\<in>A. B x)" | 
| 56166 | 37 | using open_Union [of "B ` A"] by simp | 
| 51471 | 38 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 39 | lemma open_Inter [continuous_intros, intro]: "finite S \<Longrightarrow> \<forall>T\<in>S. open T \<Longrightarrow> open (\<Inter>S)" | 
| 51471 | 40 | by (induct set: finite) auto | 
| 41 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 42 | lemma open_INT [continuous_intros, intro]: "finite A \<Longrightarrow> \<forall>x\<in>A. open (B x) \<Longrightarrow> open (\<Inter>x\<in>A. B x)" | 
| 56166 | 43 | using open_Inter [of "B ` A"] by simp | 
| 51471 | 44 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 45 | lemma openI: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 46 | assumes "\<And>x. x \<in> S \<Longrightarrow> \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 47 | shows "open S" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 48 | proof - | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 49 |   have "open (\<Union>{T. open T \<and> T \<subseteq> S})" by auto
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 50 |   moreover have "\<Union>{T. open T \<and> T \<subseteq> S} = S" by (auto dest!: assms)
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 51 | ultimately show "open S" by simp | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 52 | qed | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 53 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 54 | lemma closed_empty [continuous_intros, intro, simp]:  "closed {}"
 | 
| 51471 | 55 | unfolding closed_def by simp | 
| 56 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 57 | lemma closed_Un [continuous_intros, intro]: "closed S \<Longrightarrow> closed T \<Longrightarrow> closed (S \<union> T)" | 
| 51471 | 58 | unfolding closed_def by auto | 
| 59 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 60 | lemma closed_UNIV [continuous_intros, intro, simp]: "closed UNIV" | 
| 51471 | 61 | unfolding closed_def by simp | 
| 62 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 63 | lemma closed_Int [continuous_intros, intro]: "closed S \<Longrightarrow> closed T \<Longrightarrow> closed (S \<inter> T)" | 
| 51471 | 64 | unfolding closed_def by auto | 
| 65 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 66 | lemma closed_INT [continuous_intros, intro]: "\<forall>x\<in>A. closed (B x) \<Longrightarrow> closed (\<Inter>x\<in>A. B x)" | 
| 51471 | 67 | unfolding closed_def by auto | 
| 68 | ||
| 60585 | 69 | lemma closed_Inter [continuous_intros, intro]: "\<forall>S\<in>K. closed S \<Longrightarrow> closed (\<Inter>K)" | 
| 51471 | 70 | unfolding closed_def uminus_Inf by auto | 
| 71 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 72 | lemma closed_Union [continuous_intros, intro]: "finite S \<Longrightarrow> \<forall>T\<in>S. closed T \<Longrightarrow> closed (\<Union>S)" | 
| 51471 | 73 | by (induct set: finite) auto | 
| 74 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 75 | lemma closed_UN [continuous_intros, intro]: "finite A \<Longrightarrow> \<forall>x\<in>A. closed (B x) \<Longrightarrow> closed (\<Union>x\<in>A. B x)" | 
| 56166 | 76 | using closed_Union [of "B ` A"] by simp | 
| 51471 | 77 | |
| 78 | lemma open_closed: "open S \<longleftrightarrow> closed (- S)" | |
| 79 | unfolding closed_def by simp | |
| 80 | ||
| 81 | lemma closed_open: "closed S \<longleftrightarrow> open (- S)" | |
| 82 | unfolding closed_def by simp | |
| 83 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 84 | lemma open_Diff [continuous_intros, intro]: "open S \<Longrightarrow> closed T \<Longrightarrow> open (S - T)" | 
| 51471 | 85 | unfolding closed_open Diff_eq by (rule open_Int) | 
| 86 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 87 | lemma closed_Diff [continuous_intros, intro]: "closed S \<Longrightarrow> open T \<Longrightarrow> closed (S - T)" | 
| 51471 | 88 | unfolding open_closed Diff_eq by (rule closed_Int) | 
| 89 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 90 | lemma open_Compl [continuous_intros, intro]: "closed S \<Longrightarrow> open (- S)" | 
| 51471 | 91 | unfolding closed_open . | 
| 92 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 93 | lemma closed_Compl [continuous_intros, intro]: "open S \<Longrightarrow> closed (- S)" | 
| 51471 | 94 | unfolding open_closed . | 
| 95 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 96 | lemma open_Collect_neg: "closed {x. P x} \<Longrightarrow> open {x. \<not> P x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 97 | unfolding Collect_neg_eq by (rule open_Compl) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 98 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 99 | lemma open_Collect_conj: assumes "open {x. P x}" "open {x. Q x}" shows "open {x. P x \<and> Q x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 100 | using open_Int[OF assms] by (simp add: Int_def) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 101 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 102 | lemma open_Collect_disj: assumes "open {x. P x}" "open {x. Q x}" shows "open {x. P x \<or> Q x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 103 | using open_Un[OF assms] by (simp add: Un_def) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 104 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 105 | lemma open_Collect_ex: "(\<And>i. open {x. P i x}) \<Longrightarrow> open {x. \<exists>i. P i x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 106 |   using open_UN[of UNIV "\<lambda>i. {x. P i x}"] unfolding Collect_ex_eq by simp 
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 107 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 108 | lemma open_Collect_imp: "closed {x. P x} \<Longrightarrow> open {x. Q x} \<Longrightarrow> open {x. P x \<longrightarrow> Q x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 109 | unfolding imp_conv_disj by (intro open_Collect_disj open_Collect_neg) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 110 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 111 | lemma open_Collect_const: "open {x. P}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 112 | by (cases P) auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 113 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 114 | lemma closed_Collect_neg: "open {x. P x} \<Longrightarrow> closed {x. \<not> P x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 115 | unfolding Collect_neg_eq by (rule closed_Compl) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 116 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 117 | lemma closed_Collect_conj: assumes "closed {x. P x}" "closed {x. Q x}" shows "closed {x. P x \<and> Q x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 118 | using closed_Int[OF assms] by (simp add: Int_def) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 119 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 120 | lemma closed_Collect_disj: assumes "closed {x. P x}" "closed {x. Q x}" shows "closed {x. P x \<or> Q x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 121 | using closed_Un[OF assms] by (simp add: Un_def) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 122 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 123 | lemma closed_Collect_all: "(\<And>i. closed {x. P i x}) \<Longrightarrow> closed {x. \<forall>i. P i x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 124 |   using closed_INT[of UNIV "\<lambda>i. {x. P i x}"] unfolding Collect_all_eq by simp 
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 125 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 126 | lemma closed_Collect_imp: "open {x. P x} \<Longrightarrow> closed {x. Q x} \<Longrightarrow> closed {x. P x \<longrightarrow> Q x}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 127 | unfolding imp_conv_disj by (intro closed_Collect_disj closed_Collect_neg) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 128 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 129 | lemma closed_Collect_const: "closed {x. P}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 130 | by (cases P) auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 131 | |
| 51471 | 132 | end | 
| 133 | ||
| 60758 | 134 | subsection\<open>Hausdorff and other separation properties\<close> | 
| 51471 | 135 | |
| 136 | class t0_space = topological_space + | |
| 137 | assumes t0_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> \<not> (x \<in> U \<longleftrightarrow> y \<in> U)" | |
| 138 | ||
| 139 | class t1_space = topological_space + | |
| 140 | assumes t1_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> x \<in> U \<and> y \<notin> U" | |
| 141 | ||
| 142 | instance t1_space \<subseteq> t0_space | |
| 143 | proof qed (fast dest: t1_space) | |
| 144 | ||
| 145 | lemma separation_t1: | |
| 146 | fixes x y :: "'a::t1_space" | |
| 147 | shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> x \<in> U \<and> y \<notin> U)" | |
| 148 | using t1_space[of x y] by blast | |
| 149 | ||
| 150 | lemma closed_singleton: | |
| 151 | fixes a :: "'a::t1_space" | |
| 152 |   shows "closed {a}"
 | |
| 153 | proof - | |
| 154 |   let ?T = "\<Union>{S. open S \<and> a \<notin> S}"
 | |
| 155 | have "open ?T" by (simp add: open_Union) | |
| 156 |   also have "?T = - {a}"
 | |
| 157 | by (simp add: set_eq_iff separation_t1, auto) | |
| 158 |   finally show "closed {a}" unfolding closed_def .
 | |
| 159 | qed | |
| 160 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 161 | lemma closed_insert [continuous_intros, simp]: | 
| 51471 | 162 | fixes a :: "'a::t1_space" | 
| 163 | assumes "closed S" shows "closed (insert a S)" | |
| 164 | proof - | |
| 165 | from closed_singleton assms | |
| 166 |   have "closed ({a} \<union> S)" by (rule closed_Un)
 | |
| 167 | thus "closed (insert a S)" by simp | |
| 168 | qed | |
| 169 | ||
| 170 | lemma finite_imp_closed: | |
| 171 | fixes S :: "'a::t1_space set" | |
| 172 | shows "finite S \<Longrightarrow> closed S" | |
| 173 | by (induct set: finite, simp_all) | |
| 174 | ||
| 60758 | 175 | text \<open>T2 spaces are also known as Hausdorff spaces.\<close> | 
| 51471 | 176 | |
| 177 | class t2_space = topological_space + | |
| 178 |   assumes hausdorff: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
 | |
| 179 | ||
| 180 | instance t2_space \<subseteq> t1_space | |
| 181 | proof qed (fast dest: hausdorff) | |
| 182 | ||
| 183 | lemma separation_t2: | |
| 184 | fixes x y :: "'a::t2_space" | |
| 185 |   shows "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {})"
 | |
| 186 | using hausdorff[of x y] by blast | |
| 187 | ||
| 188 | lemma separation_t0: | |
| 189 | fixes x y :: "'a::t0_space" | |
| 190 | shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> ~(x\<in>U \<longleftrightarrow> y\<in>U))" | |
| 191 | using t0_space[of x y] by blast | |
| 192 | ||
| 60758 | 193 | text \<open>A perfect space is a topological space with no isolated points.\<close> | 
| 51471 | 194 | |
| 195 | class perfect_space = topological_space + | |
| 196 |   assumes not_open_singleton: "\<not> open {x}"
 | |
| 197 | ||
| 198 | ||
| 60758 | 199 | subsection \<open>Generators for toplogies\<close> | 
| 51471 | 200 | |
| 201 | inductive generate_topology for S where | |
| 202 | UNIV: "generate_topology S UNIV" | |
| 203 | | Int: "generate_topology S a \<Longrightarrow> generate_topology S b \<Longrightarrow> generate_topology S (a \<inter> b)" | |
| 204 | | UN: "(\<And>k. k \<in> K \<Longrightarrow> generate_topology S k) \<Longrightarrow> generate_topology S (\<Union>K)" | |
| 205 | | Basis: "s \<in> S \<Longrightarrow> generate_topology S s" | |
| 206 | ||
| 207 | hide_fact (open) UNIV Int UN Basis | |
| 208 | ||
| 209 | lemma generate_topology_Union: | |
| 210 | "(\<And>k. k \<in> I \<Longrightarrow> generate_topology S (K k)) \<Longrightarrow> generate_topology S (\<Union>k\<in>I. K k)" | |
| 56166 | 211 | using generate_topology.UN [of "K ` I"] by auto | 
| 51471 | 212 | |
| 213 | lemma topological_space_generate_topology: | |
| 214 | "class.topological_space (generate_topology S)" | |
| 61169 | 215 | by standard (auto intro: generate_topology.intros) | 
| 51471 | 216 | |
| 60758 | 217 | subsection \<open>Order topologies\<close> | 
| 51471 | 218 | |
| 219 | class order_topology = order + "open" + | |
| 220 |   assumes open_generated_order: "open = generate_topology (range (\<lambda>a. {..< a}) \<union> range (\<lambda>a. {a <..}))"
 | |
| 221 | begin | |
| 222 | ||
| 223 | subclass topological_space | |
| 224 | unfolding open_generated_order | |
| 225 | by (rule topological_space_generate_topology) | |
| 226 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 227 | lemma open_greaterThan [continuous_intros, simp]: "open {a <..}"
 | 
| 51471 | 228 | unfolding open_generated_order by (auto intro: generate_topology.Basis) | 
| 229 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 230 | lemma open_lessThan [continuous_intros, simp]: "open {..< a}"
 | 
| 51471 | 231 | unfolding open_generated_order by (auto intro: generate_topology.Basis) | 
| 232 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 233 | lemma open_greaterThanLessThan [continuous_intros, simp]: "open {a <..< b}"
 | 
| 51471 | 234 | unfolding greaterThanLessThan_eq by (simp add: open_Int) | 
| 235 | ||
| 236 | end | |
| 237 | ||
| 238 | class linorder_topology = linorder + order_topology | |
| 239 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 240 | lemma closed_atMost [continuous_intros, simp]: "closed {.. a::'a::linorder_topology}"
 | 
| 51471 | 241 | by (simp add: closed_open) | 
| 242 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 243 | lemma closed_atLeast [continuous_intros, simp]: "closed {a::'a::linorder_topology ..}"
 | 
| 51471 | 244 | by (simp add: closed_open) | 
| 245 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 246 | lemma closed_atLeastAtMost [continuous_intros, simp]: "closed {a::'a::linorder_topology .. b}"
 | 
| 51471 | 247 | proof - | 
| 248 |   have "{a .. b} = {a ..} \<inter> {.. b}"
 | |
| 249 | by auto | |
| 250 | then show ?thesis | |
| 251 | by (simp add: closed_Int) | |
| 252 | qed | |
| 253 | ||
| 254 | lemma (in linorder) less_separate: | |
| 255 | assumes "x < y" | |
| 256 |   shows "\<exists>a b. x \<in> {..< a} \<and> y \<in> {b <..} \<and> {..< a} \<inter> {b <..} = {}"
 | |
| 53381 | 257 | proof (cases "\<exists>z. x < z \<and> z < y") | 
| 258 | case True | |
| 259 | then obtain z where "x < z \<and> z < y" .. | |
| 51471 | 260 |   then have "x \<in> {..< z} \<and> y \<in> {z <..} \<and> {z <..} \<inter> {..< z} = {}"
 | 
| 261 | by auto | |
| 262 | then show ?thesis by blast | |
| 263 | next | |
| 53381 | 264 | case False | 
| 60758 | 265 |   with \<open>x < y\<close> have "x \<in> {..< y} \<and> y \<in> {x <..} \<and> {x <..} \<inter> {..< y} = {}"
 | 
| 51471 | 266 | by auto | 
| 267 | then show ?thesis by blast | |
| 268 | qed | |
| 269 | ||
| 270 | instance linorder_topology \<subseteq> t2_space | |
| 271 | proof | |
| 272 | fix x y :: 'a | |
| 273 | from less_separate[of x y] less_separate[of y x] | |
| 274 |   show "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
 | |
| 275 | by (elim neqE) (metis open_lessThan open_greaterThan Int_commute)+ | |
| 276 | qed | |
| 277 | ||
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 278 | lemma (in linorder_topology) open_right: | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 279 |   assumes "open S" "x \<in> S" and gt_ex: "x < y" shows "\<exists>b>x. {x ..< b} \<subseteq> S"
 | 
| 51471 | 280 | using assms unfolding open_generated_order | 
| 281 | proof induction | |
| 282 | case (Int A B) | |
| 283 |   then obtain a b where "a > x" "{x ..< a} \<subseteq> A"  "b > x" "{x ..< b} \<subseteq> B" by auto
 | |
| 284 | then show ?case by (auto intro!: exI[of _ "min a b"]) | |
| 285 | next | |
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 286 | case (Basis S) then show ?case by (fastforce intro: exI[of _ y] gt_ex) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 287 | qed blast+ | 
| 51471 | 288 | |
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 289 | lemma (in linorder_topology) open_left: | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 290 |   assumes "open S" "x \<in> S" and lt_ex: "y < x" shows "\<exists>b<x. {b <.. x} \<subseteq> S"
 | 
| 51471 | 291 | using assms unfolding open_generated_order | 
| 292 | proof induction | |
| 293 | case (Int A B) | |
| 294 |   then obtain a b where "a < x" "{a <.. x} \<subseteq> A"  "b < x" "{b <.. x} \<subseteq> B" by auto
 | |
| 295 | then show ?case by (auto intro!: exI[of _ "max a b"]) | |
| 296 | next | |
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 297 | case (Basis S) then show ?case by (fastforce intro: exI[of _ y] lt_ex) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 298 | qed blast+ | 
| 51471 | 299 | |
| 60758 | 300 | subsubsection \<open>Boolean is an order topology\<close> | 
| 59106 | 301 | |
| 60758 | 302 | text \<open>It also is a discrete topology, but don't have a type class for it (yet).\<close> | 
| 59106 | 303 | |
| 304 | instantiation bool :: order_topology | |
| 305 | begin | |
| 306 | ||
| 307 | definition open_bool :: "bool set \<Rightarrow> bool" where | |
| 308 |   "open_bool = generate_topology (range (\<lambda>a. {..< a}) \<union> range (\<lambda>a. {a <..}))"
 | |
| 309 | ||
| 310 | instance | |
| 311 | proof qed (rule open_bool_def) | |
| 312 | ||
| 313 | end | |
| 314 | ||
| 315 | lemma open_bool[simp, intro!]: "open (A::bool set)" | |
| 316 | proof - | |
| 317 |   have *: "{False <..} = {True}" "{..< True} = {False}"
 | |
| 318 | by auto | |
| 319 |   have "A = UNIV \<or> A = {} \<or> A = {False <..} \<or> A = {..< True}"
 | |
| 320 | using subset_UNIV[of A] unfolding UNIV_bool * by auto | |
| 321 | then show "open A" | |
| 322 | by auto | |
| 323 | qed | |
| 324 | ||
| 60758 | 325 | subsubsection \<open>Topological filters\<close> | 
| 51471 | 326 | |
| 327 | definition (in topological_space) nhds :: "'a \<Rightarrow> 'a filter" | |
| 57276 | 328 |   where "nhds a = (INF S:{S. open S \<and> a \<in> S}. principal S)"
 | 
| 51471 | 329 | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 330 | definition (in topological_space) at_within :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a filter" ("at (_) within (_)" [1000, 60] 60)
 | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 331 |   where "at a within s = inf (nhds a) (principal (s - {a}))"
 | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 332 | |
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 333 | abbreviation (in topological_space) at :: "'a \<Rightarrow> 'a filter" ("at") where
 | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 334 | "at x \<equiv> at x within (CONST UNIV)" | 
| 51471 | 335 | |
| 51473 | 336 | abbreviation (in order_topology) at_right :: "'a \<Rightarrow> 'a filter" where | 
| 51471 | 337 |   "at_right x \<equiv> at x within {x <..}"
 | 
| 338 | ||
| 51473 | 339 | abbreviation (in order_topology) at_left :: "'a \<Rightarrow> 'a filter" where | 
| 51471 | 340 |   "at_left x \<equiv> at x within {..< x}"
 | 
| 341 | ||
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 342 | lemma (in topological_space) nhds_generated_topology: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 343 |   "open = generate_topology T \<Longrightarrow> nhds x = (INF S:{S\<in>T. x \<in> S}. principal S)"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 344 | unfolding nhds_def | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 345 | proof (safe intro!: antisym INF_greatest) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 346 | fix S assume "generate_topology T S" "x \<in> S" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 347 |   then show "(INF S:{S \<in> T. x \<in> S}. principal S) \<le> principal S"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 348 | by induction | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 349 | (auto intro: INF_lower order_trans simp add: inf_principal[symmetric] simp del: inf_principal) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 350 | qed (auto intro!: INF_lower intro: generate_topology.intros) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 351 | |
| 51473 | 352 | lemma (in topological_space) eventually_nhds: | 
| 51471 | 353 | "eventually P (nhds a) \<longleftrightarrow> (\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))" | 
| 57276 | 354 | unfolding nhds_def by (subst eventually_INF_base) (auto simp: eventually_principal) | 
| 51471 | 355 | |
| 356 | lemma nhds_neq_bot [simp]: "nhds a \<noteq> bot" | |
| 357 | unfolding trivial_limit_def eventually_nhds by simp | |
| 358 | ||
| 60182 
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
 hoelzl parents: 
60172diff
changeset | 359 | lemma (in t1_space) t1_space_nhds: | 
| 
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
 hoelzl parents: 
60172diff
changeset | 360 | "x \<noteq> y \<Longrightarrow> (\<forall>\<^sub>F x in nhds x. x \<noteq> y)" | 
| 
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
 hoelzl parents: 
60172diff
changeset | 361 | by (drule t1_space) (auto simp: eventually_nhds) | 
| 
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
 hoelzl parents: 
60172diff
changeset | 362 | |
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 363 | lemma at_within_eq: "at x within s = (INF S:{S. open S \<and> x \<in> S}. principal (S \<inter> s - {x}))"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 364 | unfolding nhds_def at_within_def by (subst INF_inf_const2[symmetric]) (auto simp add: Diff_Int_distrib) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 365 | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 366 | lemma eventually_at_filter: | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 367 | "eventually P (at a within s) \<longleftrightarrow> eventually (\<lambda>x. x \<noteq> a \<longrightarrow> x \<in> s \<longrightarrow> P x) (nhds a)" | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 368 | unfolding at_within_def eventually_inf_principal by (simp add: imp_conjL[symmetric] conj_commute) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 369 | |
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 370 | lemma at_le: "s \<subseteq> t \<Longrightarrow> at x within s \<le> at x within t" | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 371 | unfolding at_within_def by (intro inf_mono) auto | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 372 | |
| 51471 | 373 | lemma eventually_at_topological: | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 374 | "eventually P (at a within s) \<longleftrightarrow> (\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. x \<noteq> a \<longrightarrow> x \<in> s \<longrightarrow> P x))" | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 375 | unfolding eventually_nhds eventually_at_filter by simp | 
| 51471 | 376 | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 377 | lemma at_within_open: "a \<in> S \<Longrightarrow> open S \<Longrightarrow> at a within S = at a" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 378 | unfolding filter_eq_iff eventually_at_topological by (metis open_Int Int_iff UNIV_I) | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 379 | |
| 61234 | 380 | lemma at_within_open_NO_MATCH: | 
| 381 | "a \<in> s \<Longrightarrow> open s \<Longrightarrow> NO_MATCH UNIV s \<Longrightarrow> at a within s = at a" | |
| 382 | by (simp only: at_within_open) | |
| 383 | ||
| 61245 | 384 | lemma at_within_nhd: | 
| 385 |   assumes "x \<in> S" "open S" "T \<inter> S - {x} = U \<inter> S - {x}"
 | |
| 386 | shows "at x within T = at x within U" | |
| 387 | unfolding filter_eq_iff eventually_at_filter | |
| 388 | proof (intro allI eventually_subst) | |
| 389 | have "eventually (\<lambda>x. x \<in> S) (nhds x)" | |
| 390 | using \<open>x \<in> S\<close> \<open>open S\<close> by (auto simp: eventually_nhds) | |
| 391 | then show "\<forall>\<^sub>F n in nhds x. (n \<noteq> x \<longrightarrow> n \<in> T \<longrightarrow> P n) = (n \<noteq> x \<longrightarrow> n \<in> U \<longrightarrow> P n)" for P | |
| 392 |     by eventually_elim (insert \<open>T \<inter> S - {x} = U \<inter> S - {x}\<close>, blast)
 | |
| 393 | qed | |
| 394 | ||
| 53859 | 395 | lemma at_within_empty [simp]: "at a within {} = bot"
 | 
| 396 | unfolding at_within_def by simp | |
| 397 | ||
| 53860 | 398 | lemma at_within_union: "at x within (S \<union> T) = sup (at x within S) (at x within T)" | 
| 399 | unfolding filter_eq_iff eventually_sup eventually_at_filter | |
| 400 | by (auto elim!: eventually_rev_mp) | |
| 401 | ||
| 51471 | 402 | lemma at_eq_bot_iff: "at a = bot \<longleftrightarrow> open {a}"
 | 
| 403 | unfolding trivial_limit_def eventually_at_topological | |
| 404 |   by (safe, case_tac "S = {a}", simp, fast, fast)
 | |
| 405 | ||
| 406 | lemma at_neq_bot [simp]: "at (a::'a::perfect_space) \<noteq> bot" | |
| 407 | by (simp add: at_eq_bot_iff not_open_singleton) | |
| 408 | ||
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 409 | lemma (in order_topology) nhds_order: "nhds x = | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 410 |   inf (INF a:{x <..}. principal {..< a}) (INF a:{..< x}. principal {a <..})"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 411 | proof - | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 412 |   have 1: "{S \<in> range lessThan \<union> range greaterThan. x \<in> S} = 
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 413 |       (\<lambda>a. {..< a}) ` {x <..} \<union> (\<lambda>a. {a <..}) ` {..< x}"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 414 | by auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 415 | show ?thesis | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 416 | unfolding nhds_generated_topology[OF open_generated_order] INF_union 1 INF_image comp_def .. | 
| 51471 | 417 | qed | 
| 418 | ||
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 419 | lemma (in linorder_topology) at_within_order: "UNIV \<noteq> {x} \<Longrightarrow> 
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 420 |   at x within s = inf (INF a:{x <..}. principal ({..< a} \<inter> s - {x}))
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 421 |                       (INF a:{..< x}. principal ({a <..} \<inter> s - {x}))"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 422 | proof (cases "{x <..} = {}" "{..< x} = {}" rule: case_split[case_product case_split])
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 423 |   assume "UNIV \<noteq> {x}" "{x<..} = {}" "{..< x} = {}"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 424 |   moreover have "UNIV = {..< x} \<union> {x} \<union> {x <..}"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 425 | by auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 426 | ultimately show ?thesis | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 427 | by auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 428 | qed (auto simp: at_within_def nhds_order Int_Diff inf_principal[symmetric] INF_inf_const2 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 429 | inf_sup_aci[where 'a="'a filter"] | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 430 | simp del: inf_principal) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 431 | |
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 432 | lemma (in linorder_topology) at_left_eq: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 433 |   "y < x \<Longrightarrow> at_left x = (INF a:{..< x}. principal {a <..< x})"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 434 | by (subst at_within_order) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 435 | (auto simp: greaterThan_Int_greaterThan greaterThanLessThan_eq[symmetric] min.absorb2 INF_constant | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 436 | intro!: INF_lower2 inf_absorb2) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 437 | |
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 438 | lemma (in linorder_topology) eventually_at_left: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 439 | "y < x \<Longrightarrow> eventually P (at_left x) \<longleftrightarrow> (\<exists>b<x. \<forall>y>b. y < x \<longrightarrow> P y)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 440 | unfolding at_left_eq by (subst eventually_INF_base) (auto simp: eventually_principal Ball_def) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 441 | |
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 442 | lemma (in linorder_topology) at_right_eq: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 443 |   "x < y \<Longrightarrow> at_right x = (INF a:{x <..}. principal {x <..< a})"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 444 | by (subst at_within_order) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 445 | (auto simp: lessThan_Int_lessThan greaterThanLessThan_eq[symmetric] max.absorb2 INF_constant Int_commute | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 446 | intro!: INF_lower2 inf_absorb1) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 447 | |
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 448 | lemma (in linorder_topology) eventually_at_right: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 449 | "x < y \<Longrightarrow> eventually P (at_right x) \<longleftrightarrow> (\<exists>b>x. \<forall>y>x. y < b \<longrightarrow> P y)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 450 | unfolding at_right_eq by (subst eventually_INF_base) (auto simp: eventually_principal Ball_def) | 
| 51471 | 451 | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 452 | lemma trivial_limit_at_right_top: "at_right (top::_::{order_top, linorder_topology}) = bot"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 453 | unfolding filter_eq_iff eventually_at_topological by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 454 | |
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 455 | lemma trivial_limit_at_left_bot: "at_left (bot::_::{order_bot, linorder_topology}) = bot"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 456 | unfolding filter_eq_iff eventually_at_topological by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 457 | |
| 51471 | 458 | lemma trivial_limit_at_left_real [simp]: | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 459 |   "\<not> trivial_limit (at_left (x::'a::{no_bot, dense_order, linorder_topology}))"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 460 | using lt_ex[of x] | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 461 | by safe (auto simp add: trivial_limit_def eventually_at_left dest: dense) | 
| 51471 | 462 | |
| 463 | lemma trivial_limit_at_right_real [simp]: | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 464 |   "\<not> trivial_limit (at_right (x::'a::{no_top, dense_order, linorder_topology}))"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 465 | using gt_ex[of x] | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 466 | by safe (auto simp add: trivial_limit_def eventually_at_right dest: dense) | 
| 51471 | 467 | |
| 468 | lemma at_eq_sup_left_right: "at (x::'a::linorder_topology) = sup (at_left x) (at_right x)" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 469 | by (auto simp: eventually_at_filter filter_eq_iff eventually_sup | 
| 51471 | 470 | elim: eventually_elim2 eventually_elim1) | 
| 471 | ||
| 472 | lemma eventually_at_split: | |
| 473 | "eventually P (at (x::'a::linorder_topology)) \<longleftrightarrow> eventually P (at_left x) \<and> eventually P (at_right x)" | |
| 474 | by (subst at_eq_sup_left_right) (simp add: eventually_sup) | |
| 475 | ||
| 60758 | 476 | subsubsection \<open>Tendsto\<close> | 
| 51471 | 477 | |
| 478 | abbreviation (in topological_space) | |
| 479 |   tendsto :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b filter \<Rightarrow> bool" (infixr "--->" 55) where
 | |
| 480 | "(f ---> l) F \<equiv> filterlim f (nhds l) F" | |
| 481 | ||
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 482 | definition (in t2_space) Lim :: "'f filter \<Rightarrow> ('f \<Rightarrow> 'a) \<Rightarrow> 'a" where
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 483 | "Lim A f = (THE l. (f ---> l) A)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 484 | |
| 51471 | 485 | lemma tendsto_eq_rhs: "(f ---> x) F \<Longrightarrow> x = y \<Longrightarrow> (f ---> y) F" | 
| 486 | by simp | |
| 487 | ||
| 57953 | 488 | named_theorems tendsto_intros "introduction rules for tendsto" | 
| 60758 | 489 | setup \<open> | 
| 51471 | 490 |   Global_Theory.add_thms_dynamic (@{binding tendsto_eq_intros},
 | 
| 57953 | 491 | fn context => | 
| 492 |       Named_Theorems.get (Context.proof_of context) @{named_theorems tendsto_intros}
 | |
| 493 |       |> map_filter (try (fn thm => @{thm tendsto_eq_rhs} OF [thm])))
 | |
| 60758 | 494 | \<close> | 
| 51471 | 495 | |
| 51473 | 496 | lemma (in topological_space) tendsto_def: | 
| 497 | "(f ---> l) F \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) F)" | |
| 57276 | 498 | unfolding nhds_def filterlim_INF filterlim_principal by auto | 
| 51471 | 499 | |
| 500 | lemma tendsto_mono: "F \<le> F' \<Longrightarrow> (f ---> l) F' \<Longrightarrow> (f ---> l) F" | |
| 501 | unfolding tendsto_def le_filter_def by fast | |
| 502 | ||
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 503 | lemma tendsto_within_subset: "(f ---> l) (at x within S) \<Longrightarrow> T \<subseteq> S \<Longrightarrow> (f ---> l) (at x within T)" | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 504 | by (blast intro: tendsto_mono at_le) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 505 | |
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 506 | lemma filterlim_at: | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 507 | "(LIM x F. f x :> at b within s) \<longleftrightarrow> (eventually (\<lambda>x. f x \<in> s \<and> f x \<noteq> b) F \<and> (f ---> b) F)" | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 508 | by (simp add: at_within_def filterlim_inf filterlim_principal conj_commute) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 509 | |
| 51473 | 510 | lemma (in topological_space) topological_tendstoI: | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 511 | "(\<And>S. open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) F) \<Longrightarrow> (f ---> l) F" | 
| 51471 | 512 | unfolding tendsto_def by auto | 
| 513 | ||
| 51473 | 514 | lemma (in topological_space) topological_tendstoD: | 
| 51471 | 515 | "(f ---> l) F \<Longrightarrow> open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) F" | 
| 516 | unfolding tendsto_def by auto | |
| 517 | ||
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 518 | lemma (in order_topology) order_tendsto_iff: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 519 | "(f ---> x) F \<longleftrightarrow> (\<forall>l<x. eventually (\<lambda>x. l < f x) F) \<and> (\<forall>u>x. eventually (\<lambda>x. f x < u) F)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 520 | unfolding nhds_order filterlim_inf filterlim_INF filterlim_principal by auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 521 | |
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 522 | lemma (in order_topology) order_tendstoI: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 523 | "(\<And>a. a < y \<Longrightarrow> eventually (\<lambda>x. a < f x) F) \<Longrightarrow> (\<And>a. y < a \<Longrightarrow> eventually (\<lambda>x. f x < a) F) \<Longrightarrow> | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 524 | (f ---> y) F" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 525 | unfolding order_tendsto_iff by auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 526 | |
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 527 | lemma (in order_topology) order_tendstoD: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 528 | assumes "(f ---> y) F" | 
| 51471 | 529 | shows "a < y \<Longrightarrow> eventually (\<lambda>x. a < f x) F" | 
| 530 | and "y < a \<Longrightarrow> eventually (\<lambda>x. f x < a) F" | |
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 531 | using assms unfolding order_tendsto_iff by auto | 
| 51471 | 532 | |
| 533 | lemma tendsto_bot [simp]: "(f ---> a) bot" | |
| 534 | unfolding tendsto_def by simp | |
| 535 | ||
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 536 | lemma (in linorder_topology) tendsto_max: | 
| 56949 | 537 | assumes X: "(X ---> x) net" | 
| 538 | assumes Y: "(Y ---> y) net" | |
| 539 | shows "((\<lambda>x. max (X x) (Y x)) ---> max x y) net" | |
| 540 | proof (rule order_tendstoI) | |
| 541 | fix a assume "a < max x y" | |
| 542 | then show "eventually (\<lambda>x. a < max (X x) (Y x)) net" | |
| 543 | using order_tendstoD(1)[OF X, of a] order_tendstoD(1)[OF Y, of a] | |
| 544 | by (auto simp: less_max_iff_disj elim: eventually_elim1) | |
| 545 | next | |
| 546 | fix a assume "max x y < a" | |
| 547 | then show "eventually (\<lambda>x. max (X x) (Y x) < a) net" | |
| 548 | using order_tendstoD(2)[OF X, of a] order_tendstoD(2)[OF Y, of a] | |
| 549 | by (auto simp: eventually_conj_iff) | |
| 550 | qed | |
| 551 | ||
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 552 | lemma (in linorder_topology) tendsto_min: | 
| 56949 | 553 | assumes X: "(X ---> x) net" | 
| 554 | assumes Y: "(Y ---> y) net" | |
| 555 | shows "((\<lambda>x. min (X x) (Y x)) ---> min x y) net" | |
| 556 | proof (rule order_tendstoI) | |
| 557 | fix a assume "a < min x y" | |
| 558 | then show "eventually (\<lambda>x. a < min (X x) (Y x)) net" | |
| 559 | using order_tendstoD(1)[OF X, of a] order_tendstoD(1)[OF Y, of a] | |
| 560 | by (auto simp: eventually_conj_iff) | |
| 561 | next | |
| 562 | fix a assume "min x y < a" | |
| 563 | then show "eventually (\<lambda>x. min (X x) (Y x) < a) net" | |
| 564 | using order_tendstoD(2)[OF X, of a] order_tendstoD(2)[OF Y, of a] | |
| 565 | by (auto simp: min_less_iff_disj elim: eventually_elim1) | |
| 566 | qed | |
| 567 | ||
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 568 | lemma tendsto_ident_at [tendsto_intros, simp, intro]: "((\<lambda>x. x) ---> a) (at a within s)" | 
| 51471 | 569 | unfolding tendsto_def eventually_at_topological by auto | 
| 570 | ||
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 571 | lemma (in topological_space) tendsto_const [tendsto_intros, simp, intro]: "((\<lambda>x. k) ---> k) F" | 
| 51471 | 572 | by (simp add: tendsto_def) | 
| 573 | ||
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 574 | lemma (in t2_space) tendsto_unique: | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 575 | assumes "F \<noteq> bot" and "(f ---> a) F" and "(f ---> b) F" | 
| 51471 | 576 | shows "a = b" | 
| 577 | proof (rule ccontr) | |
| 578 | assume "a \<noteq> b" | |
| 579 |   obtain U V where "open U" "open V" "a \<in> U" "b \<in> V" "U \<inter> V = {}"
 | |
| 60758 | 580 | using hausdorff [OF \<open>a \<noteq> b\<close>] by fast | 
| 51471 | 581 | have "eventually (\<lambda>x. f x \<in> U) F" | 
| 60758 | 582 | using \<open>(f ---> a) F\<close> \<open>open U\<close> \<open>a \<in> U\<close> by (rule topological_tendstoD) | 
| 51471 | 583 | moreover | 
| 584 | have "eventually (\<lambda>x. f x \<in> V) F" | |
| 60758 | 585 | using \<open>(f ---> b) F\<close> \<open>open V\<close> \<open>b \<in> V\<close> by (rule topological_tendstoD) | 
| 51471 | 586 | ultimately | 
| 587 | have "eventually (\<lambda>x. False) F" | |
| 588 | proof eventually_elim | |
| 589 | case (elim x) | |
| 590 | hence "f x \<in> U \<inter> V" by simp | |
| 60758 | 591 |     with \<open>U \<inter> V = {}\<close> show ?case by simp
 | 
| 51471 | 592 | qed | 
| 60758 | 593 | with \<open>\<not> trivial_limit F\<close> show "False" | 
| 51471 | 594 | by (simp add: trivial_limit_def) | 
| 595 | qed | |
| 596 | ||
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 597 | lemma (in t2_space) tendsto_const_iff: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 598 | assumes "\<not> trivial_limit F" shows "((\<lambda>x. a :: 'a) ---> b) F \<longleftrightarrow> a = b" | 
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 599 | by (auto intro!: tendsto_unique [OF assms tendsto_const]) | 
| 51471 | 600 | |
| 601 | lemma increasing_tendsto: | |
| 602 | fixes f :: "_ \<Rightarrow> 'a::order_topology" | |
| 603 | assumes bdd: "eventually (\<lambda>n. f n \<le> l) F" | |
| 604 | and en: "\<And>x. x < l \<Longrightarrow> eventually (\<lambda>n. x < f n) F" | |
| 605 | shows "(f ---> l) F" | |
| 606 | using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) | |
| 607 | ||
| 608 | lemma decreasing_tendsto: | |
| 609 | fixes f :: "_ \<Rightarrow> 'a::order_topology" | |
| 610 | assumes bdd: "eventually (\<lambda>n. l \<le> f n) F" | |
| 611 | and en: "\<And>x. l < x \<Longrightarrow> eventually (\<lambda>n. f n < x) F" | |
| 612 | shows "(f ---> l) F" | |
| 613 | using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) | |
| 614 | ||
| 615 | lemma tendsto_sandwich: | |
| 616 | fixes f g h :: "'a \<Rightarrow> 'b::order_topology" | |
| 617 | assumes ev: "eventually (\<lambda>n. f n \<le> g n) net" "eventually (\<lambda>n. g n \<le> h n) net" | |
| 618 | assumes lim: "(f ---> c) net" "(h ---> c) net" | |
| 619 | shows "(g ---> c) net" | |
| 620 | proof (rule order_tendstoI) | |
| 621 | fix a show "a < c \<Longrightarrow> eventually (\<lambda>x. a < g x) net" | |
| 622 | using order_tendstoD[OF lim(1), of a] ev by (auto elim: eventually_elim2) | |
| 623 | next | |
| 624 | fix a show "c < a \<Longrightarrow> eventually (\<lambda>x. g x < a) net" | |
| 625 | using order_tendstoD[OF lim(2), of a] ev by (auto elim: eventually_elim2) | |
| 626 | qed | |
| 627 | ||
| 628 | lemma tendsto_le: | |
| 629 | fixes f g :: "'a \<Rightarrow> 'b::linorder_topology" | |
| 630 | assumes F: "\<not> trivial_limit F" | |
| 631 | assumes x: "(f ---> x) F" and y: "(g ---> y) F" | |
| 632 | assumes ev: "eventually (\<lambda>x. g x \<le> f x) F" | |
| 633 | shows "y \<le> x" | |
| 634 | proof (rule ccontr) | |
| 635 | assume "\<not> y \<le> x" | |
| 636 |   with less_separate[of x y] obtain a b where xy: "x < a" "b < y" "{..<a} \<inter> {b<..} = {}"
 | |
| 637 | by (auto simp: not_le) | |
| 638 | then have "eventually (\<lambda>x. f x < a) F" "eventually (\<lambda>x. b < g x) F" | |
| 639 | using x y by (auto intro: order_tendstoD) | |
| 640 | with ev have "eventually (\<lambda>x. False) F" | |
| 641 | by eventually_elim (insert xy, fastforce) | |
| 642 | with F show False | |
| 643 | by (simp add: eventually_False) | |
| 644 | qed | |
| 645 | ||
| 646 | lemma tendsto_le_const: | |
| 647 | fixes f :: "'a \<Rightarrow> 'b::linorder_topology" | |
| 648 | assumes F: "\<not> trivial_limit F" | |
| 56289 | 649 | assumes x: "(f ---> x) F" and a: "eventually (\<lambda>i. a \<le> f i) F" | 
| 51471 | 650 | shows "a \<le> x" | 
| 651 | using F x tendsto_const a by (rule tendsto_le) | |
| 652 | ||
| 56289 | 653 | lemma tendsto_ge_const: | 
| 654 | fixes f :: "'a \<Rightarrow> 'b::linorder_topology" | |
| 655 | assumes F: "\<not> trivial_limit F" | |
| 656 | assumes x: "(f ---> x) F" and a: "eventually (\<lambda>i. a \<ge> f i) F" | |
| 657 | shows "a \<ge> x" | |
| 658 | by (rule tendsto_le [OF F tendsto_const x a]) | |
| 659 | ||
| 60758 | 660 | subsubsection \<open>Rules about @{const Lim}\<close>
 | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 661 | |
| 57276 | 662 | lemma tendsto_Lim: | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 663 | "\<not>(trivial_limit net) \<Longrightarrow> (f ---> l) net \<Longrightarrow> Lim net f = l" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 664 | unfolding Lim_def using tendsto_unique[of net f] by auto | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 665 | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 666 | lemma Lim_ident_at: "\<not> trivial_limit (at x within s) \<Longrightarrow> Lim (at x within s) (\<lambda>x. x) = x" | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 667 | by (rule tendsto_Lim[OF _ tendsto_ident_at]) auto | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 668 | |
| 51471 | 669 | lemma filterlim_at_bot_at_right: | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 670 | fixes f :: "'a::linorder_topology \<Rightarrow> 'b::linorder" | 
| 51471 | 671 | assumes mono: "\<And>x y. Q x \<Longrightarrow> Q y \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y" | 
| 672 | assumes bij: "\<And>x. P x \<Longrightarrow> f (g x) = x" "\<And>x. P x \<Longrightarrow> Q (g x)" | |
| 673 | assumes Q: "eventually Q (at_right a)" and bound: "\<And>b. Q b \<Longrightarrow> a < b" | |
| 674 | assumes P: "eventually P at_bot" | |
| 675 | shows "filterlim f at_bot (at_right a)" | |
| 676 | proof - | |
| 677 | from P obtain x where x: "\<And>y. y \<le> x \<Longrightarrow> P y" | |
| 678 | unfolding eventually_at_bot_linorder by auto | |
| 679 | show ?thesis | |
| 680 | proof (intro filterlim_at_bot_le[THEN iffD2] allI impI) | |
| 681 | fix z assume "z \<le> x" | |
| 682 | with x have "P z" by auto | |
| 683 | have "eventually (\<lambda>x. x \<le> g z) (at_right a)" | |
| 60758 | 684 | using bound[OF bij(2)[OF \<open>P z\<close>]] | 
| 685 | unfolding eventually_at_right[OF bound[OF bij(2)[OF \<open>P z\<close>]]] by (auto intro!: exI[of _ "g z"]) | |
| 51471 | 686 | with Q show "eventually (\<lambda>x. f x \<le> z) (at_right a)" | 
| 60758 | 687 | by eventually_elim (metis bij \<open>P z\<close> mono) | 
| 51471 | 688 | qed | 
| 689 | qed | |
| 690 | ||
| 691 | lemma filterlim_at_top_at_left: | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 692 | fixes f :: "'a::linorder_topology \<Rightarrow> 'b::linorder" | 
| 51471 | 693 | assumes mono: "\<And>x y. Q x \<Longrightarrow> Q y \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y" | 
| 694 | assumes bij: "\<And>x. P x \<Longrightarrow> f (g x) = x" "\<And>x. P x \<Longrightarrow> Q (g x)" | |
| 695 | assumes Q: "eventually Q (at_left a)" and bound: "\<And>b. Q b \<Longrightarrow> b < a" | |
| 696 | assumes P: "eventually P at_top" | |
| 697 | shows "filterlim f at_top (at_left a)" | |
| 698 | proof - | |
| 699 | from P obtain x where x: "\<And>y. x \<le> y \<Longrightarrow> P y" | |
| 700 | unfolding eventually_at_top_linorder by auto | |
| 701 | show ?thesis | |
| 702 | proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) | |
| 703 | fix z assume "x \<le> z" | |
| 704 | with x have "P z" by auto | |
| 705 | have "eventually (\<lambda>x. g z \<le> x) (at_left a)" | |
| 60758 | 706 | using bound[OF bij(2)[OF \<open>P z\<close>]] | 
| 707 | unfolding eventually_at_left[OF bound[OF bij(2)[OF \<open>P z\<close>]]] by (auto intro!: exI[of _ "g z"]) | |
| 51471 | 708 | with Q show "eventually (\<lambda>x. z \<le> f x) (at_left a)" | 
| 60758 | 709 | by eventually_elim (metis bij \<open>P z\<close> mono) | 
| 51471 | 710 | qed | 
| 711 | qed | |
| 712 | ||
| 713 | lemma filterlim_split_at: | |
| 714 | "filterlim f F (at_left x) \<Longrightarrow> filterlim f F (at_right x) \<Longrightarrow> filterlim f F (at (x::'a::linorder_topology))" | |
| 715 | by (subst at_eq_sup_left_right) (rule filterlim_sup) | |
| 716 | ||
| 717 | lemma filterlim_at_split: | |
| 718 | "filterlim f F (at (x::'a::linorder_topology)) \<longleftrightarrow> filterlim f F (at_left x) \<and> filterlim f F (at_right x)" | |
| 719 | by (subst at_eq_sup_left_right) (simp add: filterlim_def filtermap_sup) | |
| 720 | ||
| 57025 | 721 | lemma eventually_nhds_top: | 
| 722 |   fixes P :: "'a :: {order_top, linorder_topology} \<Rightarrow> bool"
 | |
| 723 | assumes "(b::'a) < top" | |
| 724 | shows "eventually P (nhds top) \<longleftrightarrow> (\<exists>b<top. (\<forall>z. b < z \<longrightarrow> P z))" | |
| 725 | unfolding eventually_nhds | |
| 726 | proof safe | |
| 727 | fix S :: "'a set" assume "open S" "top \<in> S" | |
| 60758 | 728 | note open_left[OF this \<open>b < top\<close>] | 
| 57025 | 729 | moreover assume "\<forall>s\<in>S. P s" | 
| 730 | ultimately show "\<exists>b<top. \<forall>z>b. P z" | |
| 731 | by (auto simp: subset_eq Ball_def) | |
| 732 | next | |
| 733 | fix b assume "b < top" "\<forall>z>b. P z" | |
| 734 | then show "\<exists>S. open S \<and> top \<in> S \<and> (\<forall>xa\<in>S. P xa)" | |
| 735 |     by (intro exI[of _ "{b <..}"]) auto
 | |
| 736 | qed | |
| 51471 | 737 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 738 | lemma tendsto_at_within_iff_tendsto_nhds: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 739 | "(g ---> g l) (at l within S) \<longleftrightarrow> (g ---> g l) (inf (nhds l) (principal S))" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 740 | unfolding tendsto_def eventually_at_filter eventually_inf_principal | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 741 | by (intro ext all_cong imp_cong) (auto elim!: eventually_elim1) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 742 | |
| 60758 | 743 | subsection \<open>Limits on sequences\<close> | 
| 51471 | 744 | |
| 745 | abbreviation (in topological_space) | |
| 746 | LIMSEQ :: "[nat \<Rightarrow> 'a, 'a] \<Rightarrow> bool" | |
| 747 |     ("((_)/ ----> (_))" [60, 60] 60) where
 | |
| 748 | "X ----> L \<equiv> (X ---> L) sequentially" | |
| 749 | ||
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 750 | abbreviation (in t2_space) lim :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a" where | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 751 | "lim X \<equiv> Lim sequentially X" | 
| 51471 | 752 | |
| 753 | definition (in topological_space) convergent :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" where | |
| 754 | "convergent X = (\<exists>L. X ----> L)" | |
| 755 | ||
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 756 | lemma lim_def: "lim X = (THE L. X ----> L)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 757 | unfolding Lim_def .. | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 758 | |
| 60758 | 759 | subsubsection \<open>Monotone sequences and subsequences\<close> | 
| 51471 | 760 | |
| 761 | definition | |
| 762 | monoseq :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where | |
| 60758 | 763 | --\<open>Definition of monotonicity. | 
| 51471 | 764 | The use of disjunction here complicates proofs considerably. | 
| 765 | One alternative is to add a Boolean argument to indicate the direction. | |
| 60758 | 766 | Another is to develop the notions of increasing and decreasing first.\<close> | 
| 56020 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 767 | "monoseq X = ((\<forall>m. \<forall>n\<ge>m. X m \<le> X n) \<or> (\<forall>m. \<forall>n\<ge>m. X n \<le> X m))" | 
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 768 | |
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 769 | abbreviation incseq :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where | 
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 770 | "incseq X \<equiv> mono X" | 
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 771 | |
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 772 | lemma incseq_def: "incseq X \<longleftrightarrow> (\<forall>m. \<forall>n\<ge>m. X n \<ge> X m)" | 
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 773 | unfolding mono_def .. | 
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 774 | |
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 775 | abbreviation decseq :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where | 
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 776 | "decseq X \<equiv> antimono X" | 
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 777 | |
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 778 | lemma decseq_def: "decseq X \<longleftrightarrow> (\<forall>m. \<forall>n\<ge>m. X n \<le> X m)" | 
| 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55945diff
changeset | 779 | unfolding antimono_def .. | 
| 51471 | 780 | |
| 781 | definition | |
| 782 | subseq :: "(nat \<Rightarrow> nat) \<Rightarrow> bool" where | |
| 60758 | 783 | --\<open>Definition of subsequence\<close> | 
| 51471 | 784 | "subseq f \<longleftrightarrow> (\<forall>m. \<forall>n>m. f m < f n)" | 
| 785 | ||
| 786 | lemma incseq_SucI: | |
| 787 | "(\<And>n. X n \<le> X (Suc n)) \<Longrightarrow> incseq X" | |
| 788 | using lift_Suc_mono_le[of X] | |
| 789 | by (auto simp: incseq_def) | |
| 790 | ||
| 791 | lemma incseqD: "\<And>i j. incseq f \<Longrightarrow> i \<le> j \<Longrightarrow> f i \<le> f j" | |
| 792 | by (auto simp: incseq_def) | |
| 793 | ||
| 794 | lemma incseq_SucD: "incseq A \<Longrightarrow> A i \<le> A (Suc i)" | |
| 795 | using incseqD[of A i "Suc i"] by auto | |
| 796 | ||
| 797 | lemma incseq_Suc_iff: "incseq f \<longleftrightarrow> (\<forall>n. f n \<le> f (Suc n))" | |
| 798 | by (auto intro: incseq_SucI dest: incseq_SucD) | |
| 799 | ||
| 800 | lemma incseq_const[simp, intro]: "incseq (\<lambda>x. k)" | |
| 801 | unfolding incseq_def by auto | |
| 802 | ||
| 803 | lemma decseq_SucI: | |
| 804 | "(\<And>n. X (Suc n) \<le> X n) \<Longrightarrow> decseq X" | |
| 805 | using order.lift_Suc_mono_le[OF dual_order, of X] | |
| 806 | by (auto simp: decseq_def) | |
| 807 | ||
| 808 | lemma decseqD: "\<And>i j. decseq f \<Longrightarrow> i \<le> j \<Longrightarrow> f j \<le> f i" | |
| 809 | by (auto simp: decseq_def) | |
| 810 | ||
| 811 | lemma decseq_SucD: "decseq A \<Longrightarrow> A (Suc i) \<le> A i" | |
| 812 | using decseqD[of A i "Suc i"] by auto | |
| 813 | ||
| 814 | lemma decseq_Suc_iff: "decseq f \<longleftrightarrow> (\<forall>n. f (Suc n) \<le> f n)" | |
| 815 | by (auto intro: decseq_SucI dest: decseq_SucD) | |
| 816 | ||
| 817 | lemma decseq_const[simp, intro]: "decseq (\<lambda>x. k)" | |
| 818 | unfolding decseq_def by auto | |
| 819 | ||
| 820 | lemma monoseq_iff: "monoseq X \<longleftrightarrow> incseq X \<or> decseq X" | |
| 821 | unfolding monoseq_def incseq_def decseq_def .. | |
| 822 | ||
| 823 | lemma monoseq_Suc: | |
| 824 | "monoseq X \<longleftrightarrow> (\<forall>n. X n \<le> X (Suc n)) \<or> (\<forall>n. X (Suc n) \<le> X n)" | |
| 825 | unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff .. | |
| 826 | ||
| 827 | lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X" | |
| 828 | by (simp add: monoseq_def) | |
| 829 | ||
| 830 | lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X" | |
| 831 | by (simp add: monoseq_def) | |
| 832 | ||
| 833 | lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X" | |
| 834 | by (simp add: monoseq_Suc) | |
| 835 | ||
| 836 | lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X" | |
| 837 | by (simp add: monoseq_Suc) | |
| 838 | ||
| 839 | lemma monoseq_minus: | |
| 840 | fixes a :: "nat \<Rightarrow> 'a::ordered_ab_group_add" | |
| 841 | assumes "monoseq a" | |
| 842 | shows "monoseq (\<lambda> n. - a n)" | |
| 843 | proof (cases "\<forall> m. \<forall> n \<ge> m. a m \<le> a n") | |
| 844 | case True | |
| 845 | hence "\<forall> m. \<forall> n \<ge> m. - a n \<le> - a m" by auto | |
| 846 | thus ?thesis by (rule monoI2) | |
| 847 | next | |
| 848 | case False | |
| 60758 | 849 | hence "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" using \<open>monoseq a\<close>[unfolded monoseq_def] by auto | 
| 51471 | 850 | thus ?thesis by (rule monoI1) | 
| 851 | qed | |
| 852 | ||
| 60758 | 853 | text\<open>Subsequence (alternative definition, (e.g. Hoskins)\<close> | 
| 51471 | 854 | |
| 855 | lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))" | |
| 856 | apply (simp add: subseq_def) | |
| 857 | apply (auto dest!: less_imp_Suc_add) | |
| 858 | apply (induct_tac k) | |
| 859 | apply (auto intro: less_trans) | |
| 860 | done | |
| 861 | ||
| 60758 | 862 | text\<open>for any sequence, there is a monotonic subsequence\<close> | 
| 51471 | 863 | lemma seq_monosub: | 
| 864 | fixes s :: "nat => 'a::linorder" | |
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 865 | shows "\<exists>f. subseq f \<and> monoseq (\<lambda>n. (s (f n)))" | 
| 51471 | 866 | proof cases | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 867 | assume "\<forall>n. \<exists>p>n. \<forall>m\<ge>p. s m \<le> s p" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 868 | then have "\<exists>f. \<forall>n. (\<forall>m\<ge>f n. s m \<le> s (f n)) \<and> f n < f (Suc n)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 869 | by (intro dependent_nat_choice) (auto simp: conj_commute) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 870 | then obtain f where "subseq f" and mono: "\<And>n m. f n \<le> m \<Longrightarrow> s m \<le> s (f n)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 871 | by (auto simp: subseq_Suc_iff) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 872 | moreover | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 873 | then have "incseq f" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 874 | unfolding subseq_Suc_iff incseq_Suc_iff by (auto intro: less_imp_le) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 875 | then have "monoseq (\<lambda>n. s (f n))" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 876 | by (auto simp add: incseq_def intro!: mono monoI2) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 877 | ultimately show ?thesis | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 878 | by auto | 
| 51471 | 879 | next | 
| 880 | assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. s m \<le> s p))" | |
| 881 | then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. s p < s m" by (force simp: not_le le_less) | |
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 882 | have "\<exists>f. \<forall>n. N < f n \<and> f n < f (Suc n) \<and> s (f n) \<le> s (f (Suc n))" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 883 | proof (intro dependent_nat_choice) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 884 | fix x assume "N < x" with N[of x] show "\<exists>y>N. x < y \<and> s x \<le> s y" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 885 | by (auto intro: less_trans) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 886 | qed auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 887 | then show ?thesis | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 888 | by (auto simp: monoseq_iff incseq_Suc_iff subseq_Suc_iff) | 
| 51471 | 889 | qed | 
| 890 | ||
| 891 | lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n" | |
| 892 | proof(induct n) | |
| 893 | case 0 thus ?case by simp | |
| 894 | next | |
| 895 | case (Suc n) | |
| 896 | from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps | |
| 897 | have "n < f (Suc n)" by arith | |
| 898 | thus ?case by arith | |
| 899 | qed | |
| 900 | ||
| 901 | lemma eventually_subseq: | |
| 902 | "subseq r \<Longrightarrow> eventually P sequentially \<Longrightarrow> eventually (\<lambda>n. P (r n)) sequentially" | |
| 903 | unfolding eventually_sequentially by (metis seq_suble le_trans) | |
| 904 | ||
| 51473 | 905 | lemma not_eventually_sequentiallyD: | 
| 906 | assumes P: "\<not> eventually P sequentially" | |
| 907 | shows "\<exists>r. subseq r \<and> (\<forall>n. \<not> P (r n))" | |
| 908 | proof - | |
| 909 | from P have "\<forall>n. \<exists>m\<ge>n. \<not> P m" | |
| 910 | unfolding eventually_sequentially by (simp add: not_less) | |
| 911 | then obtain r where "\<And>n. r n \<ge> n" "\<And>n. \<not> P (r n)" | |
| 912 | by (auto simp: choice_iff) | |
| 913 | then show ?thesis | |
| 914 | by (auto intro!: exI[of _ "\<lambda>n. r (((Suc \<circ> r) ^^ Suc n) 0)"] | |
| 915 | simp: less_eq_Suc_le subseq_Suc_iff) | |
| 916 | qed | |
| 917 | ||
| 51471 | 918 | lemma filterlim_subseq: "subseq f \<Longrightarrow> filterlim f sequentially sequentially" | 
| 919 | unfolding filterlim_iff by (metis eventually_subseq) | |
| 920 | ||
| 921 | lemma subseq_o: "subseq r \<Longrightarrow> subseq s \<Longrightarrow> subseq (r \<circ> s)" | |
| 922 | unfolding subseq_def by simp | |
| 923 | ||
| 924 | lemma subseq_mono: assumes "subseq r" "m < n" shows "r m < r n" | |
| 925 | using assms by (auto simp: subseq_def) | |
| 926 | ||
| 927 | lemma incseq_imp_monoseq: "incseq X \<Longrightarrow> monoseq X" | |
| 928 | by (simp add: incseq_def monoseq_def) | |
| 929 | ||
| 930 | lemma decseq_imp_monoseq: "decseq X \<Longrightarrow> monoseq X" | |
| 931 | by (simp add: decseq_def monoseq_def) | |
| 932 | ||
| 933 | lemma decseq_eq_incseq: | |
| 934 | fixes X :: "nat \<Rightarrow> 'a::ordered_ab_group_add" shows "decseq X = incseq (\<lambda>n. - X n)" | |
| 935 | by (simp add: decseq_def incseq_def) | |
| 936 | ||
| 937 | lemma INT_decseq_offset: | |
| 938 | assumes "decseq F" | |
| 939 |   shows "(\<Inter>i. F i) = (\<Inter>i\<in>{n..}. F i)"
 | |
| 940 | proof safe | |
| 941 |   fix x i assume x: "x \<in> (\<Inter>i\<in>{n..}. F i)"
 | |
| 942 | show "x \<in> F i" | |
| 943 | proof cases | |
| 944 | from x have "x \<in> F n" by auto | |
| 60758 | 945 | also assume "i \<le> n" with \<open>decseq F\<close> have "F n \<subseteq> F i" | 
| 51471 | 946 | unfolding decseq_def by simp | 
| 947 | finally show ?thesis . | |
| 948 | qed (insert x, simp) | |
| 949 | qed auto | |
| 950 | ||
| 951 | lemma LIMSEQ_const_iff: | |
| 952 | fixes k l :: "'a::t2_space" | |
| 953 | shows "(\<lambda>n. k) ----> l \<longleftrightarrow> k = l" | |
| 954 | using trivial_limit_sequentially by (rule tendsto_const_iff) | |
| 955 | ||
| 956 | lemma LIMSEQ_SUP: | |
| 957 |   "incseq X \<Longrightarrow> X ----> (SUP i. X i :: 'a :: {complete_linorder, linorder_topology})"
 | |
| 958 | by (intro increasing_tendsto) | |
| 959 | (auto simp: SUP_upper less_SUP_iff incseq_def eventually_sequentially intro: less_le_trans) | |
| 960 | ||
| 961 | lemma LIMSEQ_INF: | |
| 962 |   "decseq X \<Longrightarrow> X ----> (INF i. X i :: 'a :: {complete_linorder, linorder_topology})"
 | |
| 963 | by (intro decreasing_tendsto) | |
| 964 | (auto simp: INF_lower INF_less_iff decseq_def eventually_sequentially intro: le_less_trans) | |
| 965 | ||
| 966 | lemma LIMSEQ_ignore_initial_segment: | |
| 967 | "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a" | |
| 51474 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51473diff
changeset | 968 | unfolding tendsto_def | 
| 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51473diff
changeset | 969 | by (subst eventually_sequentially_seg[where k=k]) | 
| 51471 | 970 | |
| 971 | lemma LIMSEQ_offset: | |
| 972 | "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a" | |
| 51474 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51473diff
changeset | 973 | unfolding tendsto_def | 
| 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51473diff
changeset | 974 | by (subst (asm) eventually_sequentially_seg[where k=k]) | 
| 51471 | 975 | |
| 976 | lemma LIMSEQ_Suc: "f ----> l \<Longrightarrow> (\<lambda>n. f (Suc n)) ----> l" | |
| 977 | by (drule_tac k="Suc 0" in LIMSEQ_ignore_initial_segment, simp) | |
| 978 | ||
| 979 | lemma LIMSEQ_imp_Suc: "(\<lambda>n. f (Suc n)) ----> l \<Longrightarrow> f ----> l" | |
| 980 | by (rule_tac k="Suc 0" in LIMSEQ_offset, simp) | |
| 981 | ||
| 982 | lemma LIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) ----> l = f ----> l" | |
| 983 | by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc) | |
| 984 | ||
| 985 | lemma LIMSEQ_unique: | |
| 986 | fixes a b :: "'a::t2_space" | |
| 987 | shows "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b" | |
| 988 | using trivial_limit_sequentially by (rule tendsto_unique) | |
| 989 | ||
| 990 | lemma LIMSEQ_le_const: | |
| 991 | "\<lbrakk>X ----> (x::'a::linorder_topology); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x" | |
| 992 | using tendsto_le_const[of sequentially X x a] by (simp add: eventually_sequentially) | |
| 993 | ||
| 994 | lemma LIMSEQ_le: | |
| 995 | "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::'a::linorder_topology)" | |
| 996 | using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially) | |
| 997 | ||
| 998 | lemma LIMSEQ_le_const2: | |
| 999 | "\<lbrakk>X ----> (x::'a::linorder_topology); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a" | |
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 1000 | by (rule LIMSEQ_le[of X x "\<lambda>n. a"]) auto | 
| 51471 | 1001 | |
| 1002 | lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)" | |
| 1003 | by (simp add: convergent_def) | |
| 1004 | ||
| 1005 | lemma convergentI: "(X ----> L) ==> convergent X" | |
| 1006 | by (auto simp add: convergent_def) | |
| 1007 | ||
| 1008 | lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)" | |
| 1009 | by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) | |
| 1010 | ||
| 1011 | lemma convergent_const: "convergent (\<lambda>n. c)" | |
| 1012 | by (rule convergentI, rule tendsto_const) | |
| 1013 | ||
| 1014 | lemma monoseq_le: | |
| 1015 | "monoseq a \<Longrightarrow> a ----> (x::'a::linorder_topology) \<Longrightarrow> | |
| 1016 | ((\<forall> n. a n \<le> x) \<and> (\<forall>m. \<forall>n\<ge>m. a m \<le> a n)) \<or> ((\<forall> n. x \<le> a n) \<and> (\<forall>m. \<forall>n\<ge>m. a n \<le> a m))" | |
| 1017 | by (metis LIMSEQ_le_const LIMSEQ_le_const2 decseq_def incseq_def monoseq_iff) | |
| 1018 | ||
| 1019 | lemma LIMSEQ_subseq_LIMSEQ: | |
| 1020 | "\<lbrakk> X ----> L; subseq f \<rbrakk> \<Longrightarrow> (X o f) ----> L" | |
| 1021 | unfolding comp_def by (rule filterlim_compose[of X, OF _ filterlim_subseq]) | |
| 1022 | ||
| 1023 | lemma convergent_subseq_convergent: | |
| 1024 | "\<lbrakk>convergent X; subseq f\<rbrakk> \<Longrightarrow> convergent (X o f)" | |
| 1025 | unfolding convergent_def by (auto intro: LIMSEQ_subseq_LIMSEQ) | |
| 1026 | ||
| 1027 | lemma limI: "X ----> L ==> lim X = L" | |
| 57276 | 1028 | by (rule tendsto_Lim) (rule trivial_limit_sequentially) | 
| 51471 | 1029 | |
| 1030 | lemma lim_le: "convergent f \<Longrightarrow> (\<And>n. f n \<le> (x::'a::linorder_topology)) \<Longrightarrow> lim f \<le> x" | |
| 1031 | using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff) | |
| 1032 | ||
| 60758 | 1033 | subsubsection\<open>Increasing and Decreasing Series\<close> | 
| 51471 | 1034 | |
| 1035 | lemma incseq_le: "incseq X \<Longrightarrow> X ----> L \<Longrightarrow> X n \<le> (L::'a::linorder_topology)" | |
| 1036 | by (metis incseq_def LIMSEQ_le_const) | |
| 1037 | ||
| 1038 | lemma decseq_le: "decseq X \<Longrightarrow> X ----> L \<Longrightarrow> (L::'a::linorder_topology) \<le> X n" | |
| 1039 | by (metis decseq_def LIMSEQ_le_const2) | |
| 1040 | ||
| 60758 | 1041 | subsection \<open>First countable topologies\<close> | 
| 51473 | 1042 | |
| 1043 | class first_countable_topology = topological_space + | |
| 1044 | assumes first_countable_basis: | |
| 1045 | "\<exists>A::nat \<Rightarrow> 'a set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))" | |
| 1046 | ||
| 1047 | lemma (in first_countable_topology) countable_basis_at_decseq: | |
| 1048 | obtains A :: "nat \<Rightarrow> 'a set" where | |
| 1049 | "\<And>i. open (A i)" "\<And>i. x \<in> (A i)" | |
| 1050 | "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially" | |
| 1051 | proof atomize_elim | |
| 1052 | from first_countable_basis[of x] obtain A :: "nat \<Rightarrow> 'a set" where | |
| 1053 | nhds: "\<And>i. open (A i)" "\<And>i. x \<in> A i" | |
| 1054 | and incl: "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>i. A i \<subseteq> S" by auto | |
| 1055 | def F \<equiv> "\<lambda>n. \<Inter>i\<le>n. A i" | |
| 1056 | show "\<exists>A. (\<forall>i. open (A i)) \<and> (\<forall>i. x \<in> A i) \<and> | |
| 1057 | (\<forall>S. open S \<longrightarrow> x \<in> S \<longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially)" | |
| 1058 | proof (safe intro!: exI[of _ F]) | |
| 1059 | fix i | |
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1060 | show "open (F i)" using nhds(1) by (auto simp: F_def) | 
| 51473 | 1061 | show "x \<in> F i" using nhds(2) by (auto simp: F_def) | 
| 1062 | next | |
| 1063 | fix S assume "open S" "x \<in> S" | |
| 1064 | from incl[OF this] obtain i where "F i \<subseteq> S" unfolding F_def by auto | |
| 1065 | moreover have "\<And>j. i \<le> j \<Longrightarrow> F j \<subseteq> F i" | |
| 1066 | by (auto simp: F_def) | |
| 1067 | ultimately show "eventually (\<lambda>i. F i \<subseteq> S) sequentially" | |
| 1068 | by (auto simp: eventually_sequentially) | |
| 1069 | qed | |
| 1070 | qed | |
| 1071 | ||
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1072 | lemma (in first_countable_topology) nhds_countable: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1073 | obtains X :: "nat \<Rightarrow> 'a set" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1074 | where "decseq X" "\<And>n. open (X n)" "\<And>n. x \<in> X n" "nhds x = (INF n. principal (X n))" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1075 | proof - | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1076 | from first_countable_basis obtain A :: "nat \<Rightarrow> 'a set" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1077 | where A: "\<And>n. x \<in> A n" "\<And>n. open (A n)" "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>i. A i \<subseteq> S" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1078 | by metis | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1079 | show thesis | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1080 | proof | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1081 | show "decseq (\<lambda>n. \<Inter>i\<le>n. A i)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1082 | by (auto simp: decseq_def) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1083 | show "\<And>n. x \<in> (\<Inter>i\<le>n. A i)" "\<And>n. open (\<Inter>i\<le>n. A i)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1084 | using A by auto | 
| 60585 | 1085 | show "nhds x = (INF n. principal (\<Inter>i\<le>n. A i))" | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1086 | using A unfolding nhds_def | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1087 | apply (intro INF_eq) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1088 | apply simp_all | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1089 | apply force | 
| 60585 | 1090 | apply (intro exI[of _ "\<Inter>i\<le>n. A i" for n] conjI open_INT) | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1091 | apply auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1092 | done | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1093 | qed | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1094 | qed | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57447diff
changeset | 1095 | |
| 51473 | 1096 | lemma (in first_countable_topology) countable_basis: | 
| 1097 | obtains A :: "nat \<Rightarrow> 'a set" where | |
| 1098 | "\<And>i. open (A i)" "\<And>i. x \<in> A i" | |
| 1099 | "\<And>F. (\<forall>n. F n \<in> A n) \<Longrightarrow> F ----> x" | |
| 1100 | proof atomize_elim | |
| 53381 | 1101 | obtain A :: "nat \<Rightarrow> 'a set" where A: | 
| 1102 | "\<And>i. open (A i)" | |
| 1103 | "\<And>i. x \<in> A i" | |
| 1104 | "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially" | |
| 1105 | by (rule countable_basis_at_decseq) blast | |
| 1106 |   {
 | |
| 1107 | fix F S assume "\<forall>n. F n \<in> A n" "open S" "x \<in> S" | |
| 51473 | 1108 | with A(3)[of S] have "eventually (\<lambda>n. F n \<in> S) sequentially" | 
| 53381 | 1109 | by (auto elim: eventually_elim1 simp: subset_eq) | 
| 1110 | } | |
| 51473 | 1111 | with A show "\<exists>A. (\<forall>i. open (A i)) \<and> (\<forall>i. x \<in> A i) \<and> (\<forall>F. (\<forall>n. F n \<in> A n) \<longrightarrow> F ----> x)" | 
| 1112 | by (intro exI[of _ A]) (auto simp: tendsto_def) | |
| 1113 | qed | |
| 1114 | ||
| 1115 | lemma (in first_countable_topology) sequentially_imp_eventually_nhds_within: | |
| 1116 | assumes "\<forall>f. (\<forall>n. f n \<in> s) \<and> f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1117 | shows "eventually P (inf (nhds a) (principal s))" | 
| 51473 | 1118 | proof (rule ccontr) | 
| 53381 | 1119 | obtain A :: "nat \<Rightarrow> 'a set" where A: | 
| 1120 | "\<And>i. open (A i)" | |
| 1121 | "\<And>i. a \<in> A i" | |
| 1122 | "\<And>F. \<forall>n. F n \<in> A n \<Longrightarrow> F ----> a" | |
| 1123 | by (rule countable_basis) blast | |
| 1124 | assume "\<not> ?thesis" | |
| 51473 | 1125 | with A have P: "\<exists>F. \<forall>n. F n \<in> s \<and> F n \<in> A n \<and> \<not> P (F n)" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1126 | unfolding eventually_inf_principal eventually_nhds by (intro choice) fastforce | 
| 53381 | 1127 | then obtain F where F0: "\<forall>n. F n \<in> s" and F2: "\<forall>n. F n \<in> A n" and F3: "\<forall>n. \<not> P (F n)" | 
| 1128 | by blast | |
| 51473 | 1129 | with A have "F ----> a" by auto | 
| 1130 | hence "eventually (\<lambda>n. P (F n)) sequentially" | |
| 1131 | using assms F0 by simp | |
| 1132 | thus "False" by (simp add: F3) | |
| 1133 | qed | |
| 1134 | ||
| 1135 | lemma (in first_countable_topology) eventually_nhds_within_iff_sequentially: | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1136 | "eventually P (inf (nhds a) (principal s)) \<longleftrightarrow> | 
| 51473 | 1137 | (\<forall>f. (\<forall>n. f n \<in> s) \<and> f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially)" | 
| 1138 | proof (safe intro!: sequentially_imp_eventually_nhds_within) | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1139 | assume "eventually P (inf (nhds a) (principal s))" | 
| 51473 | 1140 | then obtain S where "open S" "a \<in> S" "\<forall>x\<in>S. x \<in> s \<longrightarrow> P x" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1141 | by (auto simp: eventually_inf_principal eventually_nhds) | 
| 51473 | 1142 | moreover fix f assume "\<forall>n. f n \<in> s" "f ----> a" | 
| 1143 | ultimately show "eventually (\<lambda>n. P (f n)) sequentially" | |
| 1144 | by (auto dest!: topological_tendstoD elim: eventually_elim1) | |
| 1145 | qed | |
| 1146 | ||
| 1147 | lemma (in first_countable_topology) eventually_nhds_iff_sequentially: | |
| 1148 | "eventually P (nhds a) \<longleftrightarrow> (\<forall>f. f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially)" | |
| 1149 | using eventually_nhds_within_iff_sequentially[of P a UNIV] by simp | |
| 1150 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1151 | lemma tendsto_at_iff_sequentially: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1152 | fixes f :: "'a :: first_countable_topology \<Rightarrow> _" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1153 |   shows "(f ---> a) (at x within s) \<longleftrightarrow> (\<forall>X. (\<forall>i. X i \<in> s - {x}) \<longrightarrow> X ----> x \<longrightarrow> ((f \<circ> X) ----> a))"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1154 | unfolding filterlim_def[of _ "nhds a"] le_filter_def eventually_filtermap at_within_def eventually_nhds_within_iff_sequentially comp_def | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1155 | by metis | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1156 | |
| 60758 | 1157 | subsection \<open>Function limit at a point\<close> | 
| 51471 | 1158 | |
| 1159 | abbreviation | |
| 1160 |   LIM :: "('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
 | |
| 1161 |         ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where
 | |
| 1162 | "f -- a --> L \<equiv> (f ---> L) (at a)" | |
| 1163 | ||
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1164 | lemma tendsto_within_open: "a \<in> S \<Longrightarrow> open S \<Longrightarrow> (f ---> l) (at a within S) \<longleftrightarrow> (f -- a --> l)" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1165 | unfolding tendsto_def by (simp add: at_within_open[where S=S]) | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1166 | |
| 51471 | 1167 | lemma LIM_const_not_eq[tendsto_intros]: | 
| 1168 | fixes a :: "'a::perfect_space" | |
| 1169 | fixes k L :: "'b::t2_space" | |
| 1170 | shows "k \<noteq> L \<Longrightarrow> \<not> (\<lambda>x. k) -- a --> L" | |
| 1171 | by (simp add: tendsto_const_iff) | |
| 1172 | ||
| 1173 | lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] | |
| 1174 | ||
| 1175 | lemma LIM_const_eq: | |
| 1176 | fixes a :: "'a::perfect_space" | |
| 1177 | fixes k L :: "'b::t2_space" | |
| 1178 | shows "(\<lambda>x. k) -- a --> L \<Longrightarrow> k = L" | |
| 1179 | by (simp add: tendsto_const_iff) | |
| 1180 | ||
| 1181 | lemma LIM_unique: | |
| 1182 | fixes a :: "'a::perfect_space" and L M :: "'b::t2_space" | |
| 1183 | shows "f -- a --> L \<Longrightarrow> f -- a --> M \<Longrightarrow> L = M" | |
| 1184 | using at_neq_bot by (rule tendsto_unique) | |
| 1185 | ||
| 60758 | 1186 | text \<open>Limits are equal for functions equal except at limit point\<close> | 
| 51471 | 1187 | |
| 1188 | lemma LIM_equal: "\<forall>x. x \<noteq> a --> (f x = g x) \<Longrightarrow> (f -- a --> l) \<longleftrightarrow> (g -- a --> l)" | |
| 1189 | unfolding tendsto_def eventually_at_topological by simp | |
| 1190 | ||
| 1191 | lemma LIM_cong: "a = b \<Longrightarrow> (\<And>x. x \<noteq> b \<Longrightarrow> f x = g x) \<Longrightarrow> l = m \<Longrightarrow> (f -- a --> l) \<longleftrightarrow> (g -- b --> m)" | |
| 1192 | by (simp add: LIM_equal) | |
| 1193 | ||
| 1194 | lemma LIM_cong_limit: "f -- x --> L \<Longrightarrow> K = L \<Longrightarrow> f -- x --> K" | |
| 1195 | by simp | |
| 1196 | ||
| 1197 | lemma tendsto_at_iff_tendsto_nhds: | |
| 1198 | "g -- l --> g l \<longleftrightarrow> (g ---> g l) (nhds l)" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1199 | unfolding tendsto_def eventually_at_filter | 
| 51471 | 1200 | by (intro ext all_cong imp_cong) (auto elim!: eventually_elim1) | 
| 1201 | ||
| 1202 | lemma tendsto_compose: | |
| 1203 | "g -- l --> g l \<Longrightarrow> (f ---> l) F \<Longrightarrow> ((\<lambda>x. g (f x)) ---> g l) F" | |
| 1204 | unfolding tendsto_at_iff_tendsto_nhds by (rule filterlim_compose[of g]) | |
| 1205 | ||
| 1206 | lemma LIM_o: "\<lbrakk>g -- l --> g l; f -- a --> l\<rbrakk> \<Longrightarrow> (g \<circ> f) -- a --> g l" | |
| 1207 | unfolding o_def by (rule tendsto_compose) | |
| 1208 | ||
| 1209 | lemma tendsto_compose_eventually: | |
| 1210 | "g -- l --> m \<Longrightarrow> (f ---> l) F \<Longrightarrow> eventually (\<lambda>x. f x \<noteq> l) F \<Longrightarrow> ((\<lambda>x. g (f x)) ---> m) F" | |
| 1211 | by (rule filterlim_compose[of g _ "at l"]) (auto simp add: filterlim_at) | |
| 1212 | ||
| 1213 | lemma LIM_compose_eventually: | |
| 1214 | assumes f: "f -- a --> b" | |
| 1215 | assumes g: "g -- b --> c" | |
| 1216 | assumes inj: "eventually (\<lambda>x. f x \<noteq> b) (at a)" | |
| 1217 | shows "(\<lambda>x. g (f x)) -- a --> c" | |
| 1218 | using g f inj by (rule tendsto_compose_eventually) | |
| 1219 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1220 | lemma tendsto_compose_filtermap: "((g \<circ> f) ---> T) F \<longleftrightarrow> (g ---> T) (filtermap f F)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1221 | by (simp add: filterlim_def filtermap_filtermap comp_def) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1222 | |
| 60758 | 1223 | subsubsection \<open>Relation of LIM and LIMSEQ\<close> | 
| 51473 | 1224 | |
| 1225 | lemma (in first_countable_topology) sequentially_imp_eventually_within: | |
| 1226 | "(\<forall>f. (\<forall>n. f n \<in> s \<and> f n \<noteq> a) \<and> f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially) \<Longrightarrow> | |
| 1227 | eventually P (at a within s)" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1228 | unfolding at_within_def | 
| 51473 | 1229 | by (intro sequentially_imp_eventually_nhds_within) auto | 
| 1230 | ||
| 1231 | lemma (in first_countable_topology) sequentially_imp_eventually_at: | |
| 1232 | "(\<forall>f. (\<forall>n. f n \<noteq> a) \<and> f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially) \<Longrightarrow> eventually P (at a)" | |
| 1233 | using assms sequentially_imp_eventually_within [where s=UNIV] by simp | |
| 1234 | ||
| 1235 | lemma LIMSEQ_SEQ_conv1: | |
| 1236 | fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" | |
| 1237 | assumes f: "f -- a --> l" | |
| 1238 | shows "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. f (S n)) ----> l" | |
| 1239 | using tendsto_compose_eventually [OF f, where F=sequentially] by simp | |
| 1240 | ||
| 1241 | lemma LIMSEQ_SEQ_conv2: | |
| 1242 | fixes f :: "'a::first_countable_topology \<Rightarrow> 'b::topological_space" | |
| 1243 | assumes "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. f (S n)) ----> l" | |
| 1244 | shows "f -- a --> l" | |
| 1245 | using assms unfolding tendsto_def [where l=l] by (simp add: sequentially_imp_eventually_at) | |
| 1246 | ||
| 1247 | lemma LIMSEQ_SEQ_conv: | |
| 1248 | "(\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> (a::'a::first_countable_topology) \<longrightarrow> (\<lambda>n. X (S n)) ----> L) = | |
| 1249 | (X -- a --> (L::'b::topological_space))" | |
| 1250 | using LIMSEQ_SEQ_conv2 LIMSEQ_SEQ_conv1 .. | |
| 1251 | ||
| 57025 | 1252 | lemma sequentially_imp_eventually_at_left: | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
60150diff
changeset | 1253 |   fixes a :: "'a :: {linorder_topology, first_countable_topology}"
 | 
| 57025 | 1254 | assumes b[simp]: "b < a" | 
| 1255 | assumes *: "\<And>f. (\<And>n. b < f n) \<Longrightarrow> (\<And>n. f n < a) \<Longrightarrow> incseq f \<Longrightarrow> f ----> a \<Longrightarrow> eventually (\<lambda>n. P (f n)) sequentially" | |
| 1256 | shows "eventually P (at_left a)" | |
| 1257 | proof (safe intro!: sequentially_imp_eventually_within) | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1258 |   fix X assume X: "\<forall>n. X n \<in> {..< a} \<and> X n \<noteq> a" "X ----> a"
 | 
| 57025 | 1259 | show "eventually (\<lambda>n. P (X n)) sequentially" | 
| 1260 | proof (rule ccontr) | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1261 | assume neg: "\<not> eventually (\<lambda>n. P (X n)) sequentially" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1262 | have "\<exists>s. \<forall>n. (\<not> P (X (s n)) \<and> b < X (s n)) \<and> (X (s n) \<le> X (s (Suc n)) \<and> Suc (s n) \<le> s (Suc n))" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1263 | proof (rule dependent_nat_choice) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1264 | have "\<not> eventually (\<lambda>n. b < X n \<longrightarrow> P (X n)) sequentially" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1265 | by (intro not_eventually_impI neg order_tendstoD(1) [OF X(2) b]) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1266 | then show "\<exists>x. \<not> P (X x) \<and> b < X x" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1267 | by (auto dest!: not_eventuallyD) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1268 | next | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1269 | fix x n | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1270 | have "\<not> eventually (\<lambda>n. Suc x \<le> n \<longrightarrow> b < X n \<longrightarrow> X x < X n \<longrightarrow> P (X n)) sequentially" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1271 | using X by (intro not_eventually_impI order_tendstoD(1)[OF X(2)] eventually_ge_at_top neg) auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1272 | then show "\<exists>n. (\<not> P (X n) \<and> b < X n) \<and> (X x \<le> X n \<and> Suc x \<le> n)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1273 | by (auto dest!: not_eventuallyD) | 
| 57025 | 1274 | qed | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1275 | then guess s .. | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1276 | then have "\<And>n. b < X (s n)" "\<And>n. X (s n) < a" "incseq (\<lambda>n. X (s n))" "(\<lambda>n. X (s n)) ----> a" "\<And>n. \<not> P (X (s n))" | 
| 60758 | 1277 | using X by (auto simp: subseq_Suc_iff Suc_le_eq incseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \<open>X ----> a\<close>, unfolded comp_def]) | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1278 | from *[OF this(1,2,3,4)] this(5) show False by auto | 
| 57025 | 1279 | qed | 
| 1280 | qed | |
| 1281 | ||
| 1282 | lemma tendsto_at_left_sequentially: | |
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
60150diff
changeset | 1283 |   fixes a :: "_ :: {linorder_topology, first_countable_topology}"
 | 
| 57025 | 1284 | assumes "b < a" | 
| 1285 | assumes *: "\<And>S. (\<And>n. S n < a) \<Longrightarrow> (\<And>n. b < S n) \<Longrightarrow> incseq S \<Longrightarrow> S ----> a \<Longrightarrow> (\<lambda>n. X (S n)) ----> L" | |
| 1286 | shows "(X ---> L) (at_left a)" | |
| 1287 | using assms unfolding tendsto_def [where l=L] | |
| 1288 | by (simp add: sequentially_imp_eventually_at_left) | |
| 1289 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1290 | lemma sequentially_imp_eventually_at_right: | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
60150diff
changeset | 1291 |   fixes a :: "'a :: {linorder_topology, first_countable_topology}"
 | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1292 | assumes b[simp]: "a < b" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1293 | assumes *: "\<And>f. (\<And>n. a < f n) \<Longrightarrow> (\<And>n. f n < b) \<Longrightarrow> decseq f \<Longrightarrow> f ----> a \<Longrightarrow> eventually (\<lambda>n. P (f n)) sequentially" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1294 | shows "eventually P (at_right a)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1295 | proof (safe intro!: sequentially_imp_eventually_within) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1296 |   fix X assume X: "\<forall>n. X n \<in> {a <..} \<and> X n \<noteq> a" "X ----> a"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1297 | show "eventually (\<lambda>n. P (X n)) sequentially" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1298 | proof (rule ccontr) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1299 | assume neg: "\<not> eventually (\<lambda>n. P (X n)) sequentially" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1300 | have "\<exists>s. \<forall>n. (\<not> P (X (s n)) \<and> X (s n) < b) \<and> (X (s (Suc n)) \<le> X (s n) \<and> Suc (s n) \<le> s (Suc n))" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1301 | proof (rule dependent_nat_choice) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1302 | have "\<not> eventually (\<lambda>n. X n < b \<longrightarrow> P (X n)) sequentially" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1303 | by (intro not_eventually_impI neg order_tendstoD(2) [OF X(2) b]) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1304 | then show "\<exists>x. \<not> P (X x) \<and> X x < b" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1305 | by (auto dest!: not_eventuallyD) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1306 | next | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1307 | fix x n | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1308 | have "\<not> eventually (\<lambda>n. Suc x \<le> n \<longrightarrow> X n < b \<longrightarrow> X n < X x \<longrightarrow> P (X n)) sequentially" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1309 | using X by (intro not_eventually_impI order_tendstoD(2)[OF X(2)] eventually_ge_at_top neg) auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1310 | then show "\<exists>n. (\<not> P (X n) \<and> X n < b) \<and> (X n \<le> X x \<and> Suc x \<le> n)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1311 | by (auto dest!: not_eventuallyD) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1312 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1313 | then guess s .. | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1314 | then have "\<And>n. a < X (s n)" "\<And>n. X (s n) < b" "decseq (\<lambda>n. X (s n))" "(\<lambda>n. X (s n)) ----> a" "\<And>n. \<not> P (X (s n))" | 
| 60758 | 1315 | using X by (auto simp: subseq_Suc_iff Suc_le_eq decseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \<open>X ----> a\<close>, unfolded comp_def]) | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1316 | from *[OF this(1,2,3,4)] this(5) show False by auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1317 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1318 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1319 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1320 | lemma tendsto_at_right_sequentially: | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
60150diff
changeset | 1321 |   fixes a :: "_ :: {linorder_topology, first_countable_topology}"
 | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1322 | assumes "a < b" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1323 | assumes *: "\<And>S. (\<And>n. a < S n) \<Longrightarrow> (\<And>n. S n < b) \<Longrightarrow> decseq S \<Longrightarrow> S ----> a \<Longrightarrow> (\<lambda>n. X (S n)) ----> L" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1324 | shows "(X ---> L) (at_right a)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1325 | using assms unfolding tendsto_def [where l=L] | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1326 | by (simp add: sequentially_imp_eventually_at_right) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1327 | |
| 60758 | 1328 | subsection \<open>Continuity\<close> | 
| 51471 | 1329 | |
| 60758 | 1330 | subsubsection \<open>Continuity on a set\<close> | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1331 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1332 | definition continuous_on :: "'a set \<Rightarrow> ('a :: topological_space \<Rightarrow> 'b :: topological_space) \<Rightarrow> bool" where
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1333 | "continuous_on s f \<longleftrightarrow> (\<forall>x\<in>s. (f ---> f x) (at x within s))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1334 | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1335 | lemma continuous_on_cong [cong]: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1336 | "s = t \<Longrightarrow> (\<And>x. x \<in> t \<Longrightarrow> f x = g x) \<Longrightarrow> continuous_on s f \<longleftrightarrow> continuous_on t g" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1337 | unfolding continuous_on_def by (intro ball_cong filterlim_cong) (auto simp: eventually_at_filter) | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1338 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1339 | lemma continuous_on_topological: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1340 | "continuous_on s f \<longleftrightarrow> | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1341 | (\<forall>x\<in>s. \<forall>B. open B \<longrightarrow> f x \<in> B \<longrightarrow> (\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> f y \<in> B)))" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1342 | unfolding continuous_on_def tendsto_def eventually_at_topological by metis | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1343 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1344 | lemma continuous_on_open_invariant: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1345 | "continuous_on s f \<longleftrightarrow> (\<forall>B. open B \<longrightarrow> (\<exists>A. open A \<and> A \<inter> s = f -` B \<inter> s))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1346 | proof safe | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1347 | fix B :: "'b set" assume "continuous_on s f" "open B" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1348 | then have "\<forall>x\<in>f -` B \<inter> s. (\<exists>A. open A \<and> x \<in> A \<and> s \<inter> A \<subseteq> f -` B)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1349 | by (auto simp: continuous_on_topological subset_eq Ball_def imp_conjL) | 
| 53381 | 1350 | then obtain A where "\<forall>x\<in>f -` B \<inter> s. open (A x) \<and> x \<in> A x \<and> s \<inter> A x \<subseteq> f -` B" | 
| 1351 | unfolding bchoice_iff .. | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1352 | then show "\<exists>A. open A \<and> A \<inter> s = f -` B \<inter> s" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1353 | by (intro exI[of _ "\<Union>x\<in>f -` B \<inter> s. A x"]) auto | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1354 | next | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1355 | assume B: "\<forall>B. open B \<longrightarrow> (\<exists>A. open A \<and> A \<inter> s = f -` B \<inter> s)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1356 | show "continuous_on s f" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1357 | unfolding continuous_on_topological | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1358 | proof safe | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1359 | fix x B assume "x \<in> s" "open B" "f x \<in> B" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1360 | with B obtain A where A: "open A" "A \<inter> s = f -` B \<inter> s" by auto | 
| 60758 | 1361 | with \<open>x \<in> s\<close> \<open>f x \<in> B\<close> show "\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> f y \<in> B)" | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1362 | by (intro exI[of _ A]) auto | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1363 | qed | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1364 | qed | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1365 | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1366 | lemma continuous_on_open_vimage: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1367 | "open s \<Longrightarrow> continuous_on s f \<longleftrightarrow> (\<forall>B. open B \<longrightarrow> open (f -` B \<inter> s))" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1368 | unfolding continuous_on_open_invariant | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1369 | by (metis open_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1370 | |
| 55734 | 1371 | corollary continuous_imp_open_vimage: | 
| 1372 | assumes "continuous_on s f" "open s" "open B" "f -` B \<subseteq> s" | |
| 1373 | shows "open (f -` B)" | |
| 1374 | by (metis assms continuous_on_open_vimage le_iff_inf) | |
| 1375 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 1376 | corollary open_vimage[continuous_intros]: | 
| 55775 | 1377 | assumes "open s" and "continuous_on UNIV f" | 
| 1378 | shows "open (f -` s)" | |
| 1379 | using assms unfolding continuous_on_open_vimage [OF open_UNIV] | |
| 1380 | by simp | |
| 1381 | ||
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1382 | lemma continuous_on_closed_invariant: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1383 | "continuous_on s f \<longleftrightarrow> (\<forall>B. closed B \<longrightarrow> (\<exists>A. closed A \<and> A \<inter> s = f -` B \<inter> s))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1384 | proof - | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1385 | have *: "\<And>P Q::'b set\<Rightarrow>bool. (\<And>A. P A \<longleftrightarrow> Q (- A)) \<Longrightarrow> (\<forall>A. P A) \<longleftrightarrow> (\<forall>A. Q A)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1386 | by (metis double_compl) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1387 | show ?thesis | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1388 | unfolding continuous_on_open_invariant by (intro *) (auto simp: open_closed[symmetric]) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1389 | qed | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1390 | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1391 | lemma continuous_on_closed_vimage: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1392 | "closed s \<Longrightarrow> continuous_on s f \<longleftrightarrow> (\<forall>B. closed B \<longrightarrow> closed (f -` B \<inter> s))" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1393 | unfolding continuous_on_closed_invariant | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1394 | by (metis closed_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1395 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61342diff
changeset | 1396 | corollary closed_vimage_Int[continuous_intros]: | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61342diff
changeset | 1397 | assumes "closed s" and "continuous_on t f" and t: "closed t" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61342diff
changeset | 1398 | shows "closed (f -` s \<inter> t)" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61342diff
changeset | 1399 | using assms unfolding continuous_on_closed_vimage [OF t] by simp | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61342diff
changeset | 1400 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 1401 | corollary closed_vimage[continuous_intros]: | 
| 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 1402 | assumes "closed s" and "continuous_on UNIV f" | 
| 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 1403 | shows "closed (f -` s)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61342diff
changeset | 1404 | using closed_vimage_Int [OF assms] by simp | 
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 1405 | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1406 | lemma continuous_on_open_Union: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1407 | "(\<And>s. s \<in> S \<Longrightarrow> open s) \<Longrightarrow> (\<And>s. s \<in> S \<Longrightarrow> continuous_on s f) \<Longrightarrow> continuous_on (\<Union>S) f" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1408 | unfolding continuous_on_def by safe (metis open_Union at_within_open UnionI) | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1409 | |
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1410 | lemma continuous_on_open_UN: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1411 | "(\<And>s. s \<in> S \<Longrightarrow> open (A s)) \<Longrightarrow> (\<And>s. s \<in> S \<Longrightarrow> continuous_on (A s) f) \<Longrightarrow> continuous_on (\<Union>s\<in>S. A s) f" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1412 | unfolding Union_image_eq[symmetric] by (rule continuous_on_open_Union) auto | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1413 | |
| 61204 | 1414 | lemma continuous_on_open_Un: | 
| 1415 | "open s \<Longrightarrow> open t \<Longrightarrow> continuous_on s f \<Longrightarrow> continuous_on t f \<Longrightarrow> continuous_on (s \<union> t) f" | |
| 1416 |   using continuous_on_open_Union [of "{s,t}"] by auto
 | |
| 1417 | ||
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1418 | lemma continuous_on_closed_Un: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1419 | "closed s \<Longrightarrow> closed t \<Longrightarrow> continuous_on s f \<Longrightarrow> continuous_on t f \<Longrightarrow> continuous_on (s \<union> t) f" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1420 | by (auto simp add: continuous_on_closed_vimage closed_Un Int_Un_distrib) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1421 | |
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1422 | lemma continuous_on_If: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1423 | assumes closed: "closed s" "closed t" and cont: "continuous_on s f" "continuous_on t g" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1424 | and P: "\<And>x. x \<in> s \<Longrightarrow> \<not> P x \<Longrightarrow> f x = g x" "\<And>x. x \<in> t \<Longrightarrow> P x \<Longrightarrow> f x = g x" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1425 | shows "continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)" (is "continuous_on _ ?h") | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1426 | proof- | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1427 | from P have "\<forall>x\<in>s. f x = ?h x" "\<forall>x\<in>t. g x = ?h x" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1428 | by auto | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1429 | with cont have "continuous_on s ?h" "continuous_on t ?h" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1430 | by simp_all | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1431 | with closed show ?thesis | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1432 | by (rule continuous_on_closed_Un) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1433 | qed | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1434 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 1435 | lemma continuous_on_id[continuous_intros]: "continuous_on s (\<lambda>x. x)" | 
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 1436 | unfolding continuous_on_def by fast | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1437 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 1438 | lemma continuous_on_const[continuous_intros]: "continuous_on s (\<lambda>x. c)" | 
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 1439 | unfolding continuous_on_def by auto | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1440 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56329diff
changeset | 1441 | lemma continuous_on_compose[continuous_intros]: | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1442 | "continuous_on s f \<Longrightarrow> continuous_on (f ` s) g \<Longrightarrow> continuous_on s (g o f)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1443 | unfolding continuous_on_topological by simp metis | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1444 | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1445 | lemma continuous_on_compose2: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1446 | "continuous_on t g \<Longrightarrow> continuous_on s f \<Longrightarrow> t = f ` s \<Longrightarrow> continuous_on s (\<lambda>x. g (f x))" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1447 | using continuous_on_compose[of s f g] by (simp add: comp_def) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1448 | |
| 60720 | 1449 | lemma continuous_on_generate_topology: | 
| 1450 | assumes *: "open = generate_topology X" | |
| 1451 | assumes **: "\<And>B. B \<in> X \<Longrightarrow> \<exists>C. open C \<and> C \<inter> A = f -` B \<inter> A" | |
| 1452 | shows "continuous_on A f" | |
| 1453 | unfolding continuous_on_open_invariant | |
| 1454 | proof safe | |
| 1455 | fix B :: "'a set" assume "open B" then show "\<exists>C. open C \<and> C \<inter> A = f -` B \<inter> A" | |
| 1456 | unfolding * | |
| 1457 | proof induction | |
| 1458 | case (UN K) | |
| 1459 | then obtain C where "\<And>k. k \<in> K \<Longrightarrow> open (C k)" "\<And>k. k \<in> K \<Longrightarrow> C k \<inter> A = f -` k \<inter> A" | |
| 1460 | by metis | |
| 1461 | then show ?case | |
| 1462 | by (intro exI[of _ "\<Union>k\<in>K. C k"]) blast | |
| 1463 | qed (auto intro: **) | |
| 1464 | qed | |
| 1465 | ||
| 1466 | lemma continuous_onI_mono: | |
| 1467 |   fixes f :: "'a::linorder_topology \<Rightarrow> 'b::{dense_order, linorder_topology}"
 | |
| 1468 | assumes "open (f`A)" | |
| 1469 | assumes mono: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y" | |
| 1470 | shows "continuous_on A f" | |
| 1471 | proof (rule continuous_on_generate_topology[OF open_generated_order], safe) | |
| 1472 | have monoD: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x < f y \<Longrightarrow> x < y" | |
| 1473 | by (auto simp: not_le[symmetric] mono) | |
| 1474 | ||
| 1475 |   { fix a b assume "a \<in> A" "f a < b"
 | |
| 1476 | moreover | |
| 1477 |     with open_right[OF \<open>open (f`A)\<close>, of "f a" b] obtain y where "f a < y" "{f a ..< y} \<subseteq> f`A"
 | |
| 1478 | by auto | |
| 1479 | moreover then obtain z where "f a < z" "z < min b y" | |
| 1480 | using dense[of "f a" "min b y"] \<open>f a < y\<close> \<open>f a < b\<close> by auto | |
| 1481 | moreover then obtain c where "z = f c" "c \<in> A" | |
| 1482 |       using \<open>{f a ..< y} \<subseteq> f`A\<close>[THEN subsetD, of z] by (auto simp: less_imp_le)
 | |
| 1483 | ultimately have "\<exists>x. x \<in> A \<and> f x < b \<and> a < x" | |
| 1484 | by (auto intro!: exI[of _ c] simp: monoD) } | |
| 1485 |   then show "\<exists>C. open C \<and> C \<inter> A = f -` {..<b} \<inter> A" for b
 | |
| 1486 |     by (intro exI[of _ "(\<Union>x\<in>{x\<in>A. f x < b}. {..< x})"])
 | |
| 1487 | (auto intro: le_less_trans[OF mono] less_imp_le) | |
| 1488 | ||
| 1489 |   { fix a b assume "a \<in> A" "b < f a"
 | |
| 1490 | moreover | |
| 1491 |     with open_left[OF \<open>open (f`A)\<close>, of "f a" b] obtain y where "y < f a" "{y <.. f a} \<subseteq> f`A"
 | |
| 1492 | by auto | |
| 1493 | moreover then obtain z where "max b y < z" "z < f a" | |
| 1494 | using dense[of "max b y" "f a"] \<open>y < f a\<close> \<open>b < f a\<close> by auto | |
| 1495 | moreover then obtain c where "z = f c" "c \<in> A" | |
| 1496 |       using \<open>{y <.. f a} \<subseteq> f`A\<close>[THEN subsetD, of z] by (auto simp: less_imp_le)
 | |
| 1497 | ultimately have "\<exists>x. x \<in> A \<and> b < f x \<and> x < a" | |
| 1498 | by (auto intro!: exI[of _ c] simp: monoD) } | |
| 1499 |   then show "\<exists>C. open C \<and> C \<inter> A = f -` {b <..} \<inter> A" for b
 | |
| 1500 |     by (intro exI[of _ "(\<Union>x\<in>{x\<in>A. b < f x}. {x <..})"])
 | |
| 1501 | (auto intro: less_le_trans[OF _ mono] less_imp_le) | |
| 1502 | qed | |
| 1503 | ||
| 60758 | 1504 | subsubsection \<open>Continuity at a point\<close> | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1505 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1506 | definition continuous :: "'a::t2_space filter \<Rightarrow> ('a \<Rightarrow> 'b::topological_space) \<Rightarrow> bool" where
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1507 | "continuous F f \<longleftrightarrow> (f ---> f (Lim F (\<lambda>x. x))) F" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1508 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1509 | lemma continuous_bot[continuous_intros, simp]: "continuous bot f" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1510 | unfolding continuous_def by auto | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1511 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1512 | lemma continuous_trivial_limit: "trivial_limit net \<Longrightarrow> continuous net f" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1513 | by simp | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1514 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1515 | lemma continuous_within: "continuous (at x within s) f \<longleftrightarrow> (f ---> f x) (at x within s)" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1516 | by (cases "trivial_limit (at x within s)") (auto simp add: Lim_ident_at continuous_def) | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1517 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1518 | lemma continuous_within_topological: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1519 | "continuous (at x within s) f \<longleftrightarrow> | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1520 | (\<forall>B. open B \<longrightarrow> f x \<in> B \<longrightarrow> (\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> f y \<in> B)))" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1521 | unfolding continuous_within tendsto_def eventually_at_topological by metis | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1522 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1523 | lemma continuous_within_compose[continuous_intros]: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1524 | "continuous (at x within s) f \<Longrightarrow> continuous (at (f x) within f ` s) g \<Longrightarrow> | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1525 | continuous (at x within s) (g o f)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1526 | by (simp add: continuous_within_topological) metis | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1527 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1528 | lemma continuous_within_compose2: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1529 | "continuous (at x within s) f \<Longrightarrow> continuous (at (f x) within f ` s) g \<Longrightarrow> | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1530 | continuous (at x within s) (\<lambda>x. g (f x))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1531 | using continuous_within_compose[of x s f g] by (simp add: comp_def) | 
| 51471 | 1532 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1533 | lemma continuous_at: "continuous (at x) f \<longleftrightarrow> f -- x --> f x" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1534 | using continuous_within[of x UNIV f] by simp | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1535 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1536 | lemma continuous_ident[continuous_intros, simp]: "continuous (at x within S) (\<lambda>x. x)" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1537 | unfolding continuous_within by (rule tendsto_ident_at) | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1538 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1539 | lemma continuous_const[continuous_intros, simp]: "continuous F (\<lambda>x. c)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1540 | unfolding continuous_def by (rule tendsto_const) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1541 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1542 | lemma continuous_on_eq_continuous_within: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1543 | "continuous_on s f \<longleftrightarrow> (\<forall>x\<in>s. continuous (at x within s) f)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1544 | unfolding continuous_on_def continuous_within .. | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1545 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1546 | abbreviation isCont :: "('a::t2_space \<Rightarrow> 'b::topological_space) \<Rightarrow> 'a \<Rightarrow> bool" where
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1547 | "isCont f a \<equiv> continuous (at a) f" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1548 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1549 | lemma isCont_def: "isCont f a \<longleftrightarrow> f -- a --> f a" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1550 | by (rule continuous_at) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1551 | |
| 60762 | 1552 | lemma continuous_at_imp_continuous_at_within: "isCont f x \<Longrightarrow> continuous (at x within s) f" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1553 | by (auto intro: tendsto_mono at_le simp: continuous_at continuous_within) | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1554 | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1555 | lemma continuous_on_eq_continuous_at: "open s \<Longrightarrow> continuous_on s f \<longleftrightarrow> (\<forall>x\<in>s. isCont f x)" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51518diff
changeset | 1556 | by (simp add: continuous_on_def continuous_at at_within_open[of _ s]) | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1557 | |
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1558 | lemma continuous_on_subset: "continuous_on s f \<Longrightarrow> t \<subseteq> s \<Longrightarrow> continuous_on t f" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1559 | unfolding continuous_on_def by (metis subset_eq tendsto_within_subset) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1560 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1561 | lemma continuous_at_imp_continuous_on: "\<forall>x\<in>s. isCont f x \<Longrightarrow> continuous_on s f" | 
| 60762 | 1562 | by (auto intro: continuous_at_imp_continuous_at_within simp: continuous_on_eq_continuous_within) | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1563 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1564 | lemma isCont_o2: "isCont f a \<Longrightarrow> isCont g (f a) \<Longrightarrow> isCont (\<lambda>x. g (f x)) a" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1565 | unfolding isCont_def by (rule tendsto_compose) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1566 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1567 | lemma isCont_o[continuous_intros]: "isCont f a \<Longrightarrow> isCont g (f a) \<Longrightarrow> isCont (g \<circ> f) a" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1568 | unfolding o_def by (rule isCont_o2) | 
| 51471 | 1569 | |
| 1570 | lemma isCont_tendsto_compose: "isCont g l \<Longrightarrow> (f ---> l) F \<Longrightarrow> ((\<lambda>x. g (f x)) ---> g l) F" | |
| 1571 | unfolding isCont_def by (rule tendsto_compose) | |
| 1572 | ||
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1573 | lemma continuous_within_compose3: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1574 | "isCont g (f x) \<Longrightarrow> continuous (at x within s) f \<Longrightarrow> continuous (at x within s) (\<lambda>x. g (f x))" | 
| 60762 | 1575 | using continuous_within_compose2[of x s f g] by (simp add: continuous_at_imp_continuous_at_within) | 
| 51471 | 1576 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1577 | lemma filtermap_nhds_open_map: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1578 | assumes cont: "isCont f a" and open_map: "\<And>S. open S \<Longrightarrow> open (f`S)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1579 | shows "filtermap f (nhds a) = nhds (f a)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1580 | unfolding filter_eq_iff | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1581 | proof safe | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1582 | fix P assume "eventually P (filtermap f (nhds a))" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1583 | then guess S unfolding eventually_filtermap eventually_nhds .. | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1584 | then show "eventually P (nhds (f a))" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1585 | unfolding eventually_nhds by (intro exI[of _ "f`S"]) (auto intro!: open_map) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1586 | qed (metis filterlim_iff tendsto_at_iff_tendsto_nhds isCont_def eventually_filtermap cont) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1587 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1588 | lemma continuous_at_split: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1589 | "continuous (at (x::'a::linorder_topology)) f = (continuous (at_left x) f \<and> continuous (at_right x) f)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1590 | by (simp add: continuous_within filterlim_at_split) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 1591 | |
| 61245 | 1592 | subsubsection \<open>Open-cover compactness\<close> | 
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1593 | |
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1594 | context topological_space | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1595 | begin | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1596 | |
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1597 | definition compact :: "'a set \<Rightarrow> bool" where | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1598 | compact_eq_heine_borel: -- "This name is used for backwards compatibility" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1599 | "compact S \<longleftrightarrow> (\<forall>C. (\<forall>c\<in>C. open c) \<and> S \<subseteq> \<Union>C \<longrightarrow> (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1600 | |
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1601 | lemma compactI: | 
| 60585 | 1602 | assumes "\<And>C. \<forall>t\<in>C. open t \<Longrightarrow> s \<subseteq> \<Union>C \<Longrightarrow> \<exists>C'. C' \<subseteq> C \<and> finite C' \<and> s \<subseteq> \<Union>C'" | 
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1603 | shows "compact s" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1604 | unfolding compact_eq_heine_borel using assms by metis | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1605 | |
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1606 | lemma compact_empty[simp]: "compact {}"
 | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1607 | by (auto intro!: compactI) | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1608 | |
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1609 | lemma compactE: | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1610 | assumes "compact s" and "\<forall>t\<in>C. open t" and "s \<subseteq> \<Union>C" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1611 | obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> \<Union>C'" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1612 | using assms unfolding compact_eq_heine_borel by metis | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1613 | |
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1614 | lemma compactE_image: | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1615 | assumes "compact s" and "\<forall>t\<in>C. open (f t)" and "s \<subseteq> (\<Union>c\<in>C. f c)" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1616 | obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> (\<Union>c\<in>C'. f c)" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1617 | using assms unfolding ball_simps[symmetric] SUP_def | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1618 | by (metis (lifting) finite_subset_image compact_eq_heine_borel[of s]) | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1619 | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1620 | lemma compact_inter_closed [intro]: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1621 | assumes "compact s" and "closed t" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1622 | shows "compact (s \<inter> t)" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1623 | proof (rule compactI) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1624 | fix C assume C: "\<forall>c\<in>C. open c" and cover: "s \<inter> t \<subseteq> \<Union>C" | 
| 60758 | 1625 |   from C \<open>closed t\<close> have "\<forall>c\<in>C \<union> {-t}. open c" by auto
 | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1626 |   moreover from cover have "s \<subseteq> \<Union>(C \<union> {-t})" by auto
 | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1627 |   ultimately have "\<exists>D\<subseteq>C \<union> {-t}. finite D \<and> s \<subseteq> \<Union>D"
 | 
| 60758 | 1628 | using \<open>compact s\<close> unfolding compact_eq_heine_borel by auto | 
| 53381 | 1629 |   then obtain D where "D \<subseteq> C \<union> {- t} \<and> finite D \<and> s \<subseteq> \<Union>D" ..
 | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1630 | then show "\<exists>D\<subseteq>C. finite D \<and> s \<inter> t \<subseteq> \<Union>D" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1631 |     by (intro exI[of _ "D - {-t}"]) auto
 | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1632 | qed | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1633 | |
| 54797 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1634 | lemma inj_setminus: "inj_on uminus (A::'a set set)" | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1635 | by (auto simp: inj_on_def) | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1636 | |
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1637 | lemma compact_fip: | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1638 | "compact U \<longleftrightarrow> | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1639 |     (\<forall>A. (\<forall>a\<in>A. closed a) \<longrightarrow> (\<forall>B \<subseteq> A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}) \<longrightarrow> U \<inter> \<Inter>A \<noteq> {})"
 | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1640 | (is "_ \<longleftrightarrow> ?R") | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1641 | proof (safe intro!: compact_eq_heine_borel[THEN iffD2]) | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1642 | fix A | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1643 | assume "compact U" | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1644 |     and A: "\<forall>a\<in>A. closed a" "U \<inter> \<Inter>A = {}"
 | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1645 |     and fi: "\<forall>B \<subseteq> A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}"
 | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1646 | from A have "(\<forall>a\<in>uminus`A. open a) \<and> U \<subseteq> \<Union>(uminus`A)" | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1647 | by auto | 
| 60758 | 1648 | with \<open>compact U\<close> obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<subseteq> \<Union>(uminus`B)" | 
| 54797 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1649 | unfolding compact_eq_heine_borel by (metis subset_image_iff) | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1650 | with fi[THEN spec, of B] show False | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1651 | by (auto dest: finite_imageD intro: inj_setminus) | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1652 | next | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1653 | fix A | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1654 | assume ?R | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1655 | assume "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A" | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1656 |   then have "U \<inter> \<Inter>(uminus`A) = {}" "\<forall>a\<in>uminus`A. closed a"
 | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1657 | by auto | 
| 60758 | 1658 |   with \<open>?R\<close> obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<inter> \<Inter>(uminus`B) = {}"
 | 
| 54797 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1659 | by (metis subset_image_iff) | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1660 | then show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T" | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1661 | by (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD) | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1662 | qed | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1663 | |
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1664 | lemma compact_imp_fip: | 
| 60585 | 1665 |   "compact s \<Longrightarrow> \<forall>t \<in> f. closed t \<Longrightarrow> \<forall>f'. finite f' \<and> f' \<subseteq> f \<longrightarrow> (s \<inter> (\<Inter>f') \<noteq> {}) \<Longrightarrow>
 | 
| 1666 |     s \<inter> (\<Inter>f) \<noteq> {}"
 | |
| 54797 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1667 | unfolding compact_fip by auto | 
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1668 | |
| 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1669 | lemma compact_imp_fip_image: | 
| 56166 | 1670 | assumes "compact s" | 
| 1671 | and P: "\<And>i. i \<in> I \<Longrightarrow> closed (f i)" | |
| 1672 |     and Q: "\<And>I'. finite I' \<Longrightarrow> I' \<subseteq> I \<Longrightarrow> (s \<inter> (\<Inter>i\<in>I'. f i) \<noteq> {})"
 | |
| 1673 |   shows "s \<inter> (\<Inter>i\<in>I. f i) \<noteq> {}"
 | |
| 1674 | proof - | |
| 60758 | 1675 | note \<open>compact s\<close> | 
| 56166 | 1676 | moreover from P have "\<forall>i \<in> f ` I. closed i" by blast | 
| 1677 |   moreover have "\<forall>A. finite A \<and> A \<subseteq> f ` I \<longrightarrow> (s \<inter> (\<Inter>A) \<noteq> {})"
 | |
| 1678 | proof (rule, rule, erule conjE) | |
| 1679 | fix A :: "'a set set" | |
| 1680 | assume "finite A" | |
| 1681 | moreover assume "A \<subseteq> f ` I" | |
| 1682 | ultimately obtain B where "B \<subseteq> I" and "finite B" and "A = f ` B" | |
| 1683 | using finite_subset_image [of A f I] by blast | |
| 1684 |     with Q [of B] show "s \<inter> \<Inter>A \<noteq> {}" by simp
 | |
| 1685 | qed | |
| 1686 |   ultimately have "s \<inter> (\<Inter>(f ` I)) \<noteq> {}" by (rule compact_imp_fip)
 | |
| 1687 | then show ?thesis by simp | |
| 1688 | qed | |
| 54797 
be020ec8560c
modernized ContNotDenum: use Set_Interval, and finite intersection property to show the nested interval property
 hoelzl parents: 
54258diff
changeset | 1689 | |
| 51471 | 1690 | end | 
| 1691 | ||
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1692 | lemma (in t2_space) compact_imp_closed: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1693 | assumes "compact s" shows "closed s" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1694 | unfolding closed_def | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1695 | proof (rule openI) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1696 | fix y assume "y \<in> - s" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1697 |   let ?C = "\<Union>x\<in>s. {u. open u \<and> x \<in> u \<and> eventually (\<lambda>y. y \<notin> u) (nhds y)}"
 | 
| 60758 | 1698 | note \<open>compact s\<close> | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1699 | moreover have "\<forall>u\<in>?C. open u" by simp | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1700 | moreover have "s \<subseteq> \<Union>?C" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1701 | proof | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1702 | fix x assume "x \<in> s" | 
| 60758 | 1703 | with \<open>y \<in> - s\<close> have "x \<noteq> y" by clarsimp | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1704 |     hence "\<exists>u v. open u \<and> open v \<and> x \<in> u \<and> y \<in> v \<and> u \<inter> v = {}"
 | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1705 | by (rule hausdorff) | 
| 60758 | 1706 | with \<open>x \<in> s\<close> show "x \<in> \<Union>?C" | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1707 | unfolding eventually_nhds by auto | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1708 | qed | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1709 | ultimately obtain D where "D \<subseteq> ?C" and "finite D" and "s \<subseteq> \<Union>D" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1710 | by (rule compactE) | 
| 60758 | 1711 | from \<open>D \<subseteq> ?C\<close> have "\<forall>x\<in>D. eventually (\<lambda>y. y \<notin> x) (nhds y)" by auto | 
| 1712 | with \<open>finite D\<close> have "eventually (\<lambda>y. y \<notin> \<Union>D) (nhds y)" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
60036diff
changeset | 1713 | by (simp add: eventually_ball_finite) | 
| 60758 | 1714 | with \<open>s \<subseteq> \<Union>D\<close> have "eventually (\<lambda>y. y \<notin> s) (nhds y)" | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1715 | by (auto elim!: eventually_mono [rotated]) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1716 | thus "\<exists>t. open t \<and> y \<in> t \<and> t \<subseteq> - s" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1717 | by (simp add: eventually_nhds subset_eq) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1718 | qed | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1719 | |
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1720 | lemma compact_continuous_image: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1721 | assumes f: "continuous_on s f" and s: "compact s" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1722 | shows "compact (f ` s)" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1723 | proof (rule compactI) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1724 | fix C assume "\<forall>c\<in>C. open c" and cover: "f`s \<subseteq> \<Union>C" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1725 | with f have "\<forall>c\<in>C. \<exists>A. open A \<and> A \<inter> s = f -` c \<inter> s" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1726 | unfolding continuous_on_open_invariant by blast | 
| 53381 | 1727 | then obtain A where A: "\<forall>c\<in>C. open (A c) \<and> A c \<inter> s = f -` c \<inter> s" | 
| 1728 | unfolding bchoice_iff .. | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1729 | with cover have "\<forall>c\<in>C. open (A c)" "s \<subseteq> (\<Union>c\<in>C. A c)" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1730 | by (fastforce simp add: subset_eq set_eq_iff)+ | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1731 | from compactE_image[OF s this] obtain D where "D \<subseteq> C" "finite D" "s \<subseteq> (\<Union>c\<in>D. A c)" . | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1732 | with A show "\<exists>D \<subseteq> C. finite D \<and> f`s \<subseteq> \<Union>D" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1733 | by (intro exI[of _ D]) (fastforce simp add: subset_eq set_eq_iff)+ | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1734 | qed | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1735 | |
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1736 | lemma continuous_on_inv: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1737 | fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1738 | assumes "continuous_on s f" "compact s" "\<forall>x\<in>s. g (f x) = x" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1739 | shows "continuous_on (f ` s) g" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1740 | unfolding continuous_on_topological | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1741 | proof (clarsimp simp add: assms(3)) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1742 | fix x :: 'a and B :: "'a set" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1743 | assume "x \<in> s" and "open B" and "x \<in> B" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1744 | have 1: "\<forall>x\<in>s. f x \<in> f ` (s - B) \<longleftrightarrow> x \<in> s - B" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1745 | using assms(3) by (auto, metis) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1746 | have "continuous_on (s - B) f" | 
| 60758 | 1747 | using \<open>continuous_on s f\<close> Diff_subset | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1748 | by (rule continuous_on_subset) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1749 | moreover have "compact (s - B)" | 
| 60758 | 1750 | using \<open>open B\<close> and \<open>compact s\<close> | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1751 | unfolding Diff_eq by (intro compact_inter_closed closed_Compl) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1752 | ultimately have "compact (f ` (s - B))" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1753 | by (rule compact_continuous_image) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1754 | hence "closed (f ` (s - B))" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1755 | by (rule compact_imp_closed) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1756 | hence "open (- f ` (s - B))" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1757 | by (rule open_Compl) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1758 | moreover have "f x \<in> - f ` (s - B)" | 
| 60758 | 1759 | using \<open>x \<in> s\<close> and \<open>x \<in> B\<close> by (simp add: 1) | 
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1760 | moreover have "\<forall>y\<in>s. f y \<in> - f ` (s - B) \<longrightarrow> y \<in> B" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1761 | by (simp add: 1) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1762 | ultimately show "\<exists>A. open A \<and> f x \<in> A \<and> (\<forall>y\<in>s. f y \<in> A \<longrightarrow> y \<in> B)" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1763 | by fast | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1764 | qed | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1765 | |
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1766 | lemma continuous_on_inv_into: | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1767 | fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1768 | assumes s: "continuous_on s f" "compact s" and f: "inj_on f s" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1769 | shows "continuous_on (f ` s) (the_inv_into s f)" | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1770 | by (rule continuous_on_inv[OF s]) (auto simp: the_inv_into_f_f[OF f]) | 
| 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1771 | |
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1772 | lemma (in linorder_topology) compact_attains_sup: | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1773 |   assumes "compact S" "S \<noteq> {}"
 | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1774 | shows "\<exists>s\<in>S. \<forall>t\<in>S. t \<le> s" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1775 | proof (rule classical) | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1776 | assume "\<not> (\<exists>s\<in>S. \<forall>t\<in>S. t \<le> s)" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1777 | then obtain t where t: "\<forall>s\<in>S. t s \<in> S" and "\<forall>s\<in>S. s < t s" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1778 | by (metis not_le) | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1779 |   then have "\<forall>s\<in>S. open {..< t s}" "S \<subseteq> (\<Union>s\<in>S. {..< t s})"
 | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1780 | by auto | 
| 60758 | 1781 |   with \<open>compact S\<close> obtain C where "C \<subseteq> S" "finite C" and C: "S \<subseteq> (\<Union>s\<in>C. {..< t s})"
 | 
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1782 | by (erule compactE_image) | 
| 60758 | 1783 |   with \<open>S \<noteq> {}\<close> have Max: "Max (t`C) \<in> t`C" and "\<forall>s\<in>t`C. s \<le> Max (t`C)"
 | 
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1784 | by (auto intro!: Max_in) | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1785 |   with C have "S \<subseteq> {..< Max (t`C)}"
 | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1786 | by (auto intro: less_le_trans simp: subset_eq) | 
| 60758 | 1787 | with t Max \<open>C \<subseteq> S\<close> show ?thesis | 
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1788 | by fastforce | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1789 | qed | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1790 | |
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1791 | lemma (in linorder_topology) compact_attains_inf: | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1792 |   assumes "compact S" "S \<noteq> {}"
 | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1793 | shows "\<exists>s\<in>S. \<forall>t\<in>S. s \<le> t" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1794 | proof (rule classical) | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1795 | assume "\<not> (\<exists>s\<in>S. \<forall>t\<in>S. s \<le> t)" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1796 | then obtain t where t: "\<forall>s\<in>S. t s \<in> S" and "\<forall>s\<in>S. t s < s" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1797 | by (metis not_le) | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1798 |   then have "\<forall>s\<in>S. open {t s <..}" "S \<subseteq> (\<Union>s\<in>S. {t s <..})"
 | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1799 | by auto | 
| 60758 | 1800 |   with \<open>compact S\<close> obtain C where "C \<subseteq> S" "finite C" and C: "S \<subseteq> (\<Union>s\<in>C. {t s <..})"
 | 
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1801 | by (erule compactE_image) | 
| 60758 | 1802 |   with \<open>S \<noteq> {}\<close> have Min: "Min (t`C) \<in> t`C" and "\<forall>s\<in>t`C. Min (t`C) \<le> s"
 | 
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1803 | by (auto intro!: Min_in) | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1804 |   with C have "S \<subseteq> {Min (t`C) <..}"
 | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1805 | by (auto intro: le_less_trans simp: subset_eq) | 
| 60758 | 1806 | with t Min \<open>C \<subseteq> S\<close> show ?thesis | 
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1807 | by fastforce | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1808 | qed | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1809 | |
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1810 | lemma continuous_attains_sup: | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1811 | fixes f :: "'a::topological_space \<Rightarrow> 'b::linorder_topology" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1812 |   shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f \<Longrightarrow> (\<exists>x\<in>s. \<forall>y\<in>s.  f y \<le> f x)"
 | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1813 | using compact_attains_sup[of "f ` s"] compact_continuous_image[of s f] by auto | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1814 | |
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1815 | lemma continuous_attains_inf: | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1816 | fixes f :: "'a::topological_space \<Rightarrow> 'b::linorder_topology" | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1817 |   shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f \<Longrightarrow> (\<exists>x\<in>s. \<forall>y\<in>s. f x \<le> f y)"
 | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1818 | using compact_attains_inf[of "f ` s"] compact_continuous_image[of s f] by auto | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1819 | |
| 60758 | 1820 | subsection \<open>Connectedness\<close> | 
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1821 | |
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1822 | context topological_space | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1823 | begin | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1824 | |
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1825 | definition "connected S \<longleftrightarrow> | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1826 |   \<not> (\<exists>A B. open A \<and> open B \<and> S \<subseteq> A \<union> B \<and> A \<inter> B \<inter> S = {} \<and> A \<inter> S \<noteq> {} \<and> B \<inter> S \<noteq> {})"
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1827 | |
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1828 | lemma connectedI: | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1829 |   "(\<And>A B. open A \<Longrightarrow> open B \<Longrightarrow> A \<inter> U \<noteq> {} \<Longrightarrow> B \<inter> U \<noteq> {} \<Longrightarrow> A \<inter> B \<inter> U = {} \<Longrightarrow> U \<subseteq> A \<union> B \<Longrightarrow> False)
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1830 | \<Longrightarrow> connected U" | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1831 | by (auto simp: connected_def) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1832 | |
| 61306 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1833 | lemma connected_empty [simp]: "connected {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1834 | by (auto intro!: connectedI) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1835 | |
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1836 | lemma connected_sing [simp]: "connected {x}"
 | 
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1837 | by (auto intro!: connectedI) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1838 | |
| 56329 | 1839 | lemma connectedD: | 
| 1840 |   "connected A \<Longrightarrow> open U \<Longrightarrow> open V \<Longrightarrow> U \<inter> V \<inter> A = {} \<Longrightarrow> A \<subseteq> U \<union> V \<Longrightarrow> U \<inter> A = {} \<or> V \<inter> A = {}" 
 | |
| 1841 | by (auto simp: connected_def) | |
| 1842 | ||
| 51479 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1843 | end | 
| 
33db4b7189af
move compact to the HOL image; prove compactness of real closed intervals; show that continuous functions attain supremum and infimum on compact sets
 hoelzl parents: 
51478diff
changeset | 1844 | |
| 61306 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1845 | lemma connected_closed: | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1846 | "connected s \<longleftrightarrow> | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1847 |      ~ (\<exists>A B. closed A \<and> closed B \<and> s \<subseteq> A \<union> B \<and> A \<inter> B \<inter> s = {} \<and> A \<inter> s \<noteq> {} \<and> B \<inter> s \<noteq> {})"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1848 | apply (simp add: connected_def del: ex_simps, safe) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1849 | apply (drule_tac x="-A" in spec) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1850 | apply (drule_tac x="-B" in spec) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1851 | apply (fastforce simp add: closed_def [symmetric]) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1852 | apply (drule_tac x="-A" in spec) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1853 | apply (drule_tac x="-B" in spec) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1854 | apply (fastforce simp add: open_closed [symmetric]) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1855 | done | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1856 | |
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1857 | |
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1858 | lemma connected_Union: | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1859 |   assumes cs: "\<And>s. s \<in> S \<Longrightarrow> connected s" and ne: "\<Inter>S \<noteq> {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1860 | shows "connected(\<Union>S)" | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1861 | proof (rule connectedI) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1862 | fix A B | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1863 |   assume A: "open A" and B: "open B" and Alap: "A \<inter> \<Union>S \<noteq> {}" and Blap: "B \<inter> \<Union>S \<noteq> {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1864 |      and disj: "A \<inter> B \<inter> \<Union>S = {}" and cover: "\<Union>S \<subseteq> A \<union> B"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1865 |   have disjs:"\<And>s. s \<in> S \<Longrightarrow> A \<inter> B \<inter> s = {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1866 | using disj by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1867 |   obtain sa where sa: "sa \<in> S" "A \<inter> sa \<noteq> {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1868 | using Alap by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1869 |   obtain sb where sb: "sb \<in> S" "B \<inter> sb \<noteq> {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1870 | using Blap by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1871 | obtain x where x: "\<And>s. s \<in> S \<Longrightarrow> x \<in> s" | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1872 | using ne by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1873 | then have "x \<in> \<Union>S" | 
| 61342 | 1874 | using \<open>sa \<in> S\<close> by blast | 
| 61306 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1875 | then have "x \<in> A \<or> x \<in> B" | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1876 | using cover by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1877 | then show False | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1878 | using cs [unfolded connected_def] | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1879 | by (metis A B IntI Sup_upper sa sb disjs x cover empty_iff subset_trans) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1880 | qed | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1881 | |
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1882 | lemma connected_Un: "\<lbrakk>connected s; connected t; s \<inter> t \<noteq> {}\<rbrakk> \<Longrightarrow> connected (s \<union> t)"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1883 |   using connected_Union [of "{s,t}"] by auto
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1884 | |
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1885 | lemma connected_diff_open_from_closed: | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1886 | assumes st: "s \<subseteq> t" and tu: "t \<subseteq> u" and s: "open s" | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1887 | and t: "closed t" and u: "connected u" and ts: "connected (t - s)" | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1888 | shows "connected(u - s)" | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1889 | proof (rule connectedI) | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1890 | fix A B | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1891 |   assume AB: "open A" "open B" "A \<inter> (u - s) \<noteq> {}" "B \<inter> (u - s) \<noteq> {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1892 |      and disj: "A \<inter> B \<inter> (u - s) = {}" and cover: "u - s \<subseteq> A \<union> B"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1893 |   then consider "A \<inter> (t - s) = {}" | "B \<inter> (t - s) = {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1894 | using st ts tu connectedD [of "t-s" "A" "B"] | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1895 | by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1896 | then show False | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1897 | proof cases | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1898 | case 1 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1899 |     then have "(A - t) \<inter> (B \<union> s) \<inter> u = {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1900 | using disj st by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1901 | moreover have "u \<subseteq> (A - t) \<union> (B \<union> s)" using 1 cover by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1902 | ultimately show False | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1903 | using connectedD [of u "A - t" "B \<union> s"] AB s t 1 u | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1904 | by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1905 | next | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1906 | case 2 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1907 |     then have "(A \<union> s) \<inter> (B - t) \<inter> u = {}"
 | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1908 | using disj st | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1909 | by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1910 | moreover have "u \<subseteq> (A \<union> s) \<union> (B - t)" using 2 cover by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1911 | ultimately show False | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1912 | using connectedD [of u "A \<union> s" "B - t"] AB s t 2 u | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1913 | by auto | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1914 | qed | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1915 | qed | 
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 1916 | |
| 59106 | 1917 | lemma connected_iff_const: | 
| 1918 | fixes S :: "'a::topological_space set" | |
| 1919 | shows "connected S \<longleftrightarrow> (\<forall>P::'a \<Rightarrow> bool. continuous_on S P \<longrightarrow> (\<exists>c. \<forall>s\<in>S. P s = c))" | |
| 1920 | proof safe | |
| 1921 | fix P :: "'a \<Rightarrow> bool" assume "connected S" "continuous_on S P" | |
| 1922 |   then have "\<And>b. \<exists>A. open A \<and> A \<inter> S = P -` {b} \<inter> S"
 | |
| 1923 | unfolding continuous_on_open_invariant by simp | |
| 1924 | from this[of True] this[of False] | |
| 1925 |   obtain t f where "open t" "open f" and *: "f \<inter> S = P -` {False} \<inter> S" "t \<inter> S = P -` {True} \<inter> S"
 | |
| 1926 | by auto | |
| 1927 |   then have "t \<inter> S = {} \<or> f \<inter> S = {}"
 | |
| 60758 | 1928 | by (intro connectedD[OF \<open>connected S\<close>]) auto | 
| 59106 | 1929 | then show "\<exists>c. \<forall>s\<in>S. P s = c" | 
| 1930 | proof (rule disjE) | |
| 1931 |     assume "t \<inter> S = {}" then show ?thesis
 | |
| 1932 | unfolding * by (intro exI[of _ False]) auto | |
| 1933 | next | |
| 1934 |     assume "f \<inter> S = {}" then show ?thesis
 | |
| 1935 | unfolding * by (intro exI[of _ True]) auto | |
| 1936 | qed | |
| 1937 | next | |
| 1938 | assume P: "\<forall>P::'a \<Rightarrow> bool. continuous_on S P \<longrightarrow> (\<exists>c. \<forall>s\<in>S. P s = c)" | |
| 1939 | show "connected S" | |
| 1940 | proof (rule connectedI) | |
| 1941 |     fix A B assume *: "open A" "open B" "A \<inter> S \<noteq> {}" "B \<inter> S \<noteq> {}" "A \<inter> B \<inter> S = {}" "S \<subseteq> A \<union> B"
 | |
| 1942 | have "continuous_on S (\<lambda>x. x \<in> A)" | |
| 1943 | unfolding continuous_on_open_invariant | |
| 1944 | proof safe | |
| 1945 | fix C :: "bool set" | |
| 1946 |       have "C = UNIV \<or> C = {True} \<or> C = {False} \<or> C = {}"
 | |
| 1947 | using subset_UNIV[of C] unfolding UNIV_bool by auto | |
| 1948 | with * show "\<exists>T. open T \<and> T \<inter> S = (\<lambda>x. x \<in> A) -` C \<inter> S" | |
| 1949 |         by (intro exI[of _ "(if True \<in> C then A else {}) \<union> (if False \<in> C then B else {})"]) auto
 | |
| 1950 | qed | |
| 1951 | from P[rule_format, OF this] obtain c where "\<And>s. s \<in> S \<Longrightarrow> (s \<in> A) = c" by blast | |
| 1952 | with * show False | |
| 1953 | by (cases c) auto | |
| 1954 | qed | |
| 1955 | qed | |
| 1956 | ||
| 1957 | lemma connectedD_const: | |
| 1958 | fixes P :: "'a::topological_space \<Rightarrow> bool" | |
| 1959 | shows "connected S \<Longrightarrow> continuous_on S P \<Longrightarrow> \<exists>c. \<forall>s\<in>S. P s = c" | |
| 1960 | unfolding connected_iff_const by auto | |
| 1961 | ||
| 1962 | lemma connectedI_const: | |
| 1963 | "(\<And>P::'a::topological_space \<Rightarrow> bool. continuous_on S P \<Longrightarrow> \<exists>c. \<forall>s\<in>S. P s = c) \<Longrightarrow> connected S" | |
| 1964 | unfolding connected_iff_const by auto | |
| 1965 | ||
| 56329 | 1966 | lemma connected_local_const: | 
| 1967 | assumes "connected A" "a \<in> A" "b \<in> A" | |
| 1968 | assumes *: "\<forall>a\<in>A. eventually (\<lambda>b. f a = f b) (at a within A)" | |
| 1969 | shows "f a = f b" | |
| 1970 | proof - | |
| 1971 | obtain S where S: "\<And>a. a \<in> A \<Longrightarrow> a \<in> S a" "\<And>a. a \<in> A \<Longrightarrow> open (S a)" | |
| 1972 | "\<And>a x. a \<in> A \<Longrightarrow> x \<in> S a \<Longrightarrow> x \<in> A \<Longrightarrow> f a = f x" | |
| 1973 | using * unfolding eventually_at_topological by metis | |
| 1974 | ||
| 1975 |   let ?P = "\<Union>b\<in>{b\<in>A. f a = f b}. S b" and ?N = "\<Union>b\<in>{b\<in>A. f a \<noteq> f b}. S b"
 | |
| 1976 |   have "?P \<inter> A = {} \<or> ?N \<inter> A = {}"
 | |
| 60758 | 1977 | using \<open>connected A\<close> S \<open>a\<in>A\<close> | 
| 56329 | 1978 | by (intro connectedD) (auto, metis) | 
| 1979 | then show "f a = f b" | |
| 1980 | proof | |
| 1981 |     assume "?N \<inter> A = {}"
 | |
| 1982 | then have "\<forall>x\<in>A. f a = f x" | |
| 1983 | using S(1) by auto | |
| 60758 | 1984 | with \<open>b\<in>A\<close> show ?thesis by auto | 
| 56329 | 1985 | next | 
| 1986 |     assume "?P \<inter> A = {}" then show ?thesis
 | |
| 60758 | 1987 | using \<open>a \<in> A\<close> S(1)[of a] by auto | 
| 56329 | 1988 | qed | 
| 1989 | qed | |
| 1990 | ||
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1991 | lemma (in linorder_topology) connectedD_interval: | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1992 | assumes "connected U" and xy: "x \<in> U" "y \<in> U" and "x \<le> z" "z \<le> y" | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1993 | shows "z \<in> U" | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1994 | proof - | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1995 |   have eq: "{..<z} \<union> {z<..} = - {z}"
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1996 | by auto | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1997 |   { assume "z \<notin> U" "x < z" "z < y"
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1998 | with xy have "\<not> connected U" | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 1999 | unfolding connected_def simp_thms | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2000 |       apply (rule_tac exI[of _ "{..< z}"])
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2001 |       apply (rule_tac exI[of _ "{z <..}"])
 | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2002 | apply (auto simp add: eq) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2003 | done } | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2004 | with assms show "z \<in> U" | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2005 | by (metis less_le) | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2006 | qed | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2007 | |
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2008 | lemma connected_continuous_image: | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2009 | assumes *: "continuous_on s f" | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2010 | assumes "connected s" | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2011 | shows "connected (f ` s)" | 
| 59106 | 2012 | proof (rule connectedI_const) | 
| 2013 | fix P :: "'b \<Rightarrow> bool" assume "continuous_on (f ` s) P" | |
| 2014 | then have "continuous_on s (P \<circ> f)" | |
| 2015 | by (rule continuous_on_compose[OF *]) | |
| 60758 | 2016 | from connectedD_const[OF \<open>connected s\<close> this] show "\<exists>c. \<forall>s\<in>f ` s. P s = c" | 
| 59106 | 2017 | by auto | 
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2018 | qed | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2019 | |
| 61306 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 2020 | |
| 
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
 paulson <lp15@cam.ac.uk> parents: 
61245diff
changeset | 2021 | section \<open>Linear Continuum Topologies\<close> | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2022 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2023 | class linear_continuum_topology = linorder_topology + linear_continuum | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2024 | begin | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2025 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2026 | lemma Inf_notin_open: | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2027 | assumes A: "open A" and bnd: "\<forall>a\<in>A. x < a" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2028 | shows "Inf A \<notin> A" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2029 | proof | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2030 | assume "Inf A \<in> A" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2031 |   then obtain b where "b < Inf A" "{b <.. Inf A} \<subseteq> A"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2032 | using open_left[of A "Inf A" x] assms by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2033 | with dense[of b "Inf A"] obtain c where "c < Inf A" "c \<in> A" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2034 | by (auto simp: subset_eq) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2035 | then show False | 
| 60758 | 2036 | using cInf_lower[OF \<open>c \<in> A\<close>] bnd by (metis not_le less_imp_le bdd_belowI) | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2037 | qed | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2038 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2039 | lemma Sup_notin_open: | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2040 | assumes A: "open A" and bnd: "\<forall>a\<in>A. a < x" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2041 | shows "Sup A \<notin> A" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2042 | proof | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2043 | assume "Sup A \<in> A" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2044 |   then obtain b where "Sup A < b" "{Sup A ..< b} \<subseteq> A"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2045 | using open_right[of A "Sup A" x] assms by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2046 | with dense[of "Sup A" b] obtain c where "Sup A < c" "c \<in> A" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2047 | by (auto simp: subset_eq) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2048 | then show False | 
| 60758 | 2049 | using cSup_upper[OF \<open>c \<in> A\<close>] bnd by (metis less_imp_le not_le bdd_aboveI) | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2050 | qed | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2051 | |
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2052 | end | 
| 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
51479diff
changeset | 2053 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2054 | instance linear_continuum_topology \<subseteq> perfect_space | 
| 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2055 | proof | 
| 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2056 | fix x :: 'a | 
| 53381 | 2057 | obtain y where "x < y \<or> y < x" | 
| 2058 | using ex_gt_or_lt [of x] .. | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2059 |   with Inf_notin_open[of "{x}" y] Sup_notin_open[of "{x}" y]
 | 
| 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2060 |   show "\<not> open {x}"
 | 
| 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2061 | by auto | 
| 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2062 | qed | 
| 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2063 | |
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2064 | lemma connectedI_interval: | 
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2065 | fixes U :: "'a :: linear_continuum_topology set" | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2066 | assumes *: "\<And>x y z. x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x \<le> z \<Longrightarrow> z \<le> y \<Longrightarrow> z \<in> U" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2067 | shows "connected U" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2068 | proof (rule connectedI) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2069 |   { fix A B assume "open A" "open B" "A \<inter> B \<inter> U = {}" "U \<subseteq> A \<union> B"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2070 | fix x y assume "x < y" "x \<in> A" "y \<in> B" "x \<in> U" "y \<in> U" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2071 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2072 |     let ?z = "Inf (B \<inter> {x <..})"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2073 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2074 | have "x \<le> ?z" "?z \<le> y" | 
| 60758 | 2075 | using \<open>y \<in> B\<close> \<open>x < y\<close> by (auto intro: cInf_lower cInf_greatest) | 
| 2076 | with \<open>x \<in> U\<close> \<open>y \<in> U\<close> have "?z \<in> U" | |
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2077 | by (rule *) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2078 |     moreover have "?z \<notin> B \<inter> {x <..}"
 | 
| 60758 | 2079 | using \<open>open B\<close> by (intro Inf_notin_open) auto | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2080 | ultimately have "?z \<in> A" | 
| 60758 | 2081 |       using \<open>x \<le> ?z\<close> \<open>A \<inter> B \<inter> U = {}\<close> \<open>x \<in> A\<close> \<open>U \<subseteq> A \<union> B\<close> by auto
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2082 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2083 |     { assume "?z < y"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2084 |       obtain a where "?z < a" "{?z ..< a} \<subseteq> A"
 | 
| 60758 | 2085 | using open_right[OF \<open>open A\<close> \<open>?z \<in> A\<close> \<open>?z < y\<close>] by auto | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2086 | moreover obtain b where "b \<in> B" "x < b" "b < min a y" | 
| 60758 | 2087 |         using cInf_less_iff[of "B \<inter> {x <..}" "min a y"] \<open>?z < a\<close> \<open>?z < y\<close> \<open>x < y\<close> \<open>y \<in> B\<close>
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2088 | by (auto intro: less_imp_le) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53215diff
changeset | 2089 | moreover have "?z \<le> b" | 
| 60758 | 2090 | using \<open>b \<in> B\<close> \<open>x < b\<close> | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
53946diff
changeset | 2091 | by (intro cInf_lower) auto | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2092 | moreover have "b \<in> U" | 
| 60758 | 2093 | using \<open>x \<le> ?z\<close> \<open>?z \<le> b\<close> \<open>b < min a y\<close> | 
| 2094 | by (intro *[OF \<open>x \<in> U\<close> \<open>y \<in> U\<close>]) (auto simp: less_imp_le) | |
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2095 | ultimately have "\<exists>b\<in>B. b \<in> A \<and> b \<in> U" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2096 | by (intro bexI[of _ b]) auto } | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2097 | then have False | 
| 60758 | 2098 |       using \<open>?z \<le> y\<close> \<open>?z \<in> A\<close> \<open>y \<in> B\<close> \<open>y \<in> U\<close> \<open>A \<inter> B \<inter> U = {}\<close> unfolding le_less by blast }
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2099 | note not_disjoint = this | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2100 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2101 |   fix A B assume AB: "open A" "open B" "U \<subseteq> A \<union> B" "A \<inter> B \<inter> U = {}"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2102 |   moreover assume "A \<inter> U \<noteq> {}" then obtain x where x: "x \<in> U" "x \<in> A" by auto
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2103 |   moreover assume "B \<inter> U \<noteq> {}" then obtain y where y: "y \<in> U" "y \<in> B" by auto
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2104 | moreover note not_disjoint[of B A y x] not_disjoint[of A B x y] | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2105 | ultimately show False by (cases x y rule: linorder_cases) auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2106 | qed | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2107 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2108 | lemma connected_iff_interval: | 
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2109 | fixes U :: "'a :: linear_continuum_topology set" | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2110 | shows "connected U \<longleftrightarrow> (\<forall>x\<in>U. \<forall>y\<in>U. \<forall>z. x \<le> z \<longrightarrow> z \<le> y \<longrightarrow> z \<in> U)" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2111 | by (auto intro: connectedI_interval dest: connectedD_interval) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2112 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2113 | lemma connected_UNIV[simp]: "connected (UNIV::'a::linear_continuum_topology set)" | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2114 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2115 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2116 | lemma connected_Ioi[simp]: "connected {a::'a::linear_continuum_topology <..}"
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2117 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2118 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2119 | lemma connected_Ici[simp]: "connected {a::'a::linear_continuum_topology ..}"
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2120 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2121 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2122 | lemma connected_Iio[simp]: "connected {..< a::'a::linear_continuum_topology}"
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2123 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2124 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2125 | lemma connected_Iic[simp]: "connected {.. a::'a::linear_continuum_topology}"
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2126 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2127 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2128 | lemma connected_Ioo[simp]: "connected {a <..< b::'a::linear_continuum_topology}"
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2129 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2130 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2131 | lemma connected_Ioc[simp]: "connected {a <.. b::'a::linear_continuum_topology}"
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2132 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2133 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2134 | lemma connected_Ico[simp]: "connected {a ..< b::'a::linear_continuum_topology}"
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2135 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2136 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2137 | lemma connected_Icc[simp]: "connected {a .. b::'a::linear_continuum_topology}"
 | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2138 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2139 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2140 | lemma connected_contains_Ioo: | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2141 | fixes A :: "'a :: linorder_topology set" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2142 |   assumes A: "connected A" "a \<in> A" "b \<in> A" shows "{a <..< b} \<subseteq> A"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2143 | using connectedD_interval[OF A] by (simp add: subset_eq Ball_def less_imp_le) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2144 | |
| 60758 | 2145 | subsection \<open>Intermediate Value Theorem\<close> | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2146 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2147 | lemma IVT': | 
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2148 | fixes f :: "'a :: linear_continuum_topology \<Rightarrow> 'b :: linorder_topology" | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2149 | assumes y: "f a \<le> y" "y \<le> f b" "a \<le> b" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2150 |   assumes *: "continuous_on {a .. b} f"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2151 | shows "\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2152 | proof - | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2153 |   have "connected {a..b}"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2154 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2155 | from connected_continuous_image[OF * this, THEN connectedD_interval, of "f a" "f b" y] y | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2156 | show ?thesis | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2157 | by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2158 | qed | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2159 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2160 | lemma IVT2': | 
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2161 | fixes f :: "'a :: linear_continuum_topology \<Rightarrow> 'b :: linorder_topology" | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2162 | assumes y: "f b \<le> y" "y \<le> f a" "a \<le> b" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2163 |   assumes *: "continuous_on {a .. b} f"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2164 | shows "\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2165 | proof - | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2166 |   have "connected {a..b}"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2167 | unfolding connected_iff_interval by auto | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2168 | from connected_continuous_image[OF * this, THEN connectedD_interval, of "f b" "f a" y] y | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2169 | show ?thesis | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2170 | by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2171 | qed | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2172 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2173 | lemma IVT: | 
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2174 | fixes f :: "'a :: linear_continuum_topology \<Rightarrow> 'b :: linorder_topology" | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2175 | shows "f a \<le> y \<Longrightarrow> y \<le> f b \<Longrightarrow> a \<le> b \<Longrightarrow> (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x) \<Longrightarrow> \<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2176 | by (rule IVT') (auto intro: continuous_at_imp_continuous_on) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2177 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2178 | lemma IVT2: | 
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2179 | fixes f :: "'a :: linear_continuum_topology \<Rightarrow> 'b :: linorder_topology" | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2180 | shows "f b \<le> y \<Longrightarrow> y \<le> f a \<Longrightarrow> a \<le> b \<Longrightarrow> (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x) \<Longrightarrow> \<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2181 | by (rule IVT2') (auto intro: continuous_at_imp_continuous_on) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2182 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2183 | lemma continuous_inj_imp_mono: | 
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 2184 | fixes f :: "'a::linear_continuum_topology \<Rightarrow> 'b :: linorder_topology" | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2185 | assumes x: "a < x" "x < b" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2186 |   assumes cont: "continuous_on {a..b} f"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2187 |   assumes inj: "inj_on f {a..b}"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2188 | shows "(f a < f x \<and> f x < f b) \<or> (f b < f x \<and> f x < f a)" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2189 | proof - | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2190 | note I = inj_on_iff[OF inj] | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2191 |   { assume "f x < f a" "f x < f b"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2192 | then obtain s t where "x \<le> s" "s \<le> b" "a \<le> t" "t \<le> x" "f s = f t" "f x < f s" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2193 | using IVT'[of f x "min (f a) (f b)" b] IVT2'[of f x "min (f a) (f b)" a] x | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2194 | by (auto simp: continuous_on_subset[OF cont] less_imp_le) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2195 | with x I have False by auto } | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2196 | moreover | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2197 |   { assume "f a < f x" "f b < f x"
 | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2198 | then obtain s t where "x \<le> s" "s \<le> b" "a \<le> t" "t \<le> x" "f s = f t" "f s < f x" | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2199 | using IVT'[of f a "max (f a) (f b)" x] IVT2'[of f b "max (f a) (f b)" x] x | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2200 | by (auto simp: continuous_on_subset[OF cont] less_imp_le) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2201 | with x I have False by auto } | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2202 | ultimately show ?thesis | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2203 | using I[of a x] I[of x b] x less_trans[OF x] by (auto simp add: le_less less_imp_neq neq_iff) | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2204 | qed | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2205 | |
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2206 | lemma continuous_at_Sup_mono: | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2207 |   fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder} \<Rightarrow> 'b :: {linorder_topology, conditionally_complete_linorder}"
 | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2208 | assumes "mono f" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2209 | assumes cont: "continuous (at_left (Sup S)) f" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2210 |   assumes S: "S \<noteq> {}" "bdd_above S"
 | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2211 | shows "f (Sup S) = (SUP s:S. f s)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2212 | proof (rule antisym) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2213 | have f: "(f ---> f (Sup S)) (at_left (Sup S))" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2214 | using cont unfolding continuous_within . | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2215 | |
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2216 | show "f (Sup S) \<le> (SUP s:S. f s)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2217 | proof cases | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2218 | assume "Sup S \<in> S" then show ?thesis | 
| 60758 | 2219 | by (rule cSUP_upper) (auto intro: bdd_above_image_mono S \<open>mono f\<close>) | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2220 | next | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2221 | assume "Sup S \<notin> S" | 
| 60758 | 2222 |     from \<open>S \<noteq> {}\<close> obtain s where "s \<in> S"
 | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2223 | by auto | 
| 60758 | 2224 | with \<open>Sup S \<notin> S\<close> S have "s < Sup S" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2225 | unfolding less_le by (blast intro: cSup_upper) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2226 | show ?thesis | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2227 | proof (rule ccontr) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2228 | assume "\<not> ?thesis" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2229 | with order_tendstoD(1)[OF f, of "SUP s:S. f s"] obtain b where "b < Sup S" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2230 | and *: "\<And>y. b < y \<Longrightarrow> y < Sup S \<Longrightarrow> (SUP s:S. f s) < f y" | 
| 60758 | 2231 | by (auto simp: not_le eventually_at_left[OF \<open>s < Sup S\<close>]) | 
| 2232 |       with \<open>S \<noteq> {}\<close> obtain c where "c \<in> S" "b < c"
 | |
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2233 | using less_cSupD[of S b] by auto | 
| 60758 | 2234 | with \<open>Sup S \<notin> S\<close> S have "c < Sup S" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2235 | unfolding less_le by (blast intro: cSup_upper) | 
| 60758 | 2236 | from *[OF \<open>b < c\<close> \<open>c < Sup S\<close>] cSUP_upper[OF \<open>c \<in> S\<close> bdd_above_image_mono[of f]] | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2237 | show False | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2238 | by (auto simp: assms) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2239 | qed | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2240 | qed | 
| 60758 | 2241 | qed (intro cSUP_least \<open>mono f\<close>[THEN monoD] cSup_upper S) | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2242 | |
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2243 | lemma continuous_at_Sup_antimono: | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2244 |   fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder} \<Rightarrow> 'b :: {linorder_topology, conditionally_complete_linorder}"
 | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2245 | assumes "antimono f" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2246 | assumes cont: "continuous (at_left (Sup S)) f" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2247 |   assumes S: "S \<noteq> {}" "bdd_above S"
 | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2248 | shows "f (Sup S) = (INF s:S. f s)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2249 | proof (rule antisym) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2250 | have f: "(f ---> f (Sup S)) (at_left (Sup S))" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2251 | using cont unfolding continuous_within . | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2252 | |
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2253 | show "(INF s:S. f s) \<le> f (Sup S)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2254 | proof cases | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2255 | assume "Sup S \<in> S" then show ?thesis | 
| 60758 | 2256 | by (intro cINF_lower) (auto intro: bdd_below_image_antimono S \<open>antimono f\<close>) | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2257 | next | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2258 | assume "Sup S \<notin> S" | 
| 60758 | 2259 |     from \<open>S \<noteq> {}\<close> obtain s where "s \<in> S"
 | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2260 | by auto | 
| 60758 | 2261 | with \<open>Sup S \<notin> S\<close> S have "s < Sup S" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2262 | unfolding less_le by (blast intro: cSup_upper) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2263 | show ?thesis | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2264 | proof (rule ccontr) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2265 | assume "\<not> ?thesis" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2266 | with order_tendstoD(2)[OF f, of "INF s:S. f s"] obtain b where "b < Sup S" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2267 | and *: "\<And>y. b < y \<Longrightarrow> y < Sup S \<Longrightarrow> f y < (INF s:S. f s)" | 
| 60758 | 2268 | by (auto simp: not_le eventually_at_left[OF \<open>s < Sup S\<close>]) | 
| 2269 |       with \<open>S \<noteq> {}\<close> obtain c where "c \<in> S" "b < c"
 | |
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2270 | using less_cSupD[of S b] by auto | 
| 60758 | 2271 | with \<open>Sup S \<notin> S\<close> S have "c < Sup S" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2272 | unfolding less_le by (blast intro: cSup_upper) | 
| 60758 | 2273 | from *[OF \<open>b < c\<close> \<open>c < Sup S\<close>] cINF_lower[OF bdd_below_image_antimono, of f S c] \<open>c \<in> S\<close> | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2274 | show False | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2275 | by (auto simp: assms) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2276 | qed | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2277 | qed | 
| 60758 | 2278 | qed (intro cINF_greatest \<open>antimono f\<close>[THEN antimonoD] cSup_upper S) | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2279 | |
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2280 | lemma continuous_at_Inf_mono: | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2281 |   fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder} \<Rightarrow> 'b :: {linorder_topology, conditionally_complete_linorder}"
 | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2282 | assumes "mono f" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2283 | assumes cont: "continuous (at_right (Inf S)) f" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2284 |   assumes S: "S \<noteq> {}" "bdd_below S"
 | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2285 | shows "f (Inf S) = (INF s:S. f s)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2286 | proof (rule antisym) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2287 | have f: "(f ---> f (Inf S)) (at_right (Inf S))" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2288 | using cont unfolding continuous_within . | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2289 | |
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2290 | show "(INF s:S. f s) \<le> f (Inf S)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2291 | proof cases | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2292 | assume "Inf S \<in> S" then show ?thesis | 
| 60758 | 2293 | by (rule cINF_lower[rotated]) (auto intro: bdd_below_image_mono S \<open>mono f\<close>) | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2294 | next | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2295 | assume "Inf S \<notin> S" | 
| 60758 | 2296 |     from \<open>S \<noteq> {}\<close> obtain s where "s \<in> S"
 | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2297 | by auto | 
| 60758 | 2298 | with \<open>Inf S \<notin> S\<close> S have "Inf S < s" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2299 | unfolding less_le by (blast intro: cInf_lower) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2300 | show ?thesis | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2301 | proof (rule ccontr) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2302 | assume "\<not> ?thesis" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2303 | with order_tendstoD(2)[OF f, of "INF s:S. f s"] obtain b where "Inf S < b" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2304 | and *: "\<And>y. Inf S < y \<Longrightarrow> y < b \<Longrightarrow> f y < (INF s:S. f s)" | 
| 60758 | 2305 | by (auto simp: not_le eventually_at_right[OF \<open>Inf S < s\<close>]) | 
| 2306 |       with \<open>S \<noteq> {}\<close> obtain c where "c \<in> S" "c < b"
 | |
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2307 | using cInf_lessD[of S b] by auto | 
| 60758 | 2308 | with \<open>Inf S \<notin> S\<close> S have "Inf S < c" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2309 | unfolding less_le by (blast intro: cInf_lower) | 
| 60758 | 2310 | from *[OF \<open>Inf S < c\<close> \<open>c < b\<close>] cINF_lower[OF bdd_below_image_mono[of f] \<open>c \<in> S\<close>] | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2311 | show False | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2312 | by (auto simp: assms) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2313 | qed | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2314 | qed | 
| 60758 | 2315 | qed (intro cINF_greatest \<open>mono f\<close>[THEN monoD] cInf_lower \<open>bdd_below S\<close> \<open>S \<noteq> {}\<close>)
 | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2316 | |
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2317 | lemma continuous_at_Inf_antimono: | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2318 |   fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder} \<Rightarrow> 'b :: {linorder_topology, conditionally_complete_linorder}"
 | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2319 | assumes "antimono f" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2320 | assumes cont: "continuous (at_right (Inf S)) f" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2321 |   assumes S: "S \<noteq> {}" "bdd_below S"
 | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2322 | shows "f (Inf S) = (SUP s:S. f s)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2323 | proof (rule antisym) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2324 | have f: "(f ---> f (Inf S)) (at_right (Inf S))" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2325 | using cont unfolding continuous_within . | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2326 | |
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2327 | show "f (Inf S) \<le> (SUP s:S. f s)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2328 | proof cases | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2329 | assume "Inf S \<in> S" then show ?thesis | 
| 60758 | 2330 | by (rule cSUP_upper) (auto intro: bdd_above_image_antimono S \<open>antimono f\<close>) | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2331 | next | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2332 | assume "Inf S \<notin> S" | 
| 60758 | 2333 |     from \<open>S \<noteq> {}\<close> obtain s where "s \<in> S"
 | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2334 | by auto | 
| 60758 | 2335 | with \<open>Inf S \<notin> S\<close> S have "Inf S < s" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2336 | unfolding less_le by (blast intro: cInf_lower) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2337 | show ?thesis | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2338 | proof (rule ccontr) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2339 | assume "\<not> ?thesis" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2340 | with order_tendstoD(1)[OF f, of "SUP s:S. f s"] obtain b where "Inf S < b" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2341 | and *: "\<And>y. Inf S < y \<Longrightarrow> y < b \<Longrightarrow> (SUP s:S. f s) < f y" | 
| 60758 | 2342 | by (auto simp: not_le eventually_at_right[OF \<open>Inf S < s\<close>]) | 
| 2343 |       with \<open>S \<noteq> {}\<close> obtain c where "c \<in> S" "c < b"
 | |
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2344 | using cInf_lessD[of S b] by auto | 
| 60758 | 2345 | with \<open>Inf S \<notin> S\<close> S have "Inf S < c" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2346 | unfolding less_le by (blast intro: cInf_lower) | 
| 60758 | 2347 | from *[OF \<open>Inf S < c\<close> \<open>c < b\<close>] cSUP_upper[OF \<open>c \<in> S\<close> bdd_above_image_antimono[of f]] | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2348 | show False | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2349 | by (auto simp: assms) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2350 | qed | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2351 | qed | 
| 60758 | 2352 | qed (intro cSUP_least \<open>antimono f\<close>[THEN antimonoD] cInf_lower S) | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59106diff
changeset | 2353 | |
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 2354 | end |