1.1 --- a/src/Doc/ProgProve/document/prelude.tex Tue Mar 26 14:38:44 2013 +0100
1.2 +++ b/src/Doc/ProgProve/document/prelude.tex Tue Mar 26 15:10:28 2013 +0100
1.3 @@ -13,15 +13,6 @@
1.4 \usepackage{mathpartir}
1.5 \usepackage{amssymb}
1.6
1.7 -% Enables fixmes
1.8 -\newif \ifDraft \Drafttrue
1.9 -
1.10 -\ifDraft
1.11 - \newcommand{\FIXME}[1]{\textbf{\textsl{FIXME: #1}}}
1.12 -\else
1.13 - \newcommand{\FIXME}[1]{\relax}
1.14 -\fi
1.15 -
1.16 \renewcommand*\descriptionlabel[1]{\hspace\labelsep \textbf{#1}\hfil}
1.17
1.18 % this should be the last package used
2.1 --- a/src/HOL/Complex_Main.thy Tue Mar 26 14:38:44 2013 +0100
2.2 +++ b/src/HOL/Complex_Main.thy Tue Mar 26 15:10:28 2013 +0100
2.3 @@ -4,10 +4,8 @@
2.4 imports
2.5 Main
2.6 Real
2.7 - SupInf
2.8 Complex
2.9 - Log
2.10 - Ln
2.11 + Transcendental
2.12 Taylor
2.13 Deriv
2.14 begin
3.1 --- /dev/null Thu Jan 01 00:00:00 1970 +0000
3.2 +++ b/src/HOL/Conditional_Complete_Lattices.thy Tue Mar 26 15:10:28 2013 +0100
3.3 @@ -0,0 +1,295 @@
3.4 +(* Title: HOL/Conditional_Complete_Lattices.thy
3.5 + Author: Amine Chaieb and L C Paulson, University of Cambridge
3.6 + Author: Johanens Hölzl, TU München
3.7 +*)
3.8 +
3.9 +header {* Conditional Complete Lattices *}
3.10 +
3.11 +theory Conditional_Complete_Lattices
3.12 +imports Main Lubs
3.13 +begin
3.14 +
3.15 +lemma Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
3.16 + by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
3.17 +
3.18 +lemma Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
3.19 + by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
3.20 +
3.21 +text {*
3.22 +
3.23 +To avoid name classes with the @{class complete_lattice}-class we prefix @{const Sup} and
3.24 +@{const Inf} in theorem names with c.
3.25 +
3.26 +*}
3.27 +
3.28 +class conditional_complete_lattice = lattice + Sup + Inf +
3.29 + assumes cInf_lower: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> z \<le> a) \<Longrightarrow> Inf X \<le> x"
3.30 + and cInf_greatest: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X"
3.31 + assumes cSup_upper: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> a \<le> z) \<Longrightarrow> x \<le> Sup X"
3.32 + and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
3.33 +begin
3.34 +
3.35 +lemma cSup_eq_maximum: (*REAL_SUP_MAX in HOL4*)
3.36 + "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
3.37 + by (blast intro: antisym cSup_upper cSup_least)
3.38 +
3.39 +lemma cInf_eq_minimum: (*REAL_INF_MIN in HOL4*)
3.40 + "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
3.41 + by (intro antisym cInf_lower[of z X z] cInf_greatest[of X z]) auto
3.42 +
3.43 +lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> (\<And>a. a \<in> S \<Longrightarrow> a \<le> z) \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
3.44 + by (metis order_trans cSup_upper cSup_least)
3.45 +
3.46 +lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> (\<And>a. a \<in> S \<Longrightarrow> z \<le> a) \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
3.47 + by (metis order_trans cInf_lower cInf_greatest)
3.48 +
3.49 +lemma cSup_singleton [simp]: "Sup {x} = x"
3.50 + by (intro cSup_eq_maximum) auto
3.51 +
3.52 +lemma cInf_singleton [simp]: "Inf {x} = x"
3.53 + by (intro cInf_eq_minimum) auto
3.54 +
3.55 +lemma cSup_upper2: (*REAL_IMP_LE_SUP in HOL4*)
3.56 + "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y \<le> Sup X"
3.57 + by (metis cSup_upper order_trans)
3.58 +
3.59 +lemma cInf_lower2:
3.60 + "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X \<le> y"
3.61 + by (metis cInf_lower order_trans)
3.62 +
3.63 +lemma cSup_upper_EX: "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> x \<le> z \<Longrightarrow> x \<le> Sup X"
3.64 + by (blast intro: cSup_upper)
3.65 +
3.66 +lemma cInf_lower_EX: "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> z \<le> x \<Longrightarrow> Inf X \<le> x"
3.67 + by (blast intro: cInf_lower)
3.68 +
3.69 +lemma cSup_eq_non_empty:
3.70 + assumes 1: "X \<noteq> {}"
3.71 + assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
3.72 + assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
3.73 + shows "Sup X = a"
3.74 + by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper)
3.75 +
3.76 +lemma cInf_eq_non_empty:
3.77 + assumes 1: "X \<noteq> {}"
3.78 + assumes 2: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
3.79 + assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
3.80 + shows "Inf X = a"
3.81 + by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
3.82 +
3.83 +lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
3.84 + by (rule cInf_eq_non_empty) (auto intro: cSup_upper cSup_least)
3.85 +
3.86 +lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
3.87 + by (rule cSup_eq_non_empty) (auto intro: cInf_lower cInf_greatest)
3.88 +
3.89 +lemma cSup_insert:
3.90 + assumes x: "X \<noteq> {}"
3.91 + and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
3.92 + shows "Sup (insert a X) = sup a (Sup X)"
3.93 +proof (intro cSup_eq_non_empty)
3.94 + fix y assume "\<And>x. x \<in> insert a X \<Longrightarrow> x \<le> y" with x show "sup a (Sup X) \<le> y" by (auto intro: cSup_least)
3.95 +qed (auto intro: le_supI2 z cSup_upper)
3.96 +
3.97 +lemma cInf_insert:
3.98 + assumes x: "X \<noteq> {}"
3.99 + and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
3.100 + shows "Inf (insert a X) = inf a (Inf X)"
3.101 +proof (intro cInf_eq_non_empty)
3.102 + fix y assume "\<And>x. x \<in> insert a X \<Longrightarrow> y \<le> x" with x show "y \<le> inf a (Inf X)" by (auto intro: cInf_greatest)
3.103 +qed (auto intro: le_infI2 z cInf_lower)
3.104 +
3.105 +lemma cSup_insert_If:
3.106 + "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
3.107 + using cSup_insert[of X z] by simp
3.108 +
3.109 +lemma cInf_insert_if:
3.110 + "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
3.111 + using cInf_insert[of X z] by simp
3.112 +
3.113 +lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
3.114 +proof (induct X arbitrary: x rule: finite_induct)
3.115 + case (insert x X y) then show ?case
3.116 + apply (cases "X = {}")
3.117 + apply simp
3.118 + apply (subst cSup_insert[of _ "Sup X"])
3.119 + apply (auto intro: le_supI2)
3.120 + done
3.121 +qed simp
3.122 +
3.123 +lemma cInf_le_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> Inf X \<le> x"
3.124 +proof (induct X arbitrary: x rule: finite_induct)
3.125 + case (insert x X y) then show ?case
3.126 + apply (cases "X = {}")
3.127 + apply simp
3.128 + apply (subst cInf_insert[of _ "Inf X"])
3.129 + apply (auto intro: le_infI2)
3.130 + done
3.131 +qed simp
3.132 +
3.133 +lemma cSup_eq_Sup_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Sup_fin X"
3.134 +proof (induct X rule: finite_ne_induct)
3.135 + case (insert x X) then show ?case
3.136 + using cSup_insert[of X "Sup_fin X" x] le_cSup_finite[of X] by simp
3.137 +qed simp
3.138 +
3.139 +lemma cInf_eq_Inf_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Inf_fin X"
3.140 +proof (induct X rule: finite_ne_induct)
3.141 + case (insert x X) then show ?case
3.142 + using cInf_insert[of X "Inf_fin X" x] cInf_le_finite[of X] by simp
3.143 +qed simp
3.144 +
3.145 +lemma cSup_atMost[simp]: "Sup {..x} = x"
3.146 + by (auto intro!: cSup_eq_maximum)
3.147 +
3.148 +lemma cSup_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Sup {y<..x} = x"
3.149 + by (auto intro!: cSup_eq_maximum)
3.150 +
3.151 +lemma cSup_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Sup {y..x} = x"
3.152 + by (auto intro!: cSup_eq_maximum)
3.153 +
3.154 +lemma cInf_atLeast[simp]: "Inf {x..} = x"
3.155 + by (auto intro!: cInf_eq_minimum)
3.156 +
3.157 +lemma cInf_atLeastLessThan[simp]: "y < x \<Longrightarrow> Inf {y..<x} = y"
3.158 + by (auto intro!: cInf_eq_minimum)
3.159 +
3.160 +lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
3.161 + by (auto intro!: cInf_eq_minimum)
3.162 +
3.163 +end
3.164 +
3.165 +instance complete_lattice \<subseteq> conditional_complete_lattice
3.166 + by default (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest)
3.167 +
3.168 +lemma isLub_cSup:
3.169 + "(S::'a :: conditional_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. S *<= b) \<Longrightarrow> isLub UNIV S (Sup S)"
3.170 + by (auto simp add: isLub_def setle_def leastP_def isUb_def
3.171 + intro!: setgeI intro: cSup_upper cSup_least)
3.172 +
3.173 +lemma cSup_eq:
3.174 + fixes a :: "'a :: {conditional_complete_lattice, no_bot}"
3.175 + assumes upper: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
3.176 + assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
3.177 + shows "Sup X = a"
3.178 +proof cases
3.179 + assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
3.180 +qed (intro cSup_eq_non_empty assms)
3.181 +
3.182 +lemma cInf_eq:
3.183 + fixes a :: "'a :: {conditional_complete_lattice, no_top}"
3.184 + assumes upper: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
3.185 + assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
3.186 + shows "Inf X = a"
3.187 +proof cases
3.188 + assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
3.189 +qed (intro cInf_eq_non_empty assms)
3.190 +
3.191 +lemma cSup_le: "(S::'a::conditional_complete_lattice set) \<noteq> {} \<Longrightarrow> S *<= b \<Longrightarrow> Sup S \<le> b"
3.192 + by (metis cSup_least setle_def)
3.193 +
3.194 +lemma cInf_ge: "(S::'a :: conditional_complete_lattice set) \<noteq> {} \<Longrightarrow> b <=* S \<Longrightarrow> Inf S \<ge> b"
3.195 + by (metis cInf_greatest setge_def)
3.196 +
3.197 +class conditional_complete_linorder = conditional_complete_lattice + linorder
3.198 +begin
3.199 +
3.200 +lemma less_cSup_iff : (*REAL_SUP_LE in HOL4*)
3.201 + "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
3.202 + by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans)
3.203 +
3.204 +lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
3.205 + by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans)
3.206 +
3.207 +lemma less_cSupE:
3.208 + assumes "y < Sup X" "X \<noteq> {}" obtains x where "x \<in> X" "y < x"
3.209 + by (metis cSup_least assms not_le that)
3.210 +
3.211 +lemma less_cSupD:
3.212 + "X \<noteq> {} \<Longrightarrow> z < Sup X \<Longrightarrow> \<exists>x\<in>X. z < x"
3.213 + by (metis less_cSup_iff not_leE)
3.214 +
3.215 +lemma cInf_lessD:
3.216 + "X \<noteq> {} \<Longrightarrow> Inf X < z \<Longrightarrow> \<exists>x\<in>X. x < z"
3.217 + by (metis cInf_less_iff not_leE)
3.218 +
3.219 +lemma complete_interval:
3.220 + assumes "a < b" and "P a" and "\<not> P b"
3.221 + shows "\<exists>c. a \<le> c \<and> c \<le> b \<and> (\<forall>x. a \<le> x \<and> x < c \<longrightarrow> P x) \<and>
3.222 + (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)"
3.223 +proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x & x < d --> P x}"], auto)
3.224 + show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
3.225 + by (rule cSup_upper [where z=b], auto)
3.226 + (metis `a < b` `\<not> P b` linear less_le)
3.227 +next
3.228 + show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
3.229 + apply (rule cSup_least)
3.230 + apply auto
3.231 + apply (metis less_le_not_le)
3.232 + apply (metis `a<b` `~ P b` linear less_le)
3.233 + done
3.234 +next
3.235 + fix x
3.236 + assume x: "a \<le> x" and lt: "x < Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
3.237 + show "P x"
3.238 + apply (rule less_cSupE [OF lt], auto)
3.239 + apply (metis less_le_not_le)
3.240 + apply (metis x)
3.241 + done
3.242 +next
3.243 + fix d
3.244 + assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x"
3.245 + thus "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
3.246 + by (rule_tac z="b" in cSup_upper, auto)
3.247 + (metis `a<b` `~ P b` linear less_le)
3.248 +qed
3.249 +
3.250 +end
3.251 +
3.252 +lemma cSup_bounds:
3.253 + fixes S :: "'a :: conditional_complete_lattice set"
3.254 + assumes Se: "S \<noteq> {}" and l: "a <=* S" and u: "S *<= b"
3.255 + shows "a \<le> Sup S \<and> Sup S \<le> b"
3.256 +proof-
3.257 + from isLub_cSup[OF Se] u have lub: "isLub UNIV S (Sup S)" by blast
3.258 + hence b: "Sup S \<le> b" using u
3.259 + by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
3.260 + from Se obtain y where y: "y \<in> S" by blast
3.261 + from lub l have "a \<le> Sup S"
3.262 + by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
3.263 + (metis le_iff_sup le_sup_iff y)
3.264 + with b show ?thesis by blast
3.265 +qed
3.266 +
3.267 +
3.268 +lemma cSup_unique: "(S::'a :: {conditional_complete_linorder, no_bot} set) *<= b \<Longrightarrow> (\<forall>b'<b. \<exists>x\<in>S. b' < x) \<Longrightarrow> Sup S = b"
3.269 + by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
3.270 +
3.271 +lemma cInf_unique: "b <=* (S::'a :: {conditional_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
3.272 + by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
3.273 +
3.274 +lemma cSup_eq_Max: "finite (X::'a::conditional_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
3.275 + using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
3.276 +
3.277 +lemma cInf_eq_Min: "finite (X::'a::conditional_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
3.278 + using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
3.279 +
3.280 +lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditional_complete_linorder, dense_linorder}} = x"
3.281 + by (auto intro!: cSup_eq_non_empty intro: dense_le)
3.282 +
3.283 +lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditional_complete_linorder, dense_linorder}} = x"
3.284 + by (auto intro!: cSup_eq intro: dense_le_bounded)
3.285 +
3.286 +lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditional_complete_linorder, dense_linorder}} = x"
3.287 + by (auto intro!: cSup_eq intro: dense_le_bounded)
3.288 +
3.289 +lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditional_complete_linorder, dense_linorder} <..} = x"
3.290 + by (auto intro!: cInf_eq intro: dense_ge)
3.291 +
3.292 +lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditional_complete_linorder, dense_linorder}} = y"
3.293 + by (auto intro!: cInf_eq intro: dense_ge_bounded)
3.294 +
3.295 +lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditional_complete_linorder, dense_linorder}} = y"
3.296 + by (auto intro!: cInf_eq intro: dense_ge_bounded)
3.297 +
3.298 +end
4.1 --- a/src/HOL/Deriv.thy Tue Mar 26 14:38:44 2013 +0100
4.2 +++ b/src/HOL/Deriv.thy Tue Mar 26 15:10:28 2013 +0100
4.3 @@ -8,7 +8,7 @@
4.4 header{* Differentiation *}
4.5
4.6 theory Deriv
4.7 -imports Lim
4.8 +imports Limits
4.9 begin
4.10
4.11 text{*Standard Definitions*}
4.12 @@ -422,178 +422,6 @@
4.13 apply (simp add: assms)
4.14 done
4.15
4.16 -
4.17 -subsection {* Nested Intervals and Bisection *}
4.18 -
4.19 -lemma nested_sequence_unique:
4.20 - assumes "\<forall>n. f n \<le> f (Suc n)" "\<forall>n. g (Suc n) \<le> g n" "\<forall>n. f n \<le> g n" "(\<lambda>n. f n - g n) ----> 0"
4.21 - shows "\<exists>l::real. ((\<forall>n. f n \<le> l) \<and> f ----> l) \<and> ((\<forall>n. l \<le> g n) \<and> g ----> l)"
4.22 -proof -
4.23 - have "incseq f" unfolding incseq_Suc_iff by fact
4.24 - have "decseq g" unfolding decseq_Suc_iff by fact
4.25 -
4.26 - { fix n
4.27 - from `decseq g` have "g n \<le> g 0" by (rule decseqD) simp
4.28 - with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f n \<le> g 0" by auto }
4.29 - then obtain u where "f ----> u" "\<forall>i. f i \<le> u"
4.30 - using incseq_convergent[OF `incseq f`] by auto
4.31 - moreover
4.32 - { fix n
4.33 - from `incseq f` have "f 0 \<le> f n" by (rule incseqD) simp
4.34 - with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f 0 \<le> g n" by simp }
4.35 - then obtain l where "g ----> l" "\<forall>i. l \<le> g i"
4.36 - using decseq_convergent[OF `decseq g`] by auto
4.37 - moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF `f ----> u` `g ----> l`]]
4.38 - ultimately show ?thesis by auto
4.39 -qed
4.40 -
4.41 -lemma Bolzano[consumes 1, case_names trans local]:
4.42 - fixes P :: "real \<Rightarrow> real \<Rightarrow> bool"
4.43 - assumes [arith]: "a \<le> b"
4.44 - assumes trans: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a c"
4.45 - assumes local: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<exists>d>0. \<forall>a b. a \<le> x \<and> x \<le> b \<and> b - a < d \<longrightarrow> P a b"
4.46 - shows "P a b"
4.47 -proof -
4.48 - def bisect \<equiv> "nat_rec (a, b) (\<lambda>n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))"
4.49 - def l \<equiv> "\<lambda>n. fst (bisect n)" and u \<equiv> "\<lambda>n. snd (bisect n)"
4.50 - have l[simp]: "l 0 = a" "\<And>n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)"
4.51 - and u[simp]: "u 0 = b" "\<And>n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)"
4.52 - by (simp_all add: l_def u_def bisect_def split: prod.split)
4.53 -
4.54 - { fix n have "l n \<le> u n" by (induct n) auto } note this[simp]
4.55 -
4.56 - have "\<exists>x. ((\<forall>n. l n \<le> x) \<and> l ----> x) \<and> ((\<forall>n. x \<le> u n) \<and> u ----> x)"
4.57 - proof (safe intro!: nested_sequence_unique)
4.58 - fix n show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" by (induct n) auto
4.59 - next
4.60 - { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) }
4.61 - then show "(\<lambda>n. l n - u n) ----> 0" by (simp add: LIMSEQ_divide_realpow_zero)
4.62 - qed fact
4.63 - then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l ----> x" "u ----> x" by auto
4.64 - obtain d where "0 < d" and d: "\<And>a b. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> b - a < d \<Longrightarrow> P a b"
4.65 - using `l 0 \<le> x` `x \<le> u 0` local[of x] by auto
4.66 -
4.67 - show "P a b"
4.68 - proof (rule ccontr)
4.69 - assume "\<not> P a b"
4.70 - { fix n have "\<not> P (l n) (u n)"
4.71 - proof (induct n)
4.72 - case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto
4.73 - qed (simp add: `\<not> P a b`) }
4.74 - moreover
4.75 - { have "eventually (\<lambda>n. x - d / 2 < l n) sequentially"
4.76 - using `0 < d` `l ----> x` by (intro order_tendstoD[of _ x]) auto
4.77 - moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially"
4.78 - using `0 < d` `u ----> x` by (intro order_tendstoD[of _ x]) auto
4.79 - ultimately have "eventually (\<lambda>n. P (l n) (u n)) sequentially"
4.80 - proof eventually_elim
4.81 - fix n assume "x - d / 2 < l n" "u n < x + d / 2"
4.82 - from add_strict_mono[OF this] have "u n - l n < d" by simp
4.83 - with x show "P (l n) (u n)" by (rule d)
4.84 - qed }
4.85 - ultimately show False by simp
4.86 - qed
4.87 -qed
4.88 -
4.89 -(*HOL style here: object-level formulations*)
4.90 -lemma IVT_objl: "(f(a::real) \<le> (y::real) & y \<le> f(b) & a \<le> b &
4.91 - (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
4.92 - --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
4.93 -apply (blast intro: IVT)
4.94 -done
4.95 -
4.96 -lemma IVT2_objl: "(f(b::real) \<le> (y::real) & y \<le> f(a) & a \<le> b &
4.97 - (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
4.98 - --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
4.99 -apply (blast intro: IVT2)
4.100 -done
4.101 -
4.102 -
4.103 -lemma compact_Icc[simp, intro]: "compact {a .. b::real}"
4.104 -proof (cases "a \<le> b", rule compactI)
4.105 - fix C assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C"
4.106 - def T == "{a .. b}"
4.107 - from C(1,3) show "\<exists>C'\<subseteq>C. finite C' \<and> {a..b} \<subseteq> \<Union>C'"
4.108 - proof (induct rule: Bolzano)
4.109 - case (trans a b c)
4.110 - then have *: "{a .. c} = {a .. b} \<union> {b .. c}" by auto
4.111 - from trans obtain C1 C2 where "C1\<subseteq>C \<and> finite C1 \<and> {a..b} \<subseteq> \<Union>C1" "C2\<subseteq>C \<and> finite C2 \<and> {b..c} \<subseteq> \<Union>C2"
4.112 - by (auto simp: *)
4.113 - with trans show ?case
4.114 - unfolding * by (intro exI[of _ "C1 \<union> C2"]) auto
4.115 - next
4.116 - case (local x)
4.117 - then have "x \<in> \<Union>C" using C by auto
4.118 - with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" by auto
4.119 - then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c"
4.120 - by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff)
4.121 - with `c \<in> C` show ?case
4.122 - by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto
4.123 - qed
4.124 -qed simp
4.125 -
4.126 -subsection {* Boundedness of continuous functions *}
4.127 -
4.128 -text{*By bisection, function continuous on closed interval is bounded above*}
4.129 -
4.130 -lemma isCont_eq_Ub:
4.131 - fixes f :: "real \<Rightarrow> 'a::linorder_topology"
4.132 - shows "a \<le> b \<Longrightarrow> \<forall>x::real. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
4.133 - \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
4.134 - using continuous_attains_sup[of "{a .. b}" f]
4.135 - apply (simp add: continuous_at_imp_continuous_on Ball_def)
4.136 - apply safe
4.137 - apply (rule_tac x="f x" in exI)
4.138 - apply auto
4.139 - done
4.140 -
4.141 -lemma isCont_eq_Lb:
4.142 - fixes f :: "real \<Rightarrow> 'a::linorder_topology"
4.143 - shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
4.144 - \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> M \<le> f x) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
4.145 - using continuous_attains_inf[of "{a .. b}" f]
4.146 - apply (simp add: continuous_at_imp_continuous_on Ball_def)
4.147 - apply safe
4.148 - apply (rule_tac x="f x" in exI)
4.149 - apply auto
4.150 - done
4.151 -
4.152 -lemma isCont_bounded:
4.153 - fixes f :: "real \<Rightarrow> 'a::linorder_topology"
4.154 - shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> \<exists>M. \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M"
4.155 - using isCont_eq_Ub[of a b f] by auto
4.156 -
4.157 -lemma isCont_has_Ub:
4.158 - fixes f :: "real \<Rightarrow> 'a::linorder_topology"
4.159 - shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
4.160 - \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M) \<and> (\<forall>N. N < M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> N < f x))"
4.161 - using isCont_eq_Ub[of a b f] by auto
4.162 -
4.163 -text{*Refine the above to existence of least upper bound*}
4.164 -
4.165 -lemma lemma_reals_complete: "((\<exists>x. x \<in> S) & (\<exists>y. isUb UNIV S (y::real))) -->
4.166 - (\<exists>t. isLub UNIV S t)"
4.167 -by (blast intro: reals_complete)
4.168 -
4.169 -
4.170 -text{*Another version.*}
4.171 -
4.172 -lemma isCont_Lb_Ub: "[|a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |]
4.173 - ==> \<exists>L M::real. (\<forall>x::real. a \<le> x & x \<le> b --> L \<le> f(x) & f(x) \<le> M) &
4.174 - (\<forall>y. L \<le> y & y \<le> M --> (\<exists>x. a \<le> x & x \<le> b & (f(x) = y)))"
4.175 -apply (frule isCont_eq_Lb)
4.176 -apply (frule_tac [2] isCont_eq_Ub)
4.177 -apply (assumption+, safe)
4.178 -apply (rule_tac x = "f x" in exI)
4.179 -apply (rule_tac x = "f xa" in exI, simp, safe)
4.180 -apply (cut_tac x = x and y = xa in linorder_linear, safe)
4.181 -apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl)
4.182 -apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe)
4.183 -apply (rule_tac [2] x = xb in exI)
4.184 -apply (rule_tac [4] x = xb in exI, simp_all)
4.185 -done
4.186 -
4.187 -
4.188 subsection {* Local extrema *}
4.189
4.190 text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*}
4.191 @@ -658,7 +486,6 @@
4.192 qed
4.193 qed
4.194
4.195 -
4.196 lemma DERIV_pos_inc_left:
4.197 fixes f :: "real => real"
4.198 shows "DERIV f x :> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x - h) < f(x)"
4.199 @@ -1081,47 +908,6 @@
4.200 by simp
4.201 qed
4.202
4.203 -text{*Continuity of inverse function*}
4.204 -
4.205 -lemma isCont_inverse_function:
4.206 - fixes f g :: "real \<Rightarrow> real"
4.207 - assumes d: "0 < d"
4.208 - and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> g (f z) = z"
4.209 - and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> isCont f z"
4.210 - shows "isCont g (f x)"
4.211 -proof -
4.212 - let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + d}"
4.213 -
4.214 - have f: "continuous_on ?D f"
4.215 - using cont by (intro continuous_at_imp_continuous_on ballI) auto
4.216 - then have g: "continuous_on (f`?D) g"
4.217 - using inj by (intro continuous_on_inv) auto
4.218 -
4.219 - from d f have "{min ?A ?B <..< max ?A ?B} \<subseteq> f ` ?D"
4.220 - by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max)
4.221 - with g have "continuous_on {min ?A ?B <..< max ?A ?B} g"
4.222 - by (rule continuous_on_subset)
4.223 - moreover
4.224 - have "(?A < f x \<and> f x < ?B) \<or> (?B < f x \<and> f x < ?A)"
4.225 - using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto
4.226 - then have "f x \<in> {min ?A ?B <..< max ?A ?B}"
4.227 - by auto
4.228 - ultimately
4.229 - show ?thesis
4.230 - by (simp add: continuous_on_eq_continuous_at)
4.231 -qed
4.232 -
4.233 -lemma isCont_inverse_function2:
4.234 - fixes f g :: "real \<Rightarrow> real" shows
4.235 - "\<lbrakk>a < x; x < b;
4.236 - \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z;
4.237 - \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z\<rbrakk>
4.238 - \<Longrightarrow> isCont g (f x)"
4.239 -apply (rule isCont_inverse_function
4.240 - [where f=f and d="min (x - a) (b - x)"])
4.241 -apply (simp_all add: abs_le_iff)
4.242 -done
4.243 -
4.244 text {* Derivative of inverse function *}
4.245
4.246 lemma DERIV_inverse_function:
4.247 @@ -1228,44 +1014,6 @@
4.248 with g'cdef f'cdef cint show ?thesis by auto
4.249 qed
4.250
4.251 -
4.252 -subsection {* Theorems about Limits *}
4.253 -
4.254 -(* need to rename second isCont_inverse *)
4.255 -
4.256 -lemma isCont_inv_fun:
4.257 - fixes f g :: "real \<Rightarrow> real"
4.258 - shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z;
4.259 - \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |]
4.260 - ==> isCont g (f x)"
4.261 -by (rule isCont_inverse_function)
4.262 -
4.263 -text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*}
4.264 -lemma LIM_fun_gt_zero:
4.265 - "[| f -- c --> (l::real); 0 < l |]
4.266 - ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> 0 < f x)"
4.267 -apply (drule (1) LIM_D, clarify)
4.268 -apply (rule_tac x = s in exI)
4.269 -apply (simp add: abs_less_iff)
4.270 -done
4.271 -
4.272 -lemma LIM_fun_less_zero:
4.273 - "[| f -- c --> (l::real); l < 0 |]
4.274 - ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x < 0)"
4.275 -apply (drule LIM_D [where r="-l"], simp, clarify)
4.276 -apply (rule_tac x = s in exI)
4.277 -apply (simp add: abs_less_iff)
4.278 -done
4.279 -
4.280 -lemma LIM_fun_not_zero:
4.281 - "[| f -- c --> (l::real); l \<noteq> 0 |]
4.282 - ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x \<noteq> 0)"
4.283 -apply (rule linorder_cases [of l 0])
4.284 -apply (drule (1) LIM_fun_less_zero, force)
4.285 -apply simp
4.286 -apply (drule (1) LIM_fun_gt_zero, force)
4.287 -done
4.288 -
4.289 lemma GMVT':
4.290 fixes f g :: "real \<Rightarrow> real"
4.291 assumes "a < b"
4.292 @@ -1284,6 +1032,9 @@
4.293 by auto
4.294 qed
4.295
4.296 +
4.297 +subsection {* L'Hopitals rule *}
4.298 +
4.299 lemma DERIV_cong_ev: "x = y \<Longrightarrow> eventually (\<lambda>x. f x = g x) (nhds x) \<Longrightarrow> u = v \<Longrightarrow>
4.300 DERIV f x :> u \<longleftrightarrow> DERIV g y :> v"
4.301 unfolding DERIV_iff2
5.1 --- a/src/HOL/Library/Diagonal_Subsequence.thy Tue Mar 26 14:38:44 2013 +0100
5.2 +++ b/src/HOL/Library/Diagonal_Subsequence.thy Tue Mar 26 15:10:28 2013 +0100
5.3 @@ -3,7 +3,7 @@
5.4 header {* Sequence of Properties on Subsequences *}
5.5
5.6 theory Diagonal_Subsequence
5.7 -imports SEQ
5.8 +imports Complex_Main
5.9 begin
5.10
5.11 locale subseqs =
6.1 --- a/src/HOL/Lim.thy Tue Mar 26 14:38:44 2013 +0100
6.2 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000
6.3 @@ -1,225 +0,0 @@
6.4 -(* Title : Lim.thy
6.5 - Author : Jacques D. Fleuriot
6.6 - Copyright : 1998 University of Cambridge
6.7 - Conversion to Isar and new proofs by Lawrence C Paulson, 2004
6.8 -*)
6.9 -
6.10 -header{* Limits and Continuity *}
6.11 -
6.12 -theory Lim
6.13 -imports SEQ
6.14 -begin
6.15 -
6.16 -subsection {* Limits of Functions *}
6.17 -
6.18 -lemma LIM_eq:
6.19 - fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
6.20 - shows "f -- a --> L =
6.21 - (\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"
6.22 -by (simp add: LIM_def dist_norm)
6.23 -
6.24 -lemma LIM_I:
6.25 - fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
6.26 - shows "(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)
6.27 - ==> f -- a --> L"
6.28 -by (simp add: LIM_eq)
6.29 -
6.30 -lemma LIM_D:
6.31 - fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
6.32 - shows "[| f -- a --> L; 0<r |]
6.33 - ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"
6.34 -by (simp add: LIM_eq)
6.35 -
6.36 -lemma LIM_offset:
6.37 - fixes a :: "'a::real_normed_vector"
6.38 - shows "f -- a --> L \<Longrightarrow> (\<lambda>x. f (x + k)) -- a - k --> L"
6.39 -apply (rule topological_tendstoI)
6.40 -apply (drule (2) topological_tendstoD)
6.41 -apply (simp only: eventually_at dist_norm)
6.42 -apply (clarify, rule_tac x=d in exI, safe)
6.43 -apply (drule_tac x="x + k" in spec)
6.44 -apply (simp add: algebra_simps)
6.45 -done
6.46 -
6.47 -lemma LIM_offset_zero:
6.48 - fixes a :: "'a::real_normed_vector"
6.49 - shows "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"
6.50 -by (drule_tac k="a" in LIM_offset, simp add: add_commute)
6.51 -
6.52 -lemma LIM_offset_zero_cancel:
6.53 - fixes a :: "'a::real_normed_vector"
6.54 - shows "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> f -- a --> L"
6.55 -by (drule_tac k="- a" in LIM_offset, simp)
6.56 -
6.57 -lemma LIM_zero:
6.58 - fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
6.59 - shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. f x - l) ---> 0) F"
6.60 -unfolding tendsto_iff dist_norm by simp
6.61 -
6.62 -lemma LIM_zero_cancel:
6.63 - fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
6.64 - shows "((\<lambda>x. f x - l) ---> 0) F \<Longrightarrow> (f ---> l) F"
6.65 -unfolding tendsto_iff dist_norm by simp
6.66 -
6.67 -lemma LIM_zero_iff:
6.68 - fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
6.69 - shows "((\<lambda>x. f x - l) ---> 0) F = (f ---> l) F"
6.70 -unfolding tendsto_iff dist_norm by simp
6.71 -
6.72 -lemma LIM_imp_LIM:
6.73 - fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
6.74 - fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector"
6.75 - assumes f: "f -- a --> l"
6.76 - assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"
6.77 - shows "g -- a --> m"
6.78 - by (rule metric_LIM_imp_LIM [OF f],
6.79 - simp add: dist_norm le)
6.80 -
6.81 -lemma LIM_equal2:
6.82 - fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
6.83 - assumes 1: "0 < R"
6.84 - assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"
6.85 - shows "g -- a --> l \<Longrightarrow> f -- a --> l"
6.86 -by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm)
6.87 -
6.88 -lemma LIM_compose2:
6.89 - fixes a :: "'a::real_normed_vector"
6.90 - assumes f: "f -- a --> b"
6.91 - assumes g: "g -- b --> c"
6.92 - assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"
6.93 - shows "(\<lambda>x. g (f x)) -- a --> c"
6.94 -by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])
6.95 -
6.96 -lemma real_LIM_sandwich_zero:
6.97 - fixes f g :: "'a::topological_space \<Rightarrow> real"
6.98 - assumes f: "f -- a --> 0"
6.99 - assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"
6.100 - assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"
6.101 - shows "g -- a --> 0"
6.102 -proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *)
6.103 - fix x assume x: "x \<noteq> a"
6.104 - have "norm (g x - 0) = g x" by (simp add: 1 x)
6.105 - also have "g x \<le> f x" by (rule 2 [OF x])
6.106 - also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self)
6.107 - also have "\<bar>f x\<bar> = norm (f x - 0)" by simp
6.108 - finally show "norm (g x - 0) \<le> norm (f x - 0)" .
6.109 -qed
6.110 -
6.111 -
6.112 -subsection {* Continuity *}
6.113 -
6.114 -lemma LIM_isCont_iff:
6.115 - fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
6.116 - shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)"
6.117 -by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])
6.118 -
6.119 -lemma isCont_iff:
6.120 - fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
6.121 - shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x"
6.122 -by (simp add: isCont_def LIM_isCont_iff)
6.123 -
6.124 -lemma isCont_LIM_compose2:
6.125 - fixes a :: "'a::real_normed_vector"
6.126 - assumes f [unfolded isCont_def]: "isCont f a"
6.127 - assumes g: "g -- f a --> l"
6.128 - assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"
6.129 - shows "(\<lambda>x. g (f x)) -- a --> l"
6.130 -by (rule LIM_compose2 [OF f g inj])
6.131 -
6.132 -
6.133 -lemma isCont_norm [simp]:
6.134 - fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
6.135 - shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"
6.136 - by (fact continuous_norm)
6.137 -
6.138 -lemma isCont_rabs [simp]:
6.139 - fixes f :: "'a::t2_space \<Rightarrow> real"
6.140 - shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
6.141 - by (fact continuous_rabs)
6.142 -
6.143 -lemma isCont_add [simp]:
6.144 - fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
6.145 - shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"
6.146 - by (fact continuous_add)
6.147 -
6.148 -lemma isCont_minus [simp]:
6.149 - fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
6.150 - shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"
6.151 - by (fact continuous_minus)
6.152 -
6.153 -lemma isCont_diff [simp]:
6.154 - fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
6.155 - shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"
6.156 - by (fact continuous_diff)
6.157 -
6.158 -lemma isCont_mult [simp]:
6.159 - fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
6.160 - shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"
6.161 - by (fact continuous_mult)
6.162 -
6.163 -lemma (in bounded_linear) isCont:
6.164 - "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
6.165 - by (fact continuous)
6.166 -
6.167 -lemma (in bounded_bilinear) isCont:
6.168 - "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"
6.169 - by (fact continuous)
6.170 -
6.171 -lemmas isCont_scaleR [simp] =
6.172 - bounded_bilinear.isCont [OF bounded_bilinear_scaleR]
6.173 -
6.174 -lemmas isCont_of_real [simp] =
6.175 - bounded_linear.isCont [OF bounded_linear_of_real]
6.176 -
6.177 -lemma isCont_power [simp]:
6.178 - fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
6.179 - shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"
6.180 - by (fact continuous_power)
6.181 -
6.182 -lemma isCont_setsum [simp]:
6.183 - fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
6.184 - shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
6.185 - by (auto intro: continuous_setsum)
6.186 -
6.187 -lemmas isCont_intros =
6.188 - isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus
6.189 - isCont_diff isCont_mult isCont_inverse isCont_divide isCont_scaleR
6.190 - isCont_of_real isCont_power isCont_sgn isCont_setsum
6.191 -
6.192 -subsection {* Uniform Continuity *}
6.193 -
6.194 -lemma (in bounded_linear) isUCont: "isUCont f"
6.195 -unfolding isUCont_def dist_norm
6.196 -proof (intro allI impI)
6.197 - fix r::real assume r: "0 < r"
6.198 - obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"
6.199 - using pos_bounded by fast
6.200 - show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"
6.201 - proof (rule exI, safe)
6.202 - from r K show "0 < r / K" by (rule divide_pos_pos)
6.203 - next
6.204 - fix x y :: 'a
6.205 - assume xy: "norm (x - y) < r / K"
6.206 - have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff)
6.207 - also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le)
6.208 - also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq)
6.209 - finally show "norm (f x - f y) < r" .
6.210 - qed
6.211 -qed
6.212 -
6.213 -lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
6.214 -by (rule isUCont [THEN isUCont_Cauchy])
6.215 -
6.216 -
6.217 -lemma LIM_less_bound:
6.218 - fixes f :: "real \<Rightarrow> real"
6.219 - assumes ev: "b < x" "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and "isCont f x"
6.220 - shows "0 \<le> f x"
6.221 -proof (rule tendsto_le_const)
6.222 - show "(f ---> f x) (at_left x)"
6.223 - using `isCont f x` by (simp add: filterlim_at_split isCont_def)
6.224 - show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)"
6.225 - using ev by (auto simp: eventually_within_less dist_real_def intro!: exI[of _ "x - b"])
6.226 -qed simp
6.227 -
6.228 -end
7.1 --- a/src/HOL/Limits.thy Tue Mar 26 14:38:44 2013 +0100
7.2 +++ b/src/HOL/Limits.thy Tue Mar 26 15:10:28 2013 +0100
7.3 @@ -1,13 +1,23 @@
7.4 -(* Title : Limits.thy
7.5 - Author : Brian Huffman
7.6 +(* Title: Limits.thy
7.7 + Author: Brian Huffman
7.8 + Author: Jacques D. Fleuriot, University of Cambridge
7.9 + Author: Lawrence C Paulson
7.10 + Author: Jeremy Avigad
7.11 +
7.12 *)
7.13
7.14 -header {* Filters and Limits *}
7.15 +header {* Limits on Real Vector Spaces *}
7.16
7.17 theory Limits
7.18 -imports RealVector
7.19 +imports Real_Vector_Spaces
7.20 begin
7.21
7.22 +(* Unfortunately eventually_within was overwritten by Multivariate_Analysis.
7.23 + Hence it was references as Limits.eventually_within, but now it is Basic_Topology.eventually_within *)
7.24 +lemmas eventually_within = eventually_within
7.25 +
7.26 +subsection {* Filter going to infinity norm *}
7.27 +
7.28 definition at_infinity :: "'a::real_normed_vector filter" where
7.29 "at_infinity = Abs_filter (\<lambda>P. \<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x)"
7.30
7.31 @@ -47,7 +57,21 @@
7.32 "at_bot \<le> (at_infinity :: real filter)"
7.33 unfolding at_infinity_eq_at_top_bot by simp
7.34
7.35 -subsection {* Boundedness *}
7.36 +subsubsection {* Boundedness *}
7.37 +
7.38 +definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where
7.39 + Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)"
7.40 +
7.41 +abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
7.42 + "Bseq X \<equiv> Bfun X sequentially"
7.43 +
7.44 +lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" ..
7.45 +
7.46 +lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
7.47 + unfolding Bfun_metric_def by (subst eventually_sequentially_seg)
7.48 +
7.49 +lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
7.50 + unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg)
7.51
7.52 lemma Bfun_def:
7.53 "Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)"
7.54 @@ -77,6 +101,154 @@
7.55 obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F"
7.56 using assms unfolding Bfun_def by fast
7.57
7.58 +lemma Cauchy_Bseq: "Cauchy X \<Longrightarrow> Bseq X"
7.59 + unfolding Cauchy_def Bfun_metric_def eventually_sequentially
7.60 + apply (erule_tac x=1 in allE)
7.61 + apply simp
7.62 + apply safe
7.63 + apply (rule_tac x="X M" in exI)
7.64 + apply (rule_tac x=1 in exI)
7.65 + apply (erule_tac x=M in allE)
7.66 + apply simp
7.67 + apply (rule_tac x=M in exI)
7.68 + apply (auto simp: dist_commute)
7.69 + done
7.70 +
7.71 +
7.72 +subsubsection {* Bounded Sequences *}
7.73 +
7.74 +lemma BseqI': "(\<And>n. norm (X n) \<le> K) \<Longrightarrow> Bseq X"
7.75 + by (intro BfunI) (auto simp: eventually_sequentially)
7.76 +
7.77 +lemma BseqI2': "\<forall>n\<ge>N. norm (X n) \<le> K \<Longrightarrow> Bseq X"
7.78 + by (intro BfunI) (auto simp: eventually_sequentially)
7.79 +
7.80 +lemma Bseq_def: "Bseq X \<longleftrightarrow> (\<exists>K>0. \<forall>n. norm (X n) \<le> K)"
7.81 + unfolding Bfun_def eventually_sequentially
7.82 +proof safe
7.83 + fix N K assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K"
7.84 + then show "\<exists>K>0. \<forall>n. norm (X n) \<le> K"
7.85 + by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] min_max.less_supI2)
7.86 + (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj)
7.87 +qed auto
7.88 +
7.89 +lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
7.90 +unfolding Bseq_def by auto
7.91 +
7.92 +lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
7.93 +by (simp add: Bseq_def)
7.94 +
7.95 +lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
7.96 +by (auto simp add: Bseq_def)
7.97 +
7.98 +lemma lemma_NBseq_def:
7.99 + "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
7.100 +proof safe
7.101 + fix K :: real
7.102 + from reals_Archimedean2 obtain n :: nat where "K < real n" ..
7.103 + then have "K \<le> real (Suc n)" by auto
7.104 + moreover assume "\<forall>m. norm (X m) \<le> K"
7.105 + ultimately have "\<forall>m. norm (X m) \<le> real (Suc n)"
7.106 + by (blast intro: order_trans)
7.107 + then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" ..
7.108 +qed (force simp add: real_of_nat_Suc)
7.109 +
7.110 +text{* alternative definition for Bseq *}
7.111 +lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
7.112 +apply (simp add: Bseq_def)
7.113 +apply (simp (no_asm) add: lemma_NBseq_def)
7.114 +done
7.115 +
7.116 +lemma lemma_NBseq_def2:
7.117 + "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
7.118 +apply (subst lemma_NBseq_def, auto)
7.119 +apply (rule_tac x = "Suc N" in exI)
7.120 +apply (rule_tac [2] x = N in exI)
7.121 +apply (auto simp add: real_of_nat_Suc)
7.122 + prefer 2 apply (blast intro: order_less_imp_le)
7.123 +apply (drule_tac x = n in spec, simp)
7.124 +done
7.125 +
7.126 +(* yet another definition for Bseq *)
7.127 +lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
7.128 +by (simp add: Bseq_def lemma_NBseq_def2)
7.129 +
7.130 +subsubsection{*A Few More Equivalence Theorems for Boundedness*}
7.131 +
7.132 +text{*alternative formulation for boundedness*}
7.133 +lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
7.134 +apply (unfold Bseq_def, safe)
7.135 +apply (rule_tac [2] x = "k + norm x" in exI)
7.136 +apply (rule_tac x = K in exI, simp)
7.137 +apply (rule exI [where x = 0], auto)
7.138 +apply (erule order_less_le_trans, simp)
7.139 +apply (drule_tac x=n in spec, fold diff_minus)
7.140 +apply (drule order_trans [OF norm_triangle_ineq2])
7.141 +apply simp
7.142 +done
7.143 +
7.144 +text{*alternative formulation for boundedness*}
7.145 +lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
7.146 +apply safe
7.147 +apply (simp add: Bseq_def, safe)
7.148 +apply (rule_tac x = "K + norm (X N)" in exI)
7.149 +apply auto
7.150 +apply (erule order_less_le_trans, simp)
7.151 +apply (rule_tac x = N in exI, safe)
7.152 +apply (drule_tac x = n in spec)
7.153 +apply (rule order_trans [OF norm_triangle_ineq], simp)
7.154 +apply (auto simp add: Bseq_iff2)
7.155 +done
7.156 +
7.157 +lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
7.158 +apply (simp add: Bseq_def)
7.159 +apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
7.160 +apply (drule_tac x = n in spec, arith)
7.161 +done
7.162 +
7.163 +subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
7.164 +
7.165 +lemma Bseq_isUb:
7.166 + "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
7.167 +by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
7.168 +
7.169 +text{* Use completeness of reals (supremum property)
7.170 + to show that any bounded sequence has a least upper bound*}
7.171 +
7.172 +lemma Bseq_isLub:
7.173 + "!!(X::nat=>real). Bseq X ==>
7.174 + \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
7.175 +by (blast intro: reals_complete Bseq_isUb)
7.176 +
7.177 +lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X"
7.178 + by (simp add: Bseq_def)
7.179 +
7.180 +lemma Bseq_eq_bounded: "range f \<subseteq> {a .. b::real} \<Longrightarrow> Bseq f"
7.181 + apply (simp add: subset_eq)
7.182 + apply (rule BseqI'[where K="max (norm a) (norm b)"])
7.183 + apply (erule_tac x=n in allE)
7.184 + apply auto
7.185 + done
7.186 +
7.187 +lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> (B::real) \<Longrightarrow> Bseq X"
7.188 + by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def)
7.189 +
7.190 +lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. (B::real) \<le> X i \<Longrightarrow> Bseq X"
7.191 + by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def)
7.192 +
7.193 +subsection {* Bounded Monotonic Sequences *}
7.194 +
7.195 +subsubsection{*A Bounded and Monotonic Sequence Converges*}
7.196 +
7.197 +(* TODO: delete *)
7.198 +(* FIXME: one use in NSA/HSEQ.thy *)
7.199 +lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
7.200 + apply (rule_tac x="X m" in exI)
7.201 + apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const])
7.202 + unfolding eventually_sequentially
7.203 + apply blast
7.204 + done
7.205 +
7.206 subsection {* Convergence to Zero *}
7.207
7.208 definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
7.209 @@ -232,6 +404,37 @@
7.210
7.211 subsubsection {* Distance and norms *}
7.212
7.213 +lemma tendsto_dist [tendsto_intros]:
7.214 + fixes l m :: "'a :: metric_space"
7.215 + assumes f: "(f ---> l) F" and g: "(g ---> m) F"
7.216 + shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) F"
7.217 +proof (rule tendstoI)
7.218 + fix e :: real assume "0 < e"
7.219 + hence e2: "0 < e/2" by simp
7.220 + from tendstoD [OF f e2] tendstoD [OF g e2]
7.221 + show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F"
7.222 + proof (eventually_elim)
7.223 + case (elim x)
7.224 + then show "dist (dist (f x) (g x)) (dist l m) < e"
7.225 + unfolding dist_real_def
7.226 + using dist_triangle2 [of "f x" "g x" "l"]
7.227 + using dist_triangle2 [of "g x" "l" "m"]
7.228 + using dist_triangle3 [of "l" "m" "f x"]
7.229 + using dist_triangle [of "f x" "m" "g x"]
7.230 + by arith
7.231 + qed
7.232 +qed
7.233 +
7.234 +lemma continuous_dist[continuous_intros]:
7.235 + fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
7.236 + shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. dist (f x) (g x))"
7.237 + unfolding continuous_def by (rule tendsto_dist)
7.238 +
7.239 +lemma continuous_on_dist[continuous_on_intros]:
7.240 + fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
7.241 + shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))"
7.242 + unfolding continuous_on_def by (auto intro: tendsto_dist)
7.243 +
7.244 lemma tendsto_norm [tendsto_intros]:
7.245 "(f ---> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> norm a) F"
7.246 unfolding norm_conv_dist by (intro tendsto_intros)
7.247 @@ -695,6 +898,7 @@
7.248 qed
7.249 qed force
7.250
7.251 +
7.252 subsection {* Relate @{const at}, @{const at_left} and @{const at_right} *}
7.253
7.254 text {*
7.255 @@ -1027,9 +1231,724 @@
7.256 qed simp
7.257
7.258
7.259 -(* Unfortunately eventually_within was overwritten by Multivariate_Analysis.
7.260 - Hence it was references as Limits.within, but now it is Basic_Topology.eventually_within *)
7.261 -lemmas eventually_within = eventually_within
7.262 +subsection {* Limits of Sequences *}
7.263 +
7.264 +lemma [trans]: "X=Y ==> Y ----> z ==> X ----> z"
7.265 + by simp
7.266 +
7.267 +lemma LIMSEQ_iff:
7.268 + fixes L :: "'a::real_normed_vector"
7.269 + shows "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
7.270 +unfolding LIMSEQ_def dist_norm ..
7.271 +
7.272 +lemma LIMSEQ_I:
7.273 + fixes L :: "'a::real_normed_vector"
7.274 + shows "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
7.275 +by (simp add: LIMSEQ_iff)
7.276 +
7.277 +lemma LIMSEQ_D:
7.278 + fixes L :: "'a::real_normed_vector"
7.279 + shows "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
7.280 +by (simp add: LIMSEQ_iff)
7.281 +
7.282 +lemma LIMSEQ_linear: "\<lbrakk> X ----> x ; l > 0 \<rbrakk> \<Longrightarrow> (\<lambda> n. X (n * l)) ----> x"
7.283 + unfolding tendsto_def eventually_sequentially
7.284 + by (metis div_le_dividend div_mult_self1_is_m le_trans nat_mult_commute)
7.285 +
7.286 +lemma Bseq_inverse_lemma:
7.287 + fixes x :: "'a::real_normed_div_algebra"
7.288 + shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
7.289 +apply (subst nonzero_norm_inverse, clarsimp)
7.290 +apply (erule (1) le_imp_inverse_le)
7.291 +done
7.292 +
7.293 +lemma Bseq_inverse:
7.294 + fixes a :: "'a::real_normed_div_algebra"
7.295 + shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
7.296 + by (rule Bfun_inverse)
7.297 +
7.298 +lemma LIMSEQ_diff_approach_zero:
7.299 + fixes L :: "'a::real_normed_vector"
7.300 + shows "g ----> L ==> (%x. f x - g x) ----> 0 ==> f ----> L"
7.301 + by (drule (1) tendsto_add, simp)
7.302 +
7.303 +lemma LIMSEQ_diff_approach_zero2:
7.304 + fixes L :: "'a::real_normed_vector"
7.305 + shows "f ----> L ==> (%x. f x - g x) ----> 0 ==> g ----> L"
7.306 + by (drule (1) tendsto_diff, simp)
7.307 +
7.308 +text{*An unbounded sequence's inverse tends to 0*}
7.309 +
7.310 +lemma LIMSEQ_inverse_zero:
7.311 + "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
7.312 + apply (rule filterlim_compose[OF tendsto_inverse_0])
7.313 + apply (simp add: filterlim_at_infinity[OF order_refl] eventually_sequentially)
7.314 + apply (metis abs_le_D1 linorder_le_cases linorder_not_le)
7.315 + done
7.316 +
7.317 +text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
7.318 +
7.319 +lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
7.320 + by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc
7.321 + filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity)
7.322 +
7.323 +text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
7.324 +infinity is now easily proved*}
7.325 +
7.326 +lemma LIMSEQ_inverse_real_of_nat_add:
7.327 + "(%n. r + inverse(real(Suc n))) ----> r"
7.328 + using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto
7.329 +
7.330 +lemma LIMSEQ_inverse_real_of_nat_add_minus:
7.331 + "(%n. r + -inverse(real(Suc n))) ----> r"
7.332 + using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]]
7.333 + by auto
7.334 +
7.335 +lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
7.336 + "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
7.337 + using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]]
7.338 + by auto
7.339 +
7.340 +subsection {* Convergence on sequences *}
7.341 +
7.342 +lemma convergent_add:
7.343 + fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
7.344 + assumes "convergent (\<lambda>n. X n)"
7.345 + assumes "convergent (\<lambda>n. Y n)"
7.346 + shows "convergent (\<lambda>n. X n + Y n)"
7.347 + using assms unfolding convergent_def by (fast intro: tendsto_add)
7.348 +
7.349 +lemma convergent_setsum:
7.350 + fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::real_normed_vector"
7.351 + assumes "\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)"
7.352 + shows "convergent (\<lambda>n. \<Sum>i\<in>A. X i n)"
7.353 +proof (cases "finite A")
7.354 + case True from this and assms show ?thesis
7.355 + by (induct A set: finite) (simp_all add: convergent_const convergent_add)
7.356 +qed (simp add: convergent_const)
7.357 +
7.358 +lemma (in bounded_linear) convergent:
7.359 + assumes "convergent (\<lambda>n. X n)"
7.360 + shows "convergent (\<lambda>n. f (X n))"
7.361 + using assms unfolding convergent_def by (fast intro: tendsto)
7.362 +
7.363 +lemma (in bounded_bilinear) convergent:
7.364 + assumes "convergent (\<lambda>n. X n)" and "convergent (\<lambda>n. Y n)"
7.365 + shows "convergent (\<lambda>n. X n ** Y n)"
7.366 + using assms unfolding convergent_def by (fast intro: tendsto)
7.367 +
7.368 +lemma convergent_minus_iff:
7.369 + fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
7.370 + shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
7.371 +apply (simp add: convergent_def)
7.372 +apply (auto dest: tendsto_minus)
7.373 +apply (drule tendsto_minus, auto)
7.374 +done
7.375 +
7.376 +
7.377 +text {* A monotone sequence converges to its least upper bound. *}
7.378 +
7.379 +lemma isLub_mono_imp_LIMSEQ:
7.380 + fixes X :: "nat \<Rightarrow> real"
7.381 + assumes u: "isLub UNIV {x. \<exists>n. X n = x} u" (* FIXME: use 'range X' *)
7.382 + assumes X: "\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n"
7.383 + shows "X ----> u"
7.384 +proof (rule LIMSEQ_I)
7.385 + have 1: "\<forall>n. X n \<le> u"
7.386 + using isLubD2 [OF u] by auto
7.387 + have "\<forall>y. (\<forall>n. X n \<le> y) \<longrightarrow> u \<le> y"
7.388 + using isLub_le_isUb [OF u] by (auto simp add: isUb_def setle_def)
7.389 + hence 2: "\<forall>y<u. \<exists>n. y < X n"
7.390 + by (metis not_le)
7.391 + fix r :: real assume "0 < r"
7.392 + hence "u - r < u" by simp
7.393 + hence "\<exists>m. u - r < X m" using 2 by simp
7.394 + then obtain m where "u - r < X m" ..
7.395 + with X have "\<forall>n\<ge>m. u - r < X n"
7.396 + by (fast intro: less_le_trans)
7.397 + hence "\<exists>m. \<forall>n\<ge>m. u - r < X n" ..
7.398 + thus "\<exists>m. \<forall>n\<ge>m. norm (X n - u) < r"
7.399 + using 1 by (simp add: diff_less_eq add_commute)
7.400 +qed
7.401 +
7.402 +text{*A standard proof of the theorem for monotone increasing sequence*}
7.403 +
7.404 +lemma Bseq_mono_convergent:
7.405 + "Bseq X \<Longrightarrow> \<forall>m. \<forall>n \<ge> m. X m \<le> X n \<Longrightarrow> convergent (X::nat=>real)"
7.406 + by (metis Bseq_isLub isLub_mono_imp_LIMSEQ convergentI)
7.407 +
7.408 +text{*Main monotonicity theorem*}
7.409 +
7.410 +lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
7.411 + by (metis monoseq_iff incseq_def decseq_eq_incseq convergent_minus_iff Bseq_minus_iff
7.412 + Bseq_mono_convergent)
7.413 +
7.414 +lemma Cauchy_iff:
7.415 + fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
7.416 + shows "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)"
7.417 + unfolding Cauchy_def dist_norm ..
7.418 +
7.419 +lemma CauchyI:
7.420 + fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
7.421 + shows "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
7.422 +by (simp add: Cauchy_iff)
7.423 +
7.424 +lemma CauchyD:
7.425 + fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
7.426 + shows "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
7.427 +by (simp add: Cauchy_iff)
7.428 +
7.429 +lemma incseq_convergent:
7.430 + fixes X :: "nat \<Rightarrow> real"
7.431 + assumes "incseq X" and "\<forall>i. X i \<le> B"
7.432 + obtains L where "X ----> L" "\<forall>i. X i \<le> L"
7.433 +proof atomize_elim
7.434 + from incseq_bounded[OF assms] `incseq X` Bseq_monoseq_convergent[of X]
7.435 + obtain L where "X ----> L"
7.436 + by (auto simp: convergent_def monoseq_def incseq_def)
7.437 + with `incseq X` show "\<exists>L. X ----> L \<and> (\<forall>i. X i \<le> L)"
7.438 + by (auto intro!: exI[of _ L] incseq_le)
7.439 +qed
7.440 +
7.441 +lemma decseq_convergent:
7.442 + fixes X :: "nat \<Rightarrow> real"
7.443 + assumes "decseq X" and "\<forall>i. B \<le> X i"
7.444 + obtains L where "X ----> L" "\<forall>i. L \<le> X i"
7.445 +proof atomize_elim
7.446 + from decseq_bounded[OF assms] `decseq X` Bseq_monoseq_convergent[of X]
7.447 + obtain L where "X ----> L"
7.448 + by (auto simp: convergent_def monoseq_def decseq_def)
7.449 + with `decseq X` show "\<exists>L. X ----> L \<and> (\<forall>i. L \<le> X i)"
7.450 + by (auto intro!: exI[of _ L] decseq_le)
7.451 +qed
7.452 +
7.453 +subsubsection {* Cauchy Sequences are Bounded *}
7.454 +
7.455 +text{*A Cauchy sequence is bounded -- this is the standard
7.456 + proof mechanization rather than the nonstandard proof*}
7.457 +
7.458 +lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
7.459 + ==> \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
7.460 +apply (clarify, drule spec, drule (1) mp)
7.461 +apply (simp only: norm_minus_commute)
7.462 +apply (drule order_le_less_trans [OF norm_triangle_ineq2])
7.463 +apply simp
7.464 +done
7.465 +
7.466 +subsection {* Power Sequences *}
7.467 +
7.468 +text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
7.469 +"x<1"}. Proof will use (NS) Cauchy equivalence for convergence and
7.470 + also fact that bounded and monotonic sequence converges.*}
7.471 +
7.472 +lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
7.473 +apply (simp add: Bseq_def)
7.474 +apply (rule_tac x = 1 in exI)
7.475 +apply (simp add: power_abs)
7.476 +apply (auto dest: power_mono)
7.477 +done
7.478 +
7.479 +lemma monoseq_realpow: fixes x :: real shows "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
7.480 +apply (clarify intro!: mono_SucI2)
7.481 +apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
7.482 +done
7.483 +
7.484 +lemma convergent_realpow:
7.485 + "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
7.486 +by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
7.487 +
7.488 +lemma LIMSEQ_inverse_realpow_zero: "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
7.489 + by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp
7.490 +
7.491 +lemma LIMSEQ_realpow_zero:
7.492 + "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
7.493 +proof cases
7.494 + assume "0 \<le> x" and "x \<noteq> 0"
7.495 + hence x0: "0 < x" by simp
7.496 + assume x1: "x < 1"
7.497 + from x0 x1 have "1 < inverse x"
7.498 + by (rule one_less_inverse)
7.499 + hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
7.500 + by (rule LIMSEQ_inverse_realpow_zero)
7.501 + thus ?thesis by (simp add: power_inverse)
7.502 +qed (rule LIMSEQ_imp_Suc, simp add: tendsto_const)
7.503 +
7.504 +lemma LIMSEQ_power_zero:
7.505 + fixes x :: "'a::{real_normed_algebra_1}"
7.506 + shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
7.507 +apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
7.508 +apply (simp only: tendsto_Zfun_iff, erule Zfun_le)
7.509 +apply (simp add: power_abs norm_power_ineq)
7.510 +done
7.511 +
7.512 +lemma LIMSEQ_divide_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. a / (x ^ n) :: real) ----> 0"
7.513 + by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp
7.514 +
7.515 +text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
7.516 +
7.517 +lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n :: real) ----> 0"
7.518 + by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
7.519 +
7.520 +lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n :: real) ----> 0"
7.521 + by (rule LIMSEQ_power_zero) simp
7.522 +
7.523 +
7.524 +subsection {* Limits of Functions *}
7.525 +
7.526 +lemma LIM_eq:
7.527 + fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
7.528 + shows "f -- a --> L =
7.529 + (\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"
7.530 +by (simp add: LIM_def dist_norm)
7.531 +
7.532 +lemma LIM_I:
7.533 + fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
7.534 + shows "(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)
7.535 + ==> f -- a --> L"
7.536 +by (simp add: LIM_eq)
7.537 +
7.538 +lemma LIM_D:
7.539 + fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
7.540 + shows "[| f -- a --> L; 0<r |]
7.541 + ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"
7.542 +by (simp add: LIM_eq)
7.543 +
7.544 +lemma LIM_offset:
7.545 + fixes a :: "'a::real_normed_vector"
7.546 + shows "f -- a --> L \<Longrightarrow> (\<lambda>x. f (x + k)) -- a - k --> L"
7.547 +apply (rule topological_tendstoI)
7.548 +apply (drule (2) topological_tendstoD)
7.549 +apply (simp only: eventually_at dist_norm)
7.550 +apply (clarify, rule_tac x=d in exI, safe)
7.551 +apply (drule_tac x="x + k" in spec)
7.552 +apply (simp add: algebra_simps)
7.553 +done
7.554 +
7.555 +lemma LIM_offset_zero:
7.556 + fixes a :: "'a::real_normed_vector"
7.557 + shows "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"
7.558 +by (drule_tac k="a" in LIM_offset, simp add: add_commute)
7.559 +
7.560 +lemma LIM_offset_zero_cancel:
7.561 + fixes a :: "'a::real_normed_vector"
7.562 + shows "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> f -- a --> L"
7.563 +by (drule_tac k="- a" in LIM_offset, simp)
7.564 +
7.565 +lemma LIM_zero:
7.566 + fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
7.567 + shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. f x - l) ---> 0) F"
7.568 +unfolding tendsto_iff dist_norm by simp
7.569 +
7.570 +lemma LIM_zero_cancel:
7.571 + fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
7.572 + shows "((\<lambda>x. f x - l) ---> 0) F \<Longrightarrow> (f ---> l) F"
7.573 +unfolding tendsto_iff dist_norm by simp
7.574 +
7.575 +lemma LIM_zero_iff:
7.576 + fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
7.577 + shows "((\<lambda>x. f x - l) ---> 0) F = (f ---> l) F"
7.578 +unfolding tendsto_iff dist_norm by simp
7.579 +
7.580 +lemma LIM_imp_LIM:
7.581 + fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
7.582 + fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector"
7.583 + assumes f: "f -- a --> l"
7.584 + assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"
7.585 + shows "g -- a --> m"
7.586 + by (rule metric_LIM_imp_LIM [OF f],
7.587 + simp add: dist_norm le)
7.588 +
7.589 +lemma LIM_equal2:
7.590 + fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
7.591 + assumes 1: "0 < R"
7.592 + assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"
7.593 + shows "g -- a --> l \<Longrightarrow> f -- a --> l"
7.594 +by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm)
7.595 +
7.596 +lemma LIM_compose2:
7.597 + fixes a :: "'a::real_normed_vector"
7.598 + assumes f: "f -- a --> b"
7.599 + assumes g: "g -- b --> c"
7.600 + assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"
7.601 + shows "(\<lambda>x. g (f x)) -- a --> c"
7.602 +by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])
7.603 +
7.604 +lemma real_LIM_sandwich_zero:
7.605 + fixes f g :: "'a::topological_space \<Rightarrow> real"
7.606 + assumes f: "f -- a --> 0"
7.607 + assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"
7.608 + assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"
7.609 + shows "g -- a --> 0"
7.610 +proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *)
7.611 + fix x assume x: "x \<noteq> a"
7.612 + have "norm (g x - 0) = g x" by (simp add: 1 x)
7.613 + also have "g x \<le> f x" by (rule 2 [OF x])
7.614 + also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self)
7.615 + also have "\<bar>f x\<bar> = norm (f x - 0)" by simp
7.616 + finally show "norm (g x - 0) \<le> norm (f x - 0)" .
7.617 +qed
7.618 +
7.619 +
7.620 +subsection {* Continuity *}
7.621 +
7.622 +lemma LIM_isCont_iff:
7.623 + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
7.624 + shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)"
7.625 +by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])
7.626 +
7.627 +lemma isCont_iff:
7.628 + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
7.629 + shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x"
7.630 +by (simp add: isCont_def LIM_isCont_iff)
7.631 +
7.632 +lemma isCont_LIM_compose2:
7.633 + fixes a :: "'a::real_normed_vector"
7.634 + assumes f [unfolded isCont_def]: "isCont f a"
7.635 + assumes g: "g -- f a --> l"
7.636 + assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"
7.637 + shows "(\<lambda>x. g (f x)) -- a --> l"
7.638 +by (rule LIM_compose2 [OF f g inj])
7.639 +
7.640 +
7.641 +lemma isCont_norm [simp]:
7.642 + fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
7.643 + shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"
7.644 + by (fact continuous_norm)
7.645 +
7.646 +lemma isCont_rabs [simp]:
7.647 + fixes f :: "'a::t2_space \<Rightarrow> real"
7.648 + shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
7.649 + by (fact continuous_rabs)
7.650 +
7.651 +lemma isCont_add [simp]:
7.652 + fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
7.653 + shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"
7.654 + by (fact continuous_add)
7.655 +
7.656 +lemma isCont_minus [simp]:
7.657 + fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
7.658 + shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"
7.659 + by (fact continuous_minus)
7.660 +
7.661 +lemma isCont_diff [simp]:
7.662 + fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
7.663 + shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"
7.664 + by (fact continuous_diff)
7.665 +
7.666 +lemma isCont_mult [simp]:
7.667 + fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
7.668 + shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"
7.669 + by (fact continuous_mult)
7.670 +
7.671 +lemma (in bounded_linear) isCont:
7.672 + "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
7.673 + by (fact continuous)
7.674 +
7.675 +lemma (in bounded_bilinear) isCont:
7.676 + "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"
7.677 + by (fact continuous)
7.678 +
7.679 +lemmas isCont_scaleR [simp] =
7.680 + bounded_bilinear.isCont [OF bounded_bilinear_scaleR]
7.681 +
7.682 +lemmas isCont_of_real [simp] =
7.683 + bounded_linear.isCont [OF bounded_linear_of_real]
7.684 +
7.685 +lemma isCont_power [simp]:
7.686 + fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
7.687 + shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"
7.688 + by (fact continuous_power)
7.689 +
7.690 +lemma isCont_setsum [simp]:
7.691 + fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
7.692 + shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
7.693 + by (auto intro: continuous_setsum)
7.694 +
7.695 +lemmas isCont_intros =
7.696 + isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus
7.697 + isCont_diff isCont_mult isCont_inverse isCont_divide isCont_scaleR
7.698 + isCont_of_real isCont_power isCont_sgn isCont_setsum
7.699 +
7.700 +subsection {* Uniform Continuity *}
7.701 +
7.702 +definition
7.703 + isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where
7.704 + "isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"
7.705 +
7.706 +lemma isUCont_isCont: "isUCont f ==> isCont f x"
7.707 +by (simp add: isUCont_def isCont_def LIM_def, force)
7.708 +
7.709 +lemma isUCont_Cauchy:
7.710 + "\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
7.711 +unfolding isUCont_def
7.712 +apply (rule metric_CauchyI)
7.713 +apply (drule_tac x=e in spec, safe)
7.714 +apply (drule_tac e=s in metric_CauchyD, safe)
7.715 +apply (rule_tac x=M in exI, simp)
7.716 +done
7.717 +
7.718 +lemma (in bounded_linear) isUCont: "isUCont f"
7.719 +unfolding isUCont_def dist_norm
7.720 +proof (intro allI impI)
7.721 + fix r::real assume r: "0 < r"
7.722 + obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"
7.723 + using pos_bounded by fast
7.724 + show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"
7.725 + proof (rule exI, safe)
7.726 + from r K show "0 < r / K" by (rule divide_pos_pos)
7.727 + next
7.728 + fix x y :: 'a
7.729 + assume xy: "norm (x - y) < r / K"
7.730 + have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff)
7.731 + also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le)
7.732 + also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq)
7.733 + finally show "norm (f x - f y) < r" .
7.734 + qed
7.735 +qed
7.736 +
7.737 +lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
7.738 +by (rule isUCont [THEN isUCont_Cauchy])
7.739 +
7.740 +lemma LIM_less_bound:
7.741 + fixes f :: "real \<Rightarrow> real"
7.742 + assumes ev: "b < x" "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and "isCont f x"
7.743 + shows "0 \<le> f x"
7.744 +proof (rule tendsto_le_const)
7.745 + show "(f ---> f x) (at_left x)"
7.746 + using `isCont f x` by (simp add: filterlim_at_split isCont_def)
7.747 + show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)"
7.748 + using ev by (auto simp: eventually_within_less dist_real_def intro!: exI[of _ "x - b"])
7.749 +qed simp
7.750 +
7.751 +
7.752 +subsection {* Nested Intervals and Bisection -- Needed for Compactness *}
7.753 +
7.754 +lemma nested_sequence_unique:
7.755 + assumes "\<forall>n. f n \<le> f (Suc n)" "\<forall>n. g (Suc n) \<le> g n" "\<forall>n. f n \<le> g n" "(\<lambda>n. f n - g n) ----> 0"
7.756 + shows "\<exists>l::real. ((\<forall>n. f n \<le> l) \<and> f ----> l) \<and> ((\<forall>n. l \<le> g n) \<and> g ----> l)"
7.757 +proof -
7.758 + have "incseq f" unfolding incseq_Suc_iff by fact
7.759 + have "decseq g" unfolding decseq_Suc_iff by fact
7.760 +
7.761 + { fix n
7.762 + from `decseq g` have "g n \<le> g 0" by (rule decseqD) simp
7.763 + with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f n \<le> g 0" by auto }
7.764 + then obtain u where "f ----> u" "\<forall>i. f i \<le> u"
7.765 + using incseq_convergent[OF `incseq f`] by auto
7.766 + moreover
7.767 + { fix n
7.768 + from `incseq f` have "f 0 \<le> f n" by (rule incseqD) simp
7.769 + with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f 0 \<le> g n" by simp }
7.770 + then obtain l where "g ----> l" "\<forall>i. l \<le> g i"
7.771 + using decseq_convergent[OF `decseq g`] by auto
7.772 + moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF `f ----> u` `g ----> l`]]
7.773 + ultimately show ?thesis by auto
7.774 +qed
7.775 +
7.776 +lemma Bolzano[consumes 1, case_names trans local]:
7.777 + fixes P :: "real \<Rightarrow> real \<Rightarrow> bool"
7.778 + assumes [arith]: "a \<le> b"
7.779 + assumes trans: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a c"
7.780 + assumes local: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<exists>d>0. \<forall>a b. a \<le> x \<and> x \<le> b \<and> b - a < d \<longrightarrow> P a b"
7.781 + shows "P a b"
7.782 +proof -
7.783 + def bisect \<equiv> "nat_rec (a, b) (\<lambda>n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))"
7.784 + def l \<equiv> "\<lambda>n. fst (bisect n)" and u \<equiv> "\<lambda>n. snd (bisect n)"
7.785 + have l[simp]: "l 0 = a" "\<And>n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)"
7.786 + and u[simp]: "u 0 = b" "\<And>n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)"
7.787 + by (simp_all add: l_def u_def bisect_def split: prod.split)
7.788 +
7.789 + { fix n have "l n \<le> u n" by (induct n) auto } note this[simp]
7.790 +
7.791 + have "\<exists>x. ((\<forall>n. l n \<le> x) \<and> l ----> x) \<and> ((\<forall>n. x \<le> u n) \<and> u ----> x)"
7.792 + proof (safe intro!: nested_sequence_unique)
7.793 + fix n show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" by (induct n) auto
7.794 + next
7.795 + { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) }
7.796 + then show "(\<lambda>n. l n - u n) ----> 0" by (simp add: LIMSEQ_divide_realpow_zero)
7.797 + qed fact
7.798 + then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l ----> x" "u ----> x" by auto
7.799 + obtain d where "0 < d" and d: "\<And>a b. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> b - a < d \<Longrightarrow> P a b"
7.800 + using `l 0 \<le> x` `x \<le> u 0` local[of x] by auto
7.801 +
7.802 + show "P a b"
7.803 + proof (rule ccontr)
7.804 + assume "\<not> P a b"
7.805 + { fix n have "\<not> P (l n) (u n)"
7.806 + proof (induct n)
7.807 + case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto
7.808 + qed (simp add: `\<not> P a b`) }
7.809 + moreover
7.810 + { have "eventually (\<lambda>n. x - d / 2 < l n) sequentially"
7.811 + using `0 < d` `l ----> x` by (intro order_tendstoD[of _ x]) auto
7.812 + moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially"
7.813 + using `0 < d` `u ----> x` by (intro order_tendstoD[of _ x]) auto
7.814 + ultimately have "eventually (\<lambda>n. P (l n) (u n)) sequentially"
7.815 + proof eventually_elim
7.816 + fix n assume "x - d / 2 < l n" "u n < x + d / 2"
7.817 + from add_strict_mono[OF this] have "u n - l n < d" by simp
7.818 + with x show "P (l n) (u n)" by (rule d)
7.819 + qed }
7.820 + ultimately show False by simp
7.821 + qed
7.822 +qed
7.823 +
7.824 +lemma compact_Icc[simp, intro]: "compact {a .. b::real}"
7.825 +proof (cases "a \<le> b", rule compactI)
7.826 + fix C assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C"
7.827 + def T == "{a .. b}"
7.828 + from C(1,3) show "\<exists>C'\<subseteq>C. finite C' \<and> {a..b} \<subseteq> \<Union>C'"
7.829 + proof (induct rule: Bolzano)
7.830 + case (trans a b c)
7.831 + then have *: "{a .. c} = {a .. b} \<union> {b .. c}" by auto
7.832 + from trans obtain C1 C2 where "C1\<subseteq>C \<and> finite C1 \<and> {a..b} \<subseteq> \<Union>C1" "C2\<subseteq>C \<and> finite C2 \<and> {b..c} \<subseteq> \<Union>C2"
7.833 + by (auto simp: *)
7.834 + with trans show ?case
7.835 + unfolding * by (intro exI[of _ "C1 \<union> C2"]) auto
7.836 + next
7.837 + case (local x)
7.838 + then have "x \<in> \<Union>C" using C by auto
7.839 + with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" by auto
7.840 + then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c"
7.841 + by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff)
7.842 + with `c \<in> C` show ?case
7.843 + by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto
7.844 + qed
7.845 +qed simp
7.846 +
7.847 +
7.848 +subsection {* Boundedness of continuous functions *}
7.849 +
7.850 +text{*By bisection, function continuous on closed interval is bounded above*}
7.851 +
7.852 +lemma isCont_eq_Ub:
7.853 + fixes f :: "real \<Rightarrow> 'a::linorder_topology"
7.854 + shows "a \<le> b \<Longrightarrow> \<forall>x::real. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
7.855 + \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
7.856 + using continuous_attains_sup[of "{a .. b}" f]
7.857 + by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
7.858 +
7.859 +lemma isCont_eq_Lb:
7.860 + fixes f :: "real \<Rightarrow> 'a::linorder_topology"
7.861 + shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
7.862 + \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> M \<le> f x) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
7.863 + using continuous_attains_inf[of "{a .. b}" f]
7.864 + by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
7.865 +
7.866 +lemma isCont_bounded:
7.867 + fixes f :: "real \<Rightarrow> 'a::linorder_topology"
7.868 + shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> \<exists>M. \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M"
7.869 + using isCont_eq_Ub[of a b f] by auto
7.870 +
7.871 +lemma isCont_has_Ub:
7.872 + fixes f :: "real \<Rightarrow> 'a::linorder_topology"
7.873 + shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
7.874 + \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M) \<and> (\<forall>N. N < M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> N < f x))"
7.875 + using isCont_eq_Ub[of a b f] by auto
7.876 +
7.877 +(*HOL style here: object-level formulations*)
7.878 +lemma IVT_objl: "(f(a::real) \<le> (y::real) & y \<le> f(b) & a \<le> b &
7.879 + (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
7.880 + --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
7.881 + by (blast intro: IVT)
7.882 +
7.883 +lemma IVT2_objl: "(f(b::real) \<le> (y::real) & y \<le> f(a) & a \<le> b &
7.884 + (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
7.885 + --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
7.886 + by (blast intro: IVT2)
7.887 +
7.888 +lemma isCont_Lb_Ub:
7.889 + fixes f :: "real \<Rightarrow> real"
7.890 + assumes "a \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x"
7.891 + shows "\<exists>L M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> L \<le> f x \<and> f x \<le> M) \<and>
7.892 + (\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> (f x = y)))"
7.893 +proof -
7.894 + obtain M where M: "a \<le> M" "M \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> f M"
7.895 + using isCont_eq_Ub[OF assms] by auto
7.896 + obtain L where L: "a \<le> L" "L \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f L \<le> f x"
7.897 + using isCont_eq_Lb[OF assms] by auto
7.898 + show ?thesis
7.899 + using IVT[of f L _ M] IVT2[of f L _ M] M L assms
7.900 + apply (rule_tac x="f L" in exI)
7.901 + apply (rule_tac x="f M" in exI)
7.902 + apply (cases "L \<le> M")
7.903 + apply (simp, metis order_trans)
7.904 + apply (simp, metis order_trans)
7.905 + done
7.906 +qed
7.907 +
7.908 +
7.909 +text{*Continuity of inverse function*}
7.910 +
7.911 +lemma isCont_inverse_function:
7.912 + fixes f g :: "real \<Rightarrow> real"
7.913 + assumes d: "0 < d"
7.914 + and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> g (f z) = z"
7.915 + and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> isCont f z"
7.916 + shows "isCont g (f x)"
7.917 +proof -
7.918 + let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + d}"
7.919 +
7.920 + have f: "continuous_on ?D f"
7.921 + using cont by (intro continuous_at_imp_continuous_on ballI) auto
7.922 + then have g: "continuous_on (f`?D) g"
7.923 + using inj by (intro continuous_on_inv) auto
7.924 +
7.925 + from d f have "{min ?A ?B <..< max ?A ?B} \<subseteq> f ` ?D"
7.926 + by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max)
7.927 + with g have "continuous_on {min ?A ?B <..< max ?A ?B} g"
7.928 + by (rule continuous_on_subset)
7.929 + moreover
7.930 + have "(?A < f x \<and> f x < ?B) \<or> (?B < f x \<and> f x < ?A)"
7.931 + using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto
7.932 + then have "f x \<in> {min ?A ?B <..< max ?A ?B}"
7.933 + by auto
7.934 + ultimately
7.935 + show ?thesis
7.936 + by (simp add: continuous_on_eq_continuous_at)
7.937 +qed
7.938 +
7.939 +lemma isCont_inverse_function2:
7.940 + fixes f g :: "real \<Rightarrow> real" shows
7.941 + "\<lbrakk>a < x; x < b;
7.942 + \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z;
7.943 + \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z\<rbrakk>
7.944 + \<Longrightarrow> isCont g (f x)"
7.945 +apply (rule isCont_inverse_function
7.946 + [where f=f and d="min (x - a) (b - x)"])
7.947 +apply (simp_all add: abs_le_iff)
7.948 +done
7.949 +
7.950 +(* need to rename second isCont_inverse *)
7.951 +
7.952 +lemma isCont_inv_fun:
7.953 + fixes f g :: "real \<Rightarrow> real"
7.954 + shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z;
7.955 + \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |]
7.956 + ==> isCont g (f x)"
7.957 +by (rule isCont_inverse_function)
7.958 +
7.959 +text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*}
7.960 +lemma LIM_fun_gt_zero:
7.961 + fixes f :: "real \<Rightarrow> real"
7.962 + shows "f -- c --> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> 0 < f x)"
7.963 +apply (drule (1) LIM_D, clarify)
7.964 +apply (rule_tac x = s in exI)
7.965 +apply (simp add: abs_less_iff)
7.966 +done
7.967 +
7.968 +lemma LIM_fun_less_zero:
7.969 + fixes f :: "real \<Rightarrow> real"
7.970 + shows "f -- c --> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x < 0)"
7.971 +apply (drule LIM_D [where r="-l"], simp, clarify)
7.972 +apply (rule_tac x = s in exI)
7.973 +apply (simp add: abs_less_iff)
7.974 +done
7.975 +
7.976 +lemma LIM_fun_not_zero:
7.977 + fixes f :: "real \<Rightarrow> real"
7.978 + shows "f -- c --> l \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x \<noteq> 0)"
7.979 + using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp add: neq_iff)
7.980
7.981 end
7.982
8.1 --- a/src/HOL/Ln.thy Tue Mar 26 14:38:44 2013 +0100
8.2 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000
8.3 @@ -1,229 +0,0 @@
8.4 -(* Title: HOL/Ln.thy
8.5 - Author: Jeremy Avigad
8.6 -*)
8.7 -
8.8 -header {* Properties of ln *}
8.9 -
8.10 -theory Ln
8.11 -imports Transcendental
8.12 -begin
8.13 -
8.14 -lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==>
8.15 - x - x^2 <= ln (1 + x)"
8.16 -proof -
8.17 - assume a: "0 <= x" and b: "x <= 1"
8.18 - have "exp (x - x^2) = exp x / exp (x^2)"
8.19 - by (rule exp_diff)
8.20 - also have "... <= (1 + x + x^2) / exp (x ^2)"
8.21 - apply (rule divide_right_mono)
8.22 - apply (rule exp_bound)
8.23 - apply (rule a, rule b)
8.24 - apply simp
8.25 - done
8.26 - also have "... <= (1 + x + x^2) / (1 + x^2)"
8.27 - apply (rule divide_left_mono)
8.28 - apply (simp add: exp_ge_add_one_self_aux)
8.29 - apply (simp add: a)
8.30 - apply (simp add: mult_pos_pos add_pos_nonneg)
8.31 - done
8.32 - also from a have "... <= 1 + x"
8.33 - by (simp add: field_simps add_strict_increasing zero_le_mult_iff)
8.34 - finally have "exp (x - x^2) <= 1 + x" .
8.35 - also have "... = exp (ln (1 + x))"
8.36 - proof -
8.37 - from a have "0 < 1 + x" by auto
8.38 - thus ?thesis
8.39 - by (auto simp only: exp_ln_iff [THEN sym])
8.40 - qed
8.41 - finally have "exp (x - x ^ 2) <= exp (ln (1 + x))" .
8.42 - thus ?thesis by (auto simp only: exp_le_cancel_iff)
8.43 -qed
8.44 -
8.45 -lemma aux5: "x < 1 ==> ln(1 - x) = - ln(1 + x / (1 - x))"
8.46 -proof -
8.47 - assume a: "x < 1"
8.48 - have "ln(1 - x) = - ln(1 / (1 - x))"
8.49 - proof -
8.50 - have "ln(1 - x) = - (- ln (1 - x))"
8.51 - by auto
8.52 - also have "- ln(1 - x) = ln 1 - ln(1 - x)"
8.53 - by simp
8.54 - also have "... = ln(1 / (1 - x))"
8.55 - apply (rule ln_div [THEN sym])
8.56 - by (insert a, auto)
8.57 - finally show ?thesis .
8.58 - qed
8.59 - also have " 1 / (1 - x) = 1 + x / (1 - x)" using a by(simp add:field_simps)
8.60 - finally show ?thesis .
8.61 -qed
8.62 -
8.63 -lemma ln_one_minus_pos_lower_bound: "0 <= x ==> x <= (1 / 2) ==>
8.64 - - x - 2 * x^2 <= ln (1 - x)"
8.65 -proof -
8.66 - assume a: "0 <= x" and b: "x <= (1 / 2)"
8.67 - from b have c: "x < 1"
8.68 - by auto
8.69 - then have "ln (1 - x) = - ln (1 + x / (1 - x))"
8.70 - by (rule aux5)
8.71 - also have "- (x / (1 - x)) <= ..."
8.72 - proof -
8.73 - have "ln (1 + x / (1 - x)) <= x / (1 - x)"
8.74 - apply (rule ln_add_one_self_le_self)
8.75 - apply (rule divide_nonneg_pos)
8.76 - by (insert a c, auto)
8.77 - thus ?thesis
8.78 - by auto
8.79 - qed
8.80 - also have "- (x / (1 - x)) = -x / (1 - x)"
8.81 - by auto
8.82 - finally have d: "- x / (1 - x) <= ln (1 - x)" .
8.83 - have "0 < 1 - x" using a b by simp
8.84 - hence e: "-x - 2 * x^2 <= - x / (1 - x)"
8.85 - using mult_right_le_one_le[of "x*x" "2*x"] a b
8.86 - by (simp add:field_simps power2_eq_square)
8.87 - from e d show "- x - 2 * x^2 <= ln (1 - x)"
8.88 - by (rule order_trans)
8.89 -qed
8.90 -
8.91 -lemma ln_add_one_self_le_self2: "-1 < x ==> ln(1 + x) <= x"
8.92 - apply (subgoal_tac "ln (1 + x) \<le> ln (exp x)", simp)
8.93 - apply (subst ln_le_cancel_iff)
8.94 - apply auto
8.95 -done
8.96 -
8.97 -lemma abs_ln_one_plus_x_minus_x_bound_nonneg:
8.98 - "0 <= x ==> x <= 1 ==> abs(ln (1 + x) - x) <= x^2"
8.99 -proof -
8.100 - assume x: "0 <= x"
8.101 - assume x1: "x <= 1"
8.102 - from x have "ln (1 + x) <= x"
8.103 - by (rule ln_add_one_self_le_self)
8.104 - then have "ln (1 + x) - x <= 0"
8.105 - by simp
8.106 - then have "abs(ln(1 + x) - x) = - (ln(1 + x) - x)"
8.107 - by (rule abs_of_nonpos)
8.108 - also have "... = x - ln (1 + x)"
8.109 - by simp
8.110 - also have "... <= x^2"
8.111 - proof -
8.112 - from x x1 have "x - x^2 <= ln (1 + x)"
8.113 - by (intro ln_one_plus_pos_lower_bound)
8.114 - thus ?thesis
8.115 - by simp
8.116 - qed
8.117 - finally show ?thesis .
8.118 -qed
8.119 -
8.120 -lemma abs_ln_one_plus_x_minus_x_bound_nonpos:
8.121 - "-(1 / 2) <= x ==> x <= 0 ==> abs(ln (1 + x) - x) <= 2 * x^2"
8.122 -proof -
8.123 - assume a: "-(1 / 2) <= x"
8.124 - assume b: "x <= 0"
8.125 - have "abs(ln (1 + x) - x) = x - ln(1 - (-x))"
8.126 - apply (subst abs_of_nonpos)
8.127 - apply simp
8.128 - apply (rule ln_add_one_self_le_self2)
8.129 - using a apply auto
8.130 - done
8.131 - also have "... <= 2 * x^2"
8.132 - apply (subgoal_tac "- (-x) - 2 * (-x)^2 <= ln (1 - (-x))")
8.133 - apply (simp add: algebra_simps)
8.134 - apply (rule ln_one_minus_pos_lower_bound)
8.135 - using a b apply auto
8.136 - done
8.137 - finally show ?thesis .
8.138 -qed
8.139 -
8.140 -lemma abs_ln_one_plus_x_minus_x_bound:
8.141 - "abs x <= 1 / 2 ==> abs(ln (1 + x) - x) <= 2 * x^2"
8.142 - apply (case_tac "0 <= x")
8.143 - apply (rule order_trans)
8.144 - apply (rule abs_ln_one_plus_x_minus_x_bound_nonneg)
8.145 - apply auto
8.146 - apply (rule abs_ln_one_plus_x_minus_x_bound_nonpos)
8.147 - apply auto
8.148 -done
8.149 -
8.150 -lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)"
8.151 -proof -
8.152 - assume x: "exp 1 <= x" "x <= y"
8.153 - moreover have "0 < exp (1::real)" by simp
8.154 - ultimately have a: "0 < x" and b: "0 < y"
8.155 - by (fast intro: less_le_trans order_trans)+
8.156 - have "x * ln y - x * ln x = x * (ln y - ln x)"
8.157 - by (simp add: algebra_simps)
8.158 - also have "... = x * ln(y / x)"
8.159 - by (simp only: ln_div a b)
8.160 - also have "y / x = (x + (y - x)) / x"
8.161 - by simp
8.162 - also have "... = 1 + (y - x) / x"
8.163 - using x a by (simp add: field_simps)
8.164 - also have "x * ln(1 + (y - x) / x) <= x * ((y - x) / x)"
8.165 - apply (rule mult_left_mono)
8.166 - apply (rule ln_add_one_self_le_self)
8.167 - apply (rule divide_nonneg_pos)
8.168 - using x a apply simp_all
8.169 - done
8.170 - also have "... = y - x" using a by simp
8.171 - also have "... = (y - x) * ln (exp 1)" by simp
8.172 - also have "... <= (y - x) * ln x"
8.173 - apply (rule mult_left_mono)
8.174 - apply (subst ln_le_cancel_iff)
8.175 - apply fact
8.176 - apply (rule a)
8.177 - apply (rule x)
8.178 - using x apply simp
8.179 - done
8.180 - also have "... = y * ln x - x * ln x"
8.181 - by (rule left_diff_distrib)
8.182 - finally have "x * ln y <= y * ln x"
8.183 - by arith
8.184 - then have "ln y <= (y * ln x) / x" using a by (simp add: field_simps)
8.185 - also have "... = y * (ln x / x)" by simp
8.186 - finally show ?thesis using b by (simp add: field_simps)
8.187 -qed
8.188 -
8.189 -lemma ln_le_minus_one:
8.190 - "0 < x \<Longrightarrow> ln x \<le> x - 1"
8.191 - using exp_ge_add_one_self[of "ln x"] by simp
8.192 -
8.193 -lemma ln_eq_minus_one:
8.194 - assumes "0 < x" "ln x = x - 1" shows "x = 1"
8.195 -proof -
8.196 - let "?l y" = "ln y - y + 1"
8.197 - have D: "\<And>x. 0 < x \<Longrightarrow> DERIV ?l x :> (1 / x - 1)"
8.198 - by (auto intro!: DERIV_intros)
8.199 -
8.200 - show ?thesis
8.201 - proof (cases rule: linorder_cases)
8.202 - assume "x < 1"
8.203 - from dense[OF `x < 1`] obtain a where "x < a" "a < 1" by blast
8.204 - from `x < a` have "?l x < ?l a"
8.205 - proof (rule DERIV_pos_imp_increasing, safe)
8.206 - fix y assume "x \<le> y" "y \<le> a"
8.207 - with `0 < x` `a < 1` have "0 < 1 / y - 1" "0 < y"
8.208 - by (auto simp: field_simps)
8.209 - with D show "\<exists>z. DERIV ?l y :> z \<and> 0 < z"
8.210 - by auto
8.211 - qed
8.212 - also have "\<dots> \<le> 0"
8.213 - using ln_le_minus_one `0 < x` `x < a` by (auto simp: field_simps)
8.214 - finally show "x = 1" using assms by auto
8.215 - next
8.216 - assume "1 < x"
8.217 - from dense[OF `1 < x`] obtain a where "1 < a" "a < x" by blast
8.218 - from `a < x` have "?l x < ?l a"
8.219 - proof (rule DERIV_neg_imp_decreasing, safe)
8.220 - fix y assume "a \<le> y" "y \<le> x"
8.221 - with `1 < a` have "1 / y - 1 < 0" "0 < y"
8.222 - by (auto simp: field_simps)
8.223 - with D show "\<exists>z. DERIV ?l y :> z \<and> z < 0"
8.224 - by blast
8.225 - qed
8.226 - also have "\<dots> \<le> 0"
8.227 - using ln_le_minus_one `1 < a` by (auto simp: field_simps)
8.228 - finally show "x = 1" using assms by auto
8.229 - qed simp
8.230 -qed
8.231 -
8.232 -end
9.1 --- a/src/HOL/Log.thy Tue Mar 26 14:38:44 2013 +0100
9.2 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000
9.3 @@ -1,421 +0,0 @@
9.4 -(* Title : Log.thy
9.5 - Author : Jacques D. Fleuriot
9.6 - Additional contributions by Jeremy Avigad
9.7 - Copyright : 2000,2001 University of Edinburgh
9.8 -*)
9.9 -
9.10 -header{*Logarithms: Standard Version*}
9.11 -
9.12 -theory Log
9.13 -imports Transcendental
9.14 -begin
9.15 -
9.16 -definition
9.17 - powr :: "[real,real] => real" (infixr "powr" 80) where
9.18 - --{*exponentation with real exponent*}
9.19 - "x powr a = exp(a * ln x)"
9.20 -
9.21 -definition
9.22 - log :: "[real,real] => real" where
9.23 - --{*logarithm of @{term x} to base @{term a}*}
9.24 - "log a x = ln x / ln a"
9.25 -
9.26 -
9.27 -lemma tendsto_log [tendsto_intros]:
9.28 - "\<lbrakk>(f ---> a) F; (g ---> b) F; 0 < a; a \<noteq> 1; 0 < b\<rbrakk> \<Longrightarrow> ((\<lambda>x. log (f x) (g x)) ---> log a b) F"
9.29 - unfolding log_def by (intro tendsto_intros) auto
9.30 -
9.31 -lemma continuous_log:
9.32 - assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\<lambda>x. x))" and "f (Lim F (\<lambda>x. x)) \<noteq> 1" and "0 < g (Lim F (\<lambda>x. x))"
9.33 - shows "continuous F (\<lambda>x. log (f x) (g x))"
9.34 - using assms unfolding continuous_def by (rule tendsto_log)
9.35 -
9.36 -lemma continuous_at_within_log[continuous_intros]:
9.37 - assumes "continuous (at a within s) f" "continuous (at a within s) g" and "0 < f a" and "f a \<noteq> 1" and "0 < g a"
9.38 - shows "continuous (at a within s) (\<lambda>x. log (f x) (g x))"
9.39 - using assms unfolding continuous_within by (rule tendsto_log)
9.40 -
9.41 -lemma isCont_log[continuous_intros, simp]:
9.42 - assumes "isCont f a" "isCont g a" "0 < f a" "f a \<noteq> 1" "0 < g a"
9.43 - shows "isCont (\<lambda>x. log (f x) (g x)) a"
9.44 - using assms unfolding continuous_at by (rule tendsto_log)
9.45 -
9.46 -lemma continuous_on_log[continuous_on_intros]:
9.47 - assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. 0 < f x" "\<forall>x\<in>s. f x \<noteq> 1" "\<forall>x\<in>s. 0 < g x"
9.48 - shows "continuous_on s (\<lambda>x. log (f x) (g x))"
9.49 - using assms unfolding continuous_on_def by (fast intro: tendsto_log)
9.50 -
9.51 -lemma powr_one_eq_one [simp]: "1 powr a = 1"
9.52 -by (simp add: powr_def)
9.53 -
9.54 -lemma powr_zero_eq_one [simp]: "x powr 0 = 1"
9.55 -by (simp add: powr_def)
9.56 -
9.57 -lemma powr_one_gt_zero_iff [simp]: "(x powr 1 = x) = (0 < x)"
9.58 -by (simp add: powr_def)
9.59 -declare powr_one_gt_zero_iff [THEN iffD2, simp]
9.60 -
9.61 -lemma powr_mult:
9.62 - "[| 0 < x; 0 < y |] ==> (x * y) powr a = (x powr a) * (y powr a)"
9.63 -by (simp add: powr_def exp_add [symmetric] ln_mult distrib_left)
9.64 -
9.65 -lemma powr_gt_zero [simp]: "0 < x powr a"
9.66 -by (simp add: powr_def)
9.67 -
9.68 -lemma powr_ge_pzero [simp]: "0 <= x powr y"
9.69 -by (rule order_less_imp_le, rule powr_gt_zero)
9.70 -
9.71 -lemma powr_not_zero [simp]: "x powr a \<noteq> 0"
9.72 -by (simp add: powr_def)
9.73 -
9.74 -lemma powr_divide:
9.75 - "[| 0 < x; 0 < y |] ==> (x / y) powr a = (x powr a)/(y powr a)"
9.76 -apply (simp add: divide_inverse positive_imp_inverse_positive powr_mult)
9.77 -apply (simp add: powr_def exp_minus [symmetric] exp_add [symmetric] ln_inverse)
9.78 -done
9.79 -
9.80 -lemma powr_divide2: "x powr a / x powr b = x powr (a - b)"
9.81 - apply (simp add: powr_def)
9.82 - apply (subst exp_diff [THEN sym])
9.83 - apply (simp add: left_diff_distrib)
9.84 -done
9.85 -
9.86 -lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)"
9.87 -by (simp add: powr_def exp_add [symmetric] distrib_right)
9.88 -
9.89 -lemma powr_mult_base:
9.90 - "0 < x \<Longrightarrow>x * x powr y = x powr (1 + y)"
9.91 -using assms by (auto simp: powr_add)
9.92 -
9.93 -lemma powr_powr: "(x powr a) powr b = x powr (a * b)"
9.94 -by (simp add: powr_def)
9.95 -
9.96 -lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a"
9.97 -by (simp add: powr_powr mult_commute)
9.98 -
9.99 -lemma powr_minus: "x powr (-a) = inverse (x powr a)"
9.100 -by (simp add: powr_def exp_minus [symmetric])
9.101 -
9.102 -lemma powr_minus_divide: "x powr (-a) = 1/(x powr a)"
9.103 -by (simp add: divide_inverse powr_minus)
9.104 -
9.105 -lemma powr_less_mono: "[| a < b; 1 < x |] ==> x powr a < x powr b"
9.106 -by (simp add: powr_def)
9.107 -
9.108 -lemma powr_less_cancel: "[| x powr a < x powr b; 1 < x |] ==> a < b"
9.109 -by (simp add: powr_def)
9.110 -
9.111 -lemma powr_less_cancel_iff [simp]: "1 < x ==> (x powr a < x powr b) = (a < b)"
9.112 -by (blast intro: powr_less_cancel powr_less_mono)
9.113 -
9.114 -lemma powr_le_cancel_iff [simp]: "1 < x ==> (x powr a \<le> x powr b) = (a \<le> b)"
9.115 -by (simp add: linorder_not_less [symmetric])
9.116 -
9.117 -lemma log_ln: "ln x = log (exp(1)) x"
9.118 -by (simp add: log_def)
9.119 -
9.120 -lemma DERIV_log: assumes "x > 0" shows "DERIV (\<lambda>y. log b y) x :> 1 / (ln b * x)"
9.121 -proof -
9.122 - def lb \<equiv> "1 / ln b"
9.123 - moreover have "DERIV (\<lambda>y. lb * ln y) x :> lb / x"
9.124 - using `x > 0` by (auto intro!: DERIV_intros)
9.125 - ultimately show ?thesis
9.126 - by (simp add: log_def)
9.127 -qed
9.128 -
9.129 -lemmas DERIV_log[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
9.130 -
9.131 -lemma powr_log_cancel [simp]:
9.132 - "[| 0 < a; a \<noteq> 1; 0 < x |] ==> a powr (log a x) = x"
9.133 -by (simp add: powr_def log_def)
9.134 -
9.135 -lemma log_powr_cancel [simp]: "[| 0 < a; a \<noteq> 1 |] ==> log a (a powr y) = y"
9.136 -by (simp add: log_def powr_def)
9.137 -
9.138 -lemma log_mult:
9.139 - "[| 0 < a; a \<noteq> 1; 0 < x; 0 < y |]
9.140 - ==> log a (x * y) = log a x + log a y"
9.141 -by (simp add: log_def ln_mult divide_inverse distrib_right)
9.142 -
9.143 -lemma log_eq_div_ln_mult_log:
9.144 - "[| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |]
9.145 - ==> log a x = (ln b/ln a) * log b x"
9.146 -by (simp add: log_def divide_inverse)
9.147 -
9.148 -text{*Base 10 logarithms*}
9.149 -lemma log_base_10_eq1: "0 < x ==> log 10 x = (ln (exp 1) / ln 10) * ln x"
9.150 -by (simp add: log_def)
9.151 -
9.152 -lemma log_base_10_eq2: "0 < x ==> log 10 x = (log 10 (exp 1)) * ln x"
9.153 -by (simp add: log_def)
9.154 -
9.155 -lemma log_one [simp]: "log a 1 = 0"
9.156 -by (simp add: log_def)
9.157 -
9.158 -lemma log_eq_one [simp]: "[| 0 < a; a \<noteq> 1 |] ==> log a a = 1"
9.159 -by (simp add: log_def)
9.160 -
9.161 -lemma log_inverse:
9.162 - "[| 0 < a; a \<noteq> 1; 0 < x |] ==> log a (inverse x) = - log a x"
9.163 -apply (rule_tac a1 = "log a x" in add_left_cancel [THEN iffD1])
9.164 -apply (simp add: log_mult [symmetric])
9.165 -done
9.166 -
9.167 -lemma log_divide:
9.168 - "[|0 < a; a \<noteq> 1; 0 < x; 0 < y|] ==> log a (x/y) = log a x - log a y"
9.169 -by (simp add: log_mult divide_inverse log_inverse)
9.170 -
9.171 -lemma log_less_cancel_iff [simp]:
9.172 - "[| 1 < a; 0 < x; 0 < y |] ==> (log a x < log a y) = (x < y)"
9.173 -apply safe
9.174 -apply (rule_tac [2] powr_less_cancel)
9.175 -apply (drule_tac a = "log a x" in powr_less_mono, auto)
9.176 -done
9.177 -
9.178 -lemma log_inj: assumes "1 < b" shows "inj_on (log b) {0 <..}"
9.179 -proof (rule inj_onI, simp)
9.180 - fix x y assume pos: "0 < x" "0 < y" and *: "log b x = log b y"
9.181 - show "x = y"
9.182 - proof (cases rule: linorder_cases)
9.183 - assume "x < y" hence "log b x < log b y"
9.184 - using log_less_cancel_iff[OF `1 < b`] pos by simp
9.185 - thus ?thesis using * by simp
9.186 - next
9.187 - assume "y < x" hence "log b y < log b x"
9.188 - using log_less_cancel_iff[OF `1 < b`] pos by simp
9.189 - thus ?thesis using * by simp
9.190 - qed simp
9.191 -qed
9.192 -
9.193 -lemma log_le_cancel_iff [simp]:
9.194 - "[| 1 < a; 0 < x; 0 < y |] ==> (log a x \<le> log a y) = (x \<le> y)"
9.195 -by (simp add: linorder_not_less [symmetric])
9.196 -
9.197 -lemma zero_less_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 < log a x \<longleftrightarrow> 1 < x"
9.198 - using log_less_cancel_iff[of a 1 x] by simp
9.199 -
9.200 -lemma zero_le_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 \<le> log a x \<longleftrightarrow> 1 \<le> x"
9.201 - using log_le_cancel_iff[of a 1 x] by simp
9.202 -
9.203 -lemma log_less_zero_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x < 0 \<longleftrightarrow> x < 1"
9.204 - using log_less_cancel_iff[of a x 1] by simp
9.205 -
9.206 -lemma log_le_zero_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x \<le> 0 \<longleftrightarrow> x \<le> 1"
9.207 - using log_le_cancel_iff[of a x 1] by simp
9.208 -
9.209 -lemma one_less_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 1 < log a x \<longleftrightarrow> a < x"
9.210 - using log_less_cancel_iff[of a a x] by simp
9.211 -
9.212 -lemma one_le_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 1 \<le> log a x \<longleftrightarrow> a \<le> x"
9.213 - using log_le_cancel_iff[of a a x] by simp
9.214 -
9.215 -lemma log_less_one_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x < 1 \<longleftrightarrow> x < a"
9.216 - using log_less_cancel_iff[of a x a] by simp
9.217 -
9.218 -lemma log_le_one_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x \<le> 1 \<longleftrightarrow> x \<le> a"
9.219 - using log_le_cancel_iff[of a x a] by simp
9.220 -
9.221 -lemma powr_realpow: "0 < x ==> x powr (real n) = x^n"
9.222 - apply (induct n, simp)
9.223 - apply (subgoal_tac "real(Suc n) = real n + 1")
9.224 - apply (erule ssubst)
9.225 - apply (subst powr_add, simp, simp)
9.226 -done
9.227 -
9.228 -lemma powr_realpow2: "0 <= x ==> 0 < n ==> x^n = (if (x = 0) then 0 else x powr (real n))"
9.229 - apply (case_tac "x = 0", simp, simp)
9.230 - apply (rule powr_realpow [THEN sym], simp)
9.231 -done
9.232 -
9.233 -lemma powr_int:
9.234 - assumes "x > 0"
9.235 - shows "x powr i = (if i \<ge> 0 then x ^ nat i else 1 / x ^ nat (-i))"
9.236 -proof cases
9.237 - assume "i < 0"
9.238 - have r: "x powr i = 1 / x powr (-i)" by (simp add: powr_minus field_simps)
9.239 - show ?thesis using `i < 0` `x > 0` by (simp add: r field_simps powr_realpow[symmetric])
9.240 -qed (simp add: assms powr_realpow[symmetric])
9.241 -
9.242 -lemma powr_numeral: "0 < x \<Longrightarrow> x powr numeral n = x^numeral n"
9.243 - using powr_realpow[of x "numeral n"] by simp
9.244 -
9.245 -lemma powr_neg_numeral: "0 < x \<Longrightarrow> x powr neg_numeral n = 1 / x^numeral n"
9.246 - using powr_int[of x "neg_numeral n"] by simp
9.247 -
9.248 -lemma root_powr_inverse:
9.249 - "0 < n \<Longrightarrow> 0 < x \<Longrightarrow> root n x = x powr (1/n)"
9.250 - by (rule real_root_pos_unique) (auto simp: powr_realpow[symmetric] powr_powr)
9.251 -
9.252 -lemma ln_powr: "0 < x ==> 0 < y ==> ln(x powr y) = y * ln x"
9.253 -by (unfold powr_def, simp)
9.254 -
9.255 -lemma log_powr: "0 < x ==> 0 \<le> y ==> log b (x powr y) = y * log b x"
9.256 - apply (case_tac "y = 0")
9.257 - apply force
9.258 - apply (auto simp add: log_def ln_powr field_simps)
9.259 -done
9.260 -
9.261 -lemma log_nat_power: "0 < x ==> log b (x^n) = real n * log b x"
9.262 - apply (subst powr_realpow [symmetric])
9.263 - apply (auto simp add: log_powr)
9.264 -done
9.265 -
9.266 -lemma ln_bound: "1 <= x ==> ln x <= x"
9.267 - apply (subgoal_tac "ln(1 + (x - 1)) <= x - 1")
9.268 - apply simp
9.269 - apply (rule ln_add_one_self_le_self, simp)
9.270 -done
9.271 -
9.272 -lemma powr_mono: "a <= b ==> 1 <= x ==> x powr a <= x powr b"
9.273 - apply (case_tac "x = 1", simp)
9.274 - apply (case_tac "a = b", simp)
9.275 - apply (rule order_less_imp_le)
9.276 - apply (rule powr_less_mono, auto)
9.277 -done
9.278 -
9.279 -lemma ge_one_powr_ge_zero: "1 <= x ==> 0 <= a ==> 1 <= x powr a"
9.280 - apply (subst powr_zero_eq_one [THEN sym])
9.281 - apply (rule powr_mono, assumption+)
9.282 -done
9.283 -
9.284 -lemma powr_less_mono2: "0 < a ==> 0 < x ==> x < y ==> x powr a <
9.285 - y powr a"
9.286 - apply (unfold powr_def)
9.287 - apply (rule exp_less_mono)
9.288 - apply (rule mult_strict_left_mono)
9.289 - apply (subst ln_less_cancel_iff, assumption)
9.290 - apply (rule order_less_trans)
9.291 - prefer 2
9.292 - apply assumption+
9.293 -done
9.294 -
9.295 -lemma powr_less_mono2_neg: "a < 0 ==> 0 < x ==> x < y ==> y powr a <
9.296 - x powr a"
9.297 - apply (unfold powr_def)
9.298 - apply (rule exp_less_mono)
9.299 - apply (rule mult_strict_left_mono_neg)
9.300 - apply (subst ln_less_cancel_iff)
9.301 - apply assumption
9.302 - apply (rule order_less_trans)
9.303 - prefer 2
9.304 - apply assumption+
9.305 -done
9.306 -
9.307 -lemma powr_mono2: "0 <= a ==> 0 < x ==> x <= y ==> x powr a <= y powr a"
9.308 - apply (case_tac "a = 0", simp)
9.309 - apply (case_tac "x = y", simp)
9.310 - apply (rule order_less_imp_le)
9.311 - apply (rule powr_less_mono2, auto)
9.312 -done
9.313 -
9.314 -lemma powr_inj:
9.315 - "0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> a powr x = a powr y \<longleftrightarrow> x = y"
9.316 - unfolding powr_def exp_inj_iff by simp
9.317 -
9.318 -lemma ln_powr_bound: "1 <= x ==> 0 < a ==> ln x <= (x powr a) / a"
9.319 - apply (rule mult_imp_le_div_pos)
9.320 - apply (assumption)
9.321 - apply (subst mult_commute)
9.322 - apply (subst ln_powr [THEN sym])
9.323 - apply auto
9.324 - apply (rule ln_bound)
9.325 - apply (erule ge_one_powr_ge_zero)
9.326 - apply (erule order_less_imp_le)
9.327 -done
9.328 -
9.329 -lemma ln_powr_bound2:
9.330 - assumes "1 < x" and "0 < a"
9.331 - shows "(ln x) powr a <= (a powr a) * x"
9.332 -proof -
9.333 - from assms have "ln x <= (x powr (1 / a)) / (1 / a)"
9.334 - apply (intro ln_powr_bound)
9.335 - apply (erule order_less_imp_le)
9.336 - apply (rule divide_pos_pos)
9.337 - apply simp_all
9.338 - done
9.339 - also have "... = a * (x powr (1 / a))"
9.340 - by simp
9.341 - finally have "(ln x) powr a <= (a * (x powr (1 / a))) powr a"
9.342 - apply (intro powr_mono2)
9.343 - apply (rule order_less_imp_le, rule assms)
9.344 - apply (rule ln_gt_zero)
9.345 - apply (rule assms)
9.346 - apply assumption
9.347 - done
9.348 - also have "... = (a powr a) * ((x powr (1 / a)) powr a)"
9.349 - apply (rule powr_mult)
9.350 - apply (rule assms)
9.351 - apply (rule powr_gt_zero)
9.352 - done
9.353 - also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)"
9.354 - by (rule powr_powr)
9.355 - also have "... = x"
9.356 - apply simp
9.357 - apply (subgoal_tac "a ~= 0")
9.358 - using assms apply auto
9.359 - done
9.360 - finally show ?thesis .
9.361 -qed
9.362 -
9.363 -lemma tendsto_powr [tendsto_intros]:
9.364 - "\<lbrakk>(f ---> a) F; (g ---> b) F; 0 < a\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x powr g x) ---> a powr b) F"
9.365 - unfolding powr_def by (intro tendsto_intros)
9.366 -
9.367 -lemma continuous_powr:
9.368 - assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\<lambda>x. x))"
9.369 - shows "continuous F (\<lambda>x. (f x) powr (g x))"
9.370 - using assms unfolding continuous_def by (rule tendsto_powr)
9.371 -
9.372 -lemma continuous_at_within_powr[continuous_intros]:
9.373 - assumes "continuous (at a within s) f" "continuous (at a within s) g" and "0 < f a"
9.374 - shows "continuous (at a within s) (\<lambda>x. (f x) powr (g x))"
9.375 - using assms unfolding continuous_within by (rule tendsto_powr)
9.376 -
9.377 -lemma isCont_powr[continuous_intros, simp]:
9.378 - assumes "isCont f a" "isCont g a" "0 < f a"
9.379 - shows "isCont (\<lambda>x. (f x) powr g x) a"
9.380 - using assms unfolding continuous_at by (rule tendsto_powr)
9.381 -
9.382 -lemma continuous_on_powr[continuous_on_intros]:
9.383 - assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. 0 < f x"
9.384 - shows "continuous_on s (\<lambda>x. (f x) powr (g x))"
9.385 - using assms unfolding continuous_on_def by (fast intro: tendsto_powr)
9.386 -
9.387 -(* FIXME: generalize by replacing d by with g x and g ---> d? *)
9.388 -lemma tendsto_zero_powrI:
9.389 - assumes "eventually (\<lambda>x. 0 < f x ) F" and "(f ---> 0) F"
9.390 - assumes "0 < d"
9.391 - shows "((\<lambda>x. f x powr d) ---> 0) F"
9.392 -proof (rule tendstoI)
9.393 - fix e :: real assume "0 < e"
9.394 - def Z \<equiv> "e powr (1 / d)"
9.395 - with `0 < e` have "0 < Z" by simp
9.396 - with assms have "eventually (\<lambda>x. 0 < f x \<and> dist (f x) 0 < Z) F"
9.397 - by (intro eventually_conj tendstoD)
9.398 - moreover
9.399 - from assms have "\<And>x. 0 < x \<and> dist x 0 < Z \<Longrightarrow> x powr d < Z powr d"
9.400 - by (intro powr_less_mono2) (auto simp: dist_real_def)
9.401 - with assms `0 < e` have "\<And>x. 0 < x \<and> dist x 0 < Z \<Longrightarrow> dist (x powr d) 0 < e"
9.402 - unfolding dist_real_def Z_def by (auto simp: powr_powr)
9.403 - ultimately
9.404 - show "eventually (\<lambda>x. dist (f x powr d) 0 < e) F" by (rule eventually_elim1)
9.405 -qed
9.406 -
9.407 -lemma tendsto_neg_powr:
9.408 - assumes "s < 0" and "LIM x F. f x :> at_top"
9.409 - shows "((\<lambda>x. f x powr s) ---> 0) F"
9.410 -proof (rule tendstoI)
9.411 - fix e :: real assume "0 < e"
9.412 - def Z \<equiv> "e powr (1 / s)"
9.413 - from assms have "eventually (\<lambda>x. Z < f x) F"
9.414 - by (simp add: filterlim_at_top_dense)
9.415 - moreover
9.416 - from assms have "\<And>x. Z < x \<Longrightarrow> x powr s < Z powr s"
9.417 - by (auto simp: Z_def intro!: powr_less_mono2_neg)
9.418 - with assms `0 < e` have "\<And>x. Z < x \<Longrightarrow> dist (x powr s) 0 < e"
9.419 - by (simp add: powr_powr Z_def dist_real_def)
9.420 - ultimately
9.421 - show "eventually (\<lambda>x. dist (f x powr s) 0 < e) F" by (rule eventually_elim1)
9.422 -qed
9.423 -
9.424 -end
10.1 --- a/src/HOL/Lubs.thy Tue Mar 26 14:38:44 2013 +0100
10.2 +++ b/src/HOL/Lubs.thy Tue Mar 26 15:10:28 2013 +0100
10.3 @@ -94,4 +94,10 @@
10.4 lemma isLub_ubs: "isLub R S x \<Longrightarrow> x <=* ubs R S"
10.5 unfolding ubs_def isLub_def by (rule leastPD2)
10.6
10.7 +lemma isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::'a::linorder)"
10.8 + apply (frule isLub_isUb)
10.9 + apply (frule_tac x = y in isLub_isUb)
10.10 + apply (blast intro!: order_antisym dest!: isLub_le_isUb)
10.11 + done
10.12 +
10.13 end
11.1 --- a/src/HOL/Metric_Spaces.thy Tue Mar 26 14:38:44 2013 +0100
11.2 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000
11.3 @@ -1,629 +0,0 @@
11.4 -(* Title: HOL/Metric_Spaces.thy
11.5 - Author: Brian Huffman
11.6 - Author: Johannes Hölzl
11.7 -*)
11.8 -
11.9 -header {* Metric Spaces *}
11.10 -
11.11 -theory Metric_Spaces
11.12 -imports RComplete Topological_Spaces
11.13 -begin
11.14 -
11.15 -
11.16 -subsection {* Metric spaces *}
11.17 -
11.18 -class dist =
11.19 - fixes dist :: "'a \<Rightarrow> 'a \<Rightarrow> real"
11.20 -
11.21 -class open_dist = "open" + dist +
11.22 - assumes open_dist: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
11.23 -
11.24 -class metric_space = open_dist +
11.25 - assumes dist_eq_0_iff [simp]: "dist x y = 0 \<longleftrightarrow> x = y"
11.26 - assumes dist_triangle2: "dist x y \<le> dist x z + dist y z"
11.27 -begin
11.28 -
11.29 -lemma dist_self [simp]: "dist x x = 0"
11.30 -by simp
11.31 -
11.32 -lemma zero_le_dist [simp]: "0 \<le> dist x y"
11.33 -using dist_triangle2 [of x x y] by simp
11.34 -
11.35 -lemma zero_less_dist_iff: "0 < dist x y \<longleftrightarrow> x \<noteq> y"
11.36 -by (simp add: less_le)
11.37 -
11.38 -lemma dist_not_less_zero [simp]: "\<not> dist x y < 0"
11.39 -by (simp add: not_less)
11.40 -
11.41 -lemma dist_le_zero_iff [simp]: "dist x y \<le> 0 \<longleftrightarrow> x = y"
11.42 -by (simp add: le_less)
11.43 -
11.44 -lemma dist_commute: "dist x y = dist y x"
11.45 -proof (rule order_antisym)
11.46 - show "dist x y \<le> dist y x"
11.47 - using dist_triangle2 [of x y x] by simp
11.48 - show "dist y x \<le> dist x y"
11.49 - using dist_triangle2 [of y x y] by simp
11.50 -qed
11.51 -
11.52 -lemma dist_triangle: "dist x z \<le> dist x y + dist y z"
11.53 -using dist_triangle2 [of x z y] by (simp add: dist_commute)
11.54 -
11.55 -lemma dist_triangle3: "dist x y \<le> dist a x + dist a y"
11.56 -using dist_triangle2 [of x y a] by (simp add: dist_commute)
11.57 -
11.58 -lemma dist_triangle_alt:
11.59 - shows "dist y z <= dist x y + dist x z"
11.60 -by (rule dist_triangle3)
11.61 -
11.62 -lemma dist_pos_lt:
11.63 - shows "x \<noteq> y ==> 0 < dist x y"
11.64 -by (simp add: zero_less_dist_iff)
11.65 -
11.66 -lemma dist_nz:
11.67 - shows "x \<noteq> y \<longleftrightarrow> 0 < dist x y"
11.68 -by (simp add: zero_less_dist_iff)
11.69 -
11.70 -lemma dist_triangle_le:
11.71 - shows "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e"
11.72 -by (rule order_trans [OF dist_triangle2])
11.73 -
11.74 -lemma dist_triangle_lt:
11.75 - shows "dist x z + dist y z < e ==> dist x y < e"
11.76 -by (rule le_less_trans [OF dist_triangle2])
11.77 -
11.78 -lemma dist_triangle_half_l:
11.79 - shows "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"
11.80 -by (rule dist_triangle_lt [where z=y], simp)
11.81 -
11.82 -lemma dist_triangle_half_r:
11.83 - shows "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"
11.84 -by (rule dist_triangle_half_l, simp_all add: dist_commute)
11.85 -
11.86 -subclass topological_space
11.87 -proof
11.88 - have "\<exists>e::real. 0 < e"
11.89 - by (fast intro: zero_less_one)
11.90 - then show "open UNIV"
11.91 - unfolding open_dist by simp
11.92 -next
11.93 - fix S T assume "open S" "open T"
11.94 - then show "open (S \<inter> T)"
11.95 - unfolding open_dist
11.96 - apply clarify
11.97 - apply (drule (1) bspec)+
11.98 - apply (clarify, rename_tac r s)
11.99 - apply (rule_tac x="min r s" in exI, simp)
11.100 - done
11.101 -next
11.102 - fix K assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)"
11.103 - unfolding open_dist by fast
11.104 -qed
11.105 -
11.106 -lemma open_ball: "open {y. dist x y < d}"
11.107 -proof (unfold open_dist, intro ballI)
11.108 - fix y assume *: "y \<in> {y. dist x y < d}"
11.109 - then show "\<exists>e>0. \<forall>z. dist z y < e \<longrightarrow> z \<in> {y. dist x y < d}"
11.110 - by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt)
11.111 -qed
11.112 -
11.113 -subclass first_countable_topology
11.114 -proof
11.115 - fix x
11.116 - show "\<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))"
11.117 - proof (safe intro!: exI[of _ "\<lambda>n. {y. dist x y < inverse (Suc n)}"])
11.118 - fix S assume "open S" "x \<in> S"
11.119 - then obtain e where "0 < e" "{y. dist x y < e} \<subseteq> S"
11.120 - by (auto simp: open_dist subset_eq dist_commute)
11.121 - moreover
11.122 - then obtain i where "inverse (Suc i) < e"
11.123 - by (auto dest!: reals_Archimedean)
11.124 - then have "{y. dist x y < inverse (Suc i)} \<subseteq> {y. dist x y < e}"
11.125 - by auto
11.126 - ultimately show "\<exists>i. {y. dist x y < inverse (Suc i)} \<subseteq> S"
11.127 - by blast
11.128 - qed (auto intro: open_ball)
11.129 -qed
11.130 -
11.131 -end
11.132 -
11.133 -instance metric_space \<subseteq> t2_space
11.134 -proof
11.135 - fix x y :: "'a::metric_space"
11.136 - assume xy: "x \<noteq> y"
11.137 - let ?U = "{y'. dist x y' < dist x y / 2}"
11.138 - let ?V = "{x'. dist y x' < dist x y / 2}"
11.139 - have th0: "\<And>d x y z. (d x z :: real) \<le> d x y + d y z \<Longrightarrow> d y z = d z y
11.140 - \<Longrightarrow> \<not>(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
11.141 - have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
11.142 - using dist_pos_lt[OF xy] th0[of dist, OF dist_triangle dist_commute]
11.143 - using open_ball[of _ "dist x y / 2"] by auto
11.144 - then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
11.145 - by blast
11.146 -qed
11.147 -
11.148 -lemma eventually_nhds_metric:
11.149 - fixes a :: "'a :: metric_space"
11.150 - shows "eventually P (nhds a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. dist x a < d \<longrightarrow> P x)"
11.151 -unfolding eventually_nhds open_dist
11.152 -apply safe
11.153 -apply fast
11.154 -apply (rule_tac x="{x. dist x a < d}" in exI, simp)
11.155 -apply clarsimp
11.156 -apply (rule_tac x="d - dist x a" in exI, clarsimp)
11.157 -apply (simp only: less_diff_eq)
11.158 -apply (erule le_less_trans [OF dist_triangle])
11.159 -done
11.160 -
11.161 -lemma eventually_at:
11.162 - fixes a :: "'a::metric_space"
11.163 - shows "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> P x)"
11.164 -unfolding at_def eventually_within eventually_nhds_metric by auto
11.165 -
11.166 -lemma eventually_within_less: (* COPY FROM Topo/eventually_within *)
11.167 - fixes a :: "'a :: metric_space"
11.168 - shows "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
11.169 - unfolding eventually_within eventually_at dist_nz by auto
11.170 -
11.171 -lemma eventually_within_le: (* COPY FROM Topo/eventually_within_le *)
11.172 - fixes a :: "'a :: metric_space"
11.173 - shows "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a <= d \<longrightarrow> P x)"
11.174 - unfolding eventually_within_less by auto (metis dense order_le_less_trans)
11.175 -
11.176 -lemma tendstoI:
11.177 - fixes l :: "'a :: metric_space"
11.178 - assumes "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F"
11.179 - shows "(f ---> l) F"
11.180 - apply (rule topological_tendstoI)
11.181 - apply (simp add: open_dist)
11.182 - apply (drule (1) bspec, clarify)
11.183 - apply (drule assms)
11.184 - apply (erule eventually_elim1, simp)
11.185 - done
11.186 -
11.187 -lemma tendstoD:
11.188 - fixes l :: "'a :: metric_space"
11.189 - shows "(f ---> l) F \<Longrightarrow> 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F"
11.190 - apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD)
11.191 - apply (clarsimp simp add: open_dist)
11.192 - apply (rule_tac x="e - dist x l" in exI, clarsimp)
11.193 - apply (simp only: less_diff_eq)
11.194 - apply (erule le_less_trans [OF dist_triangle])
11.195 - apply simp
11.196 - apply simp
11.197 - done
11.198 -
11.199 -lemma tendsto_iff:
11.200 - fixes l :: "'a :: metric_space"
11.201 - shows "(f ---> l) F \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) F)"
11.202 - using tendstoI tendstoD by fast
11.203 -
11.204 -lemma metric_tendsto_imp_tendsto:
11.205 - fixes a :: "'a :: metric_space" and b :: "'b :: metric_space"
11.206 - assumes f: "(f ---> a) F"
11.207 - assumes le: "eventually (\<lambda>x. dist (g x) b \<le> dist (f x) a) F"
11.208 - shows "(g ---> b) F"
11.209 -proof (rule tendstoI)
11.210 - fix e :: real assume "0 < e"
11.211 - with f have "eventually (\<lambda>x. dist (f x) a < e) F" by (rule tendstoD)
11.212 - with le show "eventually (\<lambda>x. dist (g x) b < e) F"
11.213 - using le_less_trans by (rule eventually_elim2)
11.214 -qed
11.215 -
11.216 -lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top"
11.217 - unfolding filterlim_at_top
11.218 - apply (intro allI)
11.219 - apply (rule_tac c="natceiling (Z + 1)" in eventually_sequentiallyI)
11.220 - apply (auto simp: natceiling_le_eq)
11.221 - done
11.222 -
11.223 -subsubsection {* Limits of Sequences *}
11.224 -
11.225 -lemma LIMSEQ_def: "X ----> (L::'a::metric_space) \<longleftrightarrow> (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)"
11.226 - unfolding tendsto_iff eventually_sequentially ..
11.227 -
11.228 -lemma LIMSEQ_iff_nz: "X ----> (L::'a::metric_space) = (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)"
11.229 - unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc)
11.230 -
11.231 -lemma metric_LIMSEQ_I:
11.232 - "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r) \<Longrightarrow> X ----> (L::'a::metric_space)"
11.233 -by (simp add: LIMSEQ_def)
11.234 -
11.235 -lemma metric_LIMSEQ_D:
11.236 - "\<lbrakk>X ----> (L::'a::metric_space); 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r"
11.237 -by (simp add: LIMSEQ_def)
11.238 -
11.239 -
11.240 -subsubsection {* Limits of Functions *}
11.241 -
11.242 -lemma LIM_def: "f -- (a::'a::metric_space) --> (L::'b::metric_space) =
11.243 - (\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & dist x a < s
11.244 - --> dist (f x) L < r)"
11.245 -unfolding tendsto_iff eventually_at ..
11.246 -
11.247 -lemma metric_LIM_I:
11.248 - "(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r)
11.249 - \<Longrightarrow> f -- (a::'a::metric_space) --> (L::'b::metric_space)"
11.250 -by (simp add: LIM_def)
11.251 -
11.252 -lemma metric_LIM_D:
11.253 - "\<lbrakk>f -- (a::'a::metric_space) --> (L::'b::metric_space); 0 < r\<rbrakk>
11.254 - \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r"
11.255 -by (simp add: LIM_def)
11.256 -
11.257 -lemma metric_LIM_imp_LIM:
11.258 - assumes f: "f -- a --> (l::'a::metric_space)"
11.259 - assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> dist (g x) m \<le> dist (f x) l"
11.260 - shows "g -- a --> (m::'b::metric_space)"
11.261 - by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp add: eventually_at_topological le)
11.262 -
11.263 -lemma metric_LIM_equal2:
11.264 - assumes 1: "0 < R"
11.265 - assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; dist x a < R\<rbrakk> \<Longrightarrow> f x = g x"
11.266 - shows "g -- a --> l \<Longrightarrow> f -- (a::'a::metric_space) --> l"
11.267 -apply (rule topological_tendstoI)
11.268 -apply (drule (2) topological_tendstoD)
11.269 -apply (simp add: eventually_at, safe)
11.270 -apply (rule_tac x="min d R" in exI, safe)
11.271 -apply (simp add: 1)
11.272 -apply (simp add: 2)
11.273 -done
11.274 -
11.275 -lemma metric_LIM_compose2:
11.276 - assumes f: "f -- (a::'a::metric_space) --> b"
11.277 - assumes g: "g -- b --> c"
11.278 - assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> b"
11.279 - shows "(\<lambda>x. g (f x)) -- a --> c"
11.280 - using g f inj [folded eventually_at]
11.281 - by (rule tendsto_compose_eventually)
11.282 -
11.283 -lemma metric_isCont_LIM_compose2:
11.284 - fixes f :: "'a :: metric_space \<Rightarrow> _"
11.285 - assumes f [unfolded isCont_def]: "isCont f a"
11.286 - assumes g: "g -- f a --> l"
11.287 - assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> f a"
11.288 - shows "(\<lambda>x. g (f x)) -- a --> l"
11.289 -by (rule metric_LIM_compose2 [OF f g inj])
11.290 -
11.291 -subsubsection {* Boundedness *}
11.292 -
11.293 -definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where
11.294 - Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)"
11.295 -
11.296 -abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
11.297 - "Bseq X \<equiv> Bfun X sequentially"
11.298 -
11.299 -lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" ..
11.300 -
11.301 -lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
11.302 - unfolding Bfun_metric_def by (subst eventually_sequentially_seg)
11.303 -
11.304 -lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
11.305 - unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg)
11.306 -
11.307 -subsection {* Complete metric spaces *}
11.308 -
11.309 -subsection {* Cauchy sequences *}
11.310 -
11.311 -definition (in metric_space) Cauchy :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" where
11.312 - "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < e)"
11.313 -
11.314 -subsection {* Cauchy Sequences *}
11.315 -
11.316 -lemma metric_CauchyI:
11.317 - "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X"
11.318 - by (simp add: Cauchy_def)
11.319 -
11.320 -lemma metric_CauchyD:
11.321 - "Cauchy X \<Longrightarrow> 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e"
11.322 - by (simp add: Cauchy_def)
11.323 -
11.324 -lemma metric_Cauchy_iff2:
11.325 - "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < inverse(real (Suc j))))"
11.326 -apply (simp add: Cauchy_def, auto)
11.327 -apply (drule reals_Archimedean, safe)
11.328 -apply (drule_tac x = n in spec, auto)
11.329 -apply (rule_tac x = M in exI, auto)
11.330 -apply (drule_tac x = m in spec, simp)
11.331 -apply (drule_tac x = na in spec, auto)
11.332 -done
11.333 -
11.334 -lemma Cauchy_subseq_Cauchy:
11.335 - "\<lbrakk> Cauchy X; subseq f \<rbrakk> \<Longrightarrow> Cauchy (X o f)"
11.336 -apply (auto simp add: Cauchy_def)
11.337 -apply (drule_tac x=e in spec, clarify)
11.338 -apply (rule_tac x=M in exI, clarify)
11.339 -apply (blast intro: le_trans [OF _ seq_suble] dest!: spec)
11.340 -done
11.341 -
11.342 -lemma Cauchy_Bseq: "Cauchy X \<Longrightarrow> Bseq X"
11.343 - unfolding Cauchy_def Bfun_metric_def eventually_sequentially
11.344 - apply (erule_tac x=1 in allE)
11.345 - apply simp
11.346 - apply safe
11.347 - apply (rule_tac x="X M" in exI)
11.348 - apply (rule_tac x=1 in exI)
11.349 - apply (erule_tac x=M in allE)
11.350 - apply simp
11.351 - apply (rule_tac x=M in exI)
11.352 - apply (auto simp: dist_commute)
11.353 - done
11.354 -
11.355 -subsubsection {* Cauchy Sequences are Convergent *}
11.356 -
11.357 -class complete_space = metric_space +
11.358 - assumes Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
11.359 -
11.360 -theorem LIMSEQ_imp_Cauchy:
11.361 - assumes X: "X ----> a" shows "Cauchy X"
11.362 -proof (rule metric_CauchyI)
11.363 - fix e::real assume "0 < e"
11.364 - hence "0 < e/2" by simp
11.365 - with X have "\<exists>N. \<forall>n\<ge>N. dist (X n) a < e/2" by (rule metric_LIMSEQ_D)
11.366 - then obtain N where N: "\<forall>n\<ge>N. dist (X n) a < e/2" ..
11.367 - show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < e"
11.368 - proof (intro exI allI impI)
11.369 - fix m assume "N \<le> m"
11.370 - hence m: "dist (X m) a < e/2" using N by fast
11.371 - fix n assume "N \<le> n"
11.372 - hence n: "dist (X n) a < e/2" using N by fast
11.373 - have "dist (X m) (X n) \<le> dist (X m) a + dist (X n) a"
11.374 - by (rule dist_triangle2)
11.375 - also from m n have "\<dots> < e" by simp
11.376 - finally show "dist (X m) (X n) < e" .
11.377 - qed
11.378 -qed
11.379 -
11.380 -lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
11.381 -unfolding convergent_def
11.382 -by (erule exE, erule LIMSEQ_imp_Cauchy)
11.383 -
11.384 -lemma Cauchy_convergent_iff:
11.385 - fixes X :: "nat \<Rightarrow> 'a::complete_space"
11.386 - shows "Cauchy X = convergent X"
11.387 -by (fast intro: Cauchy_convergent convergent_Cauchy)
11.388 -
11.389 -subsection {* Uniform Continuity *}
11.390 -
11.391 -definition
11.392 - isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where
11.393 - "isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"
11.394 -
11.395 -lemma isUCont_isCont: "isUCont f ==> isCont f x"
11.396 -by (simp add: isUCont_def isCont_def LIM_def, force)
11.397 -
11.398 -lemma isUCont_Cauchy:
11.399 - "\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
11.400 -unfolding isUCont_def
11.401 -apply (rule metric_CauchyI)
11.402 -apply (drule_tac x=e in spec, safe)
11.403 -apply (drule_tac e=s in metric_CauchyD, safe)
11.404 -apply (rule_tac x=M in exI, simp)
11.405 -done
11.406 -
11.407 -subsection {* The set of real numbers is a complete metric space *}
11.408 -
11.409 -instantiation real :: metric_space
11.410 -begin
11.411 -
11.412 -definition dist_real_def:
11.413 - "dist x y = \<bar>x - y\<bar>"
11.414 -
11.415 -definition open_real_def:
11.416 - "open (S :: real set) \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
11.417 -
11.418 -instance
11.419 - by default (auto simp: open_real_def dist_real_def)
11.420 -end
11.421 -
11.422 -instance real :: linorder_topology
11.423 -proof
11.424 - show "(open :: real set \<Rightarrow> bool) = generate_topology (range lessThan \<union> range greaterThan)"
11.425 - proof (rule ext, safe)
11.426 - fix S :: "real set" assume "open S"
11.427 - then guess f unfolding open_real_def bchoice_iff ..
11.428 - then have *: "S = (\<Union>x\<in>S. {x - f x <..} \<inter> {..< x + f x})"
11.429 - by (fastforce simp: dist_real_def)
11.430 - show "generate_topology (range lessThan \<union> range greaterThan) S"
11.431 - apply (subst *)
11.432 - apply (intro generate_topology_Union generate_topology.Int)
11.433 - apply (auto intro: generate_topology.Basis)
11.434 - done
11.435 - next
11.436 - fix S :: "real set" assume "generate_topology (range lessThan \<union> range greaterThan) S"
11.437 - moreover have "\<And>a::real. open {..<a}"
11.438 - unfolding open_real_def dist_real_def
11.439 - proof clarify
11.440 - fix x a :: real assume "x < a"
11.441 - hence "0 < a - x \<and> (\<forall>y. \<bar>y - x\<bar> < a - x \<longrightarrow> y \<in> {..<a})" by auto
11.442 - thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {..<a}" ..
11.443 - qed
11.444 - moreover have "\<And>a::real. open {a <..}"
11.445 - unfolding open_real_def dist_real_def
11.446 - proof clarify
11.447 - fix x a :: real assume "a < x"
11.448 - hence "0 < x - a \<and> (\<forall>y. \<bar>y - x\<bar> < x - a \<longrightarrow> y \<in> {a<..})" by auto
11.449 - thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {a<..}" ..
11.450 - qed
11.451 - ultimately show "open S"
11.452 - by induct auto
11.453 - qed
11.454 -qed
11.455 -
11.456 -lemmas open_real_greaterThan = open_greaterThan[where 'a=real]
11.457 -lemmas open_real_lessThan = open_lessThan[where 'a=real]
11.458 -lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real]
11.459 -lemmas closed_real_atMost = closed_atMost[where 'a=real]
11.460 -lemmas closed_real_atLeast = closed_atLeast[where 'a=real]
11.461 -lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real]
11.462 -
11.463 -text {*
11.464 -Proof that Cauchy sequences converge based on the one from
11.465 -http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
11.466 -*}
11.467 -
11.468 -text {*
11.469 - If sequence @{term "X"} is Cauchy, then its limit is the lub of
11.470 - @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
11.471 -*}
11.472 -
11.473 -lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
11.474 -by (simp add: isUbI setleI)
11.475 -
11.476 -lemma increasing_LIMSEQ:
11.477 - fixes f :: "nat \<Rightarrow> real"
11.478 - assumes inc: "\<And>n. f n \<le> f (Suc n)"
11.479 - and bdd: "\<And>n. f n \<le> l"
11.480 - and en: "\<And>e. 0 < e \<Longrightarrow> \<exists>n. l \<le> f n + e"
11.481 - shows "f ----> l"
11.482 -proof (rule increasing_tendsto)
11.483 - fix x assume "x < l"
11.484 - with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x"
11.485 - by auto
11.486 - from en[OF `0 < e`] obtain n where "l - e \<le> f n"
11.487 - by (auto simp: field_simps)
11.488 - with `e < l - x` `0 < e` have "x < f n" by simp
11.489 - with incseq_SucI[of f, OF inc] show "eventually (\<lambda>n. x < f n) sequentially"
11.490 - by (auto simp: eventually_sequentially incseq_def intro: less_le_trans)
11.491 -qed (insert bdd, auto)
11.492 -
11.493 -lemma real_Cauchy_convergent:
11.494 - fixes X :: "nat \<Rightarrow> real"
11.495 - assumes X: "Cauchy X"
11.496 - shows "convergent X"
11.497 -proof -
11.498 - def S \<equiv> "{x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
11.499 - then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S" by auto
11.500 -
11.501 - { fix N x assume N: "\<forall>n\<ge>N. X n < x"
11.502 - have "isUb UNIV S x"
11.503 - proof (rule isUb_UNIV_I)
11.504 - fix y::real assume "y \<in> S"
11.505 - hence "\<exists>M. \<forall>n\<ge>M. y < X n"
11.506 - by (simp add: S_def)
11.507 - then obtain M where "\<forall>n\<ge>M. y < X n" ..
11.508 - hence "y < X (max M N)" by simp
11.509 - also have "\<dots> < x" using N by simp
11.510 - finally show "y \<le> x"
11.511 - by (rule order_less_imp_le)
11.512 - qed }
11.513 - note bound_isUb = this
11.514 -
11.515 - have "\<exists>u. isLub UNIV S u"
11.516 - proof (rule reals_complete)
11.517 - obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < 1"
11.518 - using X[THEN metric_CauchyD, OF zero_less_one] by auto
11.519 - hence N: "\<forall>n\<ge>N. dist (X n) (X N) < 1" by simp
11.520 - show "\<exists>x. x \<in> S"
11.521 - proof
11.522 - from N have "\<forall>n\<ge>N. X N - 1 < X n"
11.523 - by (simp add: abs_diff_less_iff dist_real_def)
11.524 - thus "X N - 1 \<in> S" by (rule mem_S)
11.525 - qed
11.526 - show "\<exists>u. isUb UNIV S u"
11.527 - proof
11.528 - from N have "\<forall>n\<ge>N. X n < X N + 1"
11.529 - by (simp add: abs_diff_less_iff dist_real_def)
11.530 - thus "isUb UNIV S (X N + 1)"
11.531 - by (rule bound_isUb)
11.532 - qed
11.533 - qed
11.534 - then obtain x where x: "isLub UNIV S x" ..
11.535 - have "X ----> x"
11.536 - proof (rule metric_LIMSEQ_I)
11.537 - fix r::real assume "0 < r"
11.538 - hence r: "0 < r/2" by simp
11.539 - obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. dist (X n) (X m) < r/2"
11.540 - using metric_CauchyD [OF X r] by auto
11.541 - hence "\<forall>n\<ge>N. dist (X n) (X N) < r/2" by simp
11.542 - hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
11.543 - by (simp only: dist_real_def abs_diff_less_iff)
11.544 -
11.545 - from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
11.546 - hence "X N - r/2 \<in> S" by (rule mem_S)
11.547 - hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
11.548 -
11.549 - from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
11.550 - hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
11.551 - hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
11.552 -
11.553 - show "\<exists>N. \<forall>n\<ge>N. dist (X n) x < r"
11.554 - proof (intro exI allI impI)
11.555 - fix n assume n: "N \<le> n"
11.556 - from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
11.557 - thus "dist (X n) x < r" using 1 2
11.558 - by (simp add: abs_diff_less_iff dist_real_def)
11.559 - qed
11.560 - qed
11.561 - then show ?thesis unfolding convergent_def by auto
11.562 -qed
11.563 -
11.564 -instance real :: complete_space
11.565 - by intro_classes (rule real_Cauchy_convergent)
11.566 -
11.567 -lemma tendsto_dist [tendsto_intros]:
11.568 - fixes l m :: "'a :: metric_space"
11.569 - assumes f: "(f ---> l) F" and g: "(g ---> m) F"
11.570 - shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) F"
11.571 -proof (rule tendstoI)
11.572 - fix e :: real assume "0 < e"
11.573 - hence e2: "0 < e/2" by simp
11.574 - from tendstoD [OF f e2] tendstoD [OF g e2]
11.575 - show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F"
11.576 - proof (eventually_elim)
11.577 - case (elim x)
11.578 - then show "dist (dist (f x) (g x)) (dist l m) < e"
11.579 - unfolding dist_real_def
11.580 - using dist_triangle2 [of "f x" "g x" "l"]
11.581 - using dist_triangle2 [of "g x" "l" "m"]
11.582 - using dist_triangle3 [of "l" "m" "f x"]
11.583 - using dist_triangle [of "f x" "m" "g x"]
11.584 - by arith
11.585 - qed
11.586 -qed
11.587 -
11.588 -lemma continuous_dist[continuous_intros]:
11.589 - fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
11.590 - shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. dist (f x) (g x))"
11.591 - unfolding continuous_def by (rule tendsto_dist)
11.592 -
11.593 -lemma continuous_on_dist[continuous_on_intros]:
11.594 - fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
11.595 - shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))"
11.596 - unfolding continuous_on_def by (auto intro: tendsto_dist)
11.597 -
11.598 -lemma tendsto_at_topI_sequentially:
11.599 - fixes f :: "real \<Rightarrow> real"
11.600 - assumes mono: "mono f"
11.601 - assumes limseq: "(\<lambda>n. f (real n)) ----> y"
11.602 - shows "(f ---> y) at_top"
11.603 -proof (rule tendstoI)
11.604 - fix e :: real assume "0 < e"
11.605 - with limseq obtain N :: nat where N: "\<And>n. N \<le> n \<Longrightarrow> \<bar>f (real n) - y\<bar> < e"
11.606 - by (auto simp: LIMSEQ_def dist_real_def)
11.607 - { fix x :: real
11.608 - from ex_le_of_nat[of x] guess n ..
11.609 - note monoD[OF mono this]
11.610 - also have "f (real_of_nat n) \<le> y"
11.611 - by (rule LIMSEQ_le_const[OF limseq])
11.612 - (auto intro: exI[of _ n] monoD[OF mono] simp: real_eq_of_nat[symmetric])
11.613 - finally have "f x \<le> y" . }
11.614 - note le = this
11.615 - have "eventually (\<lambda>x. real N \<le> x) at_top"
11.616 - by (rule eventually_ge_at_top)
11.617 - then show "eventually (\<lambda>x. dist (f x) y < e) at_top"
11.618 - proof eventually_elim
11.619 - fix x assume N': "real N \<le> x"
11.620 - with N[of N] le have "y - f (real N) < e" by auto
11.621 - moreover note monoD[OF mono N']
11.622 - ultimately show "dist (f x) y < e"
11.623 - using le[of x] by (auto simp: dist_real_def field_simps)
11.624 - qed
11.625 -qed
11.626 -
11.627 -lemma Cauchy_iff2:
11.628 - "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. \<bar>X m - X n\<bar> < inverse(real (Suc j))))"
11.629 - unfolding metric_Cauchy_iff2 dist_real_def ..
11.630 -
11.631 -end
11.632 -
12.1 --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Mar 26 14:38:44 2013 +0100
12.2 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Mar 26 15:10:28 2013 +0100
12.3 @@ -468,7 +468,7 @@
12.4 u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa) /\<^sub>R (1 - u x)) \<in> V"
12.5 apply -
12.6 apply (rule as(3)[rule_format])
12.7 - unfolding RealVector.scaleR_right.setsum
12.8 + unfolding Real_Vector_Spaces.scaleR_right.setsum
12.9 using x(1) as(6) apply auto
12.10 done
12.11 then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
12.12 @@ -530,7 +530,7 @@
12.13 apply (rule_tac x="sx \<union> sy" in exI)
12.14 apply (rule_tac x="\<lambda>a. (if a\<in>sx then u * ux a else 0) + (if a\<in>sy then v * uy a else 0)" in exI)
12.15 unfolding scaleR_left_distrib setsum_addf if_smult scaleR_zero_left ** setsum_restrict_set[OF xy, symmetric]
12.16 - unfolding scaleR_scaleR[symmetric] RealVector.scaleR_right.setsum [symmetric] and setsum_right_distrib[symmetric]
12.17 + unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric] and setsum_right_distrib[symmetric]
12.18 unfolding x y
12.19 using x(1-3) y(1-3) uv apply simp
12.20 done
12.21 @@ -2848,7 +2848,7 @@
12.22 obtain e2 where e2_def: "e2>0 & cball z e2 Int affine hull (f ` S) <= (f ` S)"
12.23 using z_def rel_interior_cball[of "f ` S"] by auto
12.24 obtain K where K_def: "K>0 & (! x. norm (f x) <= norm x * K)"
12.25 - using assms RealVector.bounded_linear.pos_bounded[of f] by auto
12.26 + using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto
12.27 def e1 == "1/K" hence e1_def: "e1>0 & (! x. e1 * norm (f x) <= norm x)"
12.28 using K_def pos_le_divide_eq[of e1] by auto
12.29 def e == "e1 * e2" hence "e>0" using e1_def e2_def mult_pos_pos by auto
13.1 --- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Tue Mar 26 14:38:44 2013 +0100
13.2 +++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Tue Mar 26 15:10:28 2013 +0100
13.3 @@ -890,7 +890,7 @@
13.4 lemma Liminf_within:
13.5 fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"
13.6 shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S \<inter> ball x e - {x}). f y)"
13.7 - unfolding Liminf_def eventually_within
13.8 + unfolding Liminf_def eventually_within_less
13.9 proof (rule SUPR_eq, simp_all add: Ball_def Bex_def, safe)
13.10 fix P d assume "0 < d" "\<forall>y. y \<in> S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> P y"
13.11 then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
13.12 @@ -908,7 +908,7 @@
13.13 lemma Limsup_within:
13.14 fixes f :: "'a::metric_space => 'b::complete_lattice"
13.15 shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S \<inter> ball x e - {x}). f y)"
13.16 - unfolding Limsup_def eventually_within
13.17 + unfolding Limsup_def eventually_within_less
13.18 proof (rule INFI_eq, simp_all add: Ball_def Bex_def, safe)
13.19 fix P d assume "0 < d" "\<forall>y. y \<in> S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> P y"
13.20 then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
14.1 --- a/src/HOL/Multivariate_Analysis/Integration.thy Tue Mar 26 14:38:44 2013 +0100
14.2 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Mar 26 15:10:28 2013 +0100
14.3 @@ -8,17 +8,56 @@
14.4 "~~/src/HOL/Library/Indicator_Function"
14.5 begin
14.6
14.7 -lemma cSup_finite_ge_iff:
14.8 +lemma cSup_abs_le: (* TODO: is this really needed? *)
14.9 + fixes S :: "real set"
14.10 + shows "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>Sup S\<bar> \<le> a"
14.11 +by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2)
14.12 +
14.13 +lemma cInf_abs_ge: (* TODO: is this really needed? *)
14.14 + fixes S :: "real set"
14.15 + shows "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>Inf S\<bar> \<le> a"
14.16 +by (simp add: Inf_real_def) (rule cSup_abs_le, auto)
14.17 +
14.18 +lemma cSup_asclose: (* TODO: is this really needed? *)
14.19 fixes S :: "real set"
14.20 - assumes fS: "finite S" and Se: "S \<noteq> {}"
14.21 - shows "a \<le> Sup S \<longleftrightarrow> (\<exists> x \<in> S. a \<le> x)"
14.22 -by (metis Max_ge Se cSup_eq_Max Max_in fS linorder_not_le less_le_trans)
14.23 + assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>Sup S - l\<bar> \<le> e"
14.24 +proof-
14.25 + have th: "\<And>(x::real) l e. \<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e" by arith
14.26 + thus ?thesis using S b cSup_bounds[of S "l - e" "l+e"] unfolding th
14.27 + by (auto simp add: setge_def setle_def)
14.28 +qed
14.29 +
14.30 +lemma cInf_asclose: (* TODO: is this really needed? *)
14.31 + fixes S :: "real set"
14.32 + assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>Inf S - l\<bar> \<le> e"
14.33 +proof -
14.34 + have "\<bar>- Sup (uminus ` S) - l\<bar> = \<bar>Sup (uminus ` S) - (-l)\<bar>"
14.35 + by auto
14.36 + also have "... \<le> e"
14.37 + apply (rule cSup_asclose)
14.38 + apply (auto simp add: S)
14.39 + apply (metis abs_minus_add_cancel b add_commute diff_minus)
14.40 + done
14.41 + finally have "\<bar>- Sup (uminus ` S) - l\<bar> \<le> e" .
14.42 + thus ?thesis
14.43 + by (simp add: Inf_real_def)
14.44 +qed
14.45 +
14.46 +lemma cSup_finite_ge_iff:
14.47 + fixes S :: "real set" shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<le> Sup S \<longleftrightarrow> (\<exists>x\<in>S. a \<le> x)"
14.48 + by (metis cSup_eq_Max Max_ge_iff)
14.49
14.50 lemma cSup_finite_le_iff:
14.51 - fixes S :: "real set"
14.52 - assumes fS: "finite S" and Se: "S \<noteq> {}"
14.53 - shows "a \<ge> Sup S \<longleftrightarrow> (\<forall> x \<in> S. a \<ge> x)"
14.54 -by (metis Max_ge Se cSup_eq_Max Max_in fS order_trans)
14.55 + fixes S :: "real set" shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<ge> Sup S \<longleftrightarrow> (\<forall>x\<in>S. a \<ge> x)"
14.56 + by (metis cSup_eq_Max Max_le_iff)
14.57 +
14.58 +lemma cInf_finite_ge_iff:
14.59 + fixes S :: "real set" shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
14.60 + by (metis cInf_eq_Min Min_ge_iff)
14.61 +
14.62 +lemma cInf_finite_le_iff:
14.63 + fixes S :: "real set" shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<ge> Inf S \<longleftrightarrow> (\<exists>x\<in>S. a \<ge> x)"
14.64 + by (metis cInf_eq_Min Min_le_iff)
14.65
14.66 lemma Inf: (* rename *)
14.67 fixes S :: "real set"
14.68 @@ -6282,7 +6321,8 @@
14.69 unfolding real_norm_def abs_le_iff
14.70 apply auto
14.71 done
14.72 - show "((\<lambda>k. Inf {f j x |j. k \<le> j}) ---> g x) sequentially"
14.73 +
14.74 + show "(\<lambda>k::nat. Inf {f j x |j. k \<le> j}) ----> g x"
14.75 proof (rule LIMSEQ_I)
14.76 case goal1
14.77 hence "0<r/2" by auto
15.1 --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Mar 26 14:38:44 2013 +0100
15.2 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Mar 26 15:10:28 2013 +0100
15.3 @@ -1256,15 +1256,10 @@
15.4
15.5 text {* Some property holds "sufficiently close" to the limit point. *}
15.6
15.7 -lemma eventually_at: (* FIXME: this replaces Metric_Spaces.eventually_at *)
15.8 +lemma eventually_at2:
15.9 "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
15.10 unfolding eventually_at dist_nz by auto
15.11
15.12 -lemma eventually_within: (* FIXME: this replaces Topological_Spaces.eventually_within *)
15.13 - "eventually P (at a within S) \<longleftrightarrow>
15.14 - (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
15.15 - by (rule eventually_within_less)
15.16 -
15.17 lemma eventually_happens: "eventually P net ==> trivial_limit net \<or> (\<exists>x. P x)"
15.18 unfolding trivial_limit_def
15.19 by (auto elim: eventually_rev_mp)
15.20 @@ -1301,11 +1296,11 @@
15.21
15.22 lemma Lim_within: "(f ---> l) (at a within S) \<longleftrightarrow>
15.23 (\<forall>e >0. \<exists>d>0. \<forall>x \<in> S. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)"
15.24 - by (auto simp add: tendsto_iff eventually_within)
15.25 + by (auto simp add: tendsto_iff eventually_within_less)
15.26
15.27 lemma Lim_at: "(f ---> l) (at a) \<longleftrightarrow>
15.28 (\<forall>e >0. \<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)"
15.29 - by (auto simp add: tendsto_iff eventually_at)
15.30 + by (auto simp add: tendsto_iff eventually_at2)
15.31
15.32 lemma Lim_at_infinity:
15.33 "(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x >= b \<longrightarrow> dist (f x) l < e)"
15.34 @@ -1374,34 +1369,33 @@
15.35 by (simp add: sequentially_imp_eventually_within)
15.36
15.37 lemma Lim_right_bound:
15.38 - fixes f :: "real \<Rightarrow> real"
15.39 + fixes f :: "'a :: {linorder_topology, conditional_complete_linorder, no_top} \<Rightarrow>
15.40 + 'b::{linorder_topology, conditional_complete_linorder}"
15.41 assumes mono: "\<And>a b. a \<in> I \<Longrightarrow> b \<in> I \<Longrightarrow> x < a \<Longrightarrow> a \<le> b \<Longrightarrow> f a \<le> f b"
15.42 assumes bnd: "\<And>a. a \<in> I \<Longrightarrow> x < a \<Longrightarrow> K \<le> f a"
15.43 shows "(f ---> Inf (f ` ({x<..} \<inter> I))) (at x within ({x<..} \<inter> I))"
15.44 proof cases
15.45 assume "{x<..} \<inter> I = {}" then show ?thesis by (simp add: Lim_within_empty)
15.46 next
15.47 - assume [simp]: "{x<..} \<inter> I \<noteq> {}"
15.48 + assume e: "{x<..} \<inter> I \<noteq> {}"
15.49 show ?thesis
15.50 - proof (rule Lim_within_LIMSEQ, safe)
15.51 - fix S assume S: "\<forall>n. S n \<noteq> x \<and> S n \<in> {x <..} \<inter> I" "S ----> x"
15.52 -
15.53 - show "(\<lambda>n. f (S n)) ----> Inf (f ` ({x<..} \<inter> I))"
15.54 - proof (rule LIMSEQ_I, rule ccontr)
15.55 - fix r :: real assume "0 < r"
15.56 - with cInf_close[of "f ` ({x<..} \<inter> I)" r]
15.57 - obtain y where y: "x < y" "y \<in> I" "f y < Inf (f ` ({x <..} \<inter> I)) + r" by auto
15.58 - from `x < y` have "0 < y - x" by auto
15.59 - from S(2)[THEN LIMSEQ_D, OF this]
15.60 - obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> \<bar>S n - x\<bar> < y - x" by auto
15.61 -
15.62 - assume "\<not> (\<exists>N. \<forall>n\<ge>N. norm (f (S n) - Inf (f ` ({x<..} \<inter> I))) < r)"
15.63 - moreover have "\<And>n. Inf (f ` ({x<..} \<inter> I)) \<le> f (S n)"
15.64 - using S bnd by (intro cInf_lower[where z=K]) auto
15.65 - ultimately obtain n where n: "N \<le> n" "r + Inf (f ` ({x<..} \<inter> I)) \<le> f (S n)"
15.66 - by (auto simp: not_less field_simps)
15.67 - with N[OF n(1)] mono[OF _ `y \<in> I`, of "S n"] S(1)[THEN spec, of n] y
15.68 - show False by auto
15.69 + proof (rule order_tendstoI)
15.70 + fix a assume a: "a < Inf (f ` ({x<..} \<inter> I))"
15.71 + { fix y assume "y \<in> {x<..} \<inter> I"
15.72 + with e bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> f y"
15.73 + by (auto intro: cInf_lower)
15.74 + with a have "a < f y" by (blast intro: less_le_trans) }
15.75 + then show "eventually (\<lambda>x. a < f x) (at x within ({x<..} \<inter> I))"
15.76 + by (auto simp: Topological_Spaces.eventually_within intro: exI[of _ 1] zero_less_one)
15.77 + next
15.78 + fix a assume "Inf (f ` ({x<..} \<inter> I)) < a"
15.79 + from cInf_lessD[OF _ this] e obtain y where y: "x < y" "y \<in> I" "f y < a" by auto
15.80 + show "eventually (\<lambda>x. f x < a) (at x within ({x<..} \<inter> I))"
15.81 + unfolding within_within_eq[symmetric]
15.82 + Topological_Spaces.eventually_within[of _ _ I] eventually_at_right
15.83 + proof (safe intro!: exI[of _ y] y)
15.84 + fix z assume "x < z" "z \<in> I" "z < y"
15.85 + with mono[OF `z\<in>I` `y\<in>I`] `f y < a` show "f z < a" by (auto simp: less_imp_le)
15.86 qed
15.87 qed
15.88 qed
15.89 @@ -1602,7 +1596,7 @@
15.90 shows "(g ---> l) (at x within S)"
15.91 proof (rule Lim_transform_eventually)
15.92 show "eventually (\<lambda>x. f x = g x) (at x within S)"
15.93 - unfolding eventually_within
15.94 + unfolding eventually_within_less
15.95 using assms(1,2) by auto
15.96 show "(f ---> l) (at x within S)" by fact
15.97 qed
15.98 @@ -1613,7 +1607,7 @@
15.99 shows "(g ---> l) (at x)"
15.100 proof (rule Lim_transform_eventually)
15.101 show "eventually (\<lambda>x. f x = g x) (at x)"
15.102 - unfolding eventually_at
15.103 + unfolding eventually_at2
15.104 using assms(1,2) by auto
15.105 show "(f ---> l) (at x)" by fact
15.106 qed
15.107 @@ -1718,12 +1712,11 @@
15.108 using cInf_lower_EX[of _ S] assms by metis
15.109
15.110 fix e :: real assume "0 < e"
15.111 - then obtain x where x: "x \<in> S" "x < Inf S + e"
15.112 - using cInf_close `S \<noteq> {}` by auto
15.113 - moreover then have "x > Inf S - e" using * by auto
15.114 - ultimately have "abs (x - Inf S) < e" by (simp add: abs_diff_less_iff)
15.115 - then show "\<exists>x\<in>S. dist x (Inf S) < e"
15.116 - using x by (auto simp: dist_norm)
15.117 + then have "Inf S < Inf S + e" by simp
15.118 + with assms obtain x where "x \<in> S" "x < Inf S + e"
15.119 + by (subst (asm) cInf_less_iff[of _ B]) auto
15.120 + with * show "\<exists>x\<in>S. dist x (Inf S) < e"
15.121 + by (intro bexI[of _ x]) (auto simp add: dist_real_def)
15.122 qed
15.123
15.124 lemma closed_contains_Inf:
15.125 @@ -3790,7 +3783,7 @@
15.126 { fix x::"nat \<Rightarrow> 'a" assume x:"\<forall>n. x n \<in> s" "\<forall>e>0. eventually (\<lambda>n. dist (x n) a < e) sequentially"
15.127 fix T::"'b set" assume "open T" and "f a \<in> T"
15.128 with `?lhs` obtain d where "d>0" and d:"\<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> f x \<in> T"
15.129 - unfolding continuous_within tendsto_def eventually_within by auto
15.130 + unfolding continuous_within tendsto_def eventually_within_less by auto
15.131 have "eventually (\<lambda>n. dist (x n) a < d) sequentially"
15.132 using x(2) `d>0` by simp
15.133 hence "eventually (\<lambda>n. (f \<circ> x) n \<in> T) sequentially"
15.134 @@ -4188,8 +4181,7 @@
15.135 hence "eventually (\<lambda>y. f y \<noteq> a) (at x within s)"
15.136 using `a \<notin> U` by (fast elim: eventually_mono [rotated])
15.137 thus ?thesis
15.138 - unfolding Limits.eventually_within Metric_Spaces.eventually_at
15.139 - by (rule ex_forward, cut_tac `f x \<noteq> a`, auto simp: dist_commute)
15.140 + using `f x \<noteq> a` by (auto simp: dist_commute zero_less_dist_iff eventually_within_less)
15.141 qed
15.142
15.143 lemma continuous_at_avoid:
16.1 --- a/src/HOL/NSA/HDeriv.thy Tue Mar 26 14:38:44 2013 +0100
16.2 +++ b/src/HOL/NSA/HDeriv.thy Tue Mar 26 15:10:28 2013 +0100
16.3 @@ -7,7 +7,7 @@
16.4 header{* Differentiation (Nonstandard) *}
16.5
16.6 theory HDeriv
16.7 -imports Deriv HLim
16.8 +imports HLim
16.9 begin
16.10
16.11 text{*Nonstandard Definitions*}
17.1 --- a/src/HOL/NSA/HLim.thy Tue Mar 26 14:38:44 2013 +0100
17.2 +++ b/src/HOL/NSA/HLim.thy Tue Mar 26 15:10:28 2013 +0100
17.3 @@ -6,7 +6,7 @@
17.4 header{* Limits and Continuity (Nonstandard) *}
17.5
17.6 theory HLim
17.7 -imports Star Lim
17.8 +imports Star
17.9 begin
17.10
17.11 text{*Nonstandard Definitions*}
18.1 --- a/src/HOL/NSA/HLog.thy Tue Mar 26 14:38:44 2013 +0100
18.2 +++ b/src/HOL/NSA/HLog.thy Tue Mar 26 15:10:28 2013 +0100
18.3 @@ -6,7 +6,7 @@
18.4 header{*Logarithms: Non-Standard Version*}
18.5
18.6 theory HLog
18.7 -imports Log HTranscendental
18.8 +imports HTranscendental
18.9 begin
18.10
18.11
19.1 --- a/src/HOL/NSA/HSEQ.thy Tue Mar 26 14:38:44 2013 +0100
19.2 +++ b/src/HOL/NSA/HSEQ.thy Tue Mar 26 15:10:28 2013 +0100
19.3 @@ -9,7 +9,7 @@
19.4 header {* Sequences and Convergence (Nonstandard) *}
19.5
19.6 theory HSEQ
19.7 -imports SEQ NatStar
19.8 +imports Limits NatStar
19.9 begin
19.10
19.11 definition
20.1 --- a/src/HOL/NSA/HSeries.thy Tue Mar 26 14:38:44 2013 +0100
20.2 +++ b/src/HOL/NSA/HSeries.thy Tue Mar 26 15:10:28 2013 +0100
20.3 @@ -8,7 +8,7 @@
20.4 header{*Finite Summation and Infinite Series for Hyperreals*}
20.5
20.6 theory HSeries
20.7 -imports Series HSEQ
20.8 +imports HSEQ
20.9 begin
20.10
20.11 definition
21.1 --- a/src/HOL/NSA/HyperDef.thy Tue Mar 26 14:38:44 2013 +0100
21.2 +++ b/src/HOL/NSA/HyperDef.thy Tue Mar 26 15:10:28 2013 +0100
21.3 @@ -7,7 +7,7 @@
21.4 header{*Construction of Hyperreals Using Ultrafilters*}
21.5
21.6 theory HyperDef
21.7 -imports HyperNat Real
21.8 +imports Complex_Main HyperNat
21.9 begin
21.10
21.11 type_synonym hypreal = "real star"
22.1 --- a/src/HOL/NSA/Hyperreal.thy Tue Mar 26 14:38:44 2013 +0100
22.2 +++ b/src/HOL/NSA/Hyperreal.thy Tue Mar 26 15:10:28 2013 +0100
22.3 @@ -7,7 +7,7 @@
22.4 *)
22.5
22.6 theory Hyperreal
22.7 -imports Ln Deriv Taylor HLog
22.8 +imports HLog
22.9 begin
22.10
22.11 end
23.1 --- a/src/HOL/NSA/NSA.thy Tue Mar 26 14:38:44 2013 +0100
23.2 +++ b/src/HOL/NSA/NSA.thy Tue Mar 26 15:10:28 2013 +0100
23.3 @@ -6,7 +6,7 @@
23.4 header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*}
23.5
23.6 theory NSA
23.7 -imports HyperDef RComplete
23.8 +imports HyperDef
23.9 begin
23.10
23.11 definition
24.1 --- a/src/HOL/NSA/NSComplex.thy Tue Mar 26 14:38:44 2013 +0100
24.2 +++ b/src/HOL/NSA/NSComplex.thy Tue Mar 26 15:10:28 2013 +0100
24.3 @@ -6,7 +6,7 @@
24.4 header{*Nonstandard Complex Numbers*}
24.5
24.6 theory NSComplex
24.7 -imports Complex NSA
24.8 +imports NSA
24.9 begin
24.10
24.11 type_synonym hcomplex = "complex star"
25.1 --- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Mar 26 14:38:44 2013 +0100
25.2 +++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Mar 26 15:10:28 2013 +0100
25.3 @@ -12,7 +12,7 @@
25.4 suite. *)
25.5
25.6 theory Manual_Nits
25.7 -imports Main "~~/src/HOL/Library/Quotient_Product" RealDef
25.8 +imports Main "~~/src/HOL/Library/Quotient_Product" Real
25.9 begin
25.10
25.11 chapter {* 2. First Steps *}
26.1 --- a/src/HOL/Quickcheck_Benchmark/Find_Unused_Assms_Examples.thy Tue Mar 26 14:38:44 2013 +0100
26.2 +++ b/src/HOL/Quickcheck_Benchmark/Find_Unused_Assms_Examples.thy Tue Mar 26 15:10:28 2013 +0100
26.3 @@ -8,8 +8,7 @@
26.4
26.5 find_unused_assms Divides
26.6 find_unused_assms GCD
26.7 -find_unused_assms RealDef
26.8 -find_unused_assms RComplete
26.9 +find_unused_assms Real
26.10
26.11 section {* Set Theory *}
26.12
27.1 --- a/src/HOL/RComplete.thy Tue Mar 26 14:38:44 2013 +0100
27.2 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000
27.3 @@ -1,544 +0,0 @@
27.4 -(* Title: HOL/RComplete.thy
27.5 - Author: Jacques D. Fleuriot, University of Edinburgh
27.6 - Author: Larry Paulson, University of Cambridge
27.7 - Author: Jeremy Avigad, Carnegie Mellon University
27.8 - Author: Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen
27.9 -*)
27.10 -
27.11 -header {* Completeness of the Reals; Floor and Ceiling Functions *}
27.12 -
27.13 -theory RComplete
27.14 -imports Lubs RealDef
27.15 -begin
27.16 -
27.17 -lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
27.18 - by simp
27.19 -
27.20 -lemma abs_diff_less_iff:
27.21 - "(\<bar>x - a\<bar> < (r::'a::linordered_idom)) = (a - r < x \<and> x < a + r)"
27.22 - by auto
27.23 -
27.24 -subsection {* Completeness of Positive Reals *}
27.25 -
27.26 -text {*
27.27 - Supremum property for the set of positive reals
27.28 -
27.29 - Let @{text "P"} be a non-empty set of positive reals, with an upper
27.30 - bound @{text "y"}. Then @{text "P"} has a least upper bound
27.31 - (written @{text "S"}).
27.32 -
27.33 - FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
27.34 -*}
27.35 -
27.36 -text {* Only used in HOL/Import/HOL4Compat.thy; delete? *}
27.37 -
27.38 -lemma posreal_complete:
27.39 - fixes P :: "real set"
27.40 - assumes not_empty_P: "\<exists>x. x \<in> P"
27.41 - and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
27.42 - shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
27.43 -proof -
27.44 - from upper_bound_Ex have "\<exists>z. \<forall>x\<in>P. x \<le> z"
27.45 - by (auto intro: less_imp_le)
27.46 - from complete_real [OF not_empty_P this] obtain S
27.47 - where S1: "\<And>x. x \<in> P \<Longrightarrow> x \<le> S" and S2: "\<And>z. \<forall>x\<in>P. x \<le> z \<Longrightarrow> S \<le> z" by fast
27.48 - have "\<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
27.49 - proof
27.50 - fix y show "(\<exists>x\<in>P. y < x) = (y < S)"
27.51 - apply (cases "\<exists>x\<in>P. y < x", simp_all)
27.52 - apply (clarify, drule S1, simp)
27.53 - apply (simp add: not_less S2)
27.54 - done
27.55 - qed
27.56 - thus ?thesis ..
27.57 -qed
27.58 -
27.59 -text {*
27.60 - \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
27.61 -*}
27.62 -
27.63 -lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)"
27.64 - apply (frule isLub_isUb)
27.65 - apply (frule_tac x = y in isLub_isUb)
27.66 - apply (blast intro!: order_antisym dest!: isLub_le_isUb)
27.67 - done
27.68 -
27.69 -
27.70 -text {*
27.71 - \medskip reals Completeness (again!)
27.72 -*}
27.73 -
27.74 -lemma reals_complete:
27.75 - assumes notempty_S: "\<exists>X. X \<in> S"
27.76 - and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
27.77 - shows "\<exists>t. isLub (UNIV :: real set) S t"
27.78 -proof -
27.79 - from assms have "\<exists>X. X \<in> S" and "\<exists>Y. \<forall>x\<in>S. x \<le> Y"
27.80 - unfolding isUb_def setle_def by simp_all
27.81 - from complete_real [OF this] show ?thesis
27.82 - by (simp add: isLub_def leastP_def isUb_def setle_def setge_def)
27.83 -qed
27.84 -
27.85 -
27.86 -subsection {* The Archimedean Property of the Reals *}
27.87 -
27.88 -theorem reals_Archimedean:
27.89 - assumes x_pos: "0 < x"
27.90 - shows "\<exists>n. inverse (real (Suc n)) < x"
27.91 - unfolding real_of_nat_def using x_pos
27.92 - by (rule ex_inverse_of_nat_Suc_less)
27.93 -
27.94 -lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
27.95 - unfolding real_of_nat_def by (rule ex_less_of_nat)
27.96 -
27.97 -lemma reals_Archimedean3:
27.98 - assumes x_greater_zero: "0 < x"
27.99 - shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"
27.100 - unfolding real_of_nat_def using `0 < x`
27.101 - by (auto intro: ex_less_of_nat_mult)
27.102 -
27.103 -
27.104 -subsection{*Density of the Rational Reals in the Reals*}
27.105 -
27.106 -text{* This density proof is due to Stefan Richter and was ported by TN. The
27.107 -original source is \emph{Real Analysis} by H.L. Royden.
27.108 -It employs the Archimedean property of the reals. *}
27.109 -
27.110 -lemma Rats_dense_in_real:
27.111 - fixes x :: real
27.112 - assumes "x < y" shows "\<exists>r\<in>\<rat>. x < r \<and> r < y"
27.113 -proof -
27.114 - from `x<y` have "0 < y-x" by simp
27.115 - with reals_Archimedean obtain q::nat
27.116 - where q: "inverse (real q) < y-x" and "0 < q" by auto
27.117 - def p \<equiv> "ceiling (y * real q) - 1"
27.118 - def r \<equiv> "of_int p / real q"
27.119 - from q have "x < y - inverse (real q)" by simp
27.120 - also have "y - inverse (real q) \<le> r"
27.121 - unfolding r_def p_def
27.122 - by (simp add: le_divide_eq left_diff_distrib le_of_int_ceiling `0 < q`)
27.123 - finally have "x < r" .
27.124 - moreover have "r < y"
27.125 - unfolding r_def p_def
27.126 - by (simp add: divide_less_eq diff_less_eq `0 < q`
27.127 - less_ceiling_iff [symmetric])
27.128 - moreover from r_def have "r \<in> \<rat>" by simp
27.129 - ultimately show ?thesis by fast
27.130 -qed
27.131 -
27.132 -
27.133 -subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
27.134 -
27.135 -(* FIXME: theorems for negative numerals *)
27.136 -lemma numeral_less_real_of_int_iff [simp]:
27.137 - "((numeral n) < real (m::int)) = (numeral n < m)"
27.138 -apply auto
27.139 -apply (rule real_of_int_less_iff [THEN iffD1])
27.140 -apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
27.141 -done
27.142 -
27.143 -lemma numeral_less_real_of_int_iff2 [simp]:
27.144 - "(real (m::int) < (numeral n)) = (m < numeral n)"
27.145 -apply auto
27.146 -apply (rule real_of_int_less_iff [THEN iffD1])
27.147 -apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
27.148 -done
27.149 -
27.150 -lemma numeral_le_real_of_int_iff [simp]:
27.151 - "((numeral n) \<le> real (m::int)) = (numeral n \<le> m)"
27.152 -by (simp add: linorder_not_less [symmetric])
27.153 -
27.154 -lemma numeral_le_real_of_int_iff2 [simp]:
27.155 - "(real (m::int) \<le> (numeral n)) = (m \<le> numeral n)"
27.156 -by (simp add: linorder_not_less [symmetric])
27.157 -
27.158 -lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"
27.159 -unfolding real_of_nat_def by simp
27.160 -
27.161 -lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"
27.162 -unfolding real_of_nat_def by (simp add: floor_minus)
27.163 -
27.164 -lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"
27.165 -unfolding real_of_int_def by simp
27.166 -
27.167 -lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"
27.168 -unfolding real_of_int_def by (simp add: floor_minus)
27.169 -
27.170 -lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
27.171 -unfolding real_of_int_def by (rule floor_exists)
27.172 -
27.173 -lemma lemma_floor:
27.174 - assumes a1: "real m \<le> r" and a2: "r < real n + 1"
27.175 - shows "m \<le> (n::int)"
27.176 -proof -
27.177 - have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)
27.178 - also have "... = real (n + 1)" by simp
27.179 - finally have "m < n + 1" by (simp only: real_of_int_less_iff)
27.180 - thus ?thesis by arith
27.181 -qed
27.182 -
27.183 -lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"
27.184 -unfolding real_of_int_def by (rule of_int_floor_le)
27.185 -
27.186 -lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \<le> x"
27.187 -by (auto intro: lemma_floor)
27.188 -
27.189 -lemma real_of_int_floor_cancel [simp]:
27.190 - "(real (floor x) = x) = (\<exists>n::int. x = real n)"
27.191 - using floor_real_of_int by metis
27.192 -
27.193 -lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
27.194 - unfolding real_of_int_def using floor_unique [of n x] by simp
27.195 -
27.196 -lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"
27.197 - unfolding real_of_int_def by (rule floor_unique)
27.198 -
27.199 -lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n"
27.200 -apply (rule inj_int [THEN injD])
27.201 -apply (simp add: real_of_nat_Suc)
27.202 -apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"])
27.203 -done
27.204 -
27.205 -lemma floor_eq4: "[| real n \<le> x; x < real (Suc n) |] ==> nat(floor x) = n"
27.206 -apply (drule order_le_imp_less_or_eq)
27.207 -apply (auto intro: floor_eq3)
27.208 -done
27.209 -
27.210 -lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
27.211 - unfolding real_of_int_def using floor_correct [of r] by simp
27.212 -
27.213 -lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"
27.214 - unfolding real_of_int_def using floor_correct [of r] by simp
27.215 -
27.216 -lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"
27.217 - unfolding real_of_int_def using floor_correct [of r] by simp
27.218 -
27.219 -lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"
27.220 - unfolding real_of_int_def using floor_correct [of r] by simp
27.221 -
27.222 -lemma le_floor: "real a <= x ==> a <= floor x"
27.223 - unfolding real_of_int_def by (simp add: le_floor_iff)
27.224 -
27.225 -lemma real_le_floor: "a <= floor x ==> real a <= x"
27.226 - unfolding real_of_int_def by (simp add: le_floor_iff)
27.227 -
27.228 -lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
27.229 - unfolding real_of_int_def by (rule le_floor_iff)
27.230 -
27.231 -lemma floor_less_eq: "(floor x < a) = (x < real a)"
27.232 - unfolding real_of_int_def by (rule floor_less_iff)
27.233 -
27.234 -lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
27.235 - unfolding real_of_int_def by (rule less_floor_iff)
27.236 -
27.237 -lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
27.238 - unfolding real_of_int_def by (rule floor_le_iff)
27.239 -
27.240 -lemma floor_add [simp]: "floor (x + real a) = floor x + a"
27.241 - unfolding real_of_int_def by (rule floor_add_of_int)
27.242 -
27.243 -lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
27.244 - unfolding real_of_int_def by (rule floor_diff_of_int)
27.245 -
27.246 -lemma le_mult_floor:
27.247 - assumes "0 \<le> (a :: real)" and "0 \<le> b"
27.248 - shows "floor a * floor b \<le> floor (a * b)"
27.249 -proof -
27.250 - have "real (floor a) \<le> a"
27.251 - and "real (floor b) \<le> b" by auto
27.252 - hence "real (floor a * floor b) \<le> a * b"
27.253 - using assms by (auto intro!: mult_mono)
27.254 - also have "a * b < real (floor (a * b) + 1)" by auto
27.255 - finally show ?thesis unfolding real_of_int_less_iff by simp
27.256 -qed
27.257 -
27.258 -lemma floor_divide_eq_div:
27.259 - "floor (real a / real b) = a div b"
27.260 -proof cases
27.261 - assume "b \<noteq> 0 \<or> b dvd a"
27.262 - with real_of_int_div3[of a b] show ?thesis
27.263 - by (auto simp: real_of_int_div[symmetric] intro!: floor_eq2 real_of_int_div4 neq_le_trans)
27.264 - (metis add_left_cancel zero_neq_one real_of_int_div_aux real_of_int_inject
27.265 - real_of_int_zero_cancel right_inverse_eq div_self mod_div_trivial)
27.266 -qed (auto simp: real_of_int_div)
27.267 -
27.268 -lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
27.269 - unfolding real_of_nat_def by simp
27.270 -
27.271 -lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"
27.272 - unfolding real_of_int_def by (rule le_of_int_ceiling)
27.273 -
27.274 -lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"
27.275 - unfolding real_of_int_def by simp
27.276 -
27.277 -lemma real_of_int_ceiling_cancel [simp]:
27.278 - "(real (ceiling x) = x) = (\<exists>n::int. x = real n)"
27.279 - using ceiling_real_of_int by metis
27.280 -
27.281 -lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"
27.282 - unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
27.283 -
27.284 -lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"
27.285 - unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
27.286 -
27.287 -lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n |] ==> ceiling x = n"
27.288 - unfolding real_of_int_def using ceiling_unique [of n x] by simp
27.289 -
27.290 -lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
27.291 - unfolding real_of_int_def using ceiling_correct [of r] by simp
27.292 -
27.293 -lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"
27.294 - unfolding real_of_int_def using ceiling_correct [of r] by simp
27.295 -
27.296 -lemma ceiling_le: "x <= real a ==> ceiling x <= a"
27.297 - unfolding real_of_int_def by (simp add: ceiling_le_iff)
27.298 -
27.299 -lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"
27.300 - unfolding real_of_int_def by (simp add: ceiling_le_iff)
27.301 -
27.302 -lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
27.303 - unfolding real_of_int_def by (rule ceiling_le_iff)
27.304 -
27.305 -lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
27.306 - unfolding real_of_int_def by (rule less_ceiling_iff)
27.307 -
27.308 -lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
27.309 - unfolding real_of_int_def by (rule ceiling_less_iff)
27.310 -
27.311 -lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
27.312 - unfolding real_of_int_def by (rule le_ceiling_iff)
27.313 -
27.314 -lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
27.315 - unfolding real_of_int_def by (rule ceiling_add_of_int)
27.316 -
27.317 -lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
27.318 - unfolding real_of_int_def by (rule ceiling_diff_of_int)
27.319 -
27.320 -
27.321 -subsection {* Versions for the natural numbers *}
27.322 -
27.323 -definition
27.324 - natfloor :: "real => nat" where
27.325 - "natfloor x = nat(floor x)"
27.326 -
27.327 -definition
27.328 - natceiling :: "real => nat" where
27.329 - "natceiling x = nat(ceiling x)"
27.330 -
27.331 -lemma natfloor_zero [simp]: "natfloor 0 = 0"
27.332 - by (unfold natfloor_def, simp)
27.333 -
27.334 -lemma natfloor_one [simp]: "natfloor 1 = 1"
27.335 - by (unfold natfloor_def, simp)
27.336 -
27.337 -lemma zero_le_natfloor [simp]: "0 <= natfloor x"
27.338 - by (unfold natfloor_def, simp)
27.339 -
27.340 -lemma natfloor_numeral_eq [simp]: "natfloor (numeral n) = numeral n"
27.341 - by (unfold natfloor_def, simp)
27.342 -
27.343 -lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"
27.344 - by (unfold natfloor_def, simp)
27.345 -
27.346 -lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"
27.347 - by (unfold natfloor_def, simp)
27.348 -
27.349 -lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"
27.350 - unfolding natfloor_def by simp
27.351 -
27.352 -lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
27.353 - unfolding natfloor_def by (intro nat_mono floor_mono)
27.354 -
27.355 -lemma le_natfloor: "real x <= a ==> x <= natfloor a"
27.356 - apply (unfold natfloor_def)
27.357 - apply (subst nat_int [THEN sym])
27.358 - apply (rule nat_mono)
27.359 - apply (rule le_floor)
27.360 - apply simp
27.361 -done
27.362 -
27.363 -lemma natfloor_less_iff: "0 \<le> x \<Longrightarrow> natfloor x < n \<longleftrightarrow> x < real n"
27.364 - unfolding natfloor_def real_of_nat_def
27.365 - by (simp add: nat_less_iff floor_less_iff)
27.366 -
27.367 -lemma less_natfloor:
27.368 - assumes "0 \<le> x" and "x < real (n :: nat)"
27.369 - shows "natfloor x < n"
27.370 - using assms by (simp add: natfloor_less_iff)
27.371 -
27.372 -lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)"
27.373 - apply (rule iffI)
27.374 - apply (rule order_trans)
27.375 - prefer 2
27.376 - apply (erule real_natfloor_le)
27.377 - apply (subst real_of_nat_le_iff)
27.378 - apply assumption
27.379 - apply (erule le_natfloor)
27.380 -done
27.381 -
27.382 -lemma le_natfloor_eq_numeral [simp]:
27.383 - "~ neg((numeral n)::int) ==> 0 <= x ==>
27.384 - (numeral n <= natfloor x) = (numeral n <= x)"
27.385 - apply (subst le_natfloor_eq, assumption)
27.386 - apply simp
27.387 -done
27.388 -
27.389 -lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
27.390 - apply (case_tac "0 <= x")
27.391 - apply (subst le_natfloor_eq, assumption, simp)
27.392 - apply (rule iffI)
27.393 - apply (subgoal_tac "natfloor x <= natfloor 0")
27.394 - apply simp
27.395 - apply (rule natfloor_mono)
27.396 - apply simp
27.397 - apply simp
27.398 -done
27.399 -
27.400 -lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n"
27.401 - unfolding natfloor_def by (simp add: floor_eq2 [where n="int n"])
27.402 -
27.403 -lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"
27.404 - apply (case_tac "0 <= x")
27.405 - apply (unfold natfloor_def)
27.406 - apply simp
27.407 - apply simp_all
27.408 -done
27.409 -
27.410 -lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
27.411 -using real_natfloor_add_one_gt by (simp add: algebra_simps)
27.412 -
27.413 -lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"
27.414 - apply (subgoal_tac "z < real(natfloor z) + 1")
27.415 - apply arith
27.416 - apply (rule real_natfloor_add_one_gt)
27.417 -done
27.418 -
27.419 -lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a"
27.420 - unfolding natfloor_def
27.421 - unfolding real_of_int_of_nat_eq [symmetric] floor_add
27.422 - by (simp add: nat_add_distrib)
27.423 -
27.424 -lemma natfloor_add_numeral [simp]:
27.425 - "~neg ((numeral n)::int) ==> 0 <= x ==>
27.426 - natfloor (x + numeral n) = natfloor x + numeral n"
27.427 - by (simp add: natfloor_add [symmetric])
27.428 -
27.429 -lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1"
27.430 - by (simp add: natfloor_add [symmetric] del: One_nat_def)
27.431 -
27.432 -lemma natfloor_subtract [simp]:
27.433 - "natfloor(x - real a) = natfloor x - a"
27.434 - unfolding natfloor_def real_of_int_of_nat_eq [symmetric] floor_subtract
27.435 - by simp
27.436 -
27.437 -lemma natfloor_div_nat:
27.438 - assumes "1 <= x" and "y > 0"
27.439 - shows "natfloor (x / real y) = natfloor x div y"
27.440 -proof (rule natfloor_eq)
27.441 - have "(natfloor x) div y * y \<le> natfloor x"
27.442 - by (rule add_leD1 [where k="natfloor x mod y"], simp)
27.443 - thus "real (natfloor x div y) \<le> x / real y"
27.444 - using assms by (simp add: le_divide_eq le_natfloor_eq)
27.445 - have "natfloor x < (natfloor x) div y * y + y"
27.446 - apply (subst mod_div_equality [symmetric])
27.447 - apply (rule add_strict_left_mono)
27.448 - apply (rule mod_less_divisor)
27.449 - apply fact
27.450 - done
27.451 - thus "x / real y < real (natfloor x div y) + 1"
27.452 - using assms
27.453 - by (simp add: divide_less_eq natfloor_less_iff distrib_right)
27.454 -qed
27.455 -
27.456 -lemma le_mult_natfloor:
27.457 - shows "natfloor a * natfloor b \<le> natfloor (a * b)"
27.458 - by (cases "0 <= a & 0 <= b")
27.459 - (auto simp add: le_natfloor_eq mult_nonneg_nonneg mult_mono' real_natfloor_le natfloor_neg)
27.460 -
27.461 -lemma natceiling_zero [simp]: "natceiling 0 = 0"
27.462 - by (unfold natceiling_def, simp)
27.463 -
27.464 -lemma natceiling_one [simp]: "natceiling 1 = 1"
27.465 - by (unfold natceiling_def, simp)
27.466 -
27.467 -lemma zero_le_natceiling [simp]: "0 <= natceiling x"
27.468 - by (unfold natceiling_def, simp)
27.469 -
27.470 -lemma natceiling_numeral_eq [simp]: "natceiling (numeral n) = numeral n"
27.471 - by (unfold natceiling_def, simp)
27.472 -
27.473 -lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"
27.474 - by (unfold natceiling_def, simp)
27.475 -
27.476 -lemma real_natceiling_ge: "x <= real(natceiling x)"
27.477 - unfolding natceiling_def by (cases "x < 0", simp_all)
27.478 -
27.479 -lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"
27.480 - unfolding natceiling_def by simp
27.481 -
27.482 -lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"
27.483 - unfolding natceiling_def by (intro nat_mono ceiling_mono)
27.484 -
27.485 -lemma natceiling_le: "x <= real a ==> natceiling x <= a"
27.486 - unfolding natceiling_def real_of_nat_def
27.487 - by (simp add: nat_le_iff ceiling_le_iff)
27.488 -
27.489 -lemma natceiling_le_eq: "(natceiling x <= a) = (x <= real a)"
27.490 - unfolding natceiling_def real_of_nat_def
27.491 - by (simp add: nat_le_iff ceiling_le_iff)
27.492 -
27.493 -lemma natceiling_le_eq_numeral [simp]:
27.494 - "~ neg((numeral n)::int) ==>
27.495 - (natceiling x <= numeral n) = (x <= numeral n)"
27.496 - by (simp add: natceiling_le_eq)
27.497 -
27.498 -lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"
27.499 - unfolding natceiling_def
27.500 - by (simp add: nat_le_iff ceiling_le_iff)
27.501 -
27.502 -lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1"
27.503 - unfolding natceiling_def
27.504 - by (simp add: ceiling_eq2 [where n="int n"])
27.505 -
27.506 -lemma natceiling_add [simp]: "0 <= x ==>
27.507 - natceiling (x + real a) = natceiling x + a"
27.508 - unfolding natceiling_def
27.509 - unfolding real_of_int_of_nat_eq [symmetric] ceiling_add
27.510 - by (simp add: nat_add_distrib)
27.511 -
27.512 -lemma natceiling_add_numeral [simp]:
27.513 - "~ neg ((numeral n)::int) ==> 0 <= x ==>
27.514 - natceiling (x + numeral n) = natceiling x + numeral n"
27.515 - by (simp add: natceiling_add [symmetric])
27.516 -
27.517 -lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1"
27.518 - by (simp add: natceiling_add [symmetric] del: One_nat_def)
27.519 -
27.520 -lemma natceiling_subtract [simp]: "natceiling(x - real a) = natceiling x - a"
27.521 - unfolding natceiling_def real_of_int_of_nat_eq [symmetric] ceiling_subtract
27.522 - by simp
27.523 -
27.524 -subsection {* Exponentiation with floor *}
27.525 -
27.526 -lemma floor_power:
27.527 - assumes "x = real (floor x)"
27.528 - shows "floor (x ^ n) = floor x ^ n"
27.529 -proof -
27.530 - have *: "x ^ n = real (floor x ^ n)"
27.531 - using assms by (induct n arbitrary: x) simp_all
27.532 - show ?thesis unfolding real_of_int_inject[symmetric]
27.533 - unfolding * floor_real_of_int ..
27.534 -qed
27.535 -
27.536 -lemma natfloor_power:
27.537 - assumes "x = real (natfloor x)"
27.538 - shows "natfloor (x ^ n) = natfloor x ^ n"
27.539 -proof -
27.540 - from assms have "0 \<le> floor x" by auto
27.541 - note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \<le> floor x`]]
27.542 - from floor_power[OF this]
27.543 - show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \<le> floor x`, symmetric]
27.544 - by simp
27.545 -qed
27.546 -
27.547 -end
28.1 --- a/src/HOL/Real.thy Tue Mar 26 14:38:44 2013 +0100
28.2 +++ b/src/HOL/Real.thy Tue Mar 26 15:10:28 2013 +0100
28.3 @@ -1,7 +1,2228 @@
28.4 +(* Title: HOL/Real.thy
28.5 + Author: Jacques D. Fleuriot, University of Edinburgh, 1998
28.6 + Author: Larry Paulson, University of Cambridge
28.7 + Author: Jeremy Avigad, Carnegie Mellon University
28.8 + Author: Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen
28.9 + Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
28.10 + Construction of Cauchy Reals by Brian Huffman, 2010
28.11 +*)
28.12 +
28.13 +header {* Development of the Reals using Cauchy Sequences *}
28.14 +
28.15 theory Real
28.16 -imports RComplete RealVector
28.17 +imports Rat Conditional_Complete_Lattices
28.18 +begin
28.19 +
28.20 +text {*
28.21 + This theory contains a formalization of the real numbers as
28.22 + equivalence classes of Cauchy sequences of rationals. See
28.23 + @{file "~~/src/HOL/ex/Dedekind_Real.thy"} for an alternative
28.24 + construction using Dedekind cuts.
28.25 +*}
28.26 +
28.27 +subsection {* Preliminary lemmas *}
28.28 +
28.29 +lemma add_diff_add:
28.30 + fixes a b c d :: "'a::ab_group_add"
28.31 + shows "(a + c) - (b + d) = (a - b) + (c - d)"
28.32 + by simp
28.33 +
28.34 +lemma minus_diff_minus:
28.35 + fixes a b :: "'a::ab_group_add"
28.36 + shows "- a - - b = - (a - b)"
28.37 + by simp
28.38 +
28.39 +lemma mult_diff_mult:
28.40 + fixes x y a b :: "'a::ring"
28.41 + shows "(x * y - a * b) = x * (y - b) + (x - a) * b"
28.42 + by (simp add: algebra_simps)
28.43 +
28.44 +lemma inverse_diff_inverse:
28.45 + fixes a b :: "'a::division_ring"
28.46 + assumes "a \<noteq> 0" and "b \<noteq> 0"
28.47 + shows "inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
28.48 + using assms by (simp add: algebra_simps)
28.49 +
28.50 +lemma obtain_pos_sum:
28.51 + fixes r :: rat assumes r: "0 < r"
28.52 + obtains s t where "0 < s" and "0 < t" and "r = s + t"
28.53 +proof
28.54 + from r show "0 < r/2" by simp
28.55 + from r show "0 < r/2" by simp
28.56 + show "r = r/2 + r/2" by simp
28.57 +qed
28.58 +
28.59 +subsection {* Sequences that converge to zero *}
28.60 +
28.61 +definition
28.62 + vanishes :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
28.63 +where
28.64 + "vanishes X = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r)"
28.65 +
28.66 +lemma vanishesI: "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r) \<Longrightarrow> vanishes X"
28.67 + unfolding vanishes_def by simp
28.68 +
28.69 +lemma vanishesD: "\<lbrakk>vanishes X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r"
28.70 + unfolding vanishes_def by simp
28.71 +
28.72 +lemma vanishes_const [simp]: "vanishes (\<lambda>n. c) \<longleftrightarrow> c = 0"
28.73 + unfolding vanishes_def
28.74 + apply (cases "c = 0", auto)
28.75 + apply (rule exI [where x="\<bar>c\<bar>"], auto)
28.76 + done
28.77 +
28.78 +lemma vanishes_minus: "vanishes X \<Longrightarrow> vanishes (\<lambda>n. - X n)"
28.79 + unfolding vanishes_def by simp
28.80 +
28.81 +lemma vanishes_add:
28.82 + assumes X: "vanishes X" and Y: "vanishes Y"
28.83 + shows "vanishes (\<lambda>n. X n + Y n)"
28.84 +proof (rule vanishesI)
28.85 + fix r :: rat assume "0 < r"
28.86 + then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
28.87 + by (rule obtain_pos_sum)
28.88 + obtain i where i: "\<forall>n\<ge>i. \<bar>X n\<bar> < s"
28.89 + using vanishesD [OF X s] ..
28.90 + obtain j where j: "\<forall>n\<ge>j. \<bar>Y n\<bar> < t"
28.91 + using vanishesD [OF Y t] ..
28.92 + have "\<forall>n\<ge>max i j. \<bar>X n + Y n\<bar> < r"
28.93 + proof (clarsimp)
28.94 + fix n assume n: "i \<le> n" "j \<le> n"
28.95 + have "\<bar>X n + Y n\<bar> \<le> \<bar>X n\<bar> + \<bar>Y n\<bar>" by (rule abs_triangle_ineq)
28.96 + also have "\<dots> < s + t" by (simp add: add_strict_mono i j n)
28.97 + finally show "\<bar>X n + Y n\<bar> < r" unfolding r .
28.98 + qed
28.99 + thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n + Y n\<bar> < r" ..
28.100 +qed
28.101 +
28.102 +lemma vanishes_diff:
28.103 + assumes X: "vanishes X" and Y: "vanishes Y"
28.104 + shows "vanishes (\<lambda>n. X n - Y n)"
28.105 +unfolding diff_minus by (intro vanishes_add vanishes_minus X Y)
28.106 +
28.107 +lemma vanishes_mult_bounded:
28.108 + assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
28.109 + assumes Y: "vanishes (\<lambda>n. Y n)"
28.110 + shows "vanishes (\<lambda>n. X n * Y n)"
28.111 +proof (rule vanishesI)
28.112 + fix r :: rat assume r: "0 < r"
28.113 + obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
28.114 + using X by fast
28.115 + obtain b where b: "0 < b" "r = a * b"
28.116 + proof
28.117 + show "0 < r / a" using r a by (simp add: divide_pos_pos)
28.118 + show "r = a * (r / a)" using a by simp
28.119 + qed
28.120 + obtain k where k: "\<forall>n\<ge>k. \<bar>Y n\<bar> < b"
28.121 + using vanishesD [OF Y b(1)] ..
28.122 + have "\<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r"
28.123 + by (simp add: b(2) abs_mult mult_strict_mono' a k)
28.124 + thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r" ..
28.125 +qed
28.126 +
28.127 +subsection {* Cauchy sequences *}
28.128 +
28.129 +definition
28.130 + cauchy :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
28.131 +where
28.132 + "cauchy X \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r)"
28.133 +
28.134 +lemma cauchyI:
28.135 + "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r) \<Longrightarrow> cauchy X"
28.136 + unfolding cauchy_def by simp
28.137 +
28.138 +lemma cauchyD:
28.139 + "\<lbrakk>cauchy X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r"
28.140 + unfolding cauchy_def by simp
28.141 +
28.142 +lemma cauchy_const [simp]: "cauchy (\<lambda>n. x)"
28.143 + unfolding cauchy_def by simp
28.144 +
28.145 +lemma cauchy_add [simp]:
28.146 + assumes X: "cauchy X" and Y: "cauchy Y"
28.147 + shows "cauchy (\<lambda>n. X n + Y n)"
28.148 +proof (rule cauchyI)
28.149 + fix r :: rat assume "0 < r"
28.150 + then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
28.151 + by (rule obtain_pos_sum)
28.152 + obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
28.153 + using cauchyD [OF X s] ..
28.154 + obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
28.155 + using cauchyD [OF Y t] ..
28.156 + have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r"
28.157 + proof (clarsimp)
28.158 + fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
28.159 + have "\<bar>(X m + Y m) - (X n + Y n)\<bar> \<le> \<bar>X m - X n\<bar> + \<bar>Y m - Y n\<bar>"
28.160 + unfolding add_diff_add by (rule abs_triangle_ineq)
28.161 + also have "\<dots> < s + t"
28.162 + by (rule add_strict_mono, simp_all add: i j *)
28.163 + finally show "\<bar>(X m + Y m) - (X n + Y n)\<bar> < r" unfolding r .
28.164 + qed
28.165 + thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r" ..
28.166 +qed
28.167 +
28.168 +lemma cauchy_minus [simp]:
28.169 + assumes X: "cauchy X"
28.170 + shows "cauchy (\<lambda>n. - X n)"
28.171 +using assms unfolding cauchy_def
28.172 +unfolding minus_diff_minus abs_minus_cancel .
28.173 +
28.174 +lemma cauchy_diff [simp]:
28.175 + assumes X: "cauchy X" and Y: "cauchy Y"
28.176 + shows "cauchy (\<lambda>n. X n - Y n)"
28.177 +using assms unfolding diff_minus by simp
28.178 +
28.179 +lemma cauchy_imp_bounded:
28.180 + assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
28.181 +proof -
28.182 + obtain k where k: "\<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < 1"
28.183 + using cauchyD [OF assms zero_less_one] ..
28.184 + show "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
28.185 + proof (intro exI conjI allI)
28.186 + have "0 \<le> \<bar>X 0\<bar>" by simp
28.187 + also have "\<bar>X 0\<bar> \<le> Max (abs ` X ` {..k})" by simp
28.188 + finally have "0 \<le> Max (abs ` X ` {..k})" .
28.189 + thus "0 < Max (abs ` X ` {..k}) + 1" by simp
28.190 + next
28.191 + fix n :: nat
28.192 + show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1"
28.193 + proof (rule linorder_le_cases)
28.194 + assume "n \<le> k"
28.195 + hence "\<bar>X n\<bar> \<le> Max (abs ` X ` {..k})" by simp
28.196 + thus "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" by simp
28.197 + next
28.198 + assume "k \<le> n"
28.199 + have "\<bar>X n\<bar> = \<bar>X k + (X n - X k)\<bar>" by simp
28.200 + also have "\<bar>X k + (X n - X k)\<bar> \<le> \<bar>X k\<bar> + \<bar>X n - X k\<bar>"
28.201 + by (rule abs_triangle_ineq)
28.202 + also have "\<dots> < Max (abs ` X ` {..k}) + 1"
28.203 + by (rule add_le_less_mono, simp, simp add: k `k \<le> n`)
28.204 + finally show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" .
28.205 + qed
28.206 + qed
28.207 +qed
28.208 +
28.209 +lemma cauchy_mult [simp]:
28.210 + assumes X: "cauchy X" and Y: "cauchy Y"
28.211 + shows "cauchy (\<lambda>n. X n * Y n)"
28.212 +proof (rule cauchyI)
28.213 + fix r :: rat assume "0 < r"
28.214 + then obtain u v where u: "0 < u" and v: "0 < v" and "r = u + v"
28.215 + by (rule obtain_pos_sum)
28.216 + obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
28.217 + using cauchy_imp_bounded [OF X] by fast
28.218 + obtain b where b: "0 < b" "\<forall>n. \<bar>Y n\<bar> < b"
28.219 + using cauchy_imp_bounded [OF Y] by fast
28.220 + obtain s t where s: "0 < s" and t: "0 < t" and r: "r = a * t + s * b"
28.221 + proof
28.222 + show "0 < v/b" using v b(1) by (rule divide_pos_pos)
28.223 + show "0 < u/a" using u a(1) by (rule divide_pos_pos)
28.224 + show "r = a * (u/a) + (v/b) * b"
28.225 + using a(1) b(1) `r = u + v` by simp
28.226 + qed
28.227 + obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
28.228 + using cauchyD [OF X s] ..
28.229 + obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
28.230 + using cauchyD [OF Y t] ..
28.231 + have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>X m * Y m - X n * Y n\<bar> < r"
28.232 + proof (clarsimp)
28.233 + fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
28.234 + have "\<bar>X m * Y m - X n * Y n\<bar> = \<bar>X m * (Y m - Y n) + (X m - X n) * Y n\<bar>"
28.235 + unfolding mult_diff_mult ..
28.236 + also have "\<dots> \<le> \<bar>X m * (Y m - Y n)\<bar> + \<bar>(X m - X n) * Y n\<bar>"
28.237 + by (rule abs_triangle_ineq)
28.238 + also have "\<dots> = \<bar>X m\<bar> * \<bar>Y m - Y n\<bar> + \<bar>X m - X n\<bar> * \<bar>Y n\<bar>"
28.239 + unfolding abs_mult ..
28.240 + also have "\<dots> < a * t + s * b"
28.241 + by (simp_all add: add_strict_mono mult_strict_mono' a b i j *)
28.242 + finally show "\<bar>X m * Y m - X n * Y n\<bar> < r" unfolding r .
28.243 + qed
28.244 + thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m * Y m - X n * Y n\<bar> < r" ..
28.245 +qed
28.246 +
28.247 +lemma cauchy_not_vanishes_cases:
28.248 + assumes X: "cauchy X"
28.249 + assumes nz: "\<not> vanishes X"
28.250 + shows "\<exists>b>0. \<exists>k. (\<forall>n\<ge>k. b < - X n) \<or> (\<forall>n\<ge>k. b < X n)"
28.251 +proof -
28.252 + obtain r where "0 < r" and r: "\<forall>k. \<exists>n\<ge>k. r \<le> \<bar>X n\<bar>"
28.253 + using nz unfolding vanishes_def by (auto simp add: not_less)
28.254 + obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t"
28.255 + using `0 < r` by (rule obtain_pos_sum)
28.256 + obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
28.257 + using cauchyD [OF X s] ..
28.258 + obtain k where "i \<le> k" and "r \<le> \<bar>X k\<bar>"
28.259 + using r by fast
28.260 + have k: "\<forall>n\<ge>k. \<bar>X n - X k\<bar> < s"
28.261 + using i `i \<le> k` by auto
28.262 + have "X k \<le> - r \<or> r \<le> X k"
28.263 + using `r \<le> \<bar>X k\<bar>` by auto
28.264 + hence "(\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
28.265 + unfolding `r = s + t` using k by auto
28.266 + hence "\<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)" ..
28.267 + thus "\<exists>t>0. \<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
28.268 + using t by auto
28.269 +qed
28.270 +
28.271 +lemma cauchy_not_vanishes:
28.272 + assumes X: "cauchy X"
28.273 + assumes nz: "\<not> vanishes X"
28.274 + shows "\<exists>b>0. \<exists>k. \<forall>n\<ge>k. b < \<bar>X n\<bar>"
28.275 +using cauchy_not_vanishes_cases [OF assms]
28.276 +by clarify (rule exI, erule conjI, rule_tac x=k in exI, auto)
28.277 +
28.278 +lemma cauchy_inverse [simp]:
28.279 + assumes X: "cauchy X"
28.280 + assumes nz: "\<not> vanishes X"
28.281 + shows "cauchy (\<lambda>n. inverse (X n))"
28.282 +proof (rule cauchyI)
28.283 + fix r :: rat assume "0 < r"
28.284 + obtain b i where b: "0 < b" and i: "\<forall>n\<ge>i. b < \<bar>X n\<bar>"
28.285 + using cauchy_not_vanishes [OF X nz] by fast
28.286 + from b i have nz: "\<forall>n\<ge>i. X n \<noteq> 0" by auto
28.287 + obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b"
28.288 + proof
28.289 + show "0 < b * r * b"
28.290 + by (simp add: `0 < r` b mult_pos_pos)
28.291 + show "r = inverse b * (b * r * b) * inverse b"
28.292 + using b by simp
28.293 + qed
28.294 + obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>X m - X n\<bar> < s"
28.295 + using cauchyD [OF X s] ..
28.296 + have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>inverse (X m) - inverse (X n)\<bar> < r"
28.297 + proof (clarsimp)
28.298 + fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
28.299 + have "\<bar>inverse (X m) - inverse (X n)\<bar> =
28.300 + inverse \<bar>X m\<bar> * \<bar>X m - X n\<bar> * inverse \<bar>X n\<bar>"
28.301 + by (simp add: inverse_diff_inverse nz * abs_mult)
28.302 + also have "\<dots> < inverse b * s * inverse b"
28.303 + by (simp add: mult_strict_mono less_imp_inverse_less
28.304 + mult_pos_pos i j b * s)
28.305 + finally show "\<bar>inverse (X m) - inverse (X n)\<bar> < r" unfolding r .
28.306 + qed
28.307 + thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>inverse (X m) - inverse (X n)\<bar> < r" ..
28.308 +qed
28.309 +
28.310 +lemma vanishes_diff_inverse:
28.311 + assumes X: "cauchy X" "\<not> vanishes X"
28.312 + assumes Y: "cauchy Y" "\<not> vanishes Y"
28.313 + assumes XY: "vanishes (\<lambda>n. X n - Y n)"
28.314 + shows "vanishes (\<lambda>n. inverse (X n) - inverse (Y n))"
28.315 +proof (rule vanishesI)
28.316 + fix r :: rat assume r: "0 < r"
28.317 + obtain a i where a: "0 < a" and i: "\<forall>n\<ge>i. a < \<bar>X n\<bar>"
28.318 + using cauchy_not_vanishes [OF X] by fast
28.319 + obtain b j where b: "0 < b" and j: "\<forall>n\<ge>j. b < \<bar>Y n\<bar>"
28.320 + using cauchy_not_vanishes [OF Y] by fast
28.321 + obtain s where s: "0 < s" and "inverse a * s * inverse b = r"
28.322 + proof
28.323 + show "0 < a * r * b"
28.324 + using a r b by (simp add: mult_pos_pos)
28.325 + show "inverse a * (a * r * b) * inverse b = r"
28.326 + using a r b by simp
28.327 + qed
28.328 + obtain k where k: "\<forall>n\<ge>k. \<bar>X n - Y n\<bar> < s"
28.329 + using vanishesD [OF XY s] ..
28.330 + have "\<forall>n\<ge>max (max i j) k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r"
28.331 + proof (clarsimp)
28.332 + fix n assume n: "i \<le> n" "j \<le> n" "k \<le> n"
28.333 + have "X n \<noteq> 0" and "Y n \<noteq> 0"
28.334 + using i j a b n by auto
28.335 + hence "\<bar>inverse (X n) - inverse (Y n)\<bar> =
28.336 + inverse \<bar>X n\<bar> * \<bar>X n - Y n\<bar> * inverse \<bar>Y n\<bar>"
28.337 + by (simp add: inverse_diff_inverse abs_mult)
28.338 + also have "\<dots> < inverse a * s * inverse b"
28.339 + apply (intro mult_strict_mono' less_imp_inverse_less)
28.340 + apply (simp_all add: a b i j k n mult_nonneg_nonneg)
28.341 + done
28.342 + also note `inverse a * s * inverse b = r`
28.343 + finally show "\<bar>inverse (X n) - inverse (Y n)\<bar> < r" .
28.344 + qed
28.345 + thus "\<exists>k. \<forall>n\<ge>k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r" ..
28.346 +qed
28.347 +
28.348 +subsection {* Equivalence relation on Cauchy sequences *}
28.349 +
28.350 +definition realrel :: "(nat \<Rightarrow> rat) \<Rightarrow> (nat \<Rightarrow> rat) \<Rightarrow> bool"
28.351 + where "realrel = (\<lambda>X Y. cauchy X \<and> cauchy Y \<and> vanishes (\<lambda>n. X n - Y n))"
28.352 +
28.353 +lemma realrelI [intro?]:
28.354 + assumes "cauchy X" and "cauchy Y" and "vanishes (\<lambda>n. X n - Y n)"
28.355 + shows "realrel X Y"
28.356 + using assms unfolding realrel_def by simp
28.357 +
28.358 +lemma realrel_refl: "cauchy X \<Longrightarrow> realrel X X"
28.359 + unfolding realrel_def by simp
28.360 +
28.361 +lemma symp_realrel: "symp realrel"
28.362 + unfolding realrel_def
28.363 + by (rule sympI, clarify, drule vanishes_minus, simp)
28.364 +
28.365 +lemma transp_realrel: "transp realrel"
28.366 + unfolding realrel_def
28.367 + apply (rule transpI, clarify)
28.368 + apply (drule (1) vanishes_add)
28.369 + apply (simp add: algebra_simps)
28.370 + done
28.371 +
28.372 +lemma part_equivp_realrel: "part_equivp realrel"
28.373 + by (fast intro: part_equivpI symp_realrel transp_realrel
28.374 + realrel_refl cauchy_const)
28.375 +
28.376 +subsection {* The field of real numbers *}
28.377 +
28.378 +quotient_type real = "nat \<Rightarrow> rat" / partial: realrel
28.379 + morphisms rep_real Real
28.380 + by (rule part_equivp_realrel)
28.381 +
28.382 +lemma cr_real_eq: "pcr_real = (\<lambda>x y. cauchy x \<and> Real x = y)"
28.383 + unfolding real.pcr_cr_eq cr_real_def realrel_def by auto
28.384 +
28.385 +lemma Real_induct [induct type: real]: (* TODO: generate automatically *)
28.386 + assumes "\<And>X. cauchy X \<Longrightarrow> P (Real X)" shows "P x"
28.387 +proof (induct x)
28.388 + case (1 X)
28.389 + hence "cauchy X" by (simp add: realrel_def)
28.390 + thus "P (Real X)" by (rule assms)
28.391 +qed
28.392 +
28.393 +lemma eq_Real:
28.394 + "cauchy X \<Longrightarrow> cauchy Y \<Longrightarrow> Real X = Real Y \<longleftrightarrow> vanishes (\<lambda>n. X n - Y n)"
28.395 + using real.rel_eq_transfer
28.396 + unfolding real.pcr_cr_eq cr_real_def fun_rel_def realrel_def by simp
28.397 +
28.398 +declare real.forall_transfer [transfer_rule del]
28.399 +
28.400 +lemma forall_real_transfer [transfer_rule]: (* TODO: generate automatically *)
28.401 + "(fun_rel (fun_rel pcr_real op =) op =)
28.402 + (transfer_bforall cauchy) transfer_forall"
28.403 + using real.forall_transfer
28.404 + by (simp add: realrel_def)
28.405 +
28.406 +instantiation real :: field_inverse_zero
28.407 +begin
28.408 +
28.409 +lift_definition zero_real :: "real" is "\<lambda>n. 0"
28.410 + by (simp add: realrel_refl)
28.411 +
28.412 +lift_definition one_real :: "real" is "\<lambda>n. 1"
28.413 + by (simp add: realrel_refl)
28.414 +
28.415 +lift_definition plus_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n + Y n"
28.416 + unfolding realrel_def add_diff_add
28.417 + by (simp only: cauchy_add vanishes_add simp_thms)
28.418 +
28.419 +lift_definition uminus_real :: "real \<Rightarrow> real" is "\<lambda>X n. - X n"
28.420 + unfolding realrel_def minus_diff_minus
28.421 + by (simp only: cauchy_minus vanishes_minus simp_thms)
28.422 +
28.423 +lift_definition times_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n * Y n"
28.424 + unfolding realrel_def mult_diff_mult
28.425 + by (subst (4) mult_commute, simp only: cauchy_mult vanishes_add
28.426 + vanishes_mult_bounded cauchy_imp_bounded simp_thms)
28.427 +
28.428 +lift_definition inverse_real :: "real \<Rightarrow> real"
28.429 + is "\<lambda>X. if vanishes X then (\<lambda>n. 0) else (\<lambda>n. inverse (X n))"
28.430 +proof -
28.431 + fix X Y assume "realrel X Y"
28.432 + hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\<lambda>n. X n - Y n)"
28.433 + unfolding realrel_def by simp_all
28.434 + have "vanishes X \<longleftrightarrow> vanishes Y"
28.435 + proof
28.436 + assume "vanishes X"
28.437 + from vanishes_diff [OF this XY] show "vanishes Y" by simp
28.438 + next
28.439 + assume "vanishes Y"
28.440 + from vanishes_add [OF this XY] show "vanishes X" by simp
28.441 + qed
28.442 + thus "?thesis X Y"
28.443 + unfolding realrel_def
28.444 + by (simp add: vanishes_diff_inverse X Y XY)
28.445 +qed
28.446 +
28.447 +definition
28.448 + "x - y = (x::real) + - y"
28.449 +
28.450 +definition
28.451 + "x / y = (x::real) * inverse y"
28.452 +
28.453 +lemma add_Real:
28.454 + assumes X: "cauchy X" and Y: "cauchy Y"
28.455 + shows "Real X + Real Y = Real (\<lambda>n. X n + Y n)"
28.456 + using assms plus_real.transfer
28.457 + unfolding cr_real_eq fun_rel_def by simp
28.458 +
28.459 +lemma minus_Real:
28.460 + assumes X: "cauchy X"
28.461 + shows "- Real X = Real (\<lambda>n. - X n)"
28.462 + using assms uminus_real.transfer
28.463 + unfolding cr_real_eq fun_rel_def by simp
28.464 +
28.465 +lemma diff_Real:
28.466 + assumes X: "cauchy X" and Y: "cauchy Y"
28.467 + shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
28.468 + unfolding minus_real_def diff_minus
28.469 + by (simp add: minus_Real add_Real X Y)
28.470 +
28.471 +lemma mult_Real:
28.472 + assumes X: "cauchy X" and Y: "cauchy Y"
28.473 + shows "Real X * Real Y = Real (\<lambda>n. X n * Y n)"
28.474 + using assms times_real.transfer
28.475 + unfolding cr_real_eq fun_rel_def by simp
28.476 +
28.477 +lemma inverse_Real:
28.478 + assumes X: "cauchy X"
28.479 + shows "inverse (Real X) =
28.480 + (if vanishes X then 0 else Real (\<lambda>n. inverse (X n)))"
28.481 + using assms inverse_real.transfer zero_real.transfer
28.482 + unfolding cr_real_eq fun_rel_def by (simp split: split_if_asm, metis)
28.483 +
28.484 +instance proof
28.485 + fix a b c :: real
28.486 + show "a + b = b + a"
28.487 + by transfer (simp add: add_ac realrel_def)
28.488 + show "(a + b) + c = a + (b + c)"
28.489 + by transfer (simp add: add_ac realrel_def)
28.490 + show "0 + a = a"
28.491 + by transfer (simp add: realrel_def)
28.492 + show "- a + a = 0"
28.493 + by transfer (simp add: realrel_def)
28.494 + show "a - b = a + - b"
28.495 + by (rule minus_real_def)
28.496 + show "(a * b) * c = a * (b * c)"
28.497 + by transfer (simp add: mult_ac realrel_def)
28.498 + show "a * b = b * a"
28.499 + by transfer (simp add: mult_ac realrel_def)
28.500 + show "1 * a = a"
28.501 + by transfer (simp add: mult_ac realrel_def)
28.502 + show "(a + b) * c = a * c + b * c"
28.503 + by transfer (simp add: distrib_right realrel_def)
28.504 + show "(0\<Colon>real) \<noteq> (1\<Colon>real)"
28.505 + by transfer (simp add: realrel_def)
28.506 + show "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
28.507 + apply transfer
28.508 + apply (simp add: realrel_def)
28.509 + apply (rule vanishesI)
28.510 + apply (frule (1) cauchy_not_vanishes, clarify)
28.511 + apply (rule_tac x=k in exI, clarify)
28.512 + apply (drule_tac x=n in spec, simp)
28.513 + done
28.514 + show "a / b = a * inverse b"
28.515 + by (rule divide_real_def)
28.516 + show "inverse (0::real) = 0"
28.517 + by transfer (simp add: realrel_def)
28.518 +qed
28.519 +
28.520 +end
28.521 +
28.522 +subsection {* Positive reals *}
28.523 +
28.524 +lift_definition positive :: "real \<Rightarrow> bool"
28.525 + is "\<lambda>X. \<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
28.526 +proof -
28.527 + { fix X Y
28.528 + assume "realrel X Y"
28.529 + hence XY: "vanishes (\<lambda>n. X n - Y n)"
28.530 + unfolding realrel_def by simp_all
28.531 + assume "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
28.532 + then obtain r i where "0 < r" and i: "\<forall>n\<ge>i. r < X n"
28.533 + by fast
28.534 + obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
28.535 + using `0 < r` by (rule obtain_pos_sum)
28.536 + obtain j where j: "\<forall>n\<ge>j. \<bar>X n - Y n\<bar> < s"
28.537 + using vanishesD [OF XY s] ..
28.538 + have "\<forall>n\<ge>max i j. t < Y n"
28.539 + proof (clarsimp)
28.540 + fix n assume n: "i \<le> n" "j \<le> n"
28.541 + have "\<bar>X n - Y n\<bar> < s" and "r < X n"
28.542 + using i j n by simp_all
28.543 + thus "t < Y n" unfolding r by simp
28.544 + qed
28.545 + hence "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < Y n" using t by fast
28.546 + } note 1 = this
28.547 + fix X Y assume "realrel X Y"
28.548 + hence "realrel X Y" and "realrel Y X"
28.549 + using symp_realrel unfolding symp_def by auto
28.550 + thus "?thesis X Y"
28.551 + by (safe elim!: 1)
28.552 +qed
28.553 +
28.554 +lemma positive_Real:
28.555 + assumes X: "cauchy X"
28.556 + shows "positive (Real X) \<longleftrightarrow> (\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n)"
28.557 + using assms positive.transfer
28.558 + unfolding cr_real_eq fun_rel_def by simp
28.559 +
28.560 +lemma positive_zero: "\<not> positive 0"
28.561 + by transfer auto
28.562 +
28.563 +lemma positive_add:
28.564 + "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x + y)"
28.565 +apply transfer
28.566 +apply (clarify, rename_tac a b i j)
28.567 +apply (rule_tac x="a + b" in exI, simp)
28.568 +apply (rule_tac x="max i j" in exI, clarsimp)
28.569 +apply (simp add: add_strict_mono)
28.570 +done
28.571 +
28.572 +lemma positive_mult:
28.573 + "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x * y)"
28.574 +apply transfer
28.575 +apply (clarify, rename_tac a b i j)
28.576 +apply (rule_tac x="a * b" in exI, simp add: mult_pos_pos)
28.577 +apply (rule_tac x="max i j" in exI, clarsimp)
28.578 +apply (rule mult_strict_mono, auto)
28.579 +done
28.580 +
28.581 +lemma positive_minus:
28.582 + "\<not> positive x \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> positive (- x)"
28.583 +apply transfer
28.584 +apply (simp add: realrel_def)
28.585 +apply (drule (1) cauchy_not_vanishes_cases, safe, fast, fast)
28.586 +done
28.587 +
28.588 +instantiation real :: linordered_field_inverse_zero
28.589 +begin
28.590 +
28.591 +definition
28.592 + "x < y \<longleftrightarrow> positive (y - x)"
28.593 +
28.594 +definition
28.595 + "x \<le> (y::real) \<longleftrightarrow> x < y \<or> x = y"
28.596 +
28.597 +definition
28.598 + "abs (a::real) = (if a < 0 then - a else a)"
28.599 +
28.600 +definition
28.601 + "sgn (a::real) = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
28.602 +
28.603 +instance proof
28.604 + fix a b c :: real
28.605 + show "\<bar>a\<bar> = (if a < 0 then - a else a)"
28.606 + by (rule abs_real_def)
28.607 + show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a"
28.608 + unfolding less_eq_real_def less_real_def
28.609 + by (auto, drule (1) positive_add, simp_all add: positive_zero)
28.610 + show "a \<le> a"
28.611 + unfolding less_eq_real_def by simp
28.612 + show "a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> a \<le> c"
28.613 + unfolding less_eq_real_def less_real_def
28.614 + by (auto, drule (1) positive_add, simp add: algebra_simps)
28.615 + show "a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> a = b"
28.616 + unfolding less_eq_real_def less_real_def
28.617 + by (auto, drule (1) positive_add, simp add: positive_zero)
28.618 + show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
28.619 + unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *)
28.620 + (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \<equiv> b + - a *)
28.621 + (* Should produce c + b - (c + a) \<equiv> b - a *)
28.622 + show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
28.623 + by (rule sgn_real_def)
28.624 + show "a \<le> b \<or> b \<le> a"
28.625 + unfolding less_eq_real_def less_real_def
28.626 + by (auto dest!: positive_minus)
28.627 + show "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
28.628 + unfolding less_real_def
28.629 + by (drule (1) positive_mult, simp add: algebra_simps)
28.630 +qed
28.631 +
28.632 +end
28.633 +
28.634 +instantiation real :: distrib_lattice
28.635 +begin
28.636 +
28.637 +definition
28.638 + "(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"
28.639 +
28.640 +definition
28.641 + "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
28.642 +
28.643 +instance proof
28.644 +qed (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
28.645 +
28.646 +end
28.647 +
28.648 +lemma of_nat_Real: "of_nat x = Real (\<lambda>n. of_nat x)"
28.649 +apply (induct x)
28.650 +apply (simp add: zero_real_def)
28.651 +apply (simp add: one_real_def add_Real)
28.652 +done
28.653 +
28.654 +lemma of_int_Real: "of_int x = Real (\<lambda>n. of_int x)"
28.655 +apply (cases x rule: int_diff_cases)
28.656 +apply (simp add: of_nat_Real diff_Real)
28.657 +done
28.658 +
28.659 +lemma of_rat_Real: "of_rat x = Real (\<lambda>n. x)"
28.660 +apply (induct x)
28.661 +apply (simp add: Fract_of_int_quotient of_rat_divide)
28.662 +apply (simp add: of_int_Real divide_inverse)
28.663 +apply (simp add: inverse_Real mult_Real)
28.664 +done
28.665 +
28.666 +instance real :: archimedean_field
28.667 +proof
28.668 + fix x :: real
28.669 + show "\<exists>z. x \<le> of_int z"
28.670 + apply (induct x)
28.671 + apply (frule cauchy_imp_bounded, clarify)
28.672 + apply (rule_tac x="ceiling b + 1" in exI)
28.673 + apply (rule less_imp_le)
28.674 + apply (simp add: of_int_Real less_real_def diff_Real positive_Real)
28.675 + apply (rule_tac x=1 in exI, simp add: algebra_simps)
28.676 + apply (rule_tac x=0 in exI, clarsimp)
28.677 + apply (rule le_less_trans [OF abs_ge_self])
28.678 + apply (rule less_le_trans [OF _ le_of_int_ceiling])
28.679 + apply simp
28.680 + done
28.681 +qed
28.682 +
28.683 +instantiation real :: floor_ceiling
28.684 +begin
28.685 +
28.686 +definition [code del]:
28.687 + "floor (x::real) = (THE z. of_int z \<le> x \<and> x < of_int (z + 1))"
28.688 +
28.689 +instance proof
28.690 + fix x :: real
28.691 + show "of_int (floor x) \<le> x \<and> x < of_int (floor x + 1)"
28.692 + unfolding floor_real_def using floor_exists1 by (rule theI')
28.693 +qed
28.694 +
28.695 +end
28.696 +
28.697 +subsection {* Completeness *}
28.698 +
28.699 +lemma not_positive_Real:
28.700 + assumes X: "cauchy X"
28.701 + shows "\<not> positive (Real X) \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> r)"
28.702 +unfolding positive_Real [OF X]
28.703 +apply (auto, unfold not_less)
28.704 +apply (erule obtain_pos_sum)
28.705 +apply (drule_tac x=s in spec, simp)
28.706 +apply (drule_tac r=t in cauchyD [OF X], clarify)
28.707 +apply (drule_tac x=k in spec, clarsimp)
28.708 +apply (rule_tac x=n in exI, clarify, rename_tac m)
28.709 +apply (drule_tac x=m in spec, simp)
28.710 +apply (drule_tac x=n in spec, simp)
28.711 +apply (drule spec, drule (1) mp, clarify, rename_tac i)
28.712 +apply (rule_tac x="max i k" in exI, simp)
28.713 +done
28.714 +
28.715 +lemma le_Real:
28.716 + assumes X: "cauchy X" and Y: "cauchy Y"
28.717 + shows "Real X \<le> Real Y = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r)"
28.718 +unfolding not_less [symmetric, where 'a=real] less_real_def
28.719 +apply (simp add: diff_Real not_positive_Real X Y)
28.720 +apply (simp add: diff_le_eq add_ac)
28.721 +done
28.722 +
28.723 +lemma le_RealI:
28.724 + assumes Y: "cauchy Y"
28.725 + shows "\<forall>n. x \<le> of_rat (Y n) \<Longrightarrow> x \<le> Real Y"
28.726 +proof (induct x)
28.727 + fix X assume X: "cauchy X" and "\<forall>n. Real X \<le> of_rat (Y n)"
28.728 + hence le: "\<And>m r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. X n \<le> Y m + r"
28.729 + by (simp add: of_rat_Real le_Real)
28.730 + {
28.731 + fix r :: rat assume "0 < r"
28.732 + then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
28.733 + by (rule obtain_pos_sum)
28.734 + obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>Y m - Y n\<bar> < s"
28.735 + using cauchyD [OF Y s] ..
28.736 + obtain j where j: "\<forall>n\<ge>j. X n \<le> Y i + t"
28.737 + using le [OF t] ..
28.738 + have "\<forall>n\<ge>max i j. X n \<le> Y n + r"
28.739 + proof (clarsimp)
28.740 + fix n assume n: "i \<le> n" "j \<le> n"
28.741 + have "X n \<le> Y i + t" using n j by simp
28.742 + moreover have "\<bar>Y i - Y n\<bar> < s" using n i by simp
28.743 + ultimately show "X n \<le> Y n + r" unfolding r by simp
28.744 + qed
28.745 + hence "\<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r" ..
28.746 + }
28.747 + thus "Real X \<le> Real Y"
28.748 + by (simp add: of_rat_Real le_Real X Y)
28.749 +qed
28.750 +
28.751 +lemma Real_leI:
28.752 + assumes X: "cauchy X"
28.753 + assumes le: "\<forall>n. of_rat (X n) \<le> y"
28.754 + shows "Real X \<le> y"
28.755 +proof -
28.756 + have "- y \<le> - Real X"
28.757 + by (simp add: minus_Real X le_RealI of_rat_minus le)
28.758 + thus ?thesis by simp
28.759 +qed
28.760 +
28.761 +lemma less_RealD:
28.762 + assumes Y: "cauchy Y"
28.763 + shows "x < Real Y \<Longrightarrow> \<exists>n. x < of_rat (Y n)"
28.764 +by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y])
28.765 +
28.766 +lemma of_nat_less_two_power:
28.767 + "of_nat n < (2::'a::linordered_idom) ^ n"
28.768 +apply (induct n)
28.769 +apply simp
28.770 +apply (subgoal_tac "(1::'a) \<le> 2 ^ n")
28.771 +apply (drule (1) add_le_less_mono, simp)
28.772 +apply simp
28.773 +done
28.774 +
28.775 +lemma complete_real:
28.776 + fixes S :: "real set"
28.777 + assumes "\<exists>x. x \<in> S" and "\<exists>z. \<forall>x\<in>S. x \<le> z"
28.778 + shows "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
28.779 +proof -
28.780 + obtain x where x: "x \<in> S" using assms(1) ..
28.781 + obtain z where z: "\<forall>x\<in>S. x \<le> z" using assms(2) ..
28.782 +
28.783 + def P \<equiv> "\<lambda>x. \<forall>y\<in>S. y \<le> of_rat x"
28.784 + obtain a where a: "\<not> P a"
28.785 + proof
28.786 + have "of_int (floor (x - 1)) \<le> x - 1" by (rule of_int_floor_le)
28.787 + also have "x - 1 < x" by simp
28.788 + finally have "of_int (floor (x - 1)) < x" .
28.789 + hence "\<not> x \<le> of_int (floor (x - 1))" by (simp only: not_le)
28.790 + then show "\<not> P (of_int (floor (x - 1)))"
28.791 + unfolding P_def of_rat_of_int_eq using x by fast
28.792 + qed
28.793 + obtain b where b: "P b"
28.794 + proof
28.795 + show "P (of_int (ceiling z))"
28.796 + unfolding P_def of_rat_of_int_eq
28.797 + proof
28.798 + fix y assume "y \<in> S"
28.799 + hence "y \<le> z" using z by simp
28.800 + also have "z \<le> of_int (ceiling z)" by (rule le_of_int_ceiling)
28.801 + finally show "y \<le> of_int (ceiling z)" .
28.802 + qed
28.803 + qed
28.804 +
28.805 + def avg \<equiv> "\<lambda>x y :: rat. x/2 + y/2"
28.806 + def bisect \<equiv> "\<lambda>(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)"
28.807 + def A \<equiv> "\<lambda>n. fst ((bisect ^^ n) (a, b))"
28.808 + def B \<equiv> "\<lambda>n. snd ((bisect ^^ n) (a, b))"
28.809 + def C \<equiv> "\<lambda>n. avg (A n) (B n)"
28.810 + have A_0 [simp]: "A 0 = a" unfolding A_def by simp
28.811 + have B_0 [simp]: "B 0 = b" unfolding B_def by simp
28.812 + have A_Suc [simp]: "\<And>n. A (Suc n) = (if P (C n) then A n else C n)"
28.813 + unfolding A_def B_def C_def bisect_def split_def by simp
28.814 + have B_Suc [simp]: "\<And>n. B (Suc n) = (if P (C n) then C n else B n)"
28.815 + unfolding A_def B_def C_def bisect_def split_def by simp
28.816 +
28.817 + have width: "\<And>n. B n - A n = (b - a) / 2^n"
28.818 + apply (simp add: eq_divide_eq)
28.819 + apply (induct_tac n, simp)
28.820 + apply (simp add: C_def avg_def algebra_simps)
28.821 + done
28.822 +
28.823 + have twos: "\<And>y r :: rat. 0 < r \<Longrightarrow> \<exists>n. y / 2 ^ n < r"
28.824 + apply (simp add: divide_less_eq)
28.825 + apply (subst mult_commute)
28.826 + apply (frule_tac y=y in ex_less_of_nat_mult)
28.827 + apply clarify
28.828 + apply (rule_tac x=n in exI)
28.829 + apply (erule less_trans)
28.830 + apply (rule mult_strict_right_mono)
28.831 + apply (rule le_less_trans [OF _ of_nat_less_two_power])
28.832 + apply simp
28.833 + apply assumption
28.834 + done
28.835 +
28.836 + have PA: "\<And>n. \<not> P (A n)"
28.837 + by (induct_tac n, simp_all add: a)
28.838 + have PB: "\<And>n. P (B n)"
28.839 + by (induct_tac n, simp_all add: b)
28.840 + have ab: "a < b"
28.841 + using a b unfolding P_def
28.842 + apply (clarsimp simp add: not_le)
28.843 + apply (drule (1) bspec)
28.844 + apply (drule (1) less_le_trans)
28.845 + apply (simp add: of_rat_less)
28.846 + done
28.847 + have AB: "\<And>n. A n < B n"
28.848 + by (induct_tac n, simp add: ab, simp add: C_def avg_def)
28.849 + have A_mono: "\<And>i j. i \<le> j \<Longrightarrow> A i \<le> A j"
28.850 + apply (auto simp add: le_less [where 'a=nat])
28.851 + apply (erule less_Suc_induct)
28.852 + apply (clarsimp simp add: C_def avg_def)
28.853 + apply (simp add: add_divide_distrib [symmetric])
28.854 + apply (rule AB [THEN less_imp_le])
28.855 + apply simp
28.856 + done
28.857 + have B_mono: "\<And>i j. i \<le> j \<Longrightarrow> B j \<le> B i"
28.858 + apply (auto simp add: le_less [where 'a=nat])
28.859 + apply (erule less_Suc_induct)
28.860 + apply (clarsimp simp add: C_def avg_def)
28.861 + apply (simp add: add_divide_distrib [symmetric])
28.862 + apply (rule AB [THEN less_imp_le])
28.863 + apply simp
28.864 + done
28.865 + have cauchy_lemma:
28.866 + "\<And>X. \<forall>n. \<forall>i\<ge>n. A n \<le> X i \<and> X i \<le> B n \<Longrightarrow> cauchy X"
28.867 + apply (rule cauchyI)
28.868 + apply (drule twos [where y="b - a"])
28.869 + apply (erule exE)
28.870 + apply (rule_tac x=n in exI, clarify, rename_tac i j)
28.871 + apply (rule_tac y="B n - A n" in le_less_trans) defer
28.872 + apply (simp add: width)
28.873 + apply (drule_tac x=n in spec)
28.874 + apply (frule_tac x=i in spec, drule (1) mp)
28.875 + apply (frule_tac x=j in spec, drule (1) mp)
28.876 + apply (frule A_mono, drule B_mono)
28.877 + apply (frule A_mono, drule B_mono)
28.878 + apply arith
28.879 + done
28.880 + have "cauchy A"
28.881 + apply (rule cauchy_lemma [rule_format])
28.882 + apply (simp add: A_mono)
28.883 + apply (erule order_trans [OF less_imp_le [OF AB] B_mono])
28.884 + done
28.885 + have "cauchy B"
28.886 + apply (rule cauchy_lemma [rule_format])
28.887 + apply (simp add: B_mono)
28.888 + apply (erule order_trans [OF A_mono less_imp_le [OF AB]])
28.889 + done
28.890 + have 1: "\<forall>x\<in>S. x \<le> Real B"
28.891 + proof
28.892 + fix x assume "x \<in> S"
28.893 + then show "x \<le> Real B"
28.894 + using PB [unfolded P_def] `cauchy B`
28.895 + by (simp add: le_RealI)
28.896 + qed
28.897 + have 2: "\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> Real A \<le> z"
28.898 + apply clarify
28.899 + apply (erule contrapos_pp)
28.900 + apply (simp add: not_le)
28.901 + apply (drule less_RealD [OF `cauchy A`], clarify)
28.902 + apply (subgoal_tac "\<not> P (A n)")
28.903 + apply (simp add: P_def not_le, clarify)
28.904 + apply (erule rev_bexI)
28.905 + apply (erule (1) less_trans)
28.906 + apply (simp add: PA)
28.907 + done
28.908 + have "vanishes (\<lambda>n. (b - a) / 2 ^ n)"
28.909 + proof (rule vanishesI)
28.910 + fix r :: rat assume "0 < r"
28.911 + then obtain k where k: "\<bar>b - a\<bar> / 2 ^ k < r"
28.912 + using twos by fast
28.913 + have "\<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r"
28.914 + proof (clarify)
28.915 + fix n assume n: "k \<le> n"
28.916 + have "\<bar>(b - a) / 2 ^ n\<bar> = \<bar>b - a\<bar> / 2 ^ n"
28.917 + by simp
28.918 + also have "\<dots> \<le> \<bar>b - a\<bar> / 2 ^ k"
28.919 + using n by (simp add: divide_left_mono mult_pos_pos)
28.920 + also note k
28.921 + finally show "\<bar>(b - a) / 2 ^ n\<bar> < r" .
28.922 + qed
28.923 + thus "\<exists>k. \<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r" ..
28.924 + qed
28.925 + hence 3: "Real B = Real A"
28.926 + by (simp add: eq_Real `cauchy A` `cauchy B` width)
28.927 + show "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
28.928 + using 1 2 3 by (rule_tac x="Real B" in exI, simp)
28.929 +qed
28.930 +
28.931 +
28.932 +instantiation real :: conditional_complete_linorder
28.933 begin
28.934
28.935 +subsection{*Supremum of a set of reals*}
28.936 +
28.937 +definition
28.938 + Sup_real_def: "Sup X \<equiv> LEAST z::real. \<forall>x\<in>X. x\<le>z"
28.939 +
28.940 +definition
28.941 + Inf_real_def: "Inf (X::real set) \<equiv> - Sup (uminus ` X)"
28.942 +
28.943 +instance
28.944 +proof
28.945 + { fix z x :: real and X :: "real set"
28.946 + assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
28.947 + then obtain s where s: "\<forall>y\<in>X. y \<le> s" "\<And>z. \<forall>y\<in>X. y \<le> z \<Longrightarrow> s \<le> z"
28.948 + using complete_real[of X] by blast
28.949 + then show "x \<le> Sup X"
28.950 + unfolding Sup_real_def by (rule LeastI2_order) (auto simp: x) }
28.951 + note Sup_upper = this
28.952 +
28.953 + { fix z :: real and X :: "real set"
28.954 + assume x: "X \<noteq> {}" and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
28.955 + then obtain s where s: "\<forall>y\<in>X. y \<le> s" "\<And>z. \<forall>y\<in>X. y \<le> z \<Longrightarrow> s \<le> z"
28.956 + using complete_real[of X] by blast
28.957 + then have "Sup X = s"
28.958 + unfolding Sup_real_def by (best intro: Least_equality)
28.959 + also with s z have "... \<le> z"
28.960 + by blast
28.961 + finally show "Sup X \<le> z" . }
28.962 + note Sup_least = this
28.963 +
28.964 + { fix x z :: real and X :: "real set"
28.965 + assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
28.966 + have "-x \<le> Sup (uminus ` X)"
28.967 + by (rule Sup_upper[of _ _ "- z"]) (auto simp add: image_iff x z)
28.968 + then show "Inf X \<le> x"
28.969 + by (auto simp add: Inf_real_def) }
28.970 +
28.971 + { fix z :: real and X :: "real set"
28.972 + assume x: "X \<noteq> {}" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
28.973 + have "Sup (uminus ` X) \<le> -z"
28.974 + using x z by (force intro: Sup_least)
28.975 + then show "z \<le> Inf X"
28.976 + by (auto simp add: Inf_real_def) }
28.977 +qed
28.978 +end
28.979 +
28.980 +text {*
28.981 + \medskip Completeness properties using @{text "isUb"}, @{text "isLub"}:
28.982 +*}
28.983 +
28.984 +lemma reals_complete: "\<exists>X. X \<in> S \<Longrightarrow> \<exists>Y. isUb (UNIV::real set) S Y \<Longrightarrow> \<exists>t. isLub (UNIV :: real set) S t"
28.985 + by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro: cSup_upper)
28.986 +
28.987 +
28.988 +subsection {* Hiding implementation details *}
28.989 +
28.990 +hide_const (open) vanishes cauchy positive Real
28.991 +
28.992 +declare Real_induct [induct del]
28.993 +declare Abs_real_induct [induct del]
28.994 +declare Abs_real_cases [cases del]
28.995 +
28.996 +lemmas [transfer_rule del] =
28.997 + real.All_transfer real.Ex_transfer real.rel_eq_transfer forall_real_transfer
28.998 + zero_real.transfer one_real.transfer plus_real.transfer uminus_real.transfer
28.999 + times_real.transfer inverse_real.transfer positive.transfer real.right_unique
28.1000 + real.right_total
28.1001 +
28.1002 +subsection{*More Lemmas*}
28.1003 +
28.1004 +text {* BH: These lemmas should not be necessary; they should be
28.1005 +covered by existing simp rules and simplification procedures. *}
28.1006 +
28.1007 +lemma real_mult_left_cancel: "(c::real) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
28.1008 +by simp (* redundant with mult_cancel_left *)
28.1009 +
28.1010 +lemma real_mult_right_cancel: "(c::real) \<noteq> 0 ==> (a*c=b*c) = (a=b)"
28.1011 +by simp (* redundant with mult_cancel_right *)
28.1012 +
28.1013 +lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
28.1014 +by simp (* solved by linordered_ring_less_cancel_factor simproc *)
28.1015 +
28.1016 +lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
28.1017 +by simp (* solved by linordered_ring_le_cancel_factor simproc *)
28.1018 +
28.1019 +lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
28.1020 +by simp (* solved by linordered_ring_le_cancel_factor simproc *)
28.1021 +
28.1022 +
28.1023 +subsection {* Embedding numbers into the Reals *}
28.1024 +
28.1025 +abbreviation
28.1026 + real_of_nat :: "nat \<Rightarrow> real"
28.1027 +where
28.1028 + "real_of_nat \<equiv> of_nat"
28.1029 +
28.1030 +abbreviation
28.1031 + real_of_int :: "int \<Rightarrow> real"
28.1032 +where
28.1033 + "real_of_int \<equiv> of_int"
28.1034 +
28.1035 +abbreviation
28.1036 + real_of_rat :: "rat \<Rightarrow> real"
28.1037 +where
28.1038 + "real_of_rat \<equiv> of_rat"
28.1039 +
28.1040 +consts
28.1041 + (*overloaded constant for injecting other types into "real"*)
28.1042 + real :: "'a => real"
28.1043 +
28.1044 +defs (overloaded)
28.1045 + real_of_nat_def [code_unfold]: "real == real_of_nat"
28.1046 + real_of_int_def [code_unfold]: "real == real_of_int"
28.1047 +
28.1048 +declare [[coercion_enabled]]
28.1049 +declare [[coercion "real::nat\<Rightarrow>real"]]
28.1050 +declare [[coercion "real::int\<Rightarrow>real"]]
28.1051 +declare [[coercion "int"]]
28.1052 +
28.1053 +declare [[coercion_map map]]
28.1054 +declare [[coercion_map "% f g h x. g (h (f x))"]]
28.1055 +declare [[coercion_map "% f g (x,y) . (f x, g y)"]]
28.1056 +
28.1057 +lemma real_eq_of_nat: "real = of_nat"
28.1058 + unfolding real_of_nat_def ..
28.1059 +
28.1060 +lemma real_eq_of_int: "real = of_int"
28.1061 + unfolding real_of_int_def ..
28.1062 +
28.1063 +lemma real_of_int_zero [simp]: "real (0::int) = 0"
28.1064 +by (simp add: real_of_int_def)
28.1065 +
28.1066 +lemma real_of_one [simp]: "real (1::int) = (1::real)"
28.1067 +by (simp add: real_of_int_def)
28.1068 +
28.1069 +lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y"
28.1070 +by (simp add: real_of_int_def)
28.1071 +
28.1072 +lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)"
28.1073 +by (simp add: real_of_int_def)
28.1074 +
28.1075 +lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y"
28.1076 +by (simp add: real_of_int_def)
28.1077 +
28.1078 +lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y"
28.1079 +by (simp add: real_of_int_def)
28.1080 +
28.1081 +lemma real_of_int_power [simp]: "real (x ^ n) = real (x::int) ^ n"
28.1082 +by (simp add: real_of_int_def of_int_power)
28.1083 +
28.1084 +lemmas power_real_of_int = real_of_int_power [symmetric]
28.1085 +
28.1086 +lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))"
28.1087 + apply (subst real_eq_of_int)+
28.1088 + apply (rule of_int_setsum)
28.1089 +done
28.1090 +
28.1091 +lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) =
28.1092 + (PROD x:A. real(f x))"
28.1093 + apply (subst real_eq_of_int)+
28.1094 + apply (rule of_int_setprod)
28.1095 +done
28.1096 +
28.1097 +lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))"
28.1098 +by (simp add: real_of_int_def)
28.1099 +
28.1100 +lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)"
28.1101 +by (simp add: real_of_int_def)
28.1102 +
28.1103 +lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)"
28.1104 +by (simp add: real_of_int_def)
28.1105 +
28.1106 +lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \<le> real y) = (x \<le> y)"
28.1107 +by (simp add: real_of_int_def)
28.1108 +
28.1109 +lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)"
28.1110 +by (simp add: real_of_int_def)
28.1111 +
28.1112 +lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)"
28.1113 +by (simp add: real_of_int_def)
28.1114 +
28.1115 +lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)"
28.1116 +by (simp add: real_of_int_def)
28.1117 +
28.1118 +lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)"
28.1119 +by (simp add: real_of_int_def)
28.1120 +
28.1121 +lemma one_less_real_of_int_cancel_iff: "1 < real (i :: int) \<longleftrightarrow> 1 < i"
28.1122 + unfolding real_of_one[symmetric] real_of_int_less_iff ..
28.1123 +
28.1124 +lemma one_le_real_of_int_cancel_iff: "1 \<le> real (i :: int) \<longleftrightarrow> 1 \<le> i"
28.1125 + unfolding real_of_one[symmetric] real_of_int_le_iff ..
28.1126 +
28.1127 +lemma real_of_int_less_one_cancel_iff: "real (i :: int) < 1 \<longleftrightarrow> i < 1"
28.1128 + unfolding real_of_one[symmetric] real_of_int_less_iff ..
28.1129 +
28.1130 +lemma real_of_int_le_one_cancel_iff: "real (i :: int) \<le> 1 \<longleftrightarrow> i \<le> 1"
28.1131 + unfolding real_of_one[symmetric] real_of_int_le_iff ..
28.1132 +
28.1133 +lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))"
28.1134 +by (auto simp add: abs_if)
28.1135 +
28.1136 +lemma int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)"
28.1137 + apply (subgoal_tac "real n + 1 = real (n + 1)")
28.1138 + apply (simp del: real_of_int_add)
28.1139 + apply auto
28.1140 +done
28.1141 +
28.1142 +lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)"
28.1143 + apply (subgoal_tac "real m + 1 = real (m + 1)")
28.1144 + apply (simp del: real_of_int_add)
28.1145 + apply simp
28.1146 +done
28.1147 +
28.1148 +lemma real_of_int_div_aux: "(real (x::int)) / (real d) =
28.1149 + real (x div d) + (real (x mod d)) / (real d)"
28.1150 +proof -
28.1151 + have "x = (x div d) * d + x mod d"
28.1152 + by auto
28.1153 + then have "real x = real (x div d) * real d + real(x mod d)"
28.1154 + by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym])
28.1155 + then have "real x / real d = ... / real d"
28.1156 + by simp
28.1157 + then show ?thesis
28.1158 + by (auto simp add: add_divide_distrib algebra_simps)
28.1159 +qed
28.1160 +
28.1161 +lemma real_of_int_div: "(d :: int) dvd n ==>
28.1162 + real(n div d) = real n / real d"
28.1163 + apply (subst real_of_int_div_aux)
28.1164 + apply simp
28.1165 + apply (simp add: dvd_eq_mod_eq_0)
28.1166 +done
28.1167 +
28.1168 +lemma real_of_int_div2:
28.1169 + "0 <= real (n::int) / real (x) - real (n div x)"
28.1170 + apply (case_tac "x = 0")
28.1171 + apply simp
28.1172 + apply (case_tac "0 < x")
28.1173 + apply (simp add: algebra_simps)
28.1174 + apply (subst real_of_int_div_aux)
28.1175 + apply simp
28.1176 + apply (subst zero_le_divide_iff)
28.1177 + apply auto
28.1178 + apply (simp add: algebra_simps)
28.1179 + apply (subst real_of_int_div_aux)
28.1180 + apply simp
28.1181 + apply (subst zero_le_divide_iff)
28.1182 + apply auto
28.1183 +done
28.1184 +
28.1185 +lemma real_of_int_div3:
28.1186 + "real (n::int) / real (x) - real (n div x) <= 1"
28.1187 + apply (simp add: algebra_simps)
28.1188 + apply (subst real_of_int_div_aux)
28.1189 + apply (auto simp add: divide_le_eq intro: order_less_imp_le)
28.1190 +done
28.1191 +
28.1192 +lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x"
28.1193 +by (insert real_of_int_div2 [of n x], simp)
28.1194 +
28.1195 +lemma Ints_real_of_int [simp]: "real (x::int) \<in> Ints"
28.1196 +unfolding real_of_int_def by (rule Ints_of_int)
28.1197 +
28.1198 +
28.1199 +subsection{*Embedding the Naturals into the Reals*}
28.1200 +
28.1201 +lemma real_of_nat_zero [simp]: "real (0::nat) = 0"
28.1202 +by (simp add: real_of_nat_def)
28.1203 +
28.1204 +lemma real_of_nat_1 [simp]: "real (1::nat) = 1"
28.1205 +by (simp add: real_of_nat_def)
28.1206 +
28.1207 +lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)"
28.1208 +by (simp add: real_of_nat_def)
28.1209 +
28.1210 +lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n"
28.1211 +by (simp add: real_of_nat_def)
28.1212 +
28.1213 +(*Not for addsimps: often the LHS is used to represent a positive natural*)
28.1214 +lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)"
28.1215 +by (simp add: real_of_nat_def)
28.1216 +
28.1217 +lemma real_of_nat_less_iff [iff]:
28.1218 + "(real (n::nat) < real m) = (n < m)"
28.1219 +by (simp add: real_of_nat_def)
28.1220 +
28.1221 +lemma real_of_nat_le_iff [iff]: "(real (n::nat) \<le> real m) = (n \<le> m)"
28.1222 +by (simp add: real_of_nat_def)
28.1223 +
28.1224 +lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)"
28.1225 +by (simp add: real_of_nat_def)
28.1226 +
28.1227 +lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)"
28.1228 +by (simp add: real_of_nat_def del: of_nat_Suc)
28.1229 +
28.1230 +lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n"
28.1231 +by (simp add: real_of_nat_def of_nat_mult)
28.1232 +
28.1233 +lemma real_of_nat_power [simp]: "real (m ^ n) = real (m::nat) ^ n"
28.1234 +by (simp add: real_of_nat_def of_nat_power)
28.1235 +
28.1236 +lemmas power_real_of_nat = real_of_nat_power [symmetric]
28.1237 +
28.1238 +lemma real_of_nat_setsum [simp]: "real ((SUM x:A. f x)::nat) =
28.1239 + (SUM x:A. real(f x))"
28.1240 + apply (subst real_eq_of_nat)+
28.1241 + apply (rule of_nat_setsum)
28.1242 +done
28.1243 +
28.1244 +lemma real_of_nat_setprod [simp]: "real ((PROD x:A. f x)::nat) =
28.1245 + (PROD x:A. real(f x))"
28.1246 + apply (subst real_eq_of_nat)+
28.1247 + apply (rule of_nat_setprod)
28.1248 +done
28.1249 +
28.1250 +lemma real_of_card: "real (card A) = setsum (%x.1) A"
28.1251 + apply (subst card_eq_setsum)
28.1252 + apply (subst real_of_nat_setsum)
28.1253 + apply simp
28.1254 +done
28.1255 +
28.1256 +lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)"
28.1257 +by (simp add: real_of_nat_def)
28.1258 +
28.1259 +lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)"
28.1260 +by (simp add: real_of_nat_def)
28.1261 +
28.1262 +lemma real_of_nat_diff: "n \<le> m ==> real (m - n) = real (m::nat) - real n"
28.1263 +by (simp add: add: real_of_nat_def of_nat_diff)
28.1264 +
28.1265 +lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)"
28.1266 +by (auto simp: real_of_nat_def)
28.1267 +
28.1268 +lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \<le> 0) = (n = 0)"
28.1269 +by (simp add: add: real_of_nat_def)
28.1270 +
28.1271 +lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0"
28.1272 +by (simp add: add: real_of_nat_def)
28.1273 +
28.1274 +lemma nat_less_real_le: "((n::nat) < m) = (real n + 1 <= real m)"
28.1275 + apply (subgoal_tac "real n + 1 = real (Suc n)")
28.1276 + apply simp
28.1277 + apply (auto simp add: real_of_nat_Suc)
28.1278 +done
28.1279 +
28.1280 +lemma nat_le_real_less: "((n::nat) <= m) = (real n < real m + 1)"
28.1281 + apply (subgoal_tac "real m + 1 = real (Suc m)")
28.1282 + apply (simp add: less_Suc_eq_le)
28.1283 + apply (simp add: real_of_nat_Suc)
28.1284 +done
28.1285 +
28.1286 +lemma real_of_nat_div_aux: "(real (x::nat)) / (real d) =
28.1287 + real (x div d) + (real (x mod d)) / (real d)"
28.1288 +proof -
28.1289 + have "x = (x div d) * d + x mod d"
28.1290 + by auto
28.1291 + then have "real x = real (x div d) * real d + real(x mod d)"
28.1292 + by (simp only: real_of_nat_mult [THEN sym] real_of_nat_add [THEN sym])
28.1293 + then have "real x / real d = \<dots> / real d"
28.1294 + by simp
28.1295 + then show ?thesis
28.1296 + by (auto simp add: add_divide_distrib algebra_simps)
28.1297 +qed
28.1298 +
28.1299 +lemma real_of_nat_div: "(d :: nat) dvd n ==>
28.1300 + real(n div d) = real n / real d"
28.1301 + by (subst real_of_nat_div_aux)
28.1302 + (auto simp add: dvd_eq_mod_eq_0 [symmetric])
28.1303 +
28.1304 +lemma real_of_nat_div2:
28.1305 + "0 <= real (n::nat) / real (x) - real (n div x)"
28.1306 +apply (simp add: algebra_simps)
28.1307 +apply (subst real_of_nat_div_aux)
28.1308 +apply simp
28.1309 +apply (subst zero_le_divide_iff)
28.1310 +apply simp
28.1311 +done
28.1312 +
28.1313 +lemma real_of_nat_div3:
28.1314 + "real (n::nat) / real (x) - real (n div x) <= 1"
28.1315 +apply(case_tac "x = 0")
28.1316 +apply (simp)
28.1317 +apply (simp add: algebra_simps)
28.1318 +apply (subst real_of_nat_div_aux)
28.1319 +apply simp
28.1320 +done
28.1321 +
28.1322 +lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x"
28.1323 +by (insert real_of_nat_div2 [of n x], simp)
28.1324 +
28.1325 +lemma real_of_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n"
28.1326 +by (simp add: real_of_int_def real_of_nat_def)
28.1327 +
28.1328 +lemma real_nat_eq_real [simp]: "0 <= x ==> real(nat x) = real x"
28.1329 + apply (subgoal_tac "real(int(nat x)) = real(nat x)")
28.1330 + apply force
28.1331 + apply (simp only: real_of_int_of_nat_eq)
28.1332 +done
28.1333 +
28.1334 +lemma Nats_real_of_nat [simp]: "real (n::nat) \<in> Nats"
28.1335 +unfolding real_of_nat_def by (rule of_nat_in_Nats)
28.1336 +
28.1337 +lemma Ints_real_of_nat [simp]: "real (n::nat) \<in> Ints"
28.1338 +unfolding real_of_nat_def by (rule Ints_of_nat)
28.1339 +
28.1340 +subsection {* The Archimedean Property of the Reals *}
28.1341 +
28.1342 +theorem reals_Archimedean:
28.1343 + assumes x_pos: "0 < x"
28.1344 + shows "\<exists>n. inverse (real (Suc n)) < x"
28.1345 + unfolding real_of_nat_def using x_pos
28.1346 + by (rule ex_inverse_of_nat_Suc_less)
28.1347 +
28.1348 +lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
28.1349 + unfolding real_of_nat_def by (rule ex_less_of_nat)
28.1350 +
28.1351 +lemma reals_Archimedean3:
28.1352 + assumes x_greater_zero: "0 < x"
28.1353 + shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"
28.1354 + unfolding real_of_nat_def using `0 < x`
28.1355 + by (auto intro: ex_less_of_nat_mult)
28.1356 +
28.1357 +
28.1358 +subsection{* Rationals *}
28.1359 +
28.1360 +lemma Rats_real_nat[simp]: "real(n::nat) \<in> \<rat>"
28.1361 +by (simp add: real_eq_of_nat)
28.1362 +
28.1363 +
28.1364 +lemma Rats_eq_int_div_int:
28.1365 + "\<rat> = { real(i::int)/real(j::int) |i j. j \<noteq> 0}" (is "_ = ?S")
28.1366 +proof
28.1367 + show "\<rat> \<subseteq> ?S"
28.1368 + proof
28.1369 + fix x::real assume "x : \<rat>"
28.1370 + then obtain r where "x = of_rat r" unfolding Rats_def ..
28.1371 + have "of_rat r : ?S"
28.1372 + by (cases r)(auto simp add:of_rat_rat real_eq_of_int)
28.1373 + thus "x : ?S" using `x = of_rat r` by simp
28.1374 + qed
28.1375 +next
28.1376 + show "?S \<subseteq> \<rat>"
28.1377 + proof(auto simp:Rats_def)
28.1378 + fix i j :: int assume "j \<noteq> 0"
28.1379 + hence "real i / real j = of_rat(Fract i j)"
28.1380 + by (simp add:of_rat_rat real_eq_of_int)
28.1381 + thus "real i / real j \<in> range of_rat" by blast
28.1382 + qed
28.1383 +qed
28.1384 +
28.1385 +lemma Rats_eq_int_div_nat:
28.1386 + "\<rat> = { real(i::int)/real(n::nat) |i n. n \<noteq> 0}"
28.1387 +proof(auto simp:Rats_eq_int_div_int)
28.1388 + fix i j::int assume "j \<noteq> 0"
28.1389 + show "EX (i'::int) (n::nat). real i/real j = real i'/real n \<and> 0<n"
28.1390 + proof cases
28.1391 + assume "j>0"
28.1392 + hence "real i/real j = real i/real(nat j) \<and> 0<nat j"
28.1393 + by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
28.1394 + thus ?thesis by blast
28.1395 + next
28.1396 + assume "~ j>0"
28.1397 + hence "real i/real j = real(-i)/real(nat(-j)) \<and> 0<nat(-j)" using `j\<noteq>0`
28.1398 + by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
28.1399 + thus ?thesis by blast
28.1400 + qed
28.1401 +next
28.1402 + fix i::int and n::nat assume "0 < n"
28.1403 + hence "real i/real n = real i/real(int n) \<and> int n \<noteq> 0" by simp
28.1404 + thus "\<exists>(i'::int) j::int. real i/real n = real i'/real j \<and> j \<noteq> 0" by blast
28.1405 +qed
28.1406 +
28.1407 +lemma Rats_abs_nat_div_natE:
28.1408 + assumes "x \<in> \<rat>"
28.1409 + obtains m n :: nat
28.1410 + where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd m n = 1"
28.1411 +proof -
28.1412 + from `x \<in> \<rat>` obtain i::int and n::nat where "n \<noteq> 0" and "x = real i / real n"
28.1413 + by(auto simp add: Rats_eq_int_div_nat)
28.1414 + hence "\<bar>x\<bar> = real(nat(abs i)) / real n" by simp
28.1415 + then obtain m :: nat where x_rat: "\<bar>x\<bar> = real m / real n" by blast
28.1416 + let ?gcd = "gcd m n"
28.1417 + from `n\<noteq>0` have gcd: "?gcd \<noteq> 0" by simp
28.1418 + let ?k = "m div ?gcd"
28.1419 + let ?l = "n div ?gcd"
28.1420 + let ?gcd' = "gcd ?k ?l"
28.1421 + have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m"
28.1422 + by (rule dvd_mult_div_cancel)
28.1423 + have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
28.1424 + by (rule dvd_mult_div_cancel)
28.1425 + from `n\<noteq>0` and gcd_l have "?l \<noteq> 0" by (auto iff del: neq0_conv)
28.1426 + moreover
28.1427 + have "\<bar>x\<bar> = real ?k / real ?l"
28.1428 + proof -
28.1429 + from gcd have "real ?k / real ?l =
28.1430 + real (?gcd * ?k) / real (?gcd * ?l)" by simp
28.1431 + also from gcd_k and gcd_l have "\<dots> = real m / real n" by simp
28.1432 + also from x_rat have "\<dots> = \<bar>x\<bar>" ..
28.1433 + finally show ?thesis ..
28.1434 + qed
28.1435 + moreover
28.1436 + have "?gcd' = 1"
28.1437 + proof -
28.1438 + have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)"
28.1439 + by (rule gcd_mult_distrib_nat)
28.1440 + with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
28.1441 + with gcd show ?thesis by auto
28.1442 + qed
28.1443 + ultimately show ?thesis ..
28.1444 +qed
28.1445 +
28.1446 +subsection{*Density of the Rational Reals in the Reals*}
28.1447 +
28.1448 +text{* This density proof is due to Stefan Richter and was ported by TN. The
28.1449 +original source is \emph{Real Analysis} by H.L. Royden.
28.1450 +It employs the Archimedean property of the reals. *}
28.1451 +
28.1452 +lemma Rats_dense_in_real:
28.1453 + fixes x :: real
28.1454 + assumes "x < y" shows "\<exists>r\<in>\<rat>. x < r \<and> r < y"
28.1455 +proof -
28.1456 + from `x<y` have "0 < y-x" by simp
28.1457 + with reals_Archimedean obtain q::nat
28.1458 + where q: "inverse (real q) < y-x" and "0 < q" by auto
28.1459 + def p \<equiv> "ceiling (y * real q) - 1"
28.1460 + def r \<equiv> "of_int p / real q"
28.1461 + from q have "x < y - inverse (real q)" by simp
28.1462 + also have "y - inverse (real q) \<le> r"
28.1463 + unfolding r_def p_def
28.1464 + by (simp add: le_divide_eq left_diff_distrib le_of_int_ceiling `0 < q`)
28.1465 + finally have "x < r" .
28.1466 + moreover have "r < y"
28.1467 + unfolding r_def p_def
28.1468 + by (simp add: divide_less_eq diff_less_eq `0 < q`
28.1469 + less_ceiling_iff [symmetric])
28.1470 + moreover from r_def have "r \<in> \<rat>" by simp
28.1471 + ultimately show ?thesis by fast
28.1472 +qed
28.1473 +
28.1474 +
28.1475 +
28.1476 +subsection{*Numerals and Arithmetic*}
28.1477 +
28.1478 +lemma [code_abbrev]:
28.1479 + "real_of_int (numeral k) = numeral k"
28.1480 + "real_of_int (neg_numeral k) = neg_numeral k"
28.1481 + by simp_all
28.1482 +
28.1483 +text{*Collapse applications of @{term real} to @{term number_of}*}
28.1484 +lemma real_numeral [simp]:
28.1485 + "real (numeral v :: int) = numeral v"
28.1486 + "real (neg_numeral v :: int) = neg_numeral v"
28.1487 +by (simp_all add: real_of_int_def)
28.1488 +
28.1489 +lemma real_of_nat_numeral [simp]:
28.1490 + "real (numeral v :: nat) = numeral v"
28.1491 +by (simp add: real_of_nat_def)
28.1492 +
28.1493 +declaration {*
28.1494 + K (Lin_Arith.add_inj_thms [@{thm real_of_nat_le_iff} RS iffD2, @{thm real_of_nat_inject} RS iffD2]
28.1495 + (* not needed because x < (y::nat) can be rewritten as Suc x <= y: real_of_nat_less_iff RS iffD2 *)
28.1496 + #> Lin_Arith.add_inj_thms [@{thm real_of_int_le_iff} RS iffD2, @{thm real_of_int_inject} RS iffD2]
28.1497 + (* not needed because x < (y::int) can be rewritten as x + 1 <= y: real_of_int_less_iff RS iffD2 *)
28.1498 + #> Lin_Arith.add_simps [@{thm real_of_nat_zero}, @{thm real_of_nat_Suc}, @{thm real_of_nat_add},
28.1499 + @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one},
28.1500 + @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff},
28.1501 + @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq},
28.1502 + @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)}]
28.1503 + #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \<Rightarrow> real"})
28.1504 + #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \<Rightarrow> real"}))
28.1505 +*}
28.1506 +
28.1507 +
28.1508 +subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
28.1509 +
28.1510 +lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)"
28.1511 +by arith
28.1512 +
28.1513 +text {* FIXME: redundant with @{text add_eq_0_iff} below *}
28.1514 +lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)"
28.1515 +by auto
28.1516 +
28.1517 +lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)"
28.1518 +by auto
28.1519 +
28.1520 +lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)"
28.1521 +by auto
28.1522 +
28.1523 +lemma real_add_le_0_iff: "(x+y \<le> (0::real)) = (y \<le> -x)"
28.1524 +by auto
28.1525 +
28.1526 +lemma real_0_le_add_iff: "((0::real) \<le> x+y) = (-x \<le> y)"
28.1527 +by auto
28.1528 +
28.1529 +subsection {* Lemmas about powers *}
28.1530 +
28.1531 +text {* FIXME: declare this in Rings.thy or not at all *}
28.1532 +declare abs_mult_self [simp]
28.1533 +
28.1534 +(* used by Import/HOL/real.imp *)
28.1535 +lemma two_realpow_ge_one: "(1::real) \<le> 2 ^ n"
28.1536 +by simp
28.1537 +
28.1538 +lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n"
28.1539 +apply (induct "n")
28.1540 +apply (auto simp add: real_of_nat_Suc)
28.1541 +apply (subst mult_2)
28.1542 +apply (erule add_less_le_mono)
28.1543 +apply (rule two_realpow_ge_one)
28.1544 +done
28.1545 +
28.1546 +text {* TODO: no longer real-specific; rename and move elsewhere *}
28.1547 +lemma realpow_Suc_le_self:
28.1548 + fixes r :: "'a::linordered_semidom"
28.1549 + shows "[| 0 \<le> r; r \<le> 1 |] ==> r ^ Suc n \<le> r"
28.1550 +by (insert power_decreasing [of 1 "Suc n" r], simp)
28.1551 +
28.1552 +text {* TODO: no longer real-specific; rename and move elsewhere *}
28.1553 +lemma realpow_minus_mult:
28.1554 + fixes x :: "'a::monoid_mult"
28.1555 + shows "0 < n \<Longrightarrow> x ^ (n - 1) * x = x ^ n"
28.1556 +by (simp add: power_commutes split add: nat_diff_split)
28.1557 +
28.1558 +text {* FIXME: declare this [simp] for all types, or not at all *}
28.1559 +lemma real_two_squares_add_zero_iff [simp]:
28.1560 + "(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)"
28.1561 +by (rule sum_squares_eq_zero_iff)
28.1562 +
28.1563 +text {* FIXME: declare this [simp] for all types, or not at all *}
28.1564 +lemma realpow_two_sum_zero_iff [simp]:
28.1565 + "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)"
28.1566 +by (rule sum_power2_eq_zero_iff)
28.1567 +
28.1568 +lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))"
28.1569 +by (rule_tac y = 0 in order_trans, auto)
28.1570 +
28.1571 +lemma realpow_square_minus_le [simp]: "-(u ^ 2) \<le> (x::real) ^ 2"
28.1572 +by (auto simp add: power2_eq_square)
28.1573 +
28.1574 +
28.1575 +lemma numeral_power_le_real_of_nat_cancel_iff[simp]:
28.1576 + "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::nat) ^ n \<le> a"
28.1577 + unfolding real_of_nat_le_iff[symmetric] by simp
28.1578 +
28.1579 +lemma real_of_nat_le_numeral_power_cancel_iff[simp]:
28.1580 + "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::nat) ^ n"
28.1581 + unfolding real_of_nat_le_iff[symmetric] by simp
28.1582 +
28.1583 +lemma numeral_power_le_real_of_int_cancel_iff[simp]:
28.1584 + "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::int) ^ n \<le> a"
28.1585 + unfolding real_of_int_le_iff[symmetric] by simp
28.1586 +
28.1587 +lemma real_of_int_le_numeral_power_cancel_iff[simp]:
28.1588 + "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::int) ^ n"
28.1589 + unfolding real_of_int_le_iff[symmetric] by simp
28.1590 +
28.1591 +lemma neg_numeral_power_le_real_of_int_cancel_iff[simp]:
28.1592 + "(neg_numeral x::real) ^ n \<le> real a \<longleftrightarrow> (neg_numeral x::int) ^ n \<le> a"
28.1593 + unfolding real_of_int_le_iff[symmetric] by simp
28.1594 +
28.1595 +lemma real_of_int_le_neg_numeral_power_cancel_iff[simp]:
28.1596 + "real a \<le> (neg_numeral x::real) ^ n \<longleftrightarrow> a \<le> (neg_numeral x::int) ^ n"
28.1597 + unfolding real_of_int_le_iff[symmetric] by simp
28.1598 +
28.1599 +subsection{*Density of the Reals*}
28.1600 +
28.1601 +lemma real_lbound_gt_zero:
28.1602 + "[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2"
28.1603 +apply (rule_tac x = " (min d1 d2) /2" in exI)
28.1604 +apply (simp add: min_def)
28.1605 +done
28.1606 +
28.1607 +
28.1608 +text{*Similar results are proved in @{text Fields}*}
28.1609 +lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
28.1610 + by auto
28.1611 +
28.1612 +lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y"
28.1613 + by auto
28.1614 +
28.1615 +lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
28.1616 + by simp
28.1617 +
28.1618 +subsection{*Absolute Value Function for the Reals*}
28.1619 +
28.1620 +lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))"
28.1621 +by (simp add: abs_if)
28.1622 +
28.1623 +(* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *)
28.1624 +lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
28.1625 +by (force simp add: abs_le_iff)
28.1626 +
28.1627 +lemma abs_add_one_gt_zero: "(0::real) < 1 + abs(x)"
28.1628 +by (simp add: abs_if)
28.1629 +
28.1630 +lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
28.1631 +by (rule abs_of_nonneg [OF real_of_nat_ge_zero])
28.1632 +
28.1633 +lemma abs_add_one_not_less_self: "~ abs(x) + (1::real) < x"
28.1634 +by simp
28.1635 +
28.1636 +lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)"
28.1637 +by simp
28.1638 +
28.1639 +
28.1640 +subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
28.1641 +
28.1642 +(* FIXME: theorems for negative numerals *)
28.1643 +lemma numeral_less_real_of_int_iff [simp]:
28.1644 + "((numeral n) < real (m::int)) = (numeral n < m)"
28.1645 +apply auto
28.1646 +apply (rule real_of_int_less_iff [THEN iffD1])
28.1647 +apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
28.1648 +done
28.1649 +
28.1650 +lemma numeral_less_real_of_int_iff2 [simp]:
28.1651 + "(real (m::int) < (numeral n)) = (m < numeral n)"
28.1652 +apply auto
28.1653 +apply (rule real_of_int_less_iff [THEN iffD1])
28.1654 +apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
28.1655 +done
28.1656 +
28.1657 +lemma numeral_le_real_of_int_iff [simp]:
28.1658 + "((numeral n) \<le> real (m::int)) = (numeral n \<le> m)"
28.1659 +by (simp add: linorder_not_less [symmetric])
28.1660 +
28.1661 +lemma numeral_le_real_of_int_iff2 [simp]:
28.1662 + "(real (m::int) \<le> (numeral n)) = (m \<le> numeral n)"
28.1663 +by (simp add: linorder_not_less [symmetric])
28.1664 +
28.1665 +lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"
28.1666 +unfolding real_of_nat_def by simp
28.1667 +
28.1668 +lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"
28.1669 +unfolding real_of_nat_def by (simp add: floor_minus)
28.1670 +
28.1671 +lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"
28.1672 +unfolding real_of_int_def by simp
28.1673 +
28.1674 +lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"
28.1675 +unfolding real_of_int_def by (simp add: floor_minus)
28.1676 +
28.1677 +lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
28.1678 +unfolding real_of_int_def by (rule floor_exists)
28.1679 +
28.1680 +lemma lemma_floor:
28.1681 + assumes a1: "real m \<le> r" and a2: "r < real n + 1"
28.1682 + shows "m \<le> (n::int)"
28.1683 +proof -
28.1684 + have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)
28.1685 + also have "... = real (n + 1)" by simp
28.1686 + finally have "m < n + 1" by (simp only: real_of_int_less_iff)
28.1687 + thus ?thesis by arith
28.1688 +qed
28.1689 +
28.1690 +lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"
28.1691 +unfolding real_of_int_def by (rule of_int_floor_le)
28.1692 +
28.1693 +lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \<le> x"
28.1694 +by (auto intro: lemma_floor)
28.1695 +
28.1696 +lemma real_of_int_floor_cancel [simp]:
28.1697 + "(real (floor x) = x) = (\<exists>n::int. x = real n)"
28.1698 + using floor_real_of_int by metis
28.1699 +
28.1700 +lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
28.1701 + unfolding real_of_int_def using floor_unique [of n x] by simp
28.1702 +
28.1703 +lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"
28.1704 + unfolding real_of_int_def by (rule floor_unique)
28.1705 +
28.1706 +lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n"
28.1707 +apply (rule inj_int [THEN injD])
28.1708 +apply (simp add: real_of_nat_Suc)
28.1709 +apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"])
28.1710 +done
28.1711 +
28.1712 +lemma floor_eq4: "[| real n \<le> x; x < real (Suc n) |] ==> nat(floor x) = n"
28.1713 +apply (drule order_le_imp_less_or_eq)
28.1714 +apply (auto intro: floor_eq3)
28.1715 +done
28.1716 +
28.1717 +lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
28.1718 + unfolding real_of_int_def using floor_correct [of r] by simp
28.1719 +
28.1720 +lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"
28.1721 + unfolding real_of_int_def using floor_correct [of r] by simp
28.1722 +
28.1723 +lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"
28.1724 + unfolding real_of_int_def using floor_correct [of r] by simp
28.1725 +
28.1726 +lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"
28.1727 + unfolding real_of_int_def using floor_correct [of r] by simp
28.1728 +
28.1729 +lemma le_floor: "real a <= x ==> a <= floor x"
28.1730 + unfolding real_of_int_def by (simp add: le_floor_iff)
28.1731 +
28.1732 +lemma real_le_floor: "a <= floor x ==> real a <= x"
28.1733 + unfolding real_of_int_def by (simp add: le_floor_iff)
28.1734 +
28.1735 +lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
28.1736 + unfolding real_of_int_def by (rule le_floor_iff)
28.1737 +
28.1738 +lemma floor_less_eq: "(floor x < a) = (x < real a)"
28.1739 + unfolding real_of_int_def by (rule floor_less_iff)
28.1740 +
28.1741 +lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
28.1742 + unfolding real_of_int_def by (rule less_floor_iff)
28.1743 +
28.1744 +lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
28.1745 + unfolding real_of_int_def by (rule floor_le_iff)
28.1746 +
28.1747 +lemma floor_add [simp]: "floor (x + real a) = floor x + a"
28.1748 + unfolding real_of_int_def by (rule floor_add_of_int)
28.1749 +
28.1750 +lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
28.1751 + unfolding real_of_int_def by (rule floor_diff_of_int)
28.1752 +
28.1753 +lemma le_mult_floor:
28.1754 + assumes "0 \<le> (a :: real)" and "0 \<le> b"
28.1755 + shows "floor a * floor b \<le> floor (a * b)"
28.1756 +proof -
28.1757 + have "real (floor a) \<le> a"
28.1758 + and "real (floor b) \<le> b" by auto
28.1759 + hence "real (floor a * floor b) \<le> a * b"
28.1760 + using assms by (auto intro!: mult_mono)
28.1761 + also have "a * b < real (floor (a * b) + 1)" by auto
28.1762 + finally show ?thesis unfolding real_of_int_less_iff by simp
28.1763 +qed
28.1764 +
28.1765 +lemma floor_divide_eq_div:
28.1766 + "floor (real a / real b) = a div b"
28.1767 +proof cases
28.1768 + assume "b \<noteq> 0 \<or> b dvd a"
28.1769 + with real_of_int_div3[of a b] show ?thesis
28.1770 + by (auto simp: real_of_int_div[symmetric] intro!: floor_eq2 real_of_int_div4 neq_le_trans)
28.1771 + (metis add_left_cancel zero_neq_one real_of_int_div_aux real_of_int_inject
28.1772 + real_of_int_zero_cancel right_inverse_eq div_self mod_div_trivial)
28.1773 +qed (auto simp: real_of_int_div)
28.1774 +
28.1775 +lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
28.1776 + unfolding real_of_nat_def by simp
28.1777 +
28.1778 +lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"
28.1779 + unfolding real_of_int_def by (rule le_of_int_ceiling)
28.1780 +
28.1781 +lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"
28.1782 + unfolding real_of_int_def by simp
28.1783 +
28.1784 +lemma real_of_int_ceiling_cancel [simp]:
28.1785 + "(real (ceiling x) = x) = (\<exists>n::int. x = real n)"
28.1786 + using ceiling_real_of_int by metis
28.1787 +
28.1788 +lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"
28.1789 + unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
28.1790 +
28.1791 +lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"
28.1792 + unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
28.1793 +
28.1794 +lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n |] ==> ceiling x = n"
28.1795 + unfolding real_of_int_def using ceiling_unique [of n x] by simp
28.1796 +
28.1797 +lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
28.1798 + unfolding real_of_int_def using ceiling_correct [of r] by simp
28.1799 +
28.1800 +lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"
28.1801 + unfolding real_of_int_def using ceiling_correct [of r] by simp
28.1802 +
28.1803 +lemma ceiling_le: "x <= real a ==> ceiling x <= a"
28.1804 + unfolding real_of_int_def by (simp add: ceiling_le_iff)
28.1805 +
28.1806 +lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"
28.1807 + unfolding real_of_int_def by (simp add: ceiling_le_iff)
28.1808 +
28.1809 +lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
28.1810 + unfolding real_of_int_def by (rule ceiling_le_iff)
28.1811 +
28.1812 +lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
28.1813 + unfolding real_of_int_def by (rule less_ceiling_iff)
28.1814 +
28.1815 +lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
28.1816 + unfolding real_of_int_def by (rule ceiling_less_iff)
28.1817 +
28.1818 +lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
28.1819 + unfolding real_of_int_def by (rule le_ceiling_iff)
28.1820 +
28.1821 +lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
28.1822 + unfolding real_of_int_def by (rule ceiling_add_of_int)
28.1823 +
28.1824 +lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
28.1825 + unfolding real_of_int_def by (rule ceiling_diff_of_int)
28.1826 +
28.1827 +
28.1828 +subsubsection {* Versions for the natural numbers *}
28.1829 +
28.1830 +definition
28.1831 + natfloor :: "real => nat" where
28.1832 + "natfloor x = nat(floor x)"
28.1833 +
28.1834 +definition
28.1835 + natceiling :: "real => nat" where
28.1836 + "natceiling x = nat(ceiling x)"
28.1837 +
28.1838 +lemma natfloor_zero [simp]: "natfloor 0 = 0"
28.1839 + by (unfold natfloor_def, simp)
28.1840 +
28.1841 +lemma natfloor_one [simp]: "natfloor 1 = 1"
28.1842 + by (unfold natfloor_def, simp)
28.1843 +
28.1844 +lemma zero_le_natfloor [simp]: "0 <= natfloor x"
28.1845 + by (unfold natfloor_def, simp)
28.1846 +
28.1847 +lemma natfloor_numeral_eq [simp]: "natfloor (numeral n) = numeral n"
28.1848 + by (unfold natfloor_def, simp)
28.1849 +
28.1850 +lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"
28.1851 + by (unfold natfloor_def, simp)
28.1852 +
28.1853 +lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"
28.1854 + by (unfold natfloor_def, simp)
28.1855 +
28.1856 +lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"
28.1857 + unfolding natfloor_def by simp
28.1858 +
28.1859 +lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
28.1860 + unfolding natfloor_def by (intro nat_mono floor_mono)
28.1861 +
28.1862 +lemma le_natfloor: "real x <= a ==> x <= natfloor a"
28.1863 + apply (unfold natfloor_def)
28.1864 + apply (subst nat_int [THEN sym])
28.1865 + apply (rule nat_mono)
28.1866 + apply (rule le_floor)
28.1867 + apply simp
28.1868 +done
28.1869 +
28.1870 +lemma natfloor_less_iff: "0 \<le> x \<Longrightarrow> natfloor x < n \<longleftrightarrow> x < real n"
28.1871 + unfolding natfloor_def real_of_nat_def
28.1872 + by (simp add: nat_less_iff floor_less_iff)
28.1873 +
28.1874 +lemma less_natfloor:
28.1875 + assumes "0 \<le> x" and "x < real (n :: nat)"
28.1876 + shows "natfloor x < n"
28.1877 + using assms by (simp add: natfloor_less_iff)
28.1878 +
28.1879 +lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)"
28.1880 + apply (rule iffI)
28.1881 + apply (rule order_trans)
28.1882 + prefer 2
28.1883 + apply (erule real_natfloor_le)
28.1884 + apply (subst real_of_nat_le_iff)
28.1885 + apply assumption
28.1886 + apply (erule le_natfloor)
28.1887 +done
28.1888 +
28.1889 +lemma le_natfloor_eq_numeral [simp]:
28.1890 + "~ neg((numeral n)::int) ==> 0 <= x ==>
28.1891 + (numeral n <= natfloor x) = (numeral n <= x)"
28.1892 + apply (subst le_natfloor_eq, assumption)
28.1893 + apply simp
28.1894 +done
28.1895 +
28.1896 +lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
28.1897 + apply (case_tac "0 <= x")
28.1898 + apply (subst le_natfloor_eq, assumption, simp)
28.1899 + apply (rule iffI)
28.1900 + apply (subgoal_tac "natfloor x <= natfloor 0")
28.1901 + apply simp
28.1902 + apply (rule natfloor_mono)
28.1903 + apply simp
28.1904 + apply simp
28.1905 +done
28.1906 +
28.1907 +lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n"
28.1908 + unfolding natfloor_def by (simp add: floor_eq2 [where n="int n"])
28.1909 +
28.1910 +lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"
28.1911 + apply (case_tac "0 <= x")
28.1912 + apply (unfold natfloor_def)
28.1913 + apply simp
28.1914 + apply simp_all
28.1915 +done
28.1916 +
28.1917 +lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
28.1918 +using real_natfloor_add_one_gt by (simp add: algebra_simps)
28.1919 +
28.1920 +lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"
28.1921 + apply (subgoal_tac "z < real(natfloor z) + 1")
28.1922 + apply arith
28.1923 + apply (rule real_natfloor_add_one_gt)
28.1924 +done
28.1925 +
28.1926 +lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a"
28.1927 + unfolding natfloor_def
28.1928 + unfolding real_of_int_of_nat_eq [symmetric] floor_add
28.1929 + by (simp add: nat_add_distrib)
28.1930 +
28.1931 +lemma natfloor_add_numeral [simp]:
28.1932 + "~neg ((numeral n)::int) ==> 0 <= x ==>
28.1933 + natfloor (x + numeral n) = natfloor x + numeral n"
28.1934 + by (simp add: natfloor_add [symmetric])
28.1935 +
28.1936 +lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1"
28.1937 + by (simp add: natfloor_add [symmetric] del: One_nat_def)
28.1938 +
28.1939 +lemma natfloor_subtract [simp]:
28.1940 + "natfloor(x - real a) = natfloor x - a"
28.1941 + unfolding natfloor_def real_of_int_of_nat_eq [symmetric] floor_subtract
28.1942 + by simp
28.1943 +
28.1944 +lemma natfloor_div_nat:
28.1945 + assumes "1 <= x" and "y > 0"
28.1946 + shows "natfloor (x / real y) = natfloor x div y"
28.1947 +proof (rule natfloor_eq)
28.1948 + have "(natfloor x) div y * y \<le> natfloor x"
28.1949 + by (rule add_leD1 [where k="natfloor x mod y"], simp)
28.1950 + thus "real (natfloor x div y) \<le> x / real y"
28.1951 + using assms by (simp add: le_divide_eq le_natfloor_eq)
28.1952 + have "natfloor x < (natfloor x) div y * y + y"
28.1953 + apply (subst mod_div_equality [symmetric])
28.1954 + apply (rule add_strict_left_mono)
28.1955 + apply (rule mod_less_divisor)
28.1956 + apply fact
28.1957 + done
28.1958 + thus "x / real y < real (natfloor x div y) + 1"
28.1959 + using assms
28.1960 + by (simp add: divide_less_eq natfloor_less_iff distrib_right)
28.1961 +qed
28.1962 +
28.1963 +lemma le_mult_natfloor:
28.1964 + shows "natfloor a * natfloor b \<le> natfloor (a * b)"
28.1965 + by (cases "0 <= a & 0 <= b")
28.1966 + (auto simp add: le_natfloor_eq mult_nonneg_nonneg mult_mono' real_natfloor_le natfloor_neg)
28.1967 +
28.1968 +lemma natceiling_zero [simp]: "natceiling 0 = 0"
28.1969 + by (unfold natceiling_def, simp)
28.1970 +
28.1971 +lemma natceiling_one [simp]: "natceiling 1 = 1"
28.1972 + by (unfold natceiling_def, simp)
28.1973 +
28.1974 +lemma zero_le_natceiling [simp]: "0 <= natceiling x"
28.1975 + by (unfold natceiling_def, simp)
28.1976 +
28.1977 +lemma natceiling_numeral_eq [simp]: "natceiling (numeral n) = numeral n"
28.1978 + by (unfold natceiling_def, simp)
28.1979 +
28.1980 +lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"
28.1981 + by (unfold natceiling_def, simp)
28.1982 +
28.1983 +lemma real_natceiling_ge: "x <= real(natceiling x)"
28.1984 + unfolding natceiling_def by (cases "x < 0", simp_all)
28.1985 +
28.1986 +lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"
28.1987 + unfolding natceiling_def by simp
28.1988 +
28.1989 +lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"
28.1990 + unfolding natceiling_def by (intro nat_mono ceiling_mono)
28.1991 +
28.1992 +lemma natceiling_le: "x <= real a ==> natceiling x <= a"
28.1993 + unfolding natceiling_def real_of_nat_def
28.1994 + by (simp add: nat_le_iff ceiling_le_iff)
28.1995 +
28.1996 +lemma natceiling_le_eq: "(natceiling x <= a) = (x <= real a)"
28.1997 + unfolding natceiling_def real_of_nat_def
28.1998 + by (simp add: nat_le_iff ceiling_le_iff)
28.1999 +
28.2000 +lemma natceiling_le_eq_numeral [simp]:
28.2001 + "~ neg((numeral n)::int) ==>
28.2002 + (natceiling x <= numeral n) = (x <= numeral n)"
28.2003 + by (simp add: natceiling_le_eq)
28.2004 +
28.2005 +lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"
28.2006 + unfolding natceiling_def
28.2007 + by (simp add: nat_le_iff ceiling_le_iff)
28.2008 +
28.2009 +lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1"
28.2010 + unfolding natceiling_def
28.2011 + by (simp add: ceiling_eq2 [where n="int n"])
28.2012 +
28.2013 +lemma natceiling_add [simp]: "0 <= x ==>
28.2014 + natceiling (x + real a) = natceiling x + a"
28.2015 + unfolding natceiling_def
28.2016 + unfolding real_of_int_of_nat_eq [symmetric] ceiling_add
28.2017 + by (simp add: nat_add_distrib)
28.2018 +
28.2019 +lemma natceiling_add_numeral [simp]:
28.2020 + "~ neg ((numeral n)::int) ==> 0 <= x ==>
28.2021 + natceiling (x + numeral n) = natceiling x + numeral n"
28.2022 + by (simp add: natceiling_add [symmetric])
28.2023 +
28.2024 +lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1"
28.2025 + by (simp add: natceiling_add [symmetric] del: One_nat_def)
28.2026 +
28.2027 +lemma natceiling_subtract [simp]: "natceiling(x - real a) = natceiling x - a"
28.2028 + unfolding natceiling_def real_of_int_of_nat_eq [symmetric] ceiling_subtract
28.2029 + by simp
28.2030 +
28.2031 +subsection {* Exponentiation with floor *}
28.2032 +
28.2033 +lemma floor_power:
28.2034 + assumes "x = real (floor x)"
28.2035 + shows "floor (x ^ n) = floor x ^ n"
28.2036 +proof -
28.2037 + have *: "x ^ n = real (floor x ^ n)"
28.2038 + using assms by (induct n arbitrary: x) simp_all
28.2039 + show ?thesis unfolding real_of_int_inject[symmetric]
28.2040 + unfolding * floor_real_of_int ..
28.2041 +qed
28.2042 +
28.2043 +lemma natfloor_power:
28.2044 + assumes "x = real (natfloor x)"
28.2045 + shows "natfloor (x ^ n) = natfloor x ^ n"
28.2046 +proof -
28.2047 + from assms have "0 \<le> floor x" by auto
28.2048 + note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \<le> floor x`]]
28.2049 + from floor_power[OF this]
28.2050 + show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \<le> floor x`, symmetric]
28.2051 + by simp
28.2052 +qed
28.2053 +
28.2054 +
28.2055 +subsection {* Implementation of rational real numbers *}
28.2056 +
28.2057 +text {* Formal constructor *}
28.2058 +
28.2059 +definition Ratreal :: "rat \<Rightarrow> real" where
28.2060 + [code_abbrev, simp]: "Ratreal = of_rat"
28.2061 +
28.2062 +code_datatype Ratreal
28.2063 +
28.2064 +
28.2065 +text {* Numerals *}
28.2066 +
28.2067 +lemma [code_abbrev]:
28.2068 + "(of_rat (of_int a) :: real) = of_int a"
28.2069 + by simp
28.2070 +
28.2071 +lemma [code_abbrev]:
28.2072 + "(of_rat 0 :: real) = 0"
28.2073 + by simp
28.2074 +
28.2075 +lemma [code_abbrev]:
28.2076 + "(of_rat 1 :: real) = 1"
28.2077 + by simp
28.2078 +
28.2079 +lemma [code_abbrev]:
28.2080 + "(of_rat (numeral k) :: real) = numeral k"
28.2081 + by simp
28.2082 +
28.2083 +lemma [code_abbrev]:
28.2084 + "(of_rat (neg_numeral k) :: real) = neg_numeral k"
28.2085 + by simp
28.2086 +
28.2087 +lemma [code_post]:
28.2088 + "(of_rat (0 / r) :: real) = 0"
28.2089 + "(of_rat (r / 0) :: real) = 0"
28.2090 + "(of_rat (1 / 1) :: real) = 1"
28.2091 + "(of_rat (numeral k / 1) :: real) = numeral k"
28.2092 + "(of_rat (neg_numeral k / 1) :: real) = neg_numeral k"
28.2093 + "(of_rat (1 / numeral k) :: real) = 1 / numeral k"
28.2094 + "(of_rat (1 / neg_numeral k) :: real) = 1 / neg_numeral k"
28.2095 + "(of_rat (numeral k / numeral l) :: real) = numeral k / numeral l"
28.2096 + "(of_rat (numeral k / neg_numeral l) :: real) = numeral k / neg_numeral l"
28.2097 + "(of_rat (neg_numeral k / numeral l) :: real) = neg_numeral k / numeral l"
28.2098 + "(of_rat (neg_numeral k / neg_numeral l) :: real) = neg_numeral k / neg_numeral l"
28.2099 + by (simp_all add: of_rat_divide)
28.2100 +
28.2101 +
28.2102 +text {* Operations *}
28.2103 +
28.2104 +lemma zero_real_code [code]:
28.2105 + "0 = Ratreal 0"
28.2106 +by simp
28.2107 +
28.2108 +lemma one_real_code [code]:
28.2109 + "1 = Ratreal 1"
28.2110 +by simp
28.2111 +
28.2112 +instantiation real :: equal
28.2113 +begin
28.2114 +
28.2115 +definition "HOL.equal (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
28.2116 +
28.2117 +instance proof
28.2118 +qed (simp add: equal_real_def)
28.2119 +
28.2120 +lemma real_equal_code [code]:
28.2121 + "HOL.equal (Ratreal x) (Ratreal y) \<longleftrightarrow> HOL.equal x y"
28.2122 + by (simp add: equal_real_def equal)
28.2123 +
28.2124 +lemma [code nbe]:
28.2125 + "HOL.equal (x::real) x \<longleftrightarrow> True"
28.2126 + by (rule equal_refl)
28.2127 +
28.2128 +end
28.2129 +
28.2130 +lemma real_less_eq_code [code]: "Ratreal x \<le> Ratreal y \<longleftrightarrow> x \<le> y"
28.2131 + by (simp add: of_rat_less_eq)
28.2132 +
28.2133 +lemma real_less_code [code]: "Ratreal x < Ratreal y \<longleftrightarrow> x < y"
28.2134 + by (simp add: of_rat_less)
28.2135 +
28.2136 +lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)"
28.2137 + by (simp add: of_rat_add)
28.2138 +
28.2139 +lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)"
28.2140 + by (simp add: of_rat_mult)
28.2141 +
28.2142 +lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)"
28.2143 + by (simp add: of_rat_minus)
28.2144 +
28.2145 +lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)"
28.2146 + by (simp add: of_rat_diff)
28.2147 +
28.2148 +lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)"
28.2149 + by (simp add: of_rat_inverse)
28.2150 +
28.2151 +lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)"
28.2152 + by (simp add: of_rat_divide)
28.2153 +
28.2154 +lemma real_floor_code [code]: "floor (Ratreal x) = floor x"
28.2155 + by (metis Ratreal_def floor_le_iff floor_unique le_floor_iff of_int_floor_le of_rat_of_int_eq real_less_eq_code)
28.2156 +
28.2157 +
28.2158 +text {* Quickcheck *}
28.2159 +
28.2160 +definition (in term_syntax)
28.2161 + valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
28.2162 + [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\<cdot>} k"
28.2163 +
28.2164 +notation fcomp (infixl "\<circ>>" 60)
28.2165 +notation scomp (infixl "\<circ>\<rightarrow>" 60)
28.2166 +
28.2167 +instantiation real :: random
28.2168 +begin
28.2169 +
28.2170 +definition
28.2171 + "Quickcheck_Random.random i = Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>r. Pair (valterm_ratreal r))"
28.2172 +
28.2173 +instance ..
28.2174 +
28.2175 +end
28.2176 +
28.2177 +no_notation fcomp (infixl "\<circ>>" 60)
28.2178 +no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
28.2179 +
28.2180 +instantiation real :: exhaustive
28.2181 +begin
28.2182 +
28.2183 +definition
28.2184 + "exhaustive_real f d = Quickcheck_Exhaustive.exhaustive (%r. f (Ratreal r)) d"
28.2185 +
28.2186 +instance ..
28.2187 +
28.2188 +end
28.2189 +
28.2190 +instantiation real :: full_exhaustive
28.2191 +begin
28.2192 +
28.2193 +definition
28.2194 + "full_exhaustive_real f d = Quickcheck_Exhaustive.full_exhaustive (%r. f (valterm_ratreal r)) d"
28.2195 +
28.2196 +instance ..
28.2197 +
28.2198 +end
28.2199 +
28.2200 +instantiation real :: narrowing
28.2201 +begin
28.2202 +
28.2203 +definition
28.2204 + "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing"
28.2205 +
28.2206 +instance ..
28.2207 +
28.2208 +end
28.2209 +
28.2210 +
28.2211 +subsection {* Setup for Nitpick *}
28.2212 +
28.2213 +declaration {*
28.2214 + Nitpick_HOL.register_frac_type @{type_name real}
28.2215 + [(@{const_name zero_real_inst.zero_real}, @{const_name Nitpick.zero_frac}),
28.2216 + (@{const_name one_real_inst.one_real}, @{const_name Nitpick.one_frac}),
28.2217 + (@{const_name plus_real_inst.plus_real}, @{const_name Nitpick.plus_frac}),
28.2218 + (@{const_name times_real_inst.times_real}, @{const_name Nitpick.times_frac}),
28.2219 + (@{const_name uminus_real_inst.uminus_real}, @{const_name Nitpick.uminus_frac}),
28.2220 + (@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}),
28.2221 + (@{const_name ord_real_inst.less_real}, @{const_name Nitpick.less_frac}),
28.2222 + (@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})]
28.2223 +*}
28.2224 +
28.2225 +lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real
28.2226 + ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real
28.2227 + times_real_inst.times_real uminus_real_inst.uminus_real
28.2228 + zero_real_inst.zero_real
28.2229 +
28.2230 ML_file "Tools/SMT/smt_real.ML"
28.2231 setup SMT_Real.setup
28.2232
29.1 --- a/src/HOL/RealDef.thy Tue Mar 26 14:38:44 2013 +0100
29.2 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000
29.3 @@ -1,1704 +0,0 @@
29.4 -(* Title : HOL/RealDef.thy
29.5 - Author : Jacques D. Fleuriot, 1998
29.6 - Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
29.7 - Additional contributions by Jeremy Avigad
29.8 - Construction of Cauchy Reals by Brian Huffman, 2010
29.9 -*)
29.10 -
29.11 -header {* Development of the Reals using Cauchy Sequences *}
29.12 -
29.13 -theory RealDef
29.14 -imports Rat
29.15 -begin
29.16 -
29.17 -text {*
29.18 - This theory contains a formalization of the real numbers as
29.19 - equivalence classes of Cauchy sequences of rationals. See
29.20 - @{file "~~/src/HOL/ex/Dedekind_Real.thy"} for an alternative
29.21 - construction using Dedekind cuts.
29.22 -*}
29.23 -
29.24 -subsection {* Preliminary lemmas *}
29.25 -
29.26 -lemma add_diff_add:
29.27 - fixes a b c d :: "'a::ab_group_add"
29.28 - shows "(a + c) - (b + d) = (a - b) + (c - d)"
29.29 - by simp
29.30 -
29.31 -lemma minus_diff_minus:
29.32 - fixes a b :: "'a::ab_group_add"
29.33 - shows "- a - - b = - (a - b)"
29.34 - by simp
29.35 -
29.36 -lemma mult_diff_mult:
29.37 - fixes x y a b :: "'a::ring"
29.38 - shows "(x * y - a * b) = x * (y - b) + (x - a) * b"
29.39 - by (simp add: algebra_simps)
29.40 -
29.41 -lemma inverse_diff_inverse:
29.42 - fixes a b :: "'a::division_ring"
29.43 - assumes "a \<noteq> 0" and "b \<noteq> 0"
29.44 - shows "inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
29.45 - using assms by (simp add: algebra_simps)
29.46 -
29.47 -lemma obtain_pos_sum:
29.48 - fixes r :: rat assumes r: "0 < r"
29.49 - obtains s t where "0 < s" and "0 < t" and "r = s + t"
29.50 -proof
29.51 - from r show "0 < r/2" by simp
29.52 - from r show "0 < r/2" by simp
29.53 - show "r = r/2 + r/2" by simp
29.54 -qed
29.55 -
29.56 -subsection {* Sequences that converge to zero *}
29.57 -
29.58 -definition
29.59 - vanishes :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
29.60 -where
29.61 - "vanishes X = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r)"
29.62 -
29.63 -lemma vanishesI: "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r) \<Longrightarrow> vanishes X"
29.64 - unfolding vanishes_def by simp
29.65 -
29.66 -lemma vanishesD: "\<lbrakk>vanishes X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r"
29.67 - unfolding vanishes_def by simp
29.68 -
29.69 -lemma vanishes_const [simp]: "vanishes (\<lambda>n. c) \<longleftrightarrow> c = 0"
29.70 - unfolding vanishes_def
29.71 - apply (cases "c = 0", auto)
29.72 - apply (rule exI [where x="\<bar>c\<bar>"], auto)
29.73 - done
29.74 -
29.75 -lemma vanishes_minus: "vanishes X \<Longrightarrow> vanishes (\<lambda>n. - X n)"
29.76 - unfolding vanishes_def by simp
29.77 -
29.78 -lemma vanishes_add:
29.79 - assumes X: "vanishes X" and Y: "vanishes Y"
29.80 - shows "vanishes (\<lambda>n. X n + Y n)"
29.81 -proof (rule vanishesI)
29.82 - fix r :: rat assume "0 < r"
29.83 - then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
29.84 - by (rule obtain_pos_sum)
29.85 - obtain i where i: "\<forall>n\<ge>i. \<bar>X n\<bar> < s"
29.86 - using vanishesD [OF X s] ..
29.87 - obtain j where j: "\<forall>n\<ge>j. \<bar>Y n\<bar> < t"
29.88 - using vanishesD [OF Y t] ..
29.89 - have "\<forall>n\<ge>max i j. \<bar>X n + Y n\<bar> < r"
29.90 - proof (clarsimp)
29.91 - fix n assume n: "i \<le> n" "j \<le> n"
29.92 - have "\<bar>X n + Y n\<bar> \<le> \<bar>X n\<bar> + \<bar>Y n\<bar>" by (rule abs_triangle_ineq)
29.93 - also have "\<dots> < s + t" by (simp add: add_strict_mono i j n)
29.94 - finally show "\<bar>X n + Y n\<bar> < r" unfolding r .
29.95 - qed
29.96 - thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n + Y n\<bar> < r" ..
29.97 -qed
29.98 -
29.99 -lemma vanishes_diff:
29.100 - assumes X: "vanishes X" and Y: "vanishes Y"
29.101 - shows "vanishes (\<lambda>n. X n - Y n)"
29.102 -unfolding diff_minus by (intro vanishes_add vanishes_minus X Y)
29.103 -
29.104 -lemma vanishes_mult_bounded:
29.105 - assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
29.106 - assumes Y: "vanishes (\<lambda>n. Y n)"
29.107 - shows "vanishes (\<lambda>n. X n * Y n)"
29.108 -proof (rule vanishesI)
29.109 - fix r :: rat assume r: "0 < r"
29.110 - obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
29.111 - using X by fast
29.112 - obtain b where b: "0 < b" "r = a * b"
29.113 - proof
29.114 - show "0 < r / a" using r a by (simp add: divide_pos_pos)
29.115 - show "r = a * (r / a)" using a by simp
29.116 - qed
29.117 - obtain k where k: "\<forall>n\<ge>k. \<bar>Y n\<bar> < b"
29.118 - using vanishesD [OF Y b(1)] ..
29.119 - have "\<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r"
29.120 - by (simp add: b(2) abs_mult mult_strict_mono' a k)
29.121 - thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r" ..
29.122 -qed
29.123 -
29.124 -subsection {* Cauchy sequences *}
29.125 -
29.126 -definition
29.127 - cauchy :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
29.128 -where
29.129 - "cauchy X \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r)"
29.130 -
29.131 -lemma cauchyI:
29.132 - "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r) \<Longrightarrow> cauchy X"
29.133 - unfolding cauchy_def by simp
29.134 -
29.135 -lemma cauchyD:
29.136 - "\<lbrakk>cauchy X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r"
29.137 - unfolding cauchy_def by simp
29.138 -
29.139 -lemma cauchy_const [simp]: "cauchy (\<lambda>n. x)"
29.140 - unfolding cauchy_def by simp
29.141 -
29.142 -lemma cauchy_add [simp]:
29.143 - assumes X: "cauchy X" and Y: "cauchy Y"
29.144 - shows "cauchy (\<lambda>n. X n + Y n)"
29.145 -proof (rule cauchyI)
29.146 - fix r :: rat assume "0 < r"
29.147 - then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
29.148 - by (rule obtain_pos_sum)
29.149 - obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
29.150 - using cauchyD [OF X s] ..
29.151 - obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
29.152 - using cauchyD [OF Y t] ..
29.153 - have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r"
29.154 - proof (clarsimp)
29.155 - fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
29.156 - have "\<bar>(X m + Y m) - (X n + Y n)\<bar> \<le> \<bar>X m - X n\<bar> + \<bar>Y m - Y n\<bar>"
29.157 - unfolding add_diff_add by (rule abs_triangle_ineq)
29.158 - also have "\<dots> < s + t"
29.159 - by (rule add_strict_mono, simp_all add: i j *)
29.160 - finally show "\<bar>(X m + Y m) - (X n + Y n)\<bar> < r" unfolding r .
29.161 - qed
29.162 - thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r" ..
29.163 -qed
29.164 -
29.165 -lemma cauchy_minus [simp]:
29.166 - assumes X: "cauchy X"
29.167 - shows "cauchy (\<lambda>n. - X n)"
29.168 -using assms unfolding cauchy_def
29.169 -unfolding minus_diff_minus abs_minus_cancel .
29.170 -
29.171 -lemma cauchy_diff [simp]:
29.172 - assumes X: "cauchy X" and Y: "cauchy Y"
29.173 - shows "cauchy (\<lambda>n. X n - Y n)"
29.174 -using assms unfolding diff_minus by simp
29.175 -
29.176 -lemma cauchy_imp_bounded:
29.177 - assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
29.178 -proof -
29.179 - obtain k where k: "\<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < 1"
29.180 - using cauchyD [OF assms zero_less_one] ..
29.181 - show "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
29.182 - proof (intro exI conjI allI)
29.183 - have "0 \<le> \<bar>X 0\<bar>" by simp
29.184 - also have "\<bar>X 0\<bar> \<le> Max (abs ` X ` {..k})" by simp
29.185 - finally have "0 \<le> Max (abs ` X ` {..k})" .
29.186 - thus "0 < Max (abs ` X ` {..k}) + 1" by simp
29.187 - next
29.188 - fix n :: nat
29.189 - show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1"
29.190 - proof (rule linorder_le_cases)
29.191 - assume "n \<le> k"
29.192 - hence "\<bar>X n\<bar> \<le> Max (abs ` X ` {..k})" by simp
29.193 - thus "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" by simp
29.194 - next
29.195 - assume "k \<le> n"
29.196 - have "\<bar>X n\<bar> = \<bar>X k + (X n - X k)\<bar>" by simp
29.197 - also have "\<bar>X k + (X n - X k)\<bar> \<le> \<bar>X k\<bar> + \<bar>X n - X k\<bar>"
29.198 - by (rule abs_triangle_ineq)
29.199 - also have "\<dots> < Max (abs ` X ` {..k}) + 1"
29.200 - by (rule add_le_less_mono, simp, simp add: k `k \<le> n`)
29.201 - finally show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" .
29.202 - qed
29.203 - qed
29.204 -qed
29.205 -
29.206 -lemma cauchy_mult [simp]:
29.207 - assumes X: "cauchy X" and Y: "cauchy Y"
29.208 - shows "cauchy (\<lambda>n. X n * Y n)"
29.209 -proof (rule cauchyI)
29.210 - fix r :: rat assume "0 < r"
29.211 - then obtain u v where u: "0 < u" and v: "0 < v" and "r = u + v"
29.212 - by (rule obtain_pos_sum)
29.213 - obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
29.214 - using cauchy_imp_bounded [OF X] by fast
29.215 - obtain b where b: "0 < b" "\<forall>n. \<bar>Y n\<bar> < b"
29.216 - using cauchy_imp_bounded [OF Y] by fast
29.217 - obtain s t where s: "0 < s" and t: "0 < t" and r: "r = a * t + s * b"
29.218 - proof
29.219 - show "0 < v/b" using v b(1) by (rule divide_pos_pos)
29.220 - show "0 < u/a" using u a(1) by (rule divide_pos_pos)
29.221 - show "r = a * (u/a) + (v/b) * b"
29.222 - using a(1) b(1) `r = u + v` by simp
29.223 - qed
29.224 - obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
29.225 - using cauchyD [OF X s] ..
29.226 - obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
29.227 - using cauchyD [OF Y t] ..
29.228 - have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>X m * Y m - X n * Y n\<bar> < r"
29.229 - proof (clarsimp)
29.230 - fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
29.231 - have "\<bar>X m * Y m - X n * Y n\<bar> = \<bar>X m * (Y m - Y n) + (X m - X n) * Y n\<bar>"
29.232 - unfolding mult_diff_mult ..
29.233 - also have "\<dots> \<le> \<bar>X m * (Y m - Y n)\<bar> + \<bar>(X m - X n) * Y n\<bar>"
29.234 - by (rule abs_triangle_ineq)
29.235 - also have "\<dots> = \<bar>X m\<bar> * \<bar>Y m - Y n\<bar> + \<bar>X m - X n\<bar> * \<bar>Y n\<bar>"
29.236 - unfolding abs_mult ..
29.237 - also have "\<dots> < a * t + s * b"
29.238 - by (simp_all add: add_strict_mono mult_strict_mono' a b i j *)
29.239 - finally show "\<bar>X m * Y m - X n * Y n\<bar> < r" unfolding r .
29.240 - qed
29.241 - thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m * Y m - X n * Y n\<bar> < r" ..
29.242 -qed
29.243 -
29.244 -lemma cauchy_not_vanishes_cases:
29.245 - assumes X: "cauchy X"
29.246 - assumes nz: "\<not> vanishes X"
29.247 - shows "\<exists>b>0. \<exists>k. (\<forall>n\<ge>k. b < - X n) \<or> (\<forall>n\<ge>k. b < X n)"
29.248 -proof -
29.249 - obtain r where "0 < r" and r: "\<forall>k. \<exists>n\<ge>k. r \<le> \<bar>X n\<bar>"
29.250 - using nz unfolding vanishes_def by (auto simp add: not_less)
29.251 - obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t"
29.252 - using `0 < r` by (rule obtain_pos_sum)
29.253 - obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
29.254 - using cauchyD [OF X s] ..
29.255 - obtain k where "i \<le> k" and "r \<le> \<bar>X k\<bar>"
29.256 - using r by fast
29.257 - have k: "\<forall>n\<ge>k. \<bar>X n - X k\<bar> < s"
29.258 - using i `i \<le> k` by auto
29.259 - have "X k \<le> - r \<or> r \<le> X k"
29.260 - using `r \<le> \<bar>X k\<bar>` by auto
29.261 - hence "(\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
29.262 - unfolding `r = s + t` using k by auto
29.263 - hence "\<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)" ..
29.264 - thus "\<exists>t>0. \<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
29.265 - using t by auto
29.266 -qed
29.267 -
29.268 -lemma cauchy_not_vanishes:
29.269 - assumes X: "cauchy X"
29.270 - assumes nz: "\<not> vanishes X"
29.271 - shows "\<exists>b>0. \<exists>k. \<forall>n\<ge>k. b < \<bar>X n\<bar>"
29.272 -using cauchy_not_vanishes_cases [OF assms]
29.273 -by clarify (rule exI, erule conjI, rule_tac x=k in exI, auto)
29.274 -
29.275 -lemma cauchy_inverse [simp]:
29.276 - assumes X: "cauchy X"
29.277 - assumes nz: "\<not> vanishes X"
29.278 - shows "cauchy (\<lambda>n. inverse (X n))"
29.279 -proof (rule cauchyI)
29.280 - fix r :: rat assume "0 < r"
29.281 - obtain b i where b: "0 < b" and i: "\<forall>n\<ge>i. b < \<bar>X n\<bar>"
29.282 - using cauchy_not_vanishes [OF X nz] by fast
29.283 - from b i have nz: "\<forall>n\<ge>i. X n \<noteq> 0" by auto
29.284 - obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b"
29.285 - proof
29.286 - show "0 < b * r * b"
29.287 - by (simp add: `0 < r` b mult_pos_pos)
29.288 - show "r = inverse b * (b * r * b) * inverse b"
29.289 - using b by simp
29.290 - qed
29.291 - obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>X m - X n\<bar> < s"
29.292 - using cauchyD [OF X s] ..
29.293 - have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>inverse (X m) - inverse (X n)\<bar> < r"
29.294 - proof (clarsimp)
29.295 - fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
29.296 - have "\<bar>inverse (X m) - inverse (X n)\<bar> =
29.297 - inverse \<bar>X m\<bar> * \<bar>X m - X n\<bar> * inverse \<bar>X n\<bar>"
29.298 - by (simp add: inverse_diff_inverse nz * abs_mult)
29.299 - also have "\<dots> < inverse b * s * inverse b"
29.300 - by (simp add: mult_strict_mono less_imp_inverse_less
29.301 - mult_pos_pos i j b * s)
29.302 - finally show "\<bar>inverse (X m) - inverse (X n)\<bar> < r" unfolding r .
29.303 - qed
29.304 - thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>inverse (X m) - inverse (X n)\<bar> < r" ..
29.305 -qed
29.306 -
29.307 -lemma vanishes_diff_inverse:
29.308 - assumes X: "cauchy X" "\<not> vanishes X"
29.309 - assumes Y: "cauchy Y" "\<not> vanishes Y"
29.310 - assumes XY: "vanishes (\<lambda>n. X n - Y n)"
29.311 - shows "vanishes (\<lambda>n. inverse (X n) - inverse (Y n))"
29.312 -proof (rule vanishesI)
29.313 - fix r :: rat assume r: "0 < r"
29.314 - obtain a i where a: "0 < a" and i: "\<forall>n\<ge>i. a < \<bar>X n\<bar>"
29.315 - using cauchy_not_vanishes [OF X] by fast
29.316 - obtain b j where b: "0 < b" and j: "\<forall>n\<ge>j. b < \<bar>Y n\<bar>"
29.317 - using cauchy_not_vanishes [OF Y] by fast
29.318 - obtain s where s: "0 < s" and "inverse a * s * inverse b = r"
29.319 - proof
29.320 - show "0 < a * r * b"
29.321 - using a r b by (simp add: mult_pos_pos)
29.322 - show "inverse a * (a * r * b) * inverse b = r"
29.323 - using a r b by simp
29.324 - qed
29.325 - obtain k where k: "\<forall>n\<ge>k. \<bar>X n - Y n\<bar> < s"
29.326 - using vanishesD [OF XY s] ..
29.327 - have "\<forall>n\<ge>max (max i j) k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r"
29.328 - proof (clarsimp)
29.329 - fix n assume n: "i \<le> n" "j \<le> n" "k \<le> n"
29.330 - have "X n \<noteq> 0" and "Y n \<noteq> 0"
29.331 - using i j a b n by auto
29.332 - hence "\<bar>inverse (X n) - inverse (Y n)\<bar> =
29.333 - inverse \<bar>X n\<bar> * \<bar>X n - Y n\<bar> * inverse \<bar>Y n\<bar>"
29.334 - by (simp add: inverse_diff_inverse abs_mult)
29.335 - also have "\<dots> < inverse a * s * inverse b"
29.336 - apply (intro mult_strict_mono' less_imp_inverse_less)
29.337 - apply (simp_all add: a b i j k n mult_nonneg_nonneg)
29.338 - done
29.339 - also note `inverse a * s * inverse b = r`
29.340 - finally show "\<bar>inverse (X n) - inverse (Y n)\<bar> < r" .
29.341 - qed
29.342 - thus "\<exists>k. \<forall>n\<ge>k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r" ..
29.343 -qed
29.344 -
29.345 -subsection {* Equivalence relation on Cauchy sequences *}
29.346 -
29.347 -definition realrel :: "(nat \<Rightarrow> rat) \<Rightarrow> (nat \<Rightarrow> rat) \<Rightarrow> bool"
29.348 - where "realrel = (\<lambda>X Y. cauchy X \<and> cauchy Y \<and> vanishes (\<lambda>n. X n - Y n))"
29.349 -
29.350 -lemma realrelI [intro?]:
29.351 - assumes "cauchy X" and "cauchy Y" and "vanishes (\<lambda>n. X n - Y n)"
29.352 - shows "realrel X Y"
29.353 - using assms unfolding realrel_def by simp
29.354 -
29.355 -lemma realrel_refl: "cauchy X \<Longrightarrow> realrel X X"
29.356 - unfolding realrel_def by simp
29.357 -
29.358 -lemma symp_realrel: "symp realrel"
29.359 - unfolding realrel_def
29.360 - by (rule sympI, clarify, drule vanishes_minus, simp)
29.361 -
29.362 -lemma transp_realrel: "transp realrel"
29.363 - unfolding realrel_def
29.364 - apply (rule transpI, clarify)
29.365 - apply (drule (1) vanishes_add)
29.366 - apply (simp add: algebra_simps)
29.367 - done
29.368 -
29.369 -lemma part_equivp_realrel: "part_equivp realrel"
29.370 - by (fast intro: part_equivpI symp_realrel transp_realrel
29.371 - realrel_refl cauchy_const)
29.372 -
29.373 -subsection {* The field of real numbers *}
29.374 -
29.375 -quotient_type real = "nat \<Rightarrow> rat" / partial: realrel
29.376 - morphisms rep_real Real
29.377 - by (rule part_equivp_realrel)
29.378 -
29.379 -lemma cr_real_eq: "pcr_real = (\<lambda>x y. cauchy x \<and> Real x = y)"
29.380 - unfolding real.pcr_cr_eq cr_real_def realrel_def by auto
29.381 -
29.382 -lemma Real_induct [induct type: real]: (* TODO: generate automatically *)
29.383 - assumes "\<And>X. cauchy X \<Longrightarrow> P (Real X)" shows "P x"
29.384 -proof (induct x)
29.385 - case (1 X)
29.386 - hence "cauchy X" by (simp add: realrel_def)
29.387 - thus "P (Real X)" by (rule assms)
29.388 -qed
29.389 -
29.390 -lemma eq_Real:
29.391 - "cauchy X \<Longrightarrow> cauchy Y \<Longrightarrow> Real X = Real Y \<longleftrightarrow> vanishes (\<lambda>n. X n - Y n)"
29.392 - using real.rel_eq_transfer
29.393 - unfolding real.pcr_cr_eq cr_real_def fun_rel_def realrel_def by simp
29.394 -
29.395 -declare real.forall_transfer [transfer_rule del]
29.396 -
29.397 -lemma forall_real_transfer [transfer_rule]: (* TODO: generate automatically *)
29.398 - "(fun_rel (fun_rel pcr_real op =) op =)
29.399 - (transfer_bforall cauchy) transfer_forall"
29.400 - using real.forall_transfer
29.401 - by (simp add: realrel_def)
29.402 -
29.403 -instantiation real :: field_inverse_zero
29.404 -begin
29.405 -
29.406 -lift_definition zero_real :: "real" is "\<lambda>n. 0"
29.407 - by (simp add: realrel_refl)
29.408 -
29.409 -lift_definition one_real :: "real" is "\<lambda>n. 1"
29.410 - by (simp add: realrel_refl)
29.411 -
29.412 -lift_definition plus_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n + Y n"
29.413 - unfolding realrel_def add_diff_add
29.414 - by (simp only: cauchy_add vanishes_add simp_thms)
29.415 -
29.416 -lift_definition uminus_real :: "real \<Rightarrow> real" is "\<lambda>X n. - X n"
29.417 - unfolding realrel_def minus_diff_minus
29.418 - by (simp only: cauchy_minus vanishes_minus simp_thms)
29.419 -
29.420 -lift_definition times_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n * Y n"
29.421 - unfolding realrel_def mult_diff_mult
29.422 - by (subst (4) mult_commute, simp only: cauchy_mult vanishes_add
29.423 - vanishes_mult_bounded cauchy_imp_bounded simp_thms)
29.424 -
29.425 -lift_definition inverse_real :: "real \<Rightarrow> real"
29.426 - is "\<lambda>X. if vanishes X then (\<lambda>n. 0) else (\<lambda>n. inverse (X n))"
29.427 -proof -
29.428 - fix X Y assume "realrel X Y"
29.429 - hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\<lambda>n. X n - Y n)"
29.430 - unfolding realrel_def by simp_all
29.431 - have "vanishes X \<longleftrightarrow> vanishes Y"
29.432 - proof
29.433 - assume "vanishes X"
29.434 - from vanishes_diff [OF this XY] show "vanishes Y" by simp
29.435 - next
29.436 - assume "vanishes Y"
29.437 - from vanishes_add [OF this XY] show "vanishes X" by simp
29.438 - qed
29.439 - thus "?thesis X Y"
29.440 - unfolding realrel_def
29.441 - by (simp add: vanishes_diff_inverse X Y XY)
29.442 -qed
29.443 -
29.444 -definition
29.445 - "x - y = (x::real) + - y"
29.446 -
29.447 -definition
29.448 - "x / y = (x::real) * inverse y"
29.449 -
29.450 -lemma add_Real:
29.451 - assumes X: "cauchy X" and Y: "cauchy Y"
29.452 - shows "Real X + Real Y = Real (\<lambda>n. X n + Y n)"
29.453 - using assms plus_real.transfer
29.454 - unfolding cr_real_eq fun_rel_def by simp
29.455 -
29.456 -lemma minus_Real:
29.457 - assumes X: "cauchy X"
29.458 - shows "- Real X = Real (\<lambda>n. - X n)"
29.459 - using assms uminus_real.transfer
29.460 - unfolding cr_real_eq fun_rel_def by simp
29.461 -
29.462 -lemma diff_Real:
29.463 - assumes X: "cauchy X" and Y: "cauchy Y"
29.464 - shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
29.465 - unfolding minus_real_def diff_minus
29.466 - by (simp add: minus_Real add_Real X Y)
29.467 -
29.468 -lemma mult_Real:
29.469 - assumes X: "cauchy X" and Y: "cauchy Y"
29.470 - shows "Real X * Real Y = Real (\<lambda>n. X n * Y n)"
29.471 - using assms times_real.transfer
29.472 - unfolding cr_real_eq fun_rel_def by simp
29.473 -
29.474 -lemma inverse_Real:
29.475 - assumes X: "cauchy X"
29.476 - shows "inverse (Real X) =
29.477 - (if vanishes X then 0 else Real (\<lambda>n. inverse (X n)))"
29.478 - using assms inverse_real.transfer zero_real.transfer
29.479 - unfolding cr_real_eq fun_rel_def by (simp split: split_if_asm, metis)
29.480 -
29.481 -instance proof
29.482 - fix a b c :: real
29.483 - show "a + b = b + a"
29.484 - by transfer (simp add: add_ac realrel_def)
29.485 - show "(a + b) + c = a + (b + c)"
29.486 - by transfer (simp add: add_ac realrel_def)
29.487 - show "0 + a = a"
29.488 - by transfer (simp add: realrel_def)
29.489 - show "- a + a = 0"
29.490 - by transfer (simp add: realrel_def)
29.491 - show "a - b = a + - b"
29.492 - by (rule minus_real_def)
29.493 - show "(a * b) * c = a * (b * c)"
29.494 - by transfer (simp add: mult_ac realrel_def)
29.495 - show "a * b = b * a"
29.496 - by transfer (simp add: mult_ac realrel_def)
29.497 - show "1 * a = a"
29.498 - by transfer (simp add: mult_ac realrel_def)
29.499 - show "(a + b) * c = a * c + b * c"
29.500 - by transfer (simp add: distrib_right realrel_def)
29.501 - show "(0\<Colon>real) \<noteq> (1\<Colon>real)"
29.502 - by transfer (simp add: realrel_def)
29.503 - show "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
29.504 - apply transfer
29.505 - apply (simp add: realrel_def)
29.506 - apply (rule vanishesI)
29.507 - apply (frule (1) cauchy_not_vanishes, clarify)
29.508 - apply (rule_tac x=k in exI, clarify)
29.509 - apply (drule_tac x=n in spec, simp)
29.510 - done
29.511 - show "a / b = a * inverse b"
29.512 - by (rule divide_real_def)
29.513 - show "inverse (0::real) = 0"
29.514 - by transfer (simp add: realrel_def)
29.515 -qed
29.516 -
29.517 -end
29.518 -
29.519 -subsection {* Positive reals *}
29.520 -
29.521 -lift_definition positive :: "real \<Rightarrow> bool"
29.522 - is "\<lambda>X. \<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
29.523 -proof -
29.524 - { fix X Y
29.525 - assume "realrel X Y"
29.526 - hence XY: "vanishes (\<lambda>n. X n - Y n)"
29.527 - unfolding realrel_def by simp_all
29.528 - assume "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
29.529 - then obtain r i where "0 < r" and i: "\<forall>n\<ge>i. r < X n"
29.530 - by fast
29.531 - obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
29.532 - using `0 < r` by (rule obtain_pos_sum)
29.533 - obtain j where j: "\<forall>n\<ge>j. \<bar>X n - Y n\<bar> < s"
29.534 - using vanishesD [OF XY s] ..
29.535 - have "\<forall>n\<ge>max i j. t < Y n"
29.536 - proof (clarsimp)
29.537 - fix n assume n: "i \<le> n" "j \<le> n"
29.538 - have "\<bar>X n - Y n\<bar> < s" and "r < X n"
29.539 - using i j n by simp_all
29.540 - thus "t < Y n" unfolding r by simp
29.541 - qed
29.542 - hence "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < Y n" using t by fast
29.543 - } note 1 = this
29.544 - fix X Y assume "realrel X Y"
29.545 - hence "realrel X Y" and "realrel Y X"
29.546 - using symp_realrel unfolding symp_def by auto
29.547 - thus "?thesis X Y"
29.548 - by (safe elim!: 1)
29.549 -qed
29.550 -
29.551 -lemma positive_Real:
29.552 - assumes X: "cauchy X"
29.553 - shows "positive (Real X) \<longleftrightarrow> (\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n)"
29.554 - using assms positive.transfer
29.555 - unfolding cr_real_eq fun_rel_def by simp
29.556 -
29.557 -lemma positive_zero: "\<not> positive 0"
29.558 - by transfer auto
29.559 -
29.560 -lemma positive_add:
29.561 - "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x + y)"
29.562 -apply transfer
29.563 -apply (clarify, rename_tac a b i j)
29.564 -apply (rule_tac x="a + b" in exI, simp)
29.565 -apply (rule_tac x="max i j" in exI, clarsimp)
29.566 -apply (simp add: add_strict_mono)
29.567 -done
29.568 -
29.569 -lemma positive_mult:
29.570 - "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x * y)"
29.571 -apply transfer
29.572 -apply (clarify, rename_tac a b i j)
29.573 -apply (rule_tac x="a * b" in exI, simp add: mult_pos_pos)
29.574 -apply (rule_tac x="max i j" in exI, clarsimp)
29.575 -apply (rule mult_strict_mono, auto)
29.576 -done
29.577 -
29.578 -lemma positive_minus:
29.579 - "\<not> positive x \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> positive (- x)"
29.580 -apply transfer
29.581 -apply (simp add: realrel_def)
29.582 -apply (drule (1) cauchy_not_vanishes_cases, safe, fast, fast)
29.583 -done
29.584 -
29.585 -instantiation real :: linordered_field_inverse_zero
29.586 -begin
29.587 -
29.588 -definition
29.589 - "x < y \<longleftrightarrow> positive (y - x)"
29.590 -
29.591 -definition
29.592 - "x \<le> (y::real) \<longleftrightarrow> x < y \<or> x = y"
29.593 -
29.594 -definition
29.595 - "abs (a::real) = (if a < 0 then - a else a)"
29.596 -
29.597 -definition
29.598 - "sgn (a::real) = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
29.599 -
29.600 -instance proof
29.601 - fix a b c :: real
29.602 - show "\<bar>a\<bar> = (if a < 0 then - a else a)"
29.603 - by (rule abs_real_def)
29.604 - show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a"
29.605 - unfolding less_eq_real_def less_real_def
29.606 - by (auto, drule (1) positive_add, simp_all add: positive_zero)
29.607 - show "a \<le> a"
29.608 - unfolding less_eq_real_def by simp
29.609 - show "a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> a \<le> c"
29.610 - unfolding less_eq_real_def less_real_def
29.611 - by (auto, drule (1) positive_add, simp add: algebra_simps)
29.612 - show "a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> a = b"
29.613 - unfolding less_eq_real_def less_real_def
29.614 - by (auto, drule (1) positive_add, simp add: positive_zero)
29.615 - show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
29.616 - unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *)
29.617 - (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \<equiv> b + - a *)
29.618 - (* Should produce c + b - (c + a) \<equiv> b - a *)
29.619 - show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
29.620 - by (rule sgn_real_def)
29.621 - show "a \<le> b \<or> b \<le> a"
29.622 - unfolding less_eq_real_def less_real_def
29.623 - by (auto dest!: positive_minus)
29.624 - show "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
29.625 - unfolding less_real_def
29.626 - by (drule (1) positive_mult, simp add: algebra_simps)
29.627 -qed
29.628 -
29.629 -end
29.630 -
29.631 -instantiation real :: distrib_lattice
29.632 -begin
29.633 -
29.634 -definition
29.635 - "(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"
29.636 -
29.637 -definition
29.638 - "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
29.639 -
29.640 -instance proof
29.641 -qed (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
29.642 -
29.643 -end
29.644 -
29.645 -lemma of_nat_Real: "of_nat x = Real (\<lambda>n. of_nat x)"
29.646 -apply (induct x)
29.647 -apply (simp add: zero_real_def)
29.648 -apply (simp add: one_real_def add_Real)
29.649 -done
29.650 -
29.651 -lemma of_int_Real: "of_int x = Real (\<lambda>n. of_int x)"
29.652 -apply (cases x rule: int_diff_cases)
29.653 -apply (simp add: of_nat_Real diff_Real)
29.654 -done
29.655 -
29.656 -lemma of_rat_Real: "of_rat x = Real (\<lambda>n. x)"
29.657 -apply (induct x)
29.658 -apply (simp add: Fract_of_int_quotient of_rat_divide)
29.659 -apply (simp add: of_int_Real divide_inverse)
29.660 -apply (simp add: inverse_Real mult_Real)
29.661 -done
29.662 -
29.663 -instance real :: archimedean_field
29.664 -proof
29.665 - fix x :: real
29.666 - show "\<exists>z. x \<le> of_int z"
29.667 - apply (induct x)
29.668 - apply (frule cauchy_imp_bounded, clarify)
29.669 - apply (rule_tac x="ceiling b + 1" in exI)
29.670 - apply (rule less_imp_le)
29.671 - apply (simp add: of_int_Real less_real_def diff_Real positive_Real)
29.672 - apply (rule_tac x=1 in exI, simp add: algebra_simps)
29.673 - apply (rule_tac x=0 in exI, clarsimp)
29.674 - apply (rule le_less_trans [OF abs_ge_self])
29.675 - apply (rule less_le_trans [OF _ le_of_int_ceiling])
29.676 - apply simp
29.677 - done
29.678 -qed
29.679 -
29.680 -instantiation real :: floor_ceiling
29.681 -begin
29.682 -
29.683 -definition [code del]:
29.684 - "floor (x::real) = (THE z. of_int z \<le> x \<and> x < of_int (z + 1))"
29.685 -
29.686 -instance proof
29.687 - fix x :: real
29.688 - show "of_int (floor x) \<le> x \<and> x < of_int (floor x + 1)"
29.689 - unfolding floor_real_def using floor_exists1 by (rule theI')
29.690 -qed
29.691 -
29.692 -end
29.693 -
29.694 -subsection {* Completeness *}
29.695 -
29.696 -lemma not_positive_Real:
29.697 - assumes X: "cauchy X"
29.698 - shows "\<not> positive (Real X) \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> r)"
29.699 -unfolding positive_Real [OF X]
29.700 -apply (auto, unfold not_less)
29.701 -apply (erule obtain_pos_sum)
29.702 -apply (drule_tac x=s in spec, simp)
29.703 -apply (drule_tac r=t in cauchyD [OF X], clarify)
29.704 -apply (drule_tac x=k in spec, clarsimp)
29.705 -apply (rule_tac x=n in exI, clarify, rename_tac m)
29.706 -apply (drule_tac x=m in spec, simp)
29.707 -apply (drule_tac x=n in spec, simp)
29.708 -apply (drule spec, drule (1) mp, clarify, rename_tac i)
29.709 -apply (rule_tac x="max i k" in exI, simp)
29.710 -done
29.711 -
29.712 -lemma le_Real:
29.713 - assumes X: "cauchy X" and Y: "cauchy Y"
29.714 - shows "Real X \<le> Real Y = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r)"
29.715 -unfolding not_less [symmetric, where 'a=real] less_real_def
29.716 -apply (simp add: diff_Real not_positive_Real X Y)
29.717 -apply (simp add: diff_le_eq add_ac)
29.718 -done
29.719 -
29.720 -lemma le_RealI:
29.721 - assumes Y: "cauchy Y"
29.722 - shows "\<forall>n. x \<le> of_rat (Y n) \<Longrightarrow> x \<le> Real Y"
29.723 -proof (induct x)
29.724 - fix X assume X: "cauchy X" and "\<forall>n. Real X \<le> of_rat (Y n)"
29.725 - hence le: "\<And>m r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. X n \<le> Y m + r"
29.726 - by (simp add: of_rat_Real le_Real)
29.727 - {
29.728 - fix r :: rat assume "0 < r"
29.729 - then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
29.730 - by (rule obtain_pos_sum)
29.731 - obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>Y m - Y n\<bar> < s"
29.732 - using cauchyD [OF Y s] ..
29.733 - obtain j where j: "\<forall>n\<ge>j. X n \<le> Y i + t"
29.734 - using le [OF t] ..
29.735 - have "\<forall>n\<ge>max i j. X n \<le> Y n + r"
29.736 - proof (clarsimp)
29.737 - fix n assume n: "i \<le> n" "j \<le> n"
29.738 - have "X n \<le> Y i + t" using n j by simp
29.739 - moreover have "\<bar>Y i - Y n\<bar> < s" using n i by simp
29.740 - ultimately show "X n \<le> Y n + r" unfolding r by simp
29.741 - qed
29.742 - hence "\<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r" ..
29.743 - }
29.744 - thus "Real X \<le> Real Y"
29.745 - by (simp add: of_rat_Real le_Real X Y)
29.746 -qed
29.747 -
29.748 -lemma Real_leI:
29.749 - assumes X: "cauchy X"
29.750 - assumes le: "\<forall>n. of_rat (X n) \<le> y"
29.751 - shows "Real X \<le> y"
29.752 -proof -
29.753 - have "- y \<le> - Real X"
29.754 - by (simp add: minus_Real X le_RealI of_rat_minus le)
29.755 - thus ?thesis by simp
29.756 -qed
29.757 -
29.758 -lemma less_RealD:
29.759 - assumes Y: "cauchy Y"
29.760 - shows "x < Real Y \<Longrightarrow> \<exists>n. x < of_rat (Y n)"
29.761 -by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y])
29.762 -
29.763 -lemma of_nat_less_two_power:
29.764 - "of_nat n < (2::'a::linordered_idom) ^ n"
29.765 -apply (induct n)
29.766 -apply simp
29.767 -apply (subgoal_tac "(1::'a) \<le> 2 ^ n")
29.768 -apply (drule (1) add_le_less_mono, simp)
29.769 -apply simp
29.770 -done
29.771 -
29.772 -lemma complete_real:
29.773 - fixes S :: "real set"
29.774 - assumes "\<exists>x. x \<in> S" and "\<exists>z. \<forall>x\<in>S. x \<le> z"
29.775 - shows "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
29.776 -proof -
29.777 - obtain x where x: "x \<in> S" using assms(1) ..
29.778 - obtain z where z: "\<forall>x\<in>S. x \<le> z" using assms(2) ..
29.779 -
29.780 - def P \<equiv> "\<lambda>x. \<forall>y\<in>S. y \<le> of_rat x"
29.781 - obtain a where a: "\<not> P a"
29.782 - proof
29.783 - have "of_int (floor (x - 1)) \<le> x - 1" by (rule of_int_floor_le)
29.784 - also have "x - 1 < x" by simp
29.785 - finally have "of_int (floor (x - 1)) < x" .
29.786 - hence "\<not> x \<le> of_int (floor (x - 1))" by (simp only: not_le)
29.787 - then show "\<not> P (of_int (floor (x - 1)))"
29.788 - unfolding P_def of_rat_of_int_eq using x by fast
29.789 - qed
29.790 - obtain b where b: "P b"
29.791 - proof
29.792 - show "P (of_int (ceiling z))"
29.793 - unfolding P_def of_rat_of_int_eq
29.794 - proof
29.795 - fix y assume "y \<in> S"
29.796 - hence "y \<le> z" using z by simp
29.797 - also have "z \<le> of_int (ceiling z)" by (rule le_of_int_ceiling)
29.798 - finally show "y \<le> of_int (ceiling z)" .
29.799 - qed
29.800 - qed
29.801 -
29.802 - def avg \<equiv> "\<lambda>x y :: rat. x/2 + y/2"
29.803 - def bisect \<equiv> "\<lambda>(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)"
29.804 - def A \<equiv> "\<lambda>n. fst ((bisect ^^ n) (a, b))"
29.805 - def B \<equiv> "\<lambda>n. snd ((bisect ^^ n) (a, b))"
29.806 - def C \<equiv> "\<lambda>n. avg (A n) (B n)"
29.807 - have A_0 [simp]: "A 0 = a" unfolding A_def by simp
29.808 - have B_0 [simp]: "B 0 = b" unfolding B_def by simp
29.809 - have A_Suc [simp]: "\<And>n. A (Suc n) = (if P (C n) then A n else C n)"
29.810 - unfolding A_def B_def C_def bisect_def split_def by simp
29.811 - have B_Suc [simp]: "\<And>n. B (Suc n) = (if P (C n) then C n else B n)"
29.812 - unfolding A_def B_def C_def bisect_def split_def by simp
29.813 -
29.814 - have width: "\<And>n. B n - A n = (b - a) / 2^n"
29.815 - apply (simp add: eq_divide_eq)
29.816 - apply (induct_tac n, simp)
29.817 - apply (simp add: C_def avg_def algebra_simps)
29.818 - done
29.819 -
29.820 - have twos: "\<And>y r :: rat. 0 < r \<Longrightarrow> \<exists>n. y / 2 ^ n < r"
29.821 - apply (simp add: divide_less_eq)
29.822 - apply (subst mult_commute)
29.823 - apply (frule_tac y=y in ex_less_of_nat_mult)
29.824 - apply clarify
29.825 - apply (rule_tac x=n in exI)
29.826 - apply (erule less_trans)
29.827 - apply (rule mult_strict_right_mono)
29.828 - apply (rule le_less_trans [OF _ of_nat_less_two_power])
29.829 - apply simp
29.830 - apply assumption
29.831 - done
29.832 -
29.833 - have PA: "\<And>n. \<not> P (A n)"
29.834 - by (induct_tac n, simp_all add: a)
29.835 - have PB: "\<And>n. P (B n)"
29.836 - by (induct_tac n, simp_all add: b)
29.837 - have ab: "a < b"
29.838 - using a b unfolding P_def
29.839 - apply (clarsimp simp add: not_le)
29.840 - apply (drule (1) bspec)
29.841 - apply (drule (1) less_le_trans)
29.842 - apply (simp add: of_rat_less)
29.843 - done
29.844 - have AB: "\<And>n. A n < B n"
29.845 - by (induct_tac n, simp add: ab, simp add: C_def avg_def)
29.846 - have A_mono: "\<And>i j. i \<le> j \<Longrightarrow> A i \<le> A j"
29.847 - apply (auto simp add: le_less [where 'a=nat])
29.848 - apply (erule less_Suc_induct)
29.849 - apply (clarsimp simp add: C_def avg_def)
29.850 - apply (simp add: add_divide_distrib [symmetric])
29.851 - apply (rule AB [THEN less_imp_le])
29.852 - apply simp
29.853 - done
29.854 - have B_mono: "\<And>i j. i \<le> j \<Longrightarrow> B j \<le> B i"
29.855 - apply (auto simp add: le_less [where 'a=nat])
29.856 - apply (erule less_Suc_induct)
29.857 - apply (clarsimp simp add: C_def avg_def)
29.858 - apply (simp add: add_divide_distrib [symmetric])
29.859 - apply (rule AB [THEN less_imp_le])
29.860 - apply simp
29.861 - done
29.862 - have cauchy_lemma:
29.863 - "\<And>X. \<forall>n. \<forall>i\<ge>n. A n \<le> X i \<and> X i \<le> B n \<Longrightarrow> cauchy X"
29.864 - apply (rule cauchyI)
29.865 - apply (drule twos [where y="b - a"])
29.866 - apply (erule exE)
29.867 - apply (rule_tac x=n in exI, clarify, rename_tac i j)
29.868 - apply (rule_tac y="B n - A n" in le_less_trans) defer
29.869 - apply (simp add: width)
29.870 - apply (drule_tac x=n in spec)
29.871 - apply (frule_tac x=i in spec, drule (1) mp)
29.872 - apply (frule_tac x=j in spec, drule (1) mp)
29.873 - apply (frule A_mono, drule B_mono)
29.874 - apply (frule A_mono, drule B_mono)
29.875 - apply arith
29.876 - done
29.877 - have "cauchy A"
29.878 - apply (rule cauchy_lemma [rule_format])
29.879 - apply (simp add: A_mono)
29.880 - apply (erule order_trans [OF less_imp_le [OF AB] B_mono])
29.881 - done
29.882 - have "cauchy B"
29.883 - apply (rule cauchy_lemma [rule_format])
29.884 - apply (simp add: B_mono)
29.885 - apply (erule order_trans [OF A_mono less_imp_le [OF AB]])
29.886 - done
29.887 - have 1: "\<forall>x\<in>S. x \<le> Real B"
29.888 - proof
29.889 - fix x assume "x \<in> S"
29.890 - then show "x \<le> Real B"
29.891 - using PB [unfolded P_def] `cauchy B`
29.892 - by (simp add: le_RealI)
29.893 - qed
29.894 - have 2: "\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> Real A \<le> z"
29.895 - apply clarify
29.896 - apply (erule contrapos_pp)
29.897 - apply (simp add: not_le)
29.898 - apply (drule less_RealD [OF `cauchy A`], clarify)
29.899 - apply (subgoal_tac "\<not> P (A n)")
29.900 - apply (simp add: P_def not_le, clarify)
29.901 - apply (erule rev_bexI)
29.902 - apply (erule (1) less_trans)
29.903 - apply (simp add: PA)
29.904 - done
29.905 - have "vanishes (\<lambda>n. (b - a) / 2 ^ n)"
29.906 - proof (rule vanishesI)
29.907 - fix r :: rat assume "0 < r"
29.908 - then obtain k where k: "\<bar>b - a\<bar> / 2 ^ k < r"
29.909 - using twos by fast
29.910 - have "\<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r"
29.911 - proof (clarify)
29.912 - fix n assume n: "k \<le> n"
29.913 - have "\<bar>(b - a) / 2 ^ n\<bar> = \<bar>b - a\<bar> / 2 ^ n"
29.914 - by simp
29.915 - also have "\<dots> \<le> \<bar>b - a\<bar> / 2 ^ k"
29.916 - using n by (simp add: divide_left_mono mult_pos_pos)
29.917 - also note k
29.918 - finally show "\<bar>(b - a) / 2 ^ n\<bar> < r" .
29.919 - qed
29.920 - thus "\<exists>k. \<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r" ..
29.921 - qed
29.922 - hence 3: "Real B = Real A"
29.923 - by (simp add: eq_Real `cauchy A` `cauchy B` width)
29.924 - show "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
29.925 - using 1 2 3 by (rule_tac x="Real B" in exI, simp)
29.926 -qed
29.927 -
29.928 -subsection {* Hiding implementation details *}
29.929 -
29.930 -hide_const (open) vanishes cauchy positive Real
29.931 -
29.932 -declare Real_induct [induct del]
29.933 -declare Abs_real_induct [induct del]
29.934 -declare Abs_real_cases [cases del]
29.935 -
29.936 -lemmas [transfer_rule del] =
29.937 - real.All_transfer real.Ex_transfer real.rel_eq_transfer forall_real_transfer
29.938 - zero_real.transfer one_real.transfer plus_real.transfer uminus_real.transfer
29.939 - times_real.transfer inverse_real.transfer positive.transfer real.right_unique
29.940 - real.right_total
29.941 -
29.942 -subsection{*More Lemmas*}
29.943 -
29.944 -text {* BH: These lemmas should not be necessary; they should be
29.945 -covered by existing simp rules and simplification procedures. *}
29.946 -
29.947 -lemma real_mult_left_cancel: "(c::real) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
29.948 -by simp (* redundant with mult_cancel_left *)
29.949 -
29.950 -lemma real_mult_right_cancel: "(c::real) \<noteq> 0 ==> (a*c=b*c) = (a=b)"
29.951 -by simp (* redundant with mult_cancel_right *)
29.952 -
29.953 -lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
29.954 -by simp (* solved by linordered_ring_less_cancel_factor simproc *)
29.955 -
29.956 -lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
29.957 -by simp (* solved by linordered_ring_le_cancel_factor simproc *)
29.958 -
29.959 -lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
29.960 -by simp (* solved by linordered_ring_le_cancel_factor simproc *)
29.961 -
29.962 -
29.963 -subsection {* Embedding numbers into the Reals *}
29.964 -
29.965 -abbreviation
29.966 - real_of_nat :: "nat \<Rightarrow> real"
29.967 -where
29.968 - "real_of_nat \<equiv> of_nat"
29.969 -
29.970 -abbreviation
29.971 - real_of_int :: "int \<Rightarrow> real"
29.972 -where
29.973 - "real_of_int \<equiv> of_int"
29.974 -
29.975 -abbreviation
29.976 - real_of_rat :: "rat \<Rightarrow> real"
29.977 -where
29.978 - "real_of_rat \<equiv> of_rat"
29.979 -
29.980 -consts
29.981 - (*overloaded constant for injecting other types into "real"*)
29.982 - real :: "'a => real"
29.983 -
29.984 -defs (overloaded)
29.985 - real_of_nat_def [code_unfold]: "real == real_of_nat"
29.986 - real_of_int_def [code_unfold]: "real == real_of_int"
29.987 -
29.988 -declare [[coercion_enabled]]
29.989 -declare [[coercion "real::nat\<Rightarrow>real"]]
29.990 -declare [[coercion "real::int\<Rightarrow>real"]]
29.991 -declare [[coercion "int"]]
29.992 -
29.993 -declare [[coercion_map map]]
29.994 -declare [[coercion_map "% f g h x. g (h (f x))"]]
29.995 -declare [[coercion_map "% f g (x,y) . (f x, g y)"]]
29.996 -
29.997 -lemma real_eq_of_nat: "real = of_nat"
29.998 - unfolding real_of_nat_def ..
29.999 -
29.1000 -lemma real_eq_of_int: "real = of_int"
29.1001 - unfolding real_of_int_def ..
29.1002 -
29.1003 -lemma real_of_int_zero [simp]: "real (0::int) = 0"
29.1004 -by (simp add: real_of_int_def)
29.1005 -
29.1006 -lemma real_of_one [simp]: "real (1::int) = (1::real)"
29.1007 -by (simp add: real_of_int_def)
29.1008 -
29.1009 -lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y"
29.1010 -by (simp add: real_of_int_def)
29.1011 -
29.1012 -lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)"
29.1013 -by (simp add: real_of_int_def)
29.1014 -
29.1015 -lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y"
29.1016 -by (simp add: real_of_int_def)
29.1017 -
29.1018 -lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y"
29.1019 -by (simp add: real_of_int_def)
29.1020 -
29.1021 -lemma real_of_int_power [simp]: "real (x ^ n) = real (x::int) ^ n"
29.1022 -by (simp add: real_of_int_def of_int_power)
29.1023 -
29.1024 -lemmas power_real_of_int = real_of_int_power [symmetric]
29.1025 -
29.1026 -lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))"
29.1027 - apply (subst real_eq_of_int)+
29.1028 - apply (rule of_int_setsum)
29.1029 -done
29.1030 -
29.1031 -lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) =
29.1032 - (PROD x:A. real(f x))"
29.1033 - apply (subst real_eq_of_int)+
29.1034 - apply (rule of_int_setprod)
29.1035 -done
29.1036 -
29.1037 -lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))"
29.1038 -by (simp add: real_of_int_def)
29.1039 -
29.1040 -lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)"
29.1041 -by (simp add: real_of_int_def)
29.1042 -
29.1043 -lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)"
29.1044 -by (simp add: real_of_int_def)
29.1045 -
29.1046 -lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \<le> real y) = (x \<le> y)"
29.1047 -by (simp add: real_of_int_def)
29.1048 -
29.1049 -lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)"
29.1050 -by (simp add: real_of_int_def)
29.1051 -
29.1052 -lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)"
29.1053 -by (simp add: real_of_int_def)
29.1054 -
29.1055 -lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)"
29.1056 -by (simp add: real_of_int_def)
29.1057 -
29.1058 -lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)"
29.1059 -by (simp add: real_of_int_def)
29.1060 -
29.1061 -lemma one_less_real_of_int_cancel_iff: "1 < real (i :: int) \<longleftrightarrow> 1 < i"
29.1062 - unfolding real_of_one[symmetric] real_of_int_less_iff ..
29.1063 -
29.1064 -lemma one_le_real_of_int_cancel_iff: "1 \<le> real (i :: int) \<longleftrightarrow> 1 \<le> i"
29.1065 - unfolding real_of_one[symmetric] real_of_int_le_iff ..
29.1066 -
29.1067 -lemma real_of_int_less_one_cancel_iff: "real (i :: int) < 1 \<longleftrightarrow> i < 1"
29.1068 - unfolding real_of_one[symmetric] real_of_int_less_iff ..
29.1069 -
29.1070 -lemma real_of_int_le_one_cancel_iff: "real (i :: int) \<le> 1 \<longleftrightarrow> i \<le> 1"
29.1071 - unfolding real_of_one[symmetric] real_of_int_le_iff ..
29.1072 -
29.1073 -lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))"
29.1074 -by (auto simp add: abs_if)
29.1075 -
29.1076 -lemma int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)"
29.1077 - apply (subgoal_tac "real n + 1 = real (n + 1)")
29.1078 - apply (simp del: real_of_int_add)
29.1079 - apply auto
29.1080 -done
29.1081 -
29.1082 -lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)"
29.1083 - apply (subgoal_tac "real m + 1 = real (m + 1)")
29.1084 - apply (simp del: real_of_int_add)
29.1085 - apply simp
29.1086 -done
29.1087 -
29.1088 -lemma real_of_int_div_aux: "(real (x::int)) / (real d) =
29.1089 - real (x div d) + (real (x mod d)) / (real d)"
29.1090 -proof -
29.1091 - have "x = (x div d) * d + x mod d"
29.1092 - by auto
29.1093 - then have "real x = real (x div d) * real d + real(x mod d)"
29.1094 - by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym])
29.1095 - then have "real x / real d = ... / real d"
29.1096 - by simp
29.1097 - then show ?thesis
29.1098 - by (auto simp add: add_divide_distrib algebra_simps)
29.1099 -qed
29.1100 -
29.