--- a/src/Doc/ProgProve/document/prelude.tex Tue Mar 26 14:38:44 2013 +0100
+++ b/src/Doc/ProgProve/document/prelude.tex Tue Mar 26 15:10:28 2013 +0100
@@ -13,15 +13,6 @@
\usepackage{mathpartir}
\usepackage{amssymb}
-% Enables fixmes
-\newif \ifDraft \Drafttrue
-
-\ifDraft
- \newcommand{\FIXME}[1]{\textbf{\textsl{FIXME: #1}}}
-\else
- \newcommand{\FIXME}[1]{\relax}
-\fi
-
\renewcommand*\descriptionlabel[1]{\hspace\labelsep \textbf{#1}\hfil}
% this should be the last package used
--- a/src/HOL/Complex_Main.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Complex_Main.thy Tue Mar 26 15:10:28 2013 +0100
@@ -4,10 +4,8 @@
imports
Main
Real
- SupInf
Complex
- Log
- Ln
+ Transcendental
Taylor
Deriv
begin
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Conditional_Complete_Lattices.thy Tue Mar 26 15:10:28 2013 +0100
@@ -0,0 +1,295 @@
+(* Title: HOL/Conditional_Complete_Lattices.thy
+ Author: Amine Chaieb and L C Paulson, University of Cambridge
+ Author: Johanens Hölzl, TU München
+*)
+
+header {* Conditional Complete Lattices *}
+
+theory Conditional_Complete_Lattices
+imports Main Lubs
+begin
+
+lemma Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
+ by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
+
+lemma Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
+ by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
+
+text {*
+
+To avoid name classes with the @{class complete_lattice}-class we prefix @{const Sup} and
+@{const Inf} in theorem names with c.
+
+*}
+
+class conditional_complete_lattice = lattice + Sup + Inf +
+ assumes cInf_lower: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> z \<le> a) \<Longrightarrow> Inf X \<le> x"
+ and cInf_greatest: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X"
+ assumes cSup_upper: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> a \<le> z) \<Longrightarrow> x \<le> Sup X"
+ and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
+begin
+
+lemma cSup_eq_maximum: (*REAL_SUP_MAX in HOL4*)
+ "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
+ by (blast intro: antisym cSup_upper cSup_least)
+
+lemma cInf_eq_minimum: (*REAL_INF_MIN in HOL4*)
+ "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
+ by (intro antisym cInf_lower[of z X z] cInf_greatest[of X z]) auto
+
+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)"
+ by (metis order_trans cSup_upper cSup_least)
+
+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)"
+ by (metis order_trans cInf_lower cInf_greatest)
+
+lemma cSup_singleton [simp]: "Sup {x} = x"
+ by (intro cSup_eq_maximum) auto
+
+lemma cInf_singleton [simp]: "Inf {x} = x"
+ by (intro cInf_eq_minimum) auto
+
+lemma cSup_upper2: (*REAL_IMP_LE_SUP in HOL4*)
+ "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y \<le> Sup X"
+ by (metis cSup_upper order_trans)
+
+lemma cInf_lower2:
+ "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X \<le> y"
+ by (metis cInf_lower order_trans)
+
+lemma cSup_upper_EX: "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> x \<le> z \<Longrightarrow> x \<le> Sup X"
+ by (blast intro: cSup_upper)
+
+lemma cInf_lower_EX: "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> z \<le> x \<Longrightarrow> Inf X \<le> x"
+ by (blast intro: cInf_lower)
+
+lemma cSup_eq_non_empty:
+ assumes 1: "X \<noteq> {}"
+ assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
+ assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
+ shows "Sup X = a"
+ by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper)
+
+lemma cInf_eq_non_empty:
+ assumes 1: "X \<noteq> {}"
+ assumes 2: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
+ assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
+ shows "Inf X = a"
+ by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
+
+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}"
+ by (rule cInf_eq_non_empty) (auto intro: cSup_upper cSup_least)
+
+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}"
+ by (rule cSup_eq_non_empty) (auto intro: cInf_lower cInf_greatest)
+
+lemma cSup_insert:
+ assumes x: "X \<noteq> {}"
+ and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
+ shows "Sup (insert a X) = sup a (Sup X)"
+proof (intro cSup_eq_non_empty)
+ 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)
+qed (auto intro: le_supI2 z cSup_upper)
+
+lemma cInf_insert:
+ assumes x: "X \<noteq> {}"
+ and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
+ shows "Inf (insert a X) = inf a (Inf X)"
+proof (intro cInf_eq_non_empty)
+ 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)
+qed (auto intro: le_infI2 z cInf_lower)
+
+lemma cSup_insert_If:
+ "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
+ using cSup_insert[of X z] by simp
+
+lemma cInf_insert_if:
+ "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
+ using cInf_insert[of X z] by simp
+
+lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
+proof (induct X arbitrary: x rule: finite_induct)
+ case (insert x X y) then show ?case
+ apply (cases "X = {}")
+ apply simp
+ apply (subst cSup_insert[of _ "Sup X"])
+ apply (auto intro: le_supI2)
+ done
+qed simp
+
+lemma cInf_le_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> Inf X \<le> x"
+proof (induct X arbitrary: x rule: finite_induct)
+ case (insert x X y) then show ?case
+ apply (cases "X = {}")
+ apply simp
+ apply (subst cInf_insert[of _ "Inf X"])
+ apply (auto intro: le_infI2)
+ done
+qed simp
+
+lemma cSup_eq_Sup_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Sup_fin X"
+proof (induct X rule: finite_ne_induct)
+ case (insert x X) then show ?case
+ using cSup_insert[of X "Sup_fin X" x] le_cSup_finite[of X] by simp
+qed simp
+
+lemma cInf_eq_Inf_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Inf_fin X"
+proof (induct X rule: finite_ne_induct)
+ case (insert x X) then show ?case
+ using cInf_insert[of X "Inf_fin X" x] cInf_le_finite[of X] by simp
+qed simp
+
+lemma cSup_atMost[simp]: "Sup {..x} = x"
+ by (auto intro!: cSup_eq_maximum)
+
+lemma cSup_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Sup {y<..x} = x"
+ by (auto intro!: cSup_eq_maximum)
+
+lemma cSup_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Sup {y..x} = x"
+ by (auto intro!: cSup_eq_maximum)
+
+lemma cInf_atLeast[simp]: "Inf {x..} = x"
+ by (auto intro!: cInf_eq_minimum)
+
+lemma cInf_atLeastLessThan[simp]: "y < x \<Longrightarrow> Inf {y..<x} = y"
+ by (auto intro!: cInf_eq_minimum)
+
+lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
+ by (auto intro!: cInf_eq_minimum)
+
+end
+
+instance complete_lattice \<subseteq> conditional_complete_lattice
+ by default (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest)
+
+lemma isLub_cSup:
+ "(S::'a :: conditional_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. S *<= b) \<Longrightarrow> isLub UNIV S (Sup S)"
+ by (auto simp add: isLub_def setle_def leastP_def isUb_def
+ intro!: setgeI intro: cSup_upper cSup_least)
+
+lemma cSup_eq:
+ fixes a :: "'a :: {conditional_complete_lattice, no_bot}"
+ assumes upper: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
+ assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
+ shows "Sup X = a"
+proof cases
+ assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
+qed (intro cSup_eq_non_empty assms)
+
+lemma cInf_eq:
+ fixes a :: "'a :: {conditional_complete_lattice, no_top}"
+ assumes upper: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
+ assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
+ shows "Inf X = a"
+proof cases
+ assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
+qed (intro cInf_eq_non_empty assms)
+
+lemma cSup_le: "(S::'a::conditional_complete_lattice set) \<noteq> {} \<Longrightarrow> S *<= b \<Longrightarrow> Sup S \<le> b"
+ by (metis cSup_least setle_def)
+
+lemma cInf_ge: "(S::'a :: conditional_complete_lattice set) \<noteq> {} \<Longrightarrow> b <=* S \<Longrightarrow> Inf S \<ge> b"
+ by (metis cInf_greatest setge_def)
+
+class conditional_complete_linorder = conditional_complete_lattice + linorder
+begin
+
+lemma less_cSup_iff : (*REAL_SUP_LE in HOL4*)
+ "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
+ by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans)
+
+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)"
+ by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans)
+
+lemma less_cSupE:
+ assumes "y < Sup X" "X \<noteq> {}" obtains x where "x \<in> X" "y < x"
+ by (metis cSup_least assms not_le that)
+
+lemma less_cSupD:
+ "X \<noteq> {} \<Longrightarrow> z < Sup X \<Longrightarrow> \<exists>x\<in>X. z < x"
+ by (metis less_cSup_iff not_leE)
+
+lemma cInf_lessD:
+ "X \<noteq> {} \<Longrightarrow> Inf X < z \<Longrightarrow> \<exists>x\<in>X. x < z"
+ by (metis cInf_less_iff not_leE)
+
+lemma complete_interval:
+ assumes "a < b" and "P a" and "\<not> P b"
+ shows "\<exists>c. a \<le> c \<and> c \<le> b \<and> (\<forall>x. a \<le> x \<and> x < c \<longrightarrow> P x) \<and>
+ (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)"
+proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x & x < d --> P x}"], auto)
+ show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
+ by (rule cSup_upper [where z=b], auto)
+ (metis `a < b` `\<not> P b` linear less_le)
+next
+ show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
+ apply (rule cSup_least)
+ apply auto
+ apply (metis less_le_not_le)
+ apply (metis `a<b` `~ P b` linear less_le)
+ done
+next
+ fix x
+ assume x: "a \<le> x" and lt: "x < Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
+ show "P x"
+ apply (rule less_cSupE [OF lt], auto)
+ apply (metis less_le_not_le)
+ apply (metis x)
+ done
+next
+ fix d
+ assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x"
+ thus "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
+ by (rule_tac z="b" in cSup_upper, auto)
+ (metis `a<b` `~ P b` linear less_le)
+qed
+
+end
+
+lemma cSup_bounds:
+ fixes S :: "'a :: conditional_complete_lattice set"
+ assumes Se: "S \<noteq> {}" and l: "a <=* S" and u: "S *<= b"
+ shows "a \<le> Sup S \<and> Sup S \<le> b"
+proof-
+ from isLub_cSup[OF Se] u have lub: "isLub UNIV S (Sup S)" by blast
+ hence b: "Sup S \<le> b" using u
+ by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
+ from Se obtain y where y: "y \<in> S" by blast
+ from lub l have "a \<le> Sup S"
+ by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
+ (metis le_iff_sup le_sup_iff y)
+ with b show ?thesis by blast
+qed
+
+
+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"
+ by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
+
+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"
+ by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
+
+lemma cSup_eq_Max: "finite (X::'a::conditional_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
+ using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
+
+lemma cInf_eq_Min: "finite (X::'a::conditional_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
+ using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
+
+lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditional_complete_linorder, dense_linorder}} = x"
+ by (auto intro!: cSup_eq_non_empty intro: dense_le)
+
+lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditional_complete_linorder, dense_linorder}} = x"
+ by (auto intro!: cSup_eq intro: dense_le_bounded)
+
+lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditional_complete_linorder, dense_linorder}} = x"
+ by (auto intro!: cSup_eq intro: dense_le_bounded)
+
+lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditional_complete_linorder, dense_linorder} <..} = x"
+ by (auto intro!: cInf_eq intro: dense_ge)
+
+lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditional_complete_linorder, dense_linorder}} = y"
+ by (auto intro!: cInf_eq intro: dense_ge_bounded)
+
+lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditional_complete_linorder, dense_linorder}} = y"
+ by (auto intro!: cInf_eq intro: dense_ge_bounded)
+
+end
--- a/src/HOL/Deriv.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Deriv.thy Tue Mar 26 15:10:28 2013 +0100
@@ -8,7 +8,7 @@
header{* Differentiation *}
theory Deriv
-imports Lim
+imports Limits
begin
text{*Standard Definitions*}
@@ -422,178 +422,6 @@
apply (simp add: assms)
done
-
-subsection {* Nested Intervals and Bisection *}
-
-lemma nested_sequence_unique:
- 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"
- shows "\<exists>l::real. ((\<forall>n. f n \<le> l) \<and> f ----> l) \<and> ((\<forall>n. l \<le> g n) \<and> g ----> l)"
-proof -
- have "incseq f" unfolding incseq_Suc_iff by fact
- have "decseq g" unfolding decseq_Suc_iff by fact
-
- { fix n
- from `decseq g` have "g n \<le> g 0" by (rule decseqD) simp
- with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f n \<le> g 0" by auto }
- then obtain u where "f ----> u" "\<forall>i. f i \<le> u"
- using incseq_convergent[OF `incseq f`] by auto
- moreover
- { fix n
- from `incseq f` have "f 0 \<le> f n" by (rule incseqD) simp
- with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f 0 \<le> g n" by simp }
- then obtain l where "g ----> l" "\<forall>i. l \<le> g i"
- using decseq_convergent[OF `decseq g`] by auto
- moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF `f ----> u` `g ----> l`]]
- ultimately show ?thesis by auto
-qed
-
-lemma Bolzano[consumes 1, case_names trans local]:
- fixes P :: "real \<Rightarrow> real \<Rightarrow> bool"
- assumes [arith]: "a \<le> b"
- assumes trans: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a c"
- 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"
- shows "P a b"
-proof -
- 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))"
- def l \<equiv> "\<lambda>n. fst (bisect n)" and u \<equiv> "\<lambda>n. snd (bisect n)"
- 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)"
- 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)"
- by (simp_all add: l_def u_def bisect_def split: prod.split)
-
- { fix n have "l n \<le> u n" by (induct n) auto } note this[simp]
-
- have "\<exists>x. ((\<forall>n. l n \<le> x) \<and> l ----> x) \<and> ((\<forall>n. x \<le> u n) \<and> u ----> x)"
- proof (safe intro!: nested_sequence_unique)
- fix n show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" by (induct n) auto
- next
- { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) }
- then show "(\<lambda>n. l n - u n) ----> 0" by (simp add: LIMSEQ_divide_realpow_zero)
- qed fact
- then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l ----> x" "u ----> x" by auto
- 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"
- using `l 0 \<le> x` `x \<le> u 0` local[of x] by auto
-
- show "P a b"
- proof (rule ccontr)
- assume "\<not> P a b"
- { fix n have "\<not> P (l n) (u n)"
- proof (induct n)
- case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto
- qed (simp add: `\<not> P a b`) }
- moreover
- { have "eventually (\<lambda>n. x - d / 2 < l n) sequentially"
- using `0 < d` `l ----> x` by (intro order_tendstoD[of _ x]) auto
- moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially"
- using `0 < d` `u ----> x` by (intro order_tendstoD[of _ x]) auto
- ultimately have "eventually (\<lambda>n. P (l n) (u n)) sequentially"
- proof eventually_elim
- fix n assume "x - d / 2 < l n" "u n < x + d / 2"
- from add_strict_mono[OF this] have "u n - l n < d" by simp
- with x show "P (l n) (u n)" by (rule d)
- qed }
- ultimately show False by simp
- qed
-qed
-
-(*HOL style here: object-level formulations*)
-lemma IVT_objl: "(f(a::real) \<le> (y::real) & y \<le> f(b) & a \<le> b &
- (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
- --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
-apply (blast intro: IVT)
-done
-
-lemma IVT2_objl: "(f(b::real) \<le> (y::real) & y \<le> f(a) & a \<le> b &
- (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
- --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
-apply (blast intro: IVT2)
-done
-
-
-lemma compact_Icc[simp, intro]: "compact {a .. b::real}"
-proof (cases "a \<le> b", rule compactI)
- fix C assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C"
- def T == "{a .. b}"
- from C(1,3) show "\<exists>C'\<subseteq>C. finite C' \<and> {a..b} \<subseteq> \<Union>C'"
- proof (induct rule: Bolzano)
- case (trans a b c)
- then have *: "{a .. c} = {a .. b} \<union> {b .. c}" by auto
- 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"
- by (auto simp: *)
- with trans show ?case
- unfolding * by (intro exI[of _ "C1 \<union> C2"]) auto
- next
- case (local x)
- then have "x \<in> \<Union>C" using C by auto
- with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" by auto
- then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c"
- by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff)
- with `c \<in> C` show ?case
- by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto
- qed
-qed simp
-
-subsection {* Boundedness of continuous functions *}
-
-text{*By bisection, function continuous on closed interval is bounded above*}
-
-lemma isCont_eq_Ub:
- fixes f :: "real \<Rightarrow> 'a::linorder_topology"
- shows "a \<le> b \<Longrightarrow> \<forall>x::real. 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) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
- using continuous_attains_sup[of "{a .. b}" f]
- apply (simp add: continuous_at_imp_continuous_on Ball_def)
- apply safe
- apply (rule_tac x="f x" in exI)
- apply auto
- done
-
-lemma isCont_eq_Lb:
- fixes f :: "real \<Rightarrow> 'a::linorder_topology"
- 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> M \<le> f x) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
- using continuous_attains_inf[of "{a .. b}" f]
- apply (simp add: continuous_at_imp_continuous_on Ball_def)
- apply safe
- apply (rule_tac x="f x" in exI)
- apply auto
- done
-
-lemma isCont_bounded:
- fixes f :: "real \<Rightarrow> 'a::linorder_topology"
- 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"
- using isCont_eq_Ub[of a b f] by auto
-
-lemma isCont_has_Ub:
- fixes f :: "real \<Rightarrow> 'a::linorder_topology"
- 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) \<and> (\<forall>N. N < M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> N < f x))"
- using isCont_eq_Ub[of a b f] by auto
-
-text{*Refine the above to existence of least upper bound*}
-
-lemma lemma_reals_complete: "((\<exists>x. x \<in> S) & (\<exists>y. isUb UNIV S (y::real))) -->
- (\<exists>t. isLub UNIV S t)"
-by (blast intro: reals_complete)
-
-
-text{*Another version.*}
-
-lemma isCont_Lb_Ub: "[|a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |]
- ==> \<exists>L M::real. (\<forall>x::real. a \<le> x & x \<le> b --> L \<le> f(x) & f(x) \<le> M) &
- (\<forall>y. L \<le> y & y \<le> M --> (\<exists>x. a \<le> x & x \<le> b & (f(x) = y)))"
-apply (frule isCont_eq_Lb)
-apply (frule_tac [2] isCont_eq_Ub)
-apply (assumption+, safe)
-apply (rule_tac x = "f x" in exI)
-apply (rule_tac x = "f xa" in exI, simp, safe)
-apply (cut_tac x = x and y = xa in linorder_linear, safe)
-apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl)
-apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe)
-apply (rule_tac [2] x = xb in exI)
-apply (rule_tac [4] x = xb in exI, simp_all)
-done
-
-
subsection {* Local extrema *}
text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*}
@@ -658,7 +486,6 @@
qed
qed
-
lemma DERIV_pos_inc_left:
fixes f :: "real => real"
shows "DERIV f x :> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x - h) < f(x)"
@@ -1081,47 +908,6 @@
by simp
qed
-text{*Continuity of inverse function*}
-
-lemma isCont_inverse_function:
- fixes f g :: "real \<Rightarrow> real"
- assumes d: "0 < d"
- and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> g (f z) = z"
- and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> isCont f z"
- shows "isCont g (f x)"
-proof -
- let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + d}"
-
- have f: "continuous_on ?D f"
- using cont by (intro continuous_at_imp_continuous_on ballI) auto
- then have g: "continuous_on (f`?D) g"
- using inj by (intro continuous_on_inv) auto
-
- from d f have "{min ?A ?B <..< max ?A ?B} \<subseteq> f ` ?D"
- by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max)
- with g have "continuous_on {min ?A ?B <..< max ?A ?B} g"
- by (rule continuous_on_subset)
- moreover
- have "(?A < f x \<and> f x < ?B) \<or> (?B < f x \<and> f x < ?A)"
- using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto
- then have "f x \<in> {min ?A ?B <..< max ?A ?B}"
- by auto
- ultimately
- show ?thesis
- by (simp add: continuous_on_eq_continuous_at)
-qed
-
-lemma isCont_inverse_function2:
- fixes f g :: "real \<Rightarrow> real" shows
- "\<lbrakk>a < x; x < b;
- \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z;
- \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z\<rbrakk>
- \<Longrightarrow> isCont g (f x)"
-apply (rule isCont_inverse_function
- [where f=f and d="min (x - a) (b - x)"])
-apply (simp_all add: abs_le_iff)
-done
-
text {* Derivative of inverse function *}
lemma DERIV_inverse_function:
@@ -1228,44 +1014,6 @@
with g'cdef f'cdef cint show ?thesis by auto
qed
-
-subsection {* Theorems about Limits *}
-
-(* need to rename second isCont_inverse *)
-
-lemma isCont_inv_fun:
- fixes f g :: "real \<Rightarrow> real"
- shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z;
- \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |]
- ==> isCont g (f x)"
-by (rule isCont_inverse_function)
-
-text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*}
-lemma LIM_fun_gt_zero:
- "[| f -- c --> (l::real); 0 < l |]
- ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> 0 < f x)"
-apply (drule (1) LIM_D, clarify)
-apply (rule_tac x = s in exI)
-apply (simp add: abs_less_iff)
-done
-
-lemma LIM_fun_less_zero:
- "[| f -- c --> (l::real); l < 0 |]
- ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x < 0)"
-apply (drule LIM_D [where r="-l"], simp, clarify)
-apply (rule_tac x = s in exI)
-apply (simp add: abs_less_iff)
-done
-
-lemma LIM_fun_not_zero:
- "[| f -- c --> (l::real); l \<noteq> 0 |]
- ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x \<noteq> 0)"
-apply (rule linorder_cases [of l 0])
-apply (drule (1) LIM_fun_less_zero, force)
-apply simp
-apply (drule (1) LIM_fun_gt_zero, force)
-done
-
lemma GMVT':
fixes f g :: "real \<Rightarrow> real"
assumes "a < b"
@@ -1284,6 +1032,9 @@
by auto
qed
+
+subsection {* L'Hopitals rule *}
+
lemma DERIV_cong_ev: "x = y \<Longrightarrow> eventually (\<lambda>x. f x = g x) (nhds x) \<Longrightarrow> u = v \<Longrightarrow>
DERIV f x :> u \<longleftrightarrow> DERIV g y :> v"
unfolding DERIV_iff2
--- a/src/HOL/Library/Diagonal_Subsequence.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Library/Diagonal_Subsequence.thy Tue Mar 26 15:10:28 2013 +0100
@@ -3,7 +3,7 @@
header {* Sequence of Properties on Subsequences *}
theory Diagonal_Subsequence
-imports SEQ
+imports Complex_Main
begin
locale subseqs =
--- a/src/HOL/Lim.thy Tue Mar 26 14:38:44 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,225 +0,0 @@
-(* Title : Lim.thy
- Author : Jacques D. Fleuriot
- Copyright : 1998 University of Cambridge
- Conversion to Isar and new proofs by Lawrence C Paulson, 2004
-*)
-
-header{* Limits and Continuity *}
-
-theory Lim
-imports SEQ
-begin
-
-subsection {* Limits of Functions *}
-
-lemma LIM_eq:
- fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
- shows "f -- a --> L =
- (\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"
-by (simp add: LIM_def dist_norm)
-
-lemma LIM_I:
- fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
- shows "(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)
- ==> f -- a --> L"
-by (simp add: LIM_eq)
-
-lemma LIM_D:
- fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
- shows "[| f -- a --> L; 0<r |]
- ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"
-by (simp add: LIM_eq)
-
-lemma LIM_offset:
- fixes a :: "'a::real_normed_vector"
- shows "f -- a --> L \<Longrightarrow> (\<lambda>x. f (x + k)) -- a - k --> L"
-apply (rule topological_tendstoI)
-apply (drule (2) topological_tendstoD)
-apply (simp only: eventually_at dist_norm)
-apply (clarify, rule_tac x=d in exI, safe)
-apply (drule_tac x="x + k" in spec)
-apply (simp add: algebra_simps)
-done
-
-lemma LIM_offset_zero:
- fixes a :: "'a::real_normed_vector"
- shows "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"
-by (drule_tac k="a" in LIM_offset, simp add: add_commute)
-
-lemma LIM_offset_zero_cancel:
- fixes a :: "'a::real_normed_vector"
- shows "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> f -- a --> L"
-by (drule_tac k="- a" in LIM_offset, simp)
-
-lemma LIM_zero:
- fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
- shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. f x - l) ---> 0) F"
-unfolding tendsto_iff dist_norm by simp
-
-lemma LIM_zero_cancel:
- fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
- shows "((\<lambda>x. f x - l) ---> 0) F \<Longrightarrow> (f ---> l) F"
-unfolding tendsto_iff dist_norm by simp
-
-lemma LIM_zero_iff:
- fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
- shows "((\<lambda>x. f x - l) ---> 0) F = (f ---> l) F"
-unfolding tendsto_iff dist_norm by simp
-
-lemma LIM_imp_LIM:
- fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
- fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector"
- assumes f: "f -- a --> l"
- assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"
- shows "g -- a --> m"
- by (rule metric_LIM_imp_LIM [OF f],
- simp add: dist_norm le)
-
-lemma LIM_equal2:
- fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
- assumes 1: "0 < R"
- assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"
- shows "g -- a --> l \<Longrightarrow> f -- a --> l"
-by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm)
-
-lemma LIM_compose2:
- fixes a :: "'a::real_normed_vector"
- assumes f: "f -- a --> b"
- assumes g: "g -- b --> c"
- assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"
- shows "(\<lambda>x. g (f x)) -- a --> c"
-by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])
-
-lemma real_LIM_sandwich_zero:
- fixes f g :: "'a::topological_space \<Rightarrow> real"
- assumes f: "f -- a --> 0"
- assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"
- assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"
- shows "g -- a --> 0"
-proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *)
- fix x assume x: "x \<noteq> a"
- have "norm (g x - 0) = g x" by (simp add: 1 x)
- also have "g x \<le> f x" by (rule 2 [OF x])
- also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self)
- also have "\<bar>f x\<bar> = norm (f x - 0)" by simp
- finally show "norm (g x - 0) \<le> norm (f x - 0)" .
-qed
-
-
-subsection {* Continuity *}
-
-lemma LIM_isCont_iff:
- fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
- shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)"
-by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])
-
-lemma isCont_iff:
- fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
- shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x"
-by (simp add: isCont_def LIM_isCont_iff)
-
-lemma isCont_LIM_compose2:
- fixes a :: "'a::real_normed_vector"
- assumes f [unfolded isCont_def]: "isCont f a"
- assumes g: "g -- f a --> l"
- assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"
- shows "(\<lambda>x. g (f x)) -- a --> l"
-by (rule LIM_compose2 [OF f g inj])
-
-
-lemma isCont_norm [simp]:
- fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
- shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"
- by (fact continuous_norm)
-
-lemma isCont_rabs [simp]:
- fixes f :: "'a::t2_space \<Rightarrow> real"
- shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
- by (fact continuous_rabs)
-
-lemma isCont_add [simp]:
- fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
- shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"
- by (fact continuous_add)
-
-lemma isCont_minus [simp]:
- fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
- shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"
- by (fact continuous_minus)
-
-lemma isCont_diff [simp]:
- fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
- shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"
- by (fact continuous_diff)
-
-lemma isCont_mult [simp]:
- fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
- shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"
- by (fact continuous_mult)
-
-lemma (in bounded_linear) isCont:
- "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
- by (fact continuous)
-
-lemma (in bounded_bilinear) isCont:
- "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"
- by (fact continuous)
-
-lemmas isCont_scaleR [simp] =
- bounded_bilinear.isCont [OF bounded_bilinear_scaleR]
-
-lemmas isCont_of_real [simp] =
- bounded_linear.isCont [OF bounded_linear_of_real]
-
-lemma isCont_power [simp]:
- fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
- shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"
- by (fact continuous_power)
-
-lemma isCont_setsum [simp]:
- fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
- shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
- by (auto intro: continuous_setsum)
-
-lemmas isCont_intros =
- isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus
- isCont_diff isCont_mult isCont_inverse isCont_divide isCont_scaleR
- isCont_of_real isCont_power isCont_sgn isCont_setsum
-
-subsection {* Uniform Continuity *}
-
-lemma (in bounded_linear) isUCont: "isUCont f"
-unfolding isUCont_def dist_norm
-proof (intro allI impI)
- fix r::real assume r: "0 < r"
- obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"
- using pos_bounded by fast
- show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"
- proof (rule exI, safe)
- from r K show "0 < r / K" by (rule divide_pos_pos)
- next
- fix x y :: 'a
- assume xy: "norm (x - y) < r / K"
- have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff)
- also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le)
- also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq)
- finally show "norm (f x - f y) < r" .
- qed
-qed
-
-lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
-by (rule isUCont [THEN isUCont_Cauchy])
-
-
-lemma LIM_less_bound:
- fixes f :: "real \<Rightarrow> real"
- assumes ev: "b < x" "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and "isCont f x"
- shows "0 \<le> f x"
-proof (rule tendsto_le_const)
- show "(f ---> f x) (at_left x)"
- using `isCont f x` by (simp add: filterlim_at_split isCont_def)
- show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)"
- using ev by (auto simp: eventually_within_less dist_real_def intro!: exI[of _ "x - b"])
-qed simp
-
-end
--- a/src/HOL/Limits.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Limits.thy Tue Mar 26 15:10:28 2013 +0100
@@ -1,13 +1,23 @@
-(* Title : Limits.thy
- Author : Brian Huffman
+(* Title: Limits.thy
+ Author: Brian Huffman
+ Author: Jacques D. Fleuriot, University of Cambridge
+ Author: Lawrence C Paulson
+ Author: Jeremy Avigad
+
*)
-header {* Filters and Limits *}
+header {* Limits on Real Vector Spaces *}
theory Limits
-imports RealVector
+imports Real_Vector_Spaces
begin
+(* Unfortunately eventually_within was overwritten by Multivariate_Analysis.
+ Hence it was references as Limits.eventually_within, but now it is Basic_Topology.eventually_within *)
+lemmas eventually_within = eventually_within
+
+subsection {* Filter going to infinity norm *}
+
definition at_infinity :: "'a::real_normed_vector filter" where
"at_infinity = Abs_filter (\<lambda>P. \<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x)"
@@ -47,7 +57,21 @@
"at_bot \<le> (at_infinity :: real filter)"
unfolding at_infinity_eq_at_top_bot by simp
-subsection {* Boundedness *}
+subsubsection {* Boundedness *}
+
+definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where
+ Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)"
+
+abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
+ "Bseq X \<equiv> Bfun X sequentially"
+
+lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" ..
+
+lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
+ unfolding Bfun_metric_def by (subst eventually_sequentially_seg)
+
+lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
+ unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg)
lemma Bfun_def:
"Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)"
@@ -77,6 +101,154 @@
obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F"
using assms unfolding Bfun_def by fast
+lemma Cauchy_Bseq: "Cauchy X \<Longrightarrow> Bseq X"
+ unfolding Cauchy_def Bfun_metric_def eventually_sequentially
+ apply (erule_tac x=1 in allE)
+ apply simp
+ apply safe
+ apply (rule_tac x="X M" in exI)
+ apply (rule_tac x=1 in exI)
+ apply (erule_tac x=M in allE)
+ apply simp
+ apply (rule_tac x=M in exI)
+ apply (auto simp: dist_commute)
+ done
+
+
+subsubsection {* Bounded Sequences *}
+
+lemma BseqI': "(\<And>n. norm (X n) \<le> K) \<Longrightarrow> Bseq X"
+ by (intro BfunI) (auto simp: eventually_sequentially)
+
+lemma BseqI2': "\<forall>n\<ge>N. norm (X n) \<le> K \<Longrightarrow> Bseq X"
+ by (intro BfunI) (auto simp: eventually_sequentially)
+
+lemma Bseq_def: "Bseq X \<longleftrightarrow> (\<exists>K>0. \<forall>n. norm (X n) \<le> K)"
+ unfolding Bfun_def eventually_sequentially
+proof safe
+ fix N K assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K"
+ then show "\<exists>K>0. \<forall>n. norm (X n) \<le> K"
+ by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] min_max.less_supI2)
+ (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj)
+qed auto
+
+lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
+unfolding Bseq_def by auto
+
+lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
+by (simp add: Bseq_def)
+
+lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
+by (auto simp add: Bseq_def)
+
+lemma lemma_NBseq_def:
+ "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
+proof safe
+ fix K :: real
+ from reals_Archimedean2 obtain n :: nat where "K < real n" ..
+ then have "K \<le> real (Suc n)" by auto
+ moreover assume "\<forall>m. norm (X m) \<le> K"
+ ultimately have "\<forall>m. norm (X m) \<le> real (Suc n)"
+ by (blast intro: order_trans)
+ then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" ..
+qed (force simp add: real_of_nat_Suc)
+
+text{* alternative definition for Bseq *}
+lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
+apply (simp add: Bseq_def)
+apply (simp (no_asm) add: lemma_NBseq_def)
+done
+
+lemma lemma_NBseq_def2:
+ "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
+apply (subst lemma_NBseq_def, auto)
+apply (rule_tac x = "Suc N" in exI)
+apply (rule_tac [2] x = N in exI)
+apply (auto simp add: real_of_nat_Suc)
+ prefer 2 apply (blast intro: order_less_imp_le)
+apply (drule_tac x = n in spec, simp)
+done
+
+(* yet another definition for Bseq *)
+lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
+by (simp add: Bseq_def lemma_NBseq_def2)
+
+subsubsection{*A Few More Equivalence Theorems for Boundedness*}
+
+text{*alternative formulation for boundedness*}
+lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
+apply (unfold Bseq_def, safe)
+apply (rule_tac [2] x = "k + norm x" in exI)
+apply (rule_tac x = K in exI, simp)
+apply (rule exI [where x = 0], auto)
+apply (erule order_less_le_trans, simp)
+apply (drule_tac x=n in spec, fold diff_minus)
+apply (drule order_trans [OF norm_triangle_ineq2])
+apply simp
+done
+
+text{*alternative formulation for boundedness*}
+lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
+apply safe
+apply (simp add: Bseq_def, safe)
+apply (rule_tac x = "K + norm (X N)" in exI)
+apply auto
+apply (erule order_less_le_trans, simp)
+apply (rule_tac x = N in exI, safe)
+apply (drule_tac x = n in spec)
+apply (rule order_trans [OF norm_triangle_ineq], simp)
+apply (auto simp add: Bseq_iff2)
+done
+
+lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
+apply (simp add: Bseq_def)
+apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
+apply (drule_tac x = n in spec, arith)
+done
+
+subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
+
+lemma Bseq_isUb:
+ "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
+by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
+
+text{* Use completeness of reals (supremum property)
+ to show that any bounded sequence has a least upper bound*}
+
+lemma Bseq_isLub:
+ "!!(X::nat=>real). Bseq X ==>
+ \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
+by (blast intro: reals_complete Bseq_isUb)
+
+lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X"
+ by (simp add: Bseq_def)
+
+lemma Bseq_eq_bounded: "range f \<subseteq> {a .. b::real} \<Longrightarrow> Bseq f"
+ apply (simp add: subset_eq)
+ apply (rule BseqI'[where K="max (norm a) (norm b)"])
+ apply (erule_tac x=n in allE)
+ apply auto
+ done
+
+lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> (B::real) \<Longrightarrow> Bseq X"
+ by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def)
+
+lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. (B::real) \<le> X i \<Longrightarrow> Bseq X"
+ by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def)
+
+subsection {* Bounded Monotonic Sequences *}
+
+subsubsection{*A Bounded and Monotonic Sequence Converges*}
+
+(* TODO: delete *)
+(* FIXME: one use in NSA/HSEQ.thy *)
+lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
+ apply (rule_tac x="X m" in exI)
+ apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const])
+ unfolding eventually_sequentially
+ apply blast
+ done
+
subsection {* Convergence to Zero *}
definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
@@ -232,6 +404,37 @@
subsubsection {* Distance and norms *}
+lemma tendsto_dist [tendsto_intros]:
+ fixes l m :: "'a :: metric_space"
+ assumes f: "(f ---> l) F" and g: "(g ---> m) F"
+ shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) F"
+proof (rule tendstoI)
+ fix e :: real assume "0 < e"
+ hence e2: "0 < e/2" by simp
+ from tendstoD [OF f e2] tendstoD [OF g e2]
+ show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F"
+ proof (eventually_elim)
+ case (elim x)
+ then show "dist (dist (f x) (g x)) (dist l m) < e"
+ unfolding dist_real_def
+ using dist_triangle2 [of "f x" "g x" "l"]
+ using dist_triangle2 [of "g x" "l" "m"]
+ using dist_triangle3 [of "l" "m" "f x"]
+ using dist_triangle [of "f x" "m" "g x"]
+ by arith
+ qed
+qed
+
+lemma continuous_dist[continuous_intros]:
+ fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
+ shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. dist (f x) (g x))"
+ unfolding continuous_def by (rule tendsto_dist)
+
+lemma continuous_on_dist[continuous_on_intros]:
+ fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
+ shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))"
+ unfolding continuous_on_def by (auto intro: tendsto_dist)
+
lemma tendsto_norm [tendsto_intros]:
"(f ---> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> norm a) F"
unfolding norm_conv_dist by (intro tendsto_intros)
@@ -695,6 +898,7 @@
qed
qed force
+
subsection {* Relate @{const at}, @{const at_left} and @{const at_right} *}
text {*
@@ -1027,9 +1231,724 @@
qed simp
-(* Unfortunately eventually_within was overwritten by Multivariate_Analysis.
- Hence it was references as Limits.within, but now it is Basic_Topology.eventually_within *)
-lemmas eventually_within = eventually_within
+subsection {* Limits of Sequences *}
+
+lemma [trans]: "X=Y ==> Y ----> z ==> X ----> z"
+ by simp
+
+lemma LIMSEQ_iff:
+ fixes L :: "'a::real_normed_vector"
+ shows "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
+unfolding LIMSEQ_def dist_norm ..
+
+lemma LIMSEQ_I:
+ fixes L :: "'a::real_normed_vector"
+ shows "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
+by (simp add: LIMSEQ_iff)
+
+lemma LIMSEQ_D:
+ fixes L :: "'a::real_normed_vector"
+ shows "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
+by (simp add: LIMSEQ_iff)
+
+lemma LIMSEQ_linear: "\<lbrakk> X ----> x ; l > 0 \<rbrakk> \<Longrightarrow> (\<lambda> n. X (n * l)) ----> x"
+ unfolding tendsto_def eventually_sequentially
+ by (metis div_le_dividend div_mult_self1_is_m le_trans nat_mult_commute)
+
+lemma Bseq_inverse_lemma:
+ fixes x :: "'a::real_normed_div_algebra"
+ shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
+apply (subst nonzero_norm_inverse, clarsimp)
+apply (erule (1) le_imp_inverse_le)
+done
+
+lemma Bseq_inverse:
+ fixes a :: "'a::real_normed_div_algebra"
+ shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
+ by (rule Bfun_inverse)
+
+lemma LIMSEQ_diff_approach_zero:
+ fixes L :: "'a::real_normed_vector"
+ shows "g ----> L ==> (%x. f x - g x) ----> 0 ==> f ----> L"
+ by (drule (1) tendsto_add, simp)
+
+lemma LIMSEQ_diff_approach_zero2:
+ fixes L :: "'a::real_normed_vector"
+ shows "f ----> L ==> (%x. f x - g x) ----> 0 ==> g ----> L"
+ by (drule (1) tendsto_diff, simp)
+
+text{*An unbounded sequence's inverse tends to 0*}
+
+lemma LIMSEQ_inverse_zero:
+ "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
+ apply (rule filterlim_compose[OF tendsto_inverse_0])
+ apply (simp add: filterlim_at_infinity[OF order_refl] eventually_sequentially)
+ apply (metis abs_le_D1 linorder_le_cases linorder_not_le)
+ done
+
+text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
+
+lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
+ by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc
+ filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity)
+
+text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
+infinity is now easily proved*}
+
+lemma LIMSEQ_inverse_real_of_nat_add:
+ "(%n. r + inverse(real(Suc n))) ----> r"
+ using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto
+
+lemma LIMSEQ_inverse_real_of_nat_add_minus:
+ "(%n. r + -inverse(real(Suc n))) ----> r"
+ using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]]
+ by auto
+
+lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
+ "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
+ using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]]
+ by auto
+
+subsection {* Convergence on sequences *}
+
+lemma convergent_add:
+ fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
+ assumes "convergent (\<lambda>n. X n)"
+ assumes "convergent (\<lambda>n. Y n)"
+ shows "convergent (\<lambda>n. X n + Y n)"
+ using assms unfolding convergent_def by (fast intro: tendsto_add)
+
+lemma convergent_setsum:
+ fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::real_normed_vector"
+ assumes "\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)"
+ shows "convergent (\<lambda>n. \<Sum>i\<in>A. X i n)"
+proof (cases "finite A")
+ case True from this and assms show ?thesis
+ by (induct A set: finite) (simp_all add: convergent_const convergent_add)
+qed (simp add: convergent_const)
+
+lemma (in bounded_linear) convergent:
+ assumes "convergent (\<lambda>n. X n)"
+ shows "convergent (\<lambda>n. f (X n))"
+ using assms unfolding convergent_def by (fast intro: tendsto)
+
+lemma (in bounded_bilinear) convergent:
+ assumes "convergent (\<lambda>n. X n)" and "convergent (\<lambda>n. Y n)"
+ shows "convergent (\<lambda>n. X n ** Y n)"
+ using assms unfolding convergent_def by (fast intro: tendsto)
+
+lemma convergent_minus_iff:
+ fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
+ shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
+apply (simp add: convergent_def)
+apply (auto dest: tendsto_minus)
+apply (drule tendsto_minus, auto)
+done
+
+
+text {* A monotone sequence converges to its least upper bound. *}
+
+lemma isLub_mono_imp_LIMSEQ:
+ fixes X :: "nat \<Rightarrow> real"
+ assumes u: "isLub UNIV {x. \<exists>n. X n = x} u" (* FIXME: use 'range X' *)
+ assumes X: "\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n"
+ shows "X ----> u"
+proof (rule LIMSEQ_I)
+ have 1: "\<forall>n. X n \<le> u"
+ using isLubD2 [OF u] by auto
+ have "\<forall>y. (\<forall>n. X n \<le> y) \<longrightarrow> u \<le> y"
+ using isLub_le_isUb [OF u] by (auto simp add: isUb_def setle_def)
+ hence 2: "\<forall>y<u. \<exists>n. y < X n"
+ by (metis not_le)
+ fix r :: real assume "0 < r"
+ hence "u - r < u" by simp
+ hence "\<exists>m. u - r < X m" using 2 by simp
+ then obtain m where "u - r < X m" ..
+ with X have "\<forall>n\<ge>m. u - r < X n"
+ by (fast intro: less_le_trans)
+ hence "\<exists>m. \<forall>n\<ge>m. u - r < X n" ..
+ thus "\<exists>m. \<forall>n\<ge>m. norm (X n - u) < r"
+ using 1 by (simp add: diff_less_eq add_commute)
+qed
+
+text{*A standard proof of the theorem for monotone increasing sequence*}
+
+lemma Bseq_mono_convergent:
+ "Bseq X \<Longrightarrow> \<forall>m. \<forall>n \<ge> m. X m \<le> X n \<Longrightarrow> convergent (X::nat=>real)"
+ by (metis Bseq_isLub isLub_mono_imp_LIMSEQ convergentI)
+
+text{*Main monotonicity theorem*}
+
+lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
+ by (metis monoseq_iff incseq_def decseq_eq_incseq convergent_minus_iff Bseq_minus_iff
+ Bseq_mono_convergent)
+
+lemma Cauchy_iff:
+ fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
+ shows "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)"
+ unfolding Cauchy_def dist_norm ..
+
+lemma CauchyI:
+ fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
+ 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"
+by (simp add: Cauchy_iff)
+
+lemma CauchyD:
+ fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
+ shows "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
+by (simp add: Cauchy_iff)
+
+lemma incseq_convergent:
+ fixes X :: "nat \<Rightarrow> real"
+ assumes "incseq X" and "\<forall>i. X i \<le> B"
+ obtains L where "X ----> L" "\<forall>i. X i \<le> L"
+proof atomize_elim
+ from incseq_bounded[OF assms] `incseq X` Bseq_monoseq_convergent[of X]
+ obtain L where "X ----> L"
+ by (auto simp: convergent_def monoseq_def incseq_def)
+ with `incseq X` show "\<exists>L. X ----> L \<and> (\<forall>i. X i \<le> L)"
+ by (auto intro!: exI[of _ L] incseq_le)
+qed
+
+lemma decseq_convergent:
+ fixes X :: "nat \<Rightarrow> real"
+ assumes "decseq X" and "\<forall>i. B \<le> X i"
+ obtains L where "X ----> L" "\<forall>i. L \<le> X i"
+proof atomize_elim
+ from decseq_bounded[OF assms] `decseq X` Bseq_monoseq_convergent[of X]
+ obtain L where "X ----> L"
+ by (auto simp: convergent_def monoseq_def decseq_def)
+ with `decseq X` show "\<exists>L. X ----> L \<and> (\<forall>i. L \<le> X i)"
+ by (auto intro!: exI[of _ L] decseq_le)
+qed
+
+subsubsection {* Cauchy Sequences are Bounded *}
+
+text{*A Cauchy sequence is bounded -- this is the standard
+ proof mechanization rather than the nonstandard proof*}
+
+lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
+ ==> \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
+apply (clarify, drule spec, drule (1) mp)
+apply (simp only: norm_minus_commute)
+apply (drule order_le_less_trans [OF norm_triangle_ineq2])
+apply simp
+done
+
+subsection {* Power Sequences *}
+
+text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
+"x<1"}. Proof will use (NS) Cauchy equivalence for convergence and
+ also fact that bounded and monotonic sequence converges.*}
+
+lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
+apply (simp add: Bseq_def)
+apply (rule_tac x = 1 in exI)
+apply (simp add: power_abs)
+apply (auto dest: power_mono)
+done
+
+lemma monoseq_realpow: fixes x :: real shows "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
+apply (clarify intro!: mono_SucI2)
+apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
+done
+
+lemma convergent_realpow:
+ "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
+by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
+
+lemma LIMSEQ_inverse_realpow_zero: "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
+ by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp
+
+lemma LIMSEQ_realpow_zero:
+ "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
+proof cases
+ assume "0 \<le> x" and "x \<noteq> 0"
+ hence x0: "0 < x" by simp
+ assume x1: "x < 1"
+ from x0 x1 have "1 < inverse x"
+ by (rule one_less_inverse)
+ hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
+ by (rule LIMSEQ_inverse_realpow_zero)
+ thus ?thesis by (simp add: power_inverse)
+qed (rule LIMSEQ_imp_Suc, simp add: tendsto_const)
+
+lemma LIMSEQ_power_zero:
+ fixes x :: "'a::{real_normed_algebra_1}"
+ shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
+apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
+apply (simp only: tendsto_Zfun_iff, erule Zfun_le)
+apply (simp add: power_abs norm_power_ineq)
+done
+
+lemma LIMSEQ_divide_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. a / (x ^ n) :: real) ----> 0"
+ by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp
+
+text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
+
+lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n :: real) ----> 0"
+ by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
+
+lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n :: real) ----> 0"
+ by (rule LIMSEQ_power_zero) simp
+
+
+subsection {* Limits of Functions *}
+
+lemma LIM_eq:
+ fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
+ shows "f -- a --> L =
+ (\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"
+by (simp add: LIM_def dist_norm)
+
+lemma LIM_I:
+ fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
+ shows "(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)
+ ==> f -- a --> L"
+by (simp add: LIM_eq)
+
+lemma LIM_D:
+ fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
+ shows "[| f -- a --> L; 0<r |]
+ ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"
+by (simp add: LIM_eq)
+
+lemma LIM_offset:
+ fixes a :: "'a::real_normed_vector"
+ shows "f -- a --> L \<Longrightarrow> (\<lambda>x. f (x + k)) -- a - k --> L"
+apply (rule topological_tendstoI)
+apply (drule (2) topological_tendstoD)
+apply (simp only: eventually_at dist_norm)
+apply (clarify, rule_tac x=d in exI, safe)
+apply (drule_tac x="x + k" in spec)
+apply (simp add: algebra_simps)
+done
+
+lemma LIM_offset_zero:
+ fixes a :: "'a::real_normed_vector"
+ shows "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"
+by (drule_tac k="a" in LIM_offset, simp add: add_commute)
+
+lemma LIM_offset_zero_cancel:
+ fixes a :: "'a::real_normed_vector"
+ shows "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> f -- a --> L"
+by (drule_tac k="- a" in LIM_offset, simp)
+
+lemma LIM_zero:
+ fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
+ shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. f x - l) ---> 0) F"
+unfolding tendsto_iff dist_norm by simp
+
+lemma LIM_zero_cancel:
+ fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
+ shows "((\<lambda>x. f x - l) ---> 0) F \<Longrightarrow> (f ---> l) F"
+unfolding tendsto_iff dist_norm by simp
+
+lemma LIM_zero_iff:
+ fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+ shows "((\<lambda>x. f x - l) ---> 0) F = (f ---> l) F"
+unfolding tendsto_iff dist_norm by simp
+
+lemma LIM_imp_LIM:
+ fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
+ fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector"
+ assumes f: "f -- a --> l"
+ assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"
+ shows "g -- a --> m"
+ by (rule metric_LIM_imp_LIM [OF f],
+ simp add: dist_norm le)
+
+lemma LIM_equal2:
+ fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
+ assumes 1: "0 < R"
+ assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"
+ shows "g -- a --> l \<Longrightarrow> f -- a --> l"
+by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm)
+
+lemma LIM_compose2:
+ fixes a :: "'a::real_normed_vector"
+ assumes f: "f -- a --> b"
+ assumes g: "g -- b --> c"
+ assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"
+ shows "(\<lambda>x. g (f x)) -- a --> c"
+by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])
+
+lemma real_LIM_sandwich_zero:
+ fixes f g :: "'a::topological_space \<Rightarrow> real"
+ assumes f: "f -- a --> 0"
+ assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"
+ assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"
+ shows "g -- a --> 0"
+proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *)
+ fix x assume x: "x \<noteq> a"
+ have "norm (g x - 0) = g x" by (simp add: 1 x)
+ also have "g x \<le> f x" by (rule 2 [OF x])
+ also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self)
+ also have "\<bar>f x\<bar> = norm (f x - 0)" by simp
+ finally show "norm (g x - 0) \<le> norm (f x - 0)" .
+qed
+
+
+subsection {* Continuity *}
+
+lemma LIM_isCont_iff:
+ fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
+ shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)"
+by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])
+
+lemma isCont_iff:
+ fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
+ shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x"
+by (simp add: isCont_def LIM_isCont_iff)
+
+lemma isCont_LIM_compose2:
+ fixes a :: "'a::real_normed_vector"
+ assumes f [unfolded isCont_def]: "isCont f a"
+ assumes g: "g -- f a --> l"
+ assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"
+ shows "(\<lambda>x. g (f x)) -- a --> l"
+by (rule LIM_compose2 [OF f g inj])
+
+
+lemma isCont_norm [simp]:
+ fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+ shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"
+ by (fact continuous_norm)
+
+lemma isCont_rabs [simp]:
+ fixes f :: "'a::t2_space \<Rightarrow> real"
+ shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
+ by (fact continuous_rabs)
+
+lemma isCont_add [simp]:
+ fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+ shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"
+ by (fact continuous_add)
+
+lemma isCont_minus [simp]:
+ fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+ shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"
+ by (fact continuous_minus)
+
+lemma isCont_diff [simp]:
+ fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+ shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"
+ by (fact continuous_diff)
+
+lemma isCont_mult [simp]:
+ fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
+ shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"
+ by (fact continuous_mult)
+
+lemma (in bounded_linear) isCont:
+ "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
+ by (fact continuous)
+
+lemma (in bounded_bilinear) isCont:
+ "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"
+ by (fact continuous)
+
+lemmas isCont_scaleR [simp] =
+ bounded_bilinear.isCont [OF bounded_bilinear_scaleR]
+
+lemmas isCont_of_real [simp] =
+ bounded_linear.isCont [OF bounded_linear_of_real]
+
+lemma isCont_power [simp]:
+ fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
+ shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"
+ by (fact continuous_power)
+
+lemma isCont_setsum [simp]:
+ fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
+ shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
+ by (auto intro: continuous_setsum)
+
+lemmas isCont_intros =
+ isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus
+ isCont_diff isCont_mult isCont_inverse isCont_divide isCont_scaleR
+ isCont_of_real isCont_power isCont_sgn isCont_setsum
+
+subsection {* Uniform Continuity *}
+
+definition
+ isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where
+ "isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"
+
+lemma isUCont_isCont: "isUCont f ==> isCont f x"
+by (simp add: isUCont_def isCont_def LIM_def, force)
+
+lemma isUCont_Cauchy:
+ "\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
+unfolding isUCont_def
+apply (rule metric_CauchyI)
+apply (drule_tac x=e in spec, safe)
+apply (drule_tac e=s in metric_CauchyD, safe)
+apply (rule_tac x=M in exI, simp)
+done
+
+lemma (in bounded_linear) isUCont: "isUCont f"
+unfolding isUCont_def dist_norm
+proof (intro allI impI)
+ fix r::real assume r: "0 < r"
+ obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"
+ using pos_bounded by fast
+ show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"
+ proof (rule exI, safe)
+ from r K show "0 < r / K" by (rule divide_pos_pos)
+ next
+ fix x y :: 'a
+ assume xy: "norm (x - y) < r / K"
+ have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff)
+ also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le)
+ also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq)
+ finally show "norm (f x - f y) < r" .
+ qed
+qed
+
+lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
+by (rule isUCont [THEN isUCont_Cauchy])
+
+lemma LIM_less_bound:
+ fixes f :: "real \<Rightarrow> real"
+ assumes ev: "b < x" "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and "isCont f x"
+ shows "0 \<le> f x"
+proof (rule tendsto_le_const)
+ show "(f ---> f x) (at_left x)"
+ using `isCont f x` by (simp add: filterlim_at_split isCont_def)
+ show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)"
+ using ev by (auto simp: eventually_within_less dist_real_def intro!: exI[of _ "x - b"])
+qed simp
+
+
+subsection {* Nested Intervals and Bisection -- Needed for Compactness *}
+
+lemma nested_sequence_unique:
+ 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"
+ shows "\<exists>l::real. ((\<forall>n. f n \<le> l) \<and> f ----> l) \<and> ((\<forall>n. l \<le> g n) \<and> g ----> l)"
+proof -
+ have "incseq f" unfolding incseq_Suc_iff by fact
+ have "decseq g" unfolding decseq_Suc_iff by fact
+
+ { fix n
+ from `decseq g` have "g n \<le> g 0" by (rule decseqD) simp
+ with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f n \<le> g 0" by auto }
+ then obtain u where "f ----> u" "\<forall>i. f i \<le> u"
+ using incseq_convergent[OF `incseq f`] by auto
+ moreover
+ { fix n
+ from `incseq f` have "f 0 \<le> f n" by (rule incseqD) simp
+ with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f 0 \<le> g n" by simp }
+ then obtain l where "g ----> l" "\<forall>i. l \<le> g i"
+ using decseq_convergent[OF `decseq g`] by auto
+ moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF `f ----> u` `g ----> l`]]
+ ultimately show ?thesis by auto
+qed
+
+lemma Bolzano[consumes 1, case_names trans local]:
+ fixes P :: "real \<Rightarrow> real \<Rightarrow> bool"
+ assumes [arith]: "a \<le> b"
+ assumes trans: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a c"
+ 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"
+ shows "P a b"
+proof -
+ 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))"
+ def l \<equiv> "\<lambda>n. fst (bisect n)" and u \<equiv> "\<lambda>n. snd (bisect n)"
+ 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)"
+ 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)"
+ by (simp_all add: l_def u_def bisect_def split: prod.split)
+
+ { fix n have "l n \<le> u n" by (induct n) auto } note this[simp]
+
+ have "\<exists>x. ((\<forall>n. l n \<le> x) \<and> l ----> x) \<and> ((\<forall>n. x \<le> u n) \<and> u ----> x)"
+ proof (safe intro!: nested_sequence_unique)
+ fix n show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" by (induct n) auto
+ next
+ { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) }
+ then show "(\<lambda>n. l n - u n) ----> 0" by (simp add: LIMSEQ_divide_realpow_zero)
+ qed fact
+ then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l ----> x" "u ----> x" by auto
+ 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"
+ using `l 0 \<le> x` `x \<le> u 0` local[of x] by auto
+
+ show "P a b"
+ proof (rule ccontr)
+ assume "\<not> P a b"
+ { fix n have "\<not> P (l n) (u n)"
+ proof (induct n)
+ case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto
+ qed (simp add: `\<not> P a b`) }
+ moreover
+ { have "eventually (\<lambda>n. x - d / 2 < l n) sequentially"
+ using `0 < d` `l ----> x` by (intro order_tendstoD[of _ x]) auto
+ moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially"
+ using `0 < d` `u ----> x` by (intro order_tendstoD[of _ x]) auto
+ ultimately have "eventually (\<lambda>n. P (l n) (u n)) sequentially"
+ proof eventually_elim
+ fix n assume "x - d / 2 < l n" "u n < x + d / 2"
+ from add_strict_mono[OF this] have "u n - l n < d" by simp
+ with x show "P (l n) (u n)" by (rule d)
+ qed }
+ ultimately show False by simp
+ qed
+qed
+
+lemma compact_Icc[simp, intro]: "compact {a .. b::real}"
+proof (cases "a \<le> b", rule compactI)
+ fix C assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C"
+ def T == "{a .. b}"
+ from C(1,3) show "\<exists>C'\<subseteq>C. finite C' \<and> {a..b} \<subseteq> \<Union>C'"
+ proof (induct rule: Bolzano)
+ case (trans a b c)
+ then have *: "{a .. c} = {a .. b} \<union> {b .. c}" by auto
+ 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"
+ by (auto simp: *)
+ with trans show ?case
+ unfolding * by (intro exI[of _ "C1 \<union> C2"]) auto
+ next
+ case (local x)
+ then have "x \<in> \<Union>C" using C by auto
+ with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" by auto
+ then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c"
+ by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff)
+ with `c \<in> C` show ?case
+ by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto
+ qed
+qed simp
+
+
+subsection {* Boundedness of continuous functions *}
+
+text{*By bisection, function continuous on closed interval is bounded above*}
+
+lemma isCont_eq_Ub:
+ fixes f :: "real \<Rightarrow> 'a::linorder_topology"
+ shows "a \<le> b \<Longrightarrow> \<forall>x::real. 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) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
+ using continuous_attains_sup[of "{a .. b}" f]
+ by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
+
+lemma isCont_eq_Lb:
+ fixes f :: "real \<Rightarrow> 'a::linorder_topology"
+ 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> M \<le> f x) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
+ using continuous_attains_inf[of "{a .. b}" f]
+ by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
+
+lemma isCont_bounded:
+ fixes f :: "real \<Rightarrow> 'a::linorder_topology"
+ 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"
+ using isCont_eq_Ub[of a b f] by auto
+
+lemma isCont_has_Ub:
+ fixes f :: "real \<Rightarrow> 'a::linorder_topology"
+ 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) \<and> (\<forall>N. N < M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> N < f x))"
+ using isCont_eq_Ub[of a b f] by auto
+
+(*HOL style here: object-level formulations*)
+lemma IVT_objl: "(f(a::real) \<le> (y::real) & y \<le> f(b) & a \<le> b &
+ (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
+ --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
+ by (blast intro: IVT)
+
+lemma IVT2_objl: "(f(b::real) \<le> (y::real) & y \<le> f(a) & a \<le> b &
+ (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
+ --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
+ by (blast intro: IVT2)
+
+lemma isCont_Lb_Ub:
+ fixes f :: "real \<Rightarrow> real"
+ assumes "a \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x"
+ shows "\<exists>L M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> L \<le> f x \<and> f x \<le> M) \<and>
+ (\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> (f x = y)))"
+proof -
+ 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"
+ using isCont_eq_Ub[OF assms] by auto
+ 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"
+ using isCont_eq_Lb[OF assms] by auto
+ show ?thesis
+ using IVT[of f L _ M] IVT2[of f L _ M] M L assms
+ apply (rule_tac x="f L" in exI)
+ apply (rule_tac x="f M" in exI)
+ apply (cases "L \<le> M")
+ apply (simp, metis order_trans)
+ apply (simp, metis order_trans)
+ done
+qed
+
+
+text{*Continuity of inverse function*}
+
+lemma isCont_inverse_function:
+ fixes f g :: "real \<Rightarrow> real"
+ assumes d: "0 < d"
+ and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> g (f z) = z"
+ and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> isCont f z"
+ shows "isCont g (f x)"
+proof -
+ let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + d}"
+
+ have f: "continuous_on ?D f"
+ using cont by (intro continuous_at_imp_continuous_on ballI) auto
+ then have g: "continuous_on (f`?D) g"
+ using inj by (intro continuous_on_inv) auto
+
+ from d f have "{min ?A ?B <..< max ?A ?B} \<subseteq> f ` ?D"
+ by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max)
+ with g have "continuous_on {min ?A ?B <..< max ?A ?B} g"
+ by (rule continuous_on_subset)
+ moreover
+ have "(?A < f x \<and> f x < ?B) \<or> (?B < f x \<and> f x < ?A)"
+ using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto
+ then have "f x \<in> {min ?A ?B <..< max ?A ?B}"
+ by auto
+ ultimately
+ show ?thesis
+ by (simp add: continuous_on_eq_continuous_at)
+qed
+
+lemma isCont_inverse_function2:
+ fixes f g :: "real \<Rightarrow> real" shows
+ "\<lbrakk>a < x; x < b;
+ \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z;
+ \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z\<rbrakk>
+ \<Longrightarrow> isCont g (f x)"
+apply (rule isCont_inverse_function
+ [where f=f and d="min (x - a) (b - x)"])
+apply (simp_all add: abs_le_iff)
+done
+
+(* need to rename second isCont_inverse *)
+
+lemma isCont_inv_fun:
+ fixes f g :: "real \<Rightarrow> real"
+ shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z;
+ \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |]
+ ==> isCont g (f x)"
+by (rule isCont_inverse_function)
+
+text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*}
+lemma LIM_fun_gt_zero:
+ fixes f :: "real \<Rightarrow> real"
+ 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)"
+apply (drule (1) LIM_D, clarify)
+apply (rule_tac x = s in exI)
+apply (simp add: abs_less_iff)
+done
+
+lemma LIM_fun_less_zero:
+ fixes f :: "real \<Rightarrow> real"
+ 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)"
+apply (drule LIM_D [where r="-l"], simp, clarify)
+apply (rule_tac x = s in exI)
+apply (simp add: abs_less_iff)
+done
+
+lemma LIM_fun_not_zero:
+ fixes f :: "real \<Rightarrow> real"
+ 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)"
+ using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp add: neq_iff)
end
--- a/src/HOL/Ln.thy Tue Mar 26 14:38:44 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,229 +0,0 @@
-(* Title: HOL/Ln.thy
- Author: Jeremy Avigad
-*)
-
-header {* Properties of ln *}
-
-theory Ln
-imports Transcendental
-begin
-
-lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==>
- x - x^2 <= ln (1 + x)"
-proof -
- assume a: "0 <= x" and b: "x <= 1"
- have "exp (x - x^2) = exp x / exp (x^2)"
- by (rule exp_diff)
- also have "... <= (1 + x + x^2) / exp (x ^2)"
- apply (rule divide_right_mono)
- apply (rule exp_bound)
- apply (rule a, rule b)
- apply simp
- done
- also have "... <= (1 + x + x^2) / (1 + x^2)"
- apply (rule divide_left_mono)
- apply (simp add: exp_ge_add_one_self_aux)
- apply (simp add: a)
- apply (simp add: mult_pos_pos add_pos_nonneg)
- done
- also from a have "... <= 1 + x"
- by (simp add: field_simps add_strict_increasing zero_le_mult_iff)
- finally have "exp (x - x^2) <= 1 + x" .
- also have "... = exp (ln (1 + x))"
- proof -
- from a have "0 < 1 + x" by auto
- thus ?thesis
- by (auto simp only: exp_ln_iff [THEN sym])
- qed
- finally have "exp (x - x ^ 2) <= exp (ln (1 + x))" .
- thus ?thesis by (auto simp only: exp_le_cancel_iff)
-qed
-
-lemma aux5: "x < 1 ==> ln(1 - x) = - ln(1 + x / (1 - x))"
-proof -
- assume a: "x < 1"
- have "ln(1 - x) = - ln(1 / (1 - x))"
- proof -
- have "ln(1 - x) = - (- ln (1 - x))"
- by auto
- also have "- ln(1 - x) = ln 1 - ln(1 - x)"
- by simp
- also have "... = ln(1 / (1 - x))"
- apply (rule ln_div [THEN sym])
- by (insert a, auto)
- finally show ?thesis .
- qed
- also have " 1 / (1 - x) = 1 + x / (1 - x)" using a by(simp add:field_simps)
- finally show ?thesis .
-qed
-
-lemma ln_one_minus_pos_lower_bound: "0 <= x ==> x <= (1 / 2) ==>
- - x - 2 * x^2 <= ln (1 - x)"
-proof -
- assume a: "0 <= x" and b: "x <= (1 / 2)"
- from b have c: "x < 1"
- by auto
- then have "ln (1 - x) = - ln (1 + x / (1 - x))"
- by (rule aux5)
- also have "- (x / (1 - x)) <= ..."
- proof -
- have "ln (1 + x / (1 - x)) <= x / (1 - x)"
- apply (rule ln_add_one_self_le_self)
- apply (rule divide_nonneg_pos)
- by (insert a c, auto)
- thus ?thesis
- by auto
- qed
- also have "- (x / (1 - x)) = -x / (1 - x)"
- by auto
- finally have d: "- x / (1 - x) <= ln (1 - x)" .
- have "0 < 1 - x" using a b by simp
- hence e: "-x - 2 * x^2 <= - x / (1 - x)"
- using mult_right_le_one_le[of "x*x" "2*x"] a b
- by (simp add:field_simps power2_eq_square)
- from e d show "- x - 2 * x^2 <= ln (1 - x)"
- by (rule order_trans)
-qed
-
-lemma ln_add_one_self_le_self2: "-1 < x ==> ln(1 + x) <= x"
- apply (subgoal_tac "ln (1 + x) \<le> ln (exp x)", simp)
- apply (subst ln_le_cancel_iff)
- apply auto
-done
-
-lemma abs_ln_one_plus_x_minus_x_bound_nonneg:
- "0 <= x ==> x <= 1 ==> abs(ln (1 + x) - x) <= x^2"
-proof -
- assume x: "0 <= x"
- assume x1: "x <= 1"
- from x have "ln (1 + x) <= x"
- by (rule ln_add_one_self_le_self)
- then have "ln (1 + x) - x <= 0"
- by simp
- then have "abs(ln(1 + x) - x) = - (ln(1 + x) - x)"
- by (rule abs_of_nonpos)
- also have "... = x - ln (1 + x)"
- by simp
- also have "... <= x^2"
- proof -
- from x x1 have "x - x^2 <= ln (1 + x)"
- by (intro ln_one_plus_pos_lower_bound)
- thus ?thesis
- by simp
- qed
- finally show ?thesis .
-qed
-
-lemma abs_ln_one_plus_x_minus_x_bound_nonpos:
- "-(1 / 2) <= x ==> x <= 0 ==> abs(ln (1 + x) - x) <= 2 * x^2"
-proof -
- assume a: "-(1 / 2) <= x"
- assume b: "x <= 0"
- have "abs(ln (1 + x) - x) = x - ln(1 - (-x))"
- apply (subst abs_of_nonpos)
- apply simp
- apply (rule ln_add_one_self_le_self2)
- using a apply auto
- done
- also have "... <= 2 * x^2"
- apply (subgoal_tac "- (-x) - 2 * (-x)^2 <= ln (1 - (-x))")
- apply (simp add: algebra_simps)
- apply (rule ln_one_minus_pos_lower_bound)
- using a b apply auto
- done
- finally show ?thesis .
-qed
-
-lemma abs_ln_one_plus_x_minus_x_bound:
- "abs x <= 1 / 2 ==> abs(ln (1 + x) - x) <= 2 * x^2"
- apply (case_tac "0 <= x")
- apply (rule order_trans)
- apply (rule abs_ln_one_plus_x_minus_x_bound_nonneg)
- apply auto
- apply (rule abs_ln_one_plus_x_minus_x_bound_nonpos)
- apply auto
-done
-
-lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)"
-proof -
- assume x: "exp 1 <= x" "x <= y"
- moreover have "0 < exp (1::real)" by simp
- ultimately have a: "0 < x" and b: "0 < y"
- by (fast intro: less_le_trans order_trans)+
- have "x * ln y - x * ln x = x * (ln y - ln x)"
- by (simp add: algebra_simps)
- also have "... = x * ln(y / x)"
- by (simp only: ln_div a b)
- also have "y / x = (x + (y - x)) / x"
- by simp
- also have "... = 1 + (y - x) / x"
- using x a by (simp add: field_simps)
- also have "x * ln(1 + (y - x) / x) <= x * ((y - x) / x)"
- apply (rule mult_left_mono)
- apply (rule ln_add_one_self_le_self)
- apply (rule divide_nonneg_pos)
- using x a apply simp_all
- done
- also have "... = y - x" using a by simp
- also have "... = (y - x) * ln (exp 1)" by simp
- also have "... <= (y - x) * ln x"
- apply (rule mult_left_mono)
- apply (subst ln_le_cancel_iff)
- apply fact
- apply (rule a)
- apply (rule x)
- using x apply simp
- done
- also have "... = y * ln x - x * ln x"
- by (rule left_diff_distrib)
- finally have "x * ln y <= y * ln x"
- by arith
- then have "ln y <= (y * ln x) / x" using a by (simp add: field_simps)
- also have "... = y * (ln x / x)" by simp
- finally show ?thesis using b by (simp add: field_simps)
-qed
-
-lemma ln_le_minus_one:
- "0 < x \<Longrightarrow> ln x \<le> x - 1"
- using exp_ge_add_one_self[of "ln x"] by simp
-
-lemma ln_eq_minus_one:
- assumes "0 < x" "ln x = x - 1" shows "x = 1"
-proof -
- let "?l y" = "ln y - y + 1"
- have D: "\<And>x. 0 < x \<Longrightarrow> DERIV ?l x :> (1 / x - 1)"
- by (auto intro!: DERIV_intros)
-
- show ?thesis
- proof (cases rule: linorder_cases)
- assume "x < 1"
- from dense[OF `x < 1`] obtain a where "x < a" "a < 1" by blast
- from `x < a` have "?l x < ?l a"
- proof (rule DERIV_pos_imp_increasing, safe)
- fix y assume "x \<le> y" "y \<le> a"
- with `0 < x` `a < 1` have "0 < 1 / y - 1" "0 < y"
- by (auto simp: field_simps)
- with D show "\<exists>z. DERIV ?l y :> z \<and> 0 < z"
- by auto
- qed
- also have "\<dots> \<le> 0"
- using ln_le_minus_one `0 < x` `x < a` by (auto simp: field_simps)
- finally show "x = 1" using assms by auto
- next
- assume "1 < x"
- from dense[OF `1 < x`] obtain a where "1 < a" "a < x" by blast
- from `a < x` have "?l x < ?l a"
- proof (rule DERIV_neg_imp_decreasing, safe)
- fix y assume "a \<le> y" "y \<le> x"
- with `1 < a` have "1 / y - 1 < 0" "0 < y"
- by (auto simp: field_simps)
- with D show "\<exists>z. DERIV ?l y :> z \<and> z < 0"
- by blast
- qed
- also have "\<dots> \<le> 0"
- using ln_le_minus_one `1 < a` by (auto simp: field_simps)
- finally show "x = 1" using assms by auto
- qed simp
-qed
-
-end
--- a/src/HOL/Log.thy Tue Mar 26 14:38:44 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,421 +0,0 @@
-(* Title : Log.thy
- Author : Jacques D. Fleuriot
- Additional contributions by Jeremy Avigad
- Copyright : 2000,2001 University of Edinburgh
-*)
-
-header{*Logarithms: Standard Version*}
-
-theory Log
-imports Transcendental
-begin
-
-definition
- powr :: "[real,real] => real" (infixr "powr" 80) where
- --{*exponentation with real exponent*}
- "x powr a = exp(a * ln x)"
-
-definition
- log :: "[real,real] => real" where
- --{*logarithm of @{term x} to base @{term a}*}
- "log a x = ln x / ln a"
-
-
-lemma tendsto_log [tendsto_intros]:
- "\<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"
- unfolding log_def by (intro tendsto_intros) auto
-
-lemma continuous_log:
- 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))"
- shows "continuous F (\<lambda>x. log (f x) (g x))"
- using assms unfolding continuous_def by (rule tendsto_log)
-
-lemma continuous_at_within_log[continuous_intros]:
- 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"
- shows "continuous (at a within s) (\<lambda>x. log (f x) (g x))"
- using assms unfolding continuous_within by (rule tendsto_log)
-
-lemma isCont_log[continuous_intros, simp]:
- assumes "isCont f a" "isCont g a" "0 < f a" "f a \<noteq> 1" "0 < g a"
- shows "isCont (\<lambda>x. log (f x) (g x)) a"
- using assms unfolding continuous_at by (rule tendsto_log)
-
-lemma continuous_on_log[continuous_on_intros]:
- 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"
- shows "continuous_on s (\<lambda>x. log (f x) (g x))"
- using assms unfolding continuous_on_def by (fast intro: tendsto_log)
-
-lemma powr_one_eq_one [simp]: "1 powr a = 1"
-by (simp add: powr_def)
-
-lemma powr_zero_eq_one [simp]: "x powr 0 = 1"
-by (simp add: powr_def)
-
-lemma powr_one_gt_zero_iff [simp]: "(x powr 1 = x) = (0 < x)"
-by (simp add: powr_def)
-declare powr_one_gt_zero_iff [THEN iffD2, simp]
-
-lemma powr_mult:
- "[| 0 < x; 0 < y |] ==> (x * y) powr a = (x powr a) * (y powr a)"
-by (simp add: powr_def exp_add [symmetric] ln_mult distrib_left)
-
-lemma powr_gt_zero [simp]: "0 < x powr a"
-by (simp add: powr_def)
-
-lemma powr_ge_pzero [simp]: "0 <= x powr y"
-by (rule order_less_imp_le, rule powr_gt_zero)
-
-lemma powr_not_zero [simp]: "x powr a \<noteq> 0"
-by (simp add: powr_def)
-
-lemma powr_divide:
- "[| 0 < x; 0 < y |] ==> (x / y) powr a = (x powr a)/(y powr a)"
-apply (simp add: divide_inverse positive_imp_inverse_positive powr_mult)
-apply (simp add: powr_def exp_minus [symmetric] exp_add [symmetric] ln_inverse)
-done
-
-lemma powr_divide2: "x powr a / x powr b = x powr (a - b)"
- apply (simp add: powr_def)
- apply (subst exp_diff [THEN sym])
- apply (simp add: left_diff_distrib)
-done
-
-lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)"
-by (simp add: powr_def exp_add [symmetric] distrib_right)
-
-lemma powr_mult_base:
- "0 < x \<Longrightarrow>x * x powr y = x powr (1 + y)"
-using assms by (auto simp: powr_add)
-
-lemma powr_powr: "(x powr a) powr b = x powr (a * b)"
-by (simp add: powr_def)
-
-lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a"
-by (simp add: powr_powr mult_commute)
-
-lemma powr_minus: "x powr (-a) = inverse (x powr a)"
-by (simp add: powr_def exp_minus [symmetric])
-
-lemma powr_minus_divide: "x powr (-a) = 1/(x powr a)"
-by (simp add: divide_inverse powr_minus)
-
-lemma powr_less_mono: "[| a < b; 1 < x |] ==> x powr a < x powr b"
-by (simp add: powr_def)
-
-lemma powr_less_cancel: "[| x powr a < x powr b; 1 < x |] ==> a < b"
-by (simp add: powr_def)
-
-lemma powr_less_cancel_iff [simp]: "1 < x ==> (x powr a < x powr b) = (a < b)"
-by (blast intro: powr_less_cancel powr_less_mono)
-
-lemma powr_le_cancel_iff [simp]: "1 < x ==> (x powr a \<le> x powr b) = (a \<le> b)"
-by (simp add: linorder_not_less [symmetric])
-
-lemma log_ln: "ln x = log (exp(1)) x"
-by (simp add: log_def)
-
-lemma DERIV_log: assumes "x > 0" shows "DERIV (\<lambda>y. log b y) x :> 1 / (ln b * x)"
-proof -
- def lb \<equiv> "1 / ln b"
- moreover have "DERIV (\<lambda>y. lb * ln y) x :> lb / x"
- using `x > 0` by (auto intro!: DERIV_intros)
- ultimately show ?thesis
- by (simp add: log_def)
-qed
-
-lemmas DERIV_log[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
-
-lemma powr_log_cancel [simp]:
- "[| 0 < a; a \<noteq> 1; 0 < x |] ==> a powr (log a x) = x"
-by (simp add: powr_def log_def)
-
-lemma log_powr_cancel [simp]: "[| 0 < a; a \<noteq> 1 |] ==> log a (a powr y) = y"
-by (simp add: log_def powr_def)
-
-lemma log_mult:
- "[| 0 < a; a \<noteq> 1; 0 < x; 0 < y |]
- ==> log a (x * y) = log a x + log a y"
-by (simp add: log_def ln_mult divide_inverse distrib_right)
-
-lemma log_eq_div_ln_mult_log:
- "[| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |]
- ==> log a x = (ln b/ln a) * log b x"
-by (simp add: log_def divide_inverse)
-
-text{*Base 10 logarithms*}
-lemma log_base_10_eq1: "0 < x ==> log 10 x = (ln (exp 1) / ln 10) * ln x"
-by (simp add: log_def)
-
-lemma log_base_10_eq2: "0 < x ==> log 10 x = (log 10 (exp 1)) * ln x"
-by (simp add: log_def)
-
-lemma log_one [simp]: "log a 1 = 0"
-by (simp add: log_def)
-
-lemma log_eq_one [simp]: "[| 0 < a; a \<noteq> 1 |] ==> log a a = 1"
-by (simp add: log_def)
-
-lemma log_inverse:
- "[| 0 < a; a \<noteq> 1; 0 < x |] ==> log a (inverse x) = - log a x"
-apply (rule_tac a1 = "log a x" in add_left_cancel [THEN iffD1])
-apply (simp add: log_mult [symmetric])
-done
-
-lemma log_divide:
- "[|0 < a; a \<noteq> 1; 0 < x; 0 < y|] ==> log a (x/y) = log a x - log a y"
-by (simp add: log_mult divide_inverse log_inverse)
-
-lemma log_less_cancel_iff [simp]:
- "[| 1 < a; 0 < x; 0 < y |] ==> (log a x < log a y) = (x < y)"
-apply safe
-apply (rule_tac [2] powr_less_cancel)
-apply (drule_tac a = "log a x" in powr_less_mono, auto)
-done
-
-lemma log_inj: assumes "1 < b" shows "inj_on (log b) {0 <..}"
-proof (rule inj_onI, simp)
- fix x y assume pos: "0 < x" "0 < y" and *: "log b x = log b y"
- show "x = y"
- proof (cases rule: linorder_cases)
- assume "x < y" hence "log b x < log b y"
- using log_less_cancel_iff[OF `1 < b`] pos by simp
- thus ?thesis using * by simp
- next
- assume "y < x" hence "log b y < log b x"
- using log_less_cancel_iff[OF `1 < b`] pos by simp
- thus ?thesis using * by simp
- qed simp
-qed
-
-lemma log_le_cancel_iff [simp]:
- "[| 1 < a; 0 < x; 0 < y |] ==> (log a x \<le> log a y) = (x \<le> y)"
-by (simp add: linorder_not_less [symmetric])
-
-lemma zero_less_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 < log a x \<longleftrightarrow> 1 < x"
- using log_less_cancel_iff[of a 1 x] by simp
-
-lemma zero_le_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 \<le> log a x \<longleftrightarrow> 1 \<le> x"
- using log_le_cancel_iff[of a 1 x] by simp
-
-lemma log_less_zero_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x < 0 \<longleftrightarrow> x < 1"
- using log_less_cancel_iff[of a x 1] by simp
-
-lemma log_le_zero_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x \<le> 0 \<longleftrightarrow> x \<le> 1"
- using log_le_cancel_iff[of a x 1] by simp
-
-lemma one_less_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 1 < log a x \<longleftrightarrow> a < x"
- using log_less_cancel_iff[of a a x] by simp
-
-lemma one_le_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 1 \<le> log a x \<longleftrightarrow> a \<le> x"
- using log_le_cancel_iff[of a a x] by simp
-
-lemma log_less_one_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x < 1 \<longleftrightarrow> x < a"
- using log_less_cancel_iff[of a x a] by simp
-
-lemma log_le_one_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x \<le> 1 \<longleftrightarrow> x \<le> a"
- using log_le_cancel_iff[of a x a] by simp
-
-lemma powr_realpow: "0 < x ==> x powr (real n) = x^n"
- apply (induct n, simp)
- apply (subgoal_tac "real(Suc n) = real n + 1")
- apply (erule ssubst)
- apply (subst powr_add, simp, simp)
-done
-
-lemma powr_realpow2: "0 <= x ==> 0 < n ==> x^n = (if (x = 0) then 0 else x powr (real n))"
- apply (case_tac "x = 0", simp, simp)
- apply (rule powr_realpow [THEN sym], simp)
-done
-
-lemma powr_int:
- assumes "x > 0"
- shows "x powr i = (if i \<ge> 0 then x ^ nat i else 1 / x ^ nat (-i))"
-proof cases
- assume "i < 0"
- have r: "x powr i = 1 / x powr (-i)" by (simp add: powr_minus field_simps)
- show ?thesis using `i < 0` `x > 0` by (simp add: r field_simps powr_realpow[symmetric])
-qed (simp add: assms powr_realpow[symmetric])
-
-lemma powr_numeral: "0 < x \<Longrightarrow> x powr numeral n = x^numeral n"
- using powr_realpow[of x "numeral n"] by simp
-
-lemma powr_neg_numeral: "0 < x \<Longrightarrow> x powr neg_numeral n = 1 / x^numeral n"
- using powr_int[of x "neg_numeral n"] by simp
-
-lemma root_powr_inverse:
- "0 < n \<Longrightarrow> 0 < x \<Longrightarrow> root n x = x powr (1/n)"
- by (rule real_root_pos_unique) (auto simp: powr_realpow[symmetric] powr_powr)
-
-lemma ln_powr: "0 < x ==> 0 < y ==> ln(x powr y) = y * ln x"
-by (unfold powr_def, simp)
-
-lemma log_powr: "0 < x ==> 0 \<le> y ==> log b (x powr y) = y * log b x"
- apply (case_tac "y = 0")
- apply force
- apply (auto simp add: log_def ln_powr field_simps)
-done
-
-lemma log_nat_power: "0 < x ==> log b (x^n) = real n * log b x"
- apply (subst powr_realpow [symmetric])
- apply (auto simp add: log_powr)
-done
-
-lemma ln_bound: "1 <= x ==> ln x <= x"
- apply (subgoal_tac "ln(1 + (x - 1)) <= x - 1")
- apply simp
- apply (rule ln_add_one_self_le_self, simp)
-done
-
-lemma powr_mono: "a <= b ==> 1 <= x ==> x powr a <= x powr b"
- apply (case_tac "x = 1", simp)
- apply (case_tac "a = b", simp)
- apply (rule order_less_imp_le)
- apply (rule powr_less_mono, auto)
-done
-
-lemma ge_one_powr_ge_zero: "1 <= x ==> 0 <= a ==> 1 <= x powr a"
- apply (subst powr_zero_eq_one [THEN sym])
- apply (rule powr_mono, assumption+)
-done
-
-lemma powr_less_mono2: "0 < a ==> 0 < x ==> x < y ==> x powr a <
- y powr a"
- apply (unfold powr_def)
- apply (rule exp_less_mono)
- apply (rule mult_strict_left_mono)
- apply (subst ln_less_cancel_iff, assumption)
- apply (rule order_less_trans)
- prefer 2
- apply assumption+
-done
-
-lemma powr_less_mono2_neg: "a < 0 ==> 0 < x ==> x < y ==> y powr a <
- x powr a"
- apply (unfold powr_def)
- apply (rule exp_less_mono)
- apply (rule mult_strict_left_mono_neg)
- apply (subst ln_less_cancel_iff)
- apply assumption
- apply (rule order_less_trans)
- prefer 2
- apply assumption+
-done
-
-lemma powr_mono2: "0 <= a ==> 0 < x ==> x <= y ==> x powr a <= y powr a"
- apply (case_tac "a = 0", simp)
- apply (case_tac "x = y", simp)
- apply (rule order_less_imp_le)
- apply (rule powr_less_mono2, auto)
-done
-
-lemma powr_inj:
- "0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> a powr x = a powr y \<longleftrightarrow> x = y"
- unfolding powr_def exp_inj_iff by simp
-
-lemma ln_powr_bound: "1 <= x ==> 0 < a ==> ln x <= (x powr a) / a"
- apply (rule mult_imp_le_div_pos)
- apply (assumption)
- apply (subst mult_commute)
- apply (subst ln_powr [THEN sym])
- apply auto
- apply (rule ln_bound)
- apply (erule ge_one_powr_ge_zero)
- apply (erule order_less_imp_le)
-done
-
-lemma ln_powr_bound2:
- assumes "1 < x" and "0 < a"
- shows "(ln x) powr a <= (a powr a) * x"
-proof -
- from assms have "ln x <= (x powr (1 / a)) / (1 / a)"
- apply (intro ln_powr_bound)
- apply (erule order_less_imp_le)
- apply (rule divide_pos_pos)
- apply simp_all
- done
- also have "... = a * (x powr (1 / a))"
- by simp
- finally have "(ln x) powr a <= (a * (x powr (1 / a))) powr a"
- apply (intro powr_mono2)
- apply (rule order_less_imp_le, rule assms)
- apply (rule ln_gt_zero)
- apply (rule assms)
- apply assumption
- done
- also have "... = (a powr a) * ((x powr (1 / a)) powr a)"
- apply (rule powr_mult)
- apply (rule assms)
- apply (rule powr_gt_zero)
- done
- also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)"
- by (rule powr_powr)
- also have "... = x"
- apply simp
- apply (subgoal_tac "a ~= 0")
- using assms apply auto
- done
- finally show ?thesis .
-qed
-
-lemma tendsto_powr [tendsto_intros]:
- "\<lbrakk>(f ---> a) F; (g ---> b) F; 0 < a\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x powr g x) ---> a powr b) F"
- unfolding powr_def by (intro tendsto_intros)
-
-lemma continuous_powr:
- assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\<lambda>x. x))"
- shows "continuous F (\<lambda>x. (f x) powr (g x))"
- using assms unfolding continuous_def by (rule tendsto_powr)
-
-lemma continuous_at_within_powr[continuous_intros]:
- assumes "continuous (at a within s) f" "continuous (at a within s) g" and "0 < f a"
- shows "continuous (at a within s) (\<lambda>x. (f x) powr (g x))"
- using assms unfolding continuous_within by (rule tendsto_powr)
-
-lemma isCont_powr[continuous_intros, simp]:
- assumes "isCont f a" "isCont g a" "0 < f a"
- shows "isCont (\<lambda>x. (f x) powr g x) a"
- using assms unfolding continuous_at by (rule tendsto_powr)
-
-lemma continuous_on_powr[continuous_on_intros]:
- assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. 0 < f x"
- shows "continuous_on s (\<lambda>x. (f x) powr (g x))"
- using assms unfolding continuous_on_def by (fast intro: tendsto_powr)
-
-(* FIXME: generalize by replacing d by with g x and g ---> d? *)
-lemma tendsto_zero_powrI:
- assumes "eventually (\<lambda>x. 0 < f x ) F" and "(f ---> 0) F"
- assumes "0 < d"
- shows "((\<lambda>x. f x powr d) ---> 0) F"
-proof (rule tendstoI)
- fix e :: real assume "0 < e"
- def Z \<equiv> "e powr (1 / d)"
- with `0 < e` have "0 < Z" by simp
- with assms have "eventually (\<lambda>x. 0 < f x \<and> dist (f x) 0 < Z) F"
- by (intro eventually_conj tendstoD)
- moreover
- from assms have "\<And>x. 0 < x \<and> dist x 0 < Z \<Longrightarrow> x powr d < Z powr d"
- by (intro powr_less_mono2) (auto simp: dist_real_def)
- with assms `0 < e` have "\<And>x. 0 < x \<and> dist x 0 < Z \<Longrightarrow> dist (x powr d) 0 < e"
- unfolding dist_real_def Z_def by (auto simp: powr_powr)
- ultimately
- show "eventually (\<lambda>x. dist (f x powr d) 0 < e) F" by (rule eventually_elim1)
-qed
-
-lemma tendsto_neg_powr:
- assumes "s < 0" and "LIM x F. f x :> at_top"
- shows "((\<lambda>x. f x powr s) ---> 0) F"
-proof (rule tendstoI)
- fix e :: real assume "0 < e"
- def Z \<equiv> "e powr (1 / s)"
- from assms have "eventually (\<lambda>x. Z < f x) F"
- by (simp add: filterlim_at_top_dense)
- moreover
- from assms have "\<And>x. Z < x \<Longrightarrow> x powr s < Z powr s"
- by (auto simp: Z_def intro!: powr_less_mono2_neg)
- with assms `0 < e` have "\<And>x. Z < x \<Longrightarrow> dist (x powr s) 0 < e"
- by (simp add: powr_powr Z_def dist_real_def)
- ultimately
- show "eventually (\<lambda>x. dist (f x powr s) 0 < e) F" by (rule eventually_elim1)
-qed
-
-end
--- a/src/HOL/Lubs.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Lubs.thy Tue Mar 26 15:10:28 2013 +0100
@@ -94,4 +94,10 @@
lemma isLub_ubs: "isLub R S x \<Longrightarrow> x <=* ubs R S"
unfolding ubs_def isLub_def by (rule leastPD2)
+lemma isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::'a::linorder)"
+ apply (frule isLub_isUb)
+ apply (frule_tac x = y in isLub_isUb)
+ apply (blast intro!: order_antisym dest!: isLub_le_isUb)
+ done
+
end
--- a/src/HOL/Metric_Spaces.thy Tue Mar 26 14:38:44 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,629 +0,0 @@
-(* Title: HOL/Metric_Spaces.thy
- Author: Brian Huffman
- Author: Johannes Hölzl
-*)
-
-header {* Metric Spaces *}
-
-theory Metric_Spaces
-imports RComplete Topological_Spaces
-begin
-
-
-subsection {* Metric spaces *}
-
-class dist =
- fixes dist :: "'a \<Rightarrow> 'a \<Rightarrow> real"
-
-class open_dist = "open" + dist +
- assumes open_dist: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
-
-class metric_space = open_dist +
- assumes dist_eq_0_iff [simp]: "dist x y = 0 \<longleftrightarrow> x = y"
- assumes dist_triangle2: "dist x y \<le> dist x z + dist y z"
-begin
-
-lemma dist_self [simp]: "dist x x = 0"
-by simp
-
-lemma zero_le_dist [simp]: "0 \<le> dist x y"
-using dist_triangle2 [of x x y] by simp
-
-lemma zero_less_dist_iff: "0 < dist x y \<longleftrightarrow> x \<noteq> y"
-by (simp add: less_le)
-
-lemma dist_not_less_zero [simp]: "\<not> dist x y < 0"
-by (simp add: not_less)
-
-lemma dist_le_zero_iff [simp]: "dist x y \<le> 0 \<longleftrightarrow> x = y"
-by (simp add: le_less)
-
-lemma dist_commute: "dist x y = dist y x"
-proof (rule order_antisym)
- show "dist x y \<le> dist y x"
- using dist_triangle2 [of x y x] by simp
- show "dist y x \<le> dist x y"
- using dist_triangle2 [of y x y] by simp
-qed
-
-lemma dist_triangle: "dist x z \<le> dist x y + dist y z"
-using dist_triangle2 [of x z y] by (simp add: dist_commute)
-
-lemma dist_triangle3: "dist x y \<le> dist a x + dist a y"
-using dist_triangle2 [of x y a] by (simp add: dist_commute)
-
-lemma dist_triangle_alt:
- shows "dist y z <= dist x y + dist x z"
-by (rule dist_triangle3)
-
-lemma dist_pos_lt:
- shows "x \<noteq> y ==> 0 < dist x y"
-by (simp add: zero_less_dist_iff)
-
-lemma dist_nz:
- shows "x \<noteq> y \<longleftrightarrow> 0 < dist x y"
-by (simp add: zero_less_dist_iff)
-
-lemma dist_triangle_le:
- shows "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e"
-by (rule order_trans [OF dist_triangle2])
-
-lemma dist_triangle_lt:
- shows "dist x z + dist y z < e ==> dist x y < e"
-by (rule le_less_trans [OF dist_triangle2])
-
-lemma dist_triangle_half_l:
- shows "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"
-by (rule dist_triangle_lt [where z=y], simp)
-
-lemma dist_triangle_half_r:
- shows "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"
-by (rule dist_triangle_half_l, simp_all add: dist_commute)
-
-subclass topological_space
-proof
- have "\<exists>e::real. 0 < e"
- by (fast intro: zero_less_one)
- then show "open UNIV"
- unfolding open_dist by simp
-next
- fix S T assume "open S" "open T"
- then show "open (S \<inter> T)"
- unfolding open_dist
- apply clarify
- apply (drule (1) bspec)+
- apply (clarify, rename_tac r s)
- apply (rule_tac x="min r s" in exI, simp)
- done
-next
- fix K assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)"
- unfolding open_dist by fast
-qed
-
-lemma open_ball: "open {y. dist x y < d}"
-proof (unfold open_dist, intro ballI)
- fix y assume *: "y \<in> {y. dist x y < d}"
- then show "\<exists>e>0. \<forall>z. dist z y < e \<longrightarrow> z \<in> {y. dist x y < d}"
- by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt)
-qed
-
-subclass first_countable_topology
-proof
- fix x
- 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))"
- proof (safe intro!: exI[of _ "\<lambda>n. {y. dist x y < inverse (Suc n)}"])
- fix S assume "open S" "x \<in> S"
- then obtain e where "0 < e" "{y. dist x y < e} \<subseteq> S"
- by (auto simp: open_dist subset_eq dist_commute)
- moreover
- then obtain i where "inverse (Suc i) < e"
- by (auto dest!: reals_Archimedean)
- then have "{y. dist x y < inverse (Suc i)} \<subseteq> {y. dist x y < e}"
- by auto
- ultimately show "\<exists>i. {y. dist x y < inverse (Suc i)} \<subseteq> S"
- by blast
- qed (auto intro: open_ball)
-qed
-
-end
-
-instance metric_space \<subseteq> t2_space
-proof
- fix x y :: "'a::metric_space"
- assume xy: "x \<noteq> y"
- let ?U = "{y'. dist x y' < dist x y / 2}"
- let ?V = "{x'. dist y x' < dist x y / 2}"
- 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
- \<Longrightarrow> \<not>(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
- have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
- using dist_pos_lt[OF xy] th0[of dist, OF dist_triangle dist_commute]
- using open_ball[of _ "dist x y / 2"] by auto
- then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
- by blast
-qed
-
-lemma eventually_nhds_metric:
- fixes a :: "'a :: metric_space"
- shows "eventually P (nhds a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. dist x a < d \<longrightarrow> P x)"
-unfolding eventually_nhds open_dist
-apply safe
-apply fast
-apply (rule_tac x="{x. dist x a < d}" in exI, simp)
-apply clarsimp
-apply (rule_tac x="d - dist x a" in exI, clarsimp)
-apply (simp only: less_diff_eq)
-apply (erule le_less_trans [OF dist_triangle])
-done
-
-lemma eventually_at:
- fixes a :: "'a::metric_space"
- shows "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> P x)"
-unfolding at_def eventually_within eventually_nhds_metric by auto
-
-lemma eventually_within_less: (* COPY FROM Topo/eventually_within *)
- fixes a :: "'a :: metric_space"
- 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)"
- unfolding eventually_within eventually_at dist_nz by auto
-
-lemma eventually_within_le: (* COPY FROM Topo/eventually_within_le *)
- fixes a :: "'a :: metric_space"
- 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)"
- unfolding eventually_within_less by auto (metis dense order_le_less_trans)
-
-lemma tendstoI:
- fixes l :: "'a :: metric_space"
- assumes "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F"
- shows "(f ---> l) F"
- apply (rule topological_tendstoI)
- apply (simp add: open_dist)
- apply (drule (1) bspec, clarify)
- apply (drule assms)
- apply (erule eventually_elim1, simp)
- done
-
-lemma tendstoD:
- fixes l :: "'a :: metric_space"
- shows "(f ---> l) F \<Longrightarrow> 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F"
- apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD)
- apply (clarsimp simp add: open_dist)
- apply (rule_tac x="e - dist x l" in exI, clarsimp)
- apply (simp only: less_diff_eq)
- apply (erule le_less_trans [OF dist_triangle])
- apply simp
- apply simp
- done
-
-lemma tendsto_iff:
- fixes l :: "'a :: metric_space"
- shows "(f ---> l) F \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) F)"
- using tendstoI tendstoD by fast
-
-lemma metric_tendsto_imp_tendsto:
- fixes a :: "'a :: metric_space" and b :: "'b :: metric_space"
- assumes f: "(f ---> a) F"
- assumes le: "eventually (\<lambda>x. dist (g x) b \<le> dist (f x) a) F"
- shows "(g ---> b) F"
-proof (rule tendstoI)
- fix e :: real assume "0 < e"
- with f have "eventually (\<lambda>x. dist (f x) a < e) F" by (rule tendstoD)
- with le show "eventually (\<lambda>x. dist (g x) b < e) F"
- using le_less_trans by (rule eventually_elim2)
-qed
-
-lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top"
- unfolding filterlim_at_top
- apply (intro allI)
- apply (rule_tac c="natceiling (Z + 1)" in eventually_sequentiallyI)
- apply (auto simp: natceiling_le_eq)
- done
-
-subsubsection {* Limits of Sequences *}
-
-lemma LIMSEQ_def: "X ----> (L::'a::metric_space) \<longleftrightarrow> (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)"
- unfolding tendsto_iff eventually_sequentially ..
-
-lemma LIMSEQ_iff_nz: "X ----> (L::'a::metric_space) = (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)"
- unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc)
-
-lemma metric_LIMSEQ_I:
- "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r) \<Longrightarrow> X ----> (L::'a::metric_space)"
-by (simp add: LIMSEQ_def)
-
-lemma metric_LIMSEQ_D:
- "\<lbrakk>X ----> (L::'a::metric_space); 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r"
-by (simp add: LIMSEQ_def)
-
-
-subsubsection {* Limits of Functions *}
-
-lemma LIM_def: "f -- (a::'a::metric_space) --> (L::'b::metric_space) =
- (\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & dist x a < s
- --> dist (f x) L < r)"
-unfolding tendsto_iff eventually_at ..
-
-lemma metric_LIM_I:
- "(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r)
- \<Longrightarrow> f -- (a::'a::metric_space) --> (L::'b::metric_space)"
-by (simp add: LIM_def)
-
-lemma metric_LIM_D:
- "\<lbrakk>f -- (a::'a::metric_space) --> (L::'b::metric_space); 0 < r\<rbrakk>
- \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r"
-by (simp add: LIM_def)
-
-lemma metric_LIM_imp_LIM:
- assumes f: "f -- a --> (l::'a::metric_space)"
- assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> dist (g x) m \<le> dist (f x) l"
- shows "g -- a --> (m::'b::metric_space)"
- by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp add: eventually_at_topological le)
-
-lemma metric_LIM_equal2:
- assumes 1: "0 < R"
- assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; dist x a < R\<rbrakk> \<Longrightarrow> f x = g x"
- shows "g -- a --> l \<Longrightarrow> f -- (a::'a::metric_space) --> l"
-apply (rule topological_tendstoI)
-apply (drule (2) topological_tendstoD)
-apply (simp add: eventually_at, safe)
-apply (rule_tac x="min d R" in exI, safe)
-apply (simp add: 1)
-apply (simp add: 2)
-done
-
-lemma metric_LIM_compose2:
- assumes f: "f -- (a::'a::metric_space) --> b"
- assumes g: "g -- b --> c"
- assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> b"
- shows "(\<lambda>x. g (f x)) -- a --> c"
- using g f inj [folded eventually_at]
- by (rule tendsto_compose_eventually)
-
-lemma metric_isCont_LIM_compose2:
- fixes f :: "'a :: metric_space \<Rightarrow> _"
- assumes f [unfolded isCont_def]: "isCont f a"
- assumes g: "g -- f a --> l"
- assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> f a"
- shows "(\<lambda>x. g (f x)) -- a --> l"
-by (rule metric_LIM_compose2 [OF f g inj])
-
-subsubsection {* Boundedness *}
-
-definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where
- Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)"
-
-abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
- "Bseq X \<equiv> Bfun X sequentially"
-
-lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" ..
-
-lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
- unfolding Bfun_metric_def by (subst eventually_sequentially_seg)
-
-lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
- unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg)
-
-subsection {* Complete metric spaces *}
-
-subsection {* Cauchy sequences *}
-
-definition (in metric_space) Cauchy :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" where
- "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < e)"
-
-subsection {* Cauchy Sequences *}
-
-lemma metric_CauchyI:
- "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X"
- by (simp add: Cauchy_def)
-
-lemma metric_CauchyD:
- "Cauchy X \<Longrightarrow> 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e"
- by (simp add: Cauchy_def)
-
-lemma metric_Cauchy_iff2:
- "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < inverse(real (Suc j))))"
-apply (simp add: Cauchy_def, auto)
-apply (drule reals_Archimedean, safe)
-apply (drule_tac x = n in spec, auto)
-apply (rule_tac x = M in exI, auto)
-apply (drule_tac x = m in spec, simp)
-apply (drule_tac x = na in spec, auto)
-done
-
-lemma Cauchy_subseq_Cauchy:
- "\<lbrakk> Cauchy X; subseq f \<rbrakk> \<Longrightarrow> Cauchy (X o f)"
-apply (auto simp add: Cauchy_def)
-apply (drule_tac x=e in spec, clarify)
-apply (rule_tac x=M in exI, clarify)
-apply (blast intro: le_trans [OF _ seq_suble] dest!: spec)
-done
-
-lemma Cauchy_Bseq: "Cauchy X \<Longrightarrow> Bseq X"
- unfolding Cauchy_def Bfun_metric_def eventually_sequentially
- apply (erule_tac x=1 in allE)
- apply simp
- apply safe
- apply (rule_tac x="X M" in exI)
- apply (rule_tac x=1 in exI)
- apply (erule_tac x=M in allE)
- apply simp
- apply (rule_tac x=M in exI)
- apply (auto simp: dist_commute)
- done
-
-subsubsection {* Cauchy Sequences are Convergent *}
-
-class complete_space = metric_space +
- assumes Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
-
-theorem LIMSEQ_imp_Cauchy:
- assumes X: "X ----> a" shows "Cauchy X"
-proof (rule metric_CauchyI)
- fix e::real assume "0 < e"
- hence "0 < e/2" by simp
- with X have "\<exists>N. \<forall>n\<ge>N. dist (X n) a < e/2" by (rule metric_LIMSEQ_D)
- then obtain N where N: "\<forall>n\<ge>N. dist (X n) a < e/2" ..
- show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < e"
- proof (intro exI allI impI)
- fix m assume "N \<le> m"
- hence m: "dist (X m) a < e/2" using N by fast
- fix n assume "N \<le> n"
- hence n: "dist (X n) a < e/2" using N by fast
- have "dist (X m) (X n) \<le> dist (X m) a + dist (X n) a"
- by (rule dist_triangle2)
- also from m n have "\<dots> < e" by simp
- finally show "dist (X m) (X n) < e" .
- qed
-qed
-
-lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
-unfolding convergent_def
-by (erule exE, erule LIMSEQ_imp_Cauchy)
-
-lemma Cauchy_convergent_iff:
- fixes X :: "nat \<Rightarrow> 'a::complete_space"
- shows "Cauchy X = convergent X"
-by (fast intro: Cauchy_convergent convergent_Cauchy)
-
-subsection {* Uniform Continuity *}
-
-definition
- isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where
- "isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"
-
-lemma isUCont_isCont: "isUCont f ==> isCont f x"
-by (simp add: isUCont_def isCont_def LIM_def, force)
-
-lemma isUCont_Cauchy:
- "\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
-unfolding isUCont_def
-apply (rule metric_CauchyI)
-apply (drule_tac x=e in spec, safe)
-apply (drule_tac e=s in metric_CauchyD, safe)
-apply (rule_tac x=M in exI, simp)
-done
-
-subsection {* The set of real numbers is a complete metric space *}
-
-instantiation real :: metric_space
-begin
-
-definition dist_real_def:
- "dist x y = \<bar>x - y\<bar>"
-
-definition open_real_def:
- "open (S :: real set) \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
-
-instance
- by default (auto simp: open_real_def dist_real_def)
-end
-
-instance real :: linorder_topology
-proof
- show "(open :: real set \<Rightarrow> bool) = generate_topology (range lessThan \<union> range greaterThan)"
- proof (rule ext, safe)
- fix S :: "real set" assume "open S"
- then guess f unfolding open_real_def bchoice_iff ..
- then have *: "S = (\<Union>x\<in>S. {x - f x <..} \<inter> {..< x + f x})"
- by (fastforce simp: dist_real_def)
- show "generate_topology (range lessThan \<union> range greaterThan) S"
- apply (subst *)
- apply (intro generate_topology_Union generate_topology.Int)
- apply (auto intro: generate_topology.Basis)
- done
- next
- fix S :: "real set" assume "generate_topology (range lessThan \<union> range greaterThan) S"
- moreover have "\<And>a::real. open {..<a}"
- unfolding open_real_def dist_real_def
- proof clarify
- fix x a :: real assume "x < a"
- hence "0 < a - x \<and> (\<forall>y. \<bar>y - x\<bar> < a - x \<longrightarrow> y \<in> {..<a})" by auto
- thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {..<a}" ..
- qed
- moreover have "\<And>a::real. open {a <..}"
- unfolding open_real_def dist_real_def
- proof clarify
- fix x a :: real assume "a < x"
- hence "0 < x - a \<and> (\<forall>y. \<bar>y - x\<bar> < x - a \<longrightarrow> y \<in> {a<..})" by auto
- thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {a<..}" ..
- qed
- ultimately show "open S"
- by induct auto
- qed
-qed
-
-lemmas open_real_greaterThan = open_greaterThan[where 'a=real]
-lemmas open_real_lessThan = open_lessThan[where 'a=real]
-lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real]
-lemmas closed_real_atMost = closed_atMost[where 'a=real]
-lemmas closed_real_atLeast = closed_atLeast[where 'a=real]
-lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real]
-
-text {*
-Proof that Cauchy sequences converge based on the one from
-http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
-*}
-
-text {*
- If sequence @{term "X"} is Cauchy, then its limit is the lub of
- @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
-*}
-
-lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
-by (simp add: isUbI setleI)
-
-lemma increasing_LIMSEQ:
- fixes f :: "nat \<Rightarrow> real"
- assumes inc: "\<And>n. f n \<le> f (Suc n)"
- and bdd: "\<And>n. f n \<le> l"
- and en: "\<And>e. 0 < e \<Longrightarrow> \<exists>n. l \<le> f n + e"
- shows "f ----> l"
-proof (rule increasing_tendsto)
- fix x assume "x < l"
- with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x"
- by auto
- from en[OF `0 < e`] obtain n where "l - e \<le> f n"
- by (auto simp: field_simps)
- with `e < l - x` `0 < e` have "x < f n" by simp
- with incseq_SucI[of f, OF inc] show "eventually (\<lambda>n. x < f n) sequentially"
- by (auto simp: eventually_sequentially incseq_def intro: less_le_trans)
-qed (insert bdd, auto)
-
-lemma real_Cauchy_convergent:
- fixes X :: "nat \<Rightarrow> real"
- assumes X: "Cauchy X"
- shows "convergent X"
-proof -
- def S \<equiv> "{x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
- then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S" by auto
-
- { fix N x assume N: "\<forall>n\<ge>N. X n < x"
- have "isUb UNIV S x"
- proof (rule isUb_UNIV_I)
- fix y::real assume "y \<in> S"
- hence "\<exists>M. \<forall>n\<ge>M. y < X n"
- by (simp add: S_def)
- then obtain M where "\<forall>n\<ge>M. y < X n" ..
- hence "y < X (max M N)" by simp
- also have "\<dots> < x" using N by simp
- finally show "y \<le> x"
- by (rule order_less_imp_le)
- qed }
- note bound_isUb = this
-
- have "\<exists>u. isLub UNIV S u"
- proof (rule reals_complete)
- obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < 1"
- using X[THEN metric_CauchyD, OF zero_less_one] by auto
- hence N: "\<forall>n\<ge>N. dist (X n) (X N) < 1" by simp
- show "\<exists>x. x \<in> S"
- proof
- from N have "\<forall>n\<ge>N. X N - 1 < X n"
- by (simp add: abs_diff_less_iff dist_real_def)
- thus "X N - 1 \<in> S" by (rule mem_S)
- qed
- show "\<exists>u. isUb UNIV S u"
- proof
- from N have "\<forall>n\<ge>N. X n < X N + 1"
- by (simp add: abs_diff_less_iff dist_real_def)
- thus "isUb UNIV S (X N + 1)"
- by (rule bound_isUb)
- qed
- qed
- then obtain x where x: "isLub UNIV S x" ..
- have "X ----> x"
- proof (rule metric_LIMSEQ_I)
- fix r::real assume "0 < r"
- hence r: "0 < r/2" by simp
- obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. dist (X n) (X m) < r/2"
- using metric_CauchyD [OF X r] by auto
- hence "\<forall>n\<ge>N. dist (X n) (X N) < r/2" by simp
- hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
- by (simp only: dist_real_def abs_diff_less_iff)
-
- from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
- hence "X N - r/2 \<in> S" by (rule mem_S)
- hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
-
- from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
- hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
- hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
-
- show "\<exists>N. \<forall>n\<ge>N. dist (X n) x < r"
- proof (intro exI allI impI)
- fix n assume n: "N \<le> n"
- from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
- thus "dist (X n) x < r" using 1 2
- by (simp add: abs_diff_less_iff dist_real_def)
- qed
- qed
- then show ?thesis unfolding convergent_def by auto
-qed
-
-instance real :: complete_space
- by intro_classes (rule real_Cauchy_convergent)
-
-lemma tendsto_dist [tendsto_intros]:
- fixes l m :: "'a :: metric_space"
- assumes f: "(f ---> l) F" and g: "(g ---> m) F"
- shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) F"
-proof (rule tendstoI)
- fix e :: real assume "0 < e"
- hence e2: "0 < e/2" by simp
- from tendstoD [OF f e2] tendstoD [OF g e2]
- show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F"
- proof (eventually_elim)
- case (elim x)
- then show "dist (dist (f x) (g x)) (dist l m) < e"
- unfolding dist_real_def
- using dist_triangle2 [of "f x" "g x" "l"]
- using dist_triangle2 [of "g x" "l" "m"]
- using dist_triangle3 [of "l" "m" "f x"]
- using dist_triangle [of "f x" "m" "g x"]
- by arith
- qed
-qed
-
-lemma continuous_dist[continuous_intros]:
- fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
- shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. dist (f x) (g x))"
- unfolding continuous_def by (rule tendsto_dist)
-
-lemma continuous_on_dist[continuous_on_intros]:
- fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
- shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))"
- unfolding continuous_on_def by (auto intro: tendsto_dist)
-
-lemma tendsto_at_topI_sequentially:
- fixes f :: "real \<Rightarrow> real"
- assumes mono: "mono f"
- assumes limseq: "(\<lambda>n. f (real n)) ----> y"
- shows "(f ---> y) at_top"
-proof (rule tendstoI)
- fix e :: real assume "0 < e"
- with limseq obtain N :: nat where N: "\<And>n. N \<le> n \<Longrightarrow> \<bar>f (real n) - y\<bar> < e"
- by (auto simp: LIMSEQ_def dist_real_def)
- { fix x :: real
- from ex_le_of_nat[of x] guess n ..
- note monoD[OF mono this]
- also have "f (real_of_nat n) \<le> y"
- by (rule LIMSEQ_le_const[OF limseq])
- (auto intro: exI[of _ n] monoD[OF mono] simp: real_eq_of_nat[symmetric])
- finally have "f x \<le> y" . }
- note le = this
- have "eventually (\<lambda>x. real N \<le> x) at_top"
- by (rule eventually_ge_at_top)
- then show "eventually (\<lambda>x. dist (f x) y < e) at_top"
- proof eventually_elim
- fix x assume N': "real N \<le> x"
- with N[of N] le have "y - f (real N) < e" by auto
- moreover note monoD[OF mono N']
- ultimately show "dist (f x) y < e"
- using le[of x] by (auto simp: dist_real_def field_simps)
- qed
-qed
-
-lemma Cauchy_iff2:
- "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. \<bar>X m - X n\<bar> < inverse(real (Suc j))))"
- unfolding metric_Cauchy_iff2 dist_real_def ..
-
-end
-
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Mar 26 15:10:28 2013 +0100
@@ -468,7 +468,7 @@
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"
apply -
apply (rule as(3)[rule_format])
- unfolding RealVector.scaleR_right.setsum
+ unfolding Real_Vector_Spaces.scaleR_right.setsum
using x(1) as(6) apply auto
done
then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
@@ -530,7 +530,7 @@
apply (rule_tac x="sx \<union> sy" in exI)
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)
unfolding scaleR_left_distrib setsum_addf if_smult scaleR_zero_left ** setsum_restrict_set[OF xy, symmetric]
- unfolding scaleR_scaleR[symmetric] RealVector.scaleR_right.setsum [symmetric] and setsum_right_distrib[symmetric]
+ unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric] and setsum_right_distrib[symmetric]
unfolding x y
using x(1-3) y(1-3) uv apply simp
done
@@ -2848,7 +2848,7 @@
obtain e2 where e2_def: "e2>0 & cball z e2 Int affine hull (f ` S) <= (f ` S)"
using z_def rel_interior_cball[of "f ` S"] by auto
obtain K where K_def: "K>0 & (! x. norm (f x) <= norm x * K)"
- using assms RealVector.bounded_linear.pos_bounded[of f] by auto
+ using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto
def e1 == "1/K" hence e1_def: "e1>0 & (! x. e1 * norm (f x) <= norm x)"
using K_def pos_le_divide_eq[of e1] by auto
def e == "e1 * e2" hence "e>0" using e1_def e2_def mult_pos_pos by auto
--- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Tue Mar 26 15:10:28 2013 +0100
@@ -890,7 +890,7 @@
lemma Liminf_within:
fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"
shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S \<inter> ball x e - {x}). f y)"
- unfolding Liminf_def eventually_within
+ unfolding Liminf_def eventually_within_less
proof (rule SUPR_eq, simp_all add: Ball_def Bex_def, safe)
fix P d assume "0 < d" "\<forall>y. y \<in> S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> P y"
then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
@@ -908,7 +908,7 @@
lemma Limsup_within:
fixes f :: "'a::metric_space => 'b::complete_lattice"
shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S \<inter> ball x e - {x}). f y)"
- unfolding Limsup_def eventually_within
+ unfolding Limsup_def eventually_within_less
proof (rule INFI_eq, simp_all add: Ball_def Bex_def, safe)
fix P d assume "0 < d" "\<forall>y. y \<in> S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> P y"
then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
--- a/src/HOL/Multivariate_Analysis/Integration.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Mar 26 15:10:28 2013 +0100
@@ -8,17 +8,56 @@
"~~/src/HOL/Library/Indicator_Function"
begin
-lemma cSup_finite_ge_iff:
+lemma cSup_abs_le: (* TODO: is this really needed? *)
+ fixes S :: "real set"
+ shows "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>Sup S\<bar> \<le> a"
+by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2)
+
+lemma cInf_abs_ge: (* TODO: is this really needed? *)
+ fixes S :: "real set"
+ shows "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>Inf S\<bar> \<le> a"
+by (simp add: Inf_real_def) (rule cSup_abs_le, auto)
+
+lemma cSup_asclose: (* TODO: is this really needed? *)
fixes S :: "real set"
- assumes fS: "finite S" and Se: "S \<noteq> {}"
- shows "a \<le> Sup S \<longleftrightarrow> (\<exists> x \<in> S. a \<le> x)"
-by (metis Max_ge Se cSup_eq_Max Max_in fS linorder_not_le less_le_trans)
+ assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>Sup S - l\<bar> \<le> e"
+proof-
+ 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
+ thus ?thesis using S b cSup_bounds[of S "l - e" "l+e"] unfolding th
+ by (auto simp add: setge_def setle_def)
+qed
+
+lemma cInf_asclose: (* TODO: is this really needed? *)
+ fixes S :: "real set"
+ assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>Inf S - l\<bar> \<le> e"
+proof -
+ have "\<bar>- Sup (uminus ` S) - l\<bar> = \<bar>Sup (uminus ` S) - (-l)\<bar>"
+ by auto
+ also have "... \<le> e"
+ apply (rule cSup_asclose)
+ apply (auto simp add: S)
+ apply (metis abs_minus_add_cancel b add_commute diff_minus)
+ done
+ finally have "\<bar>- Sup (uminus ` S) - l\<bar> \<le> e" .
+ thus ?thesis
+ by (simp add: Inf_real_def)
+qed
+
+lemma cSup_finite_ge_iff:
+ fixes S :: "real set" shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<le> Sup S \<longleftrightarrow> (\<exists>x\<in>S. a \<le> x)"
+ by (metis cSup_eq_Max Max_ge_iff)
lemma cSup_finite_le_iff:
- fixes S :: "real set"
- assumes fS: "finite S" and Se: "S \<noteq> {}"
- shows "a \<ge> Sup S \<longleftrightarrow> (\<forall> x \<in> S. a \<ge> x)"
-by (metis Max_ge Se cSup_eq_Max Max_in fS order_trans)
+ fixes S :: "real set" shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<ge> Sup S \<longleftrightarrow> (\<forall>x\<in>S. a \<ge> x)"
+ by (metis cSup_eq_Max Max_le_iff)
+
+lemma cInf_finite_ge_iff:
+ fixes S :: "real set" shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
+ by (metis cInf_eq_Min Min_ge_iff)
+
+lemma cInf_finite_le_iff:
+ fixes S :: "real set" shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<ge> Inf S \<longleftrightarrow> (\<exists>x\<in>S. a \<ge> x)"
+ by (metis cInf_eq_Min Min_le_iff)
lemma Inf: (* rename *)
fixes S :: "real set"
@@ -6282,7 +6321,8 @@
unfolding real_norm_def abs_le_iff
apply auto
done
- show "((\<lambda>k. Inf {f j x |j. k \<le> j}) ---> g x) sequentially"
+
+ show "(\<lambda>k::nat. Inf {f j x |j. k \<le> j}) ----> g x"
proof (rule LIMSEQ_I)
case goal1
hence "0<r/2" by auto
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Mar 26 15:10:28 2013 +0100
@@ -1256,15 +1256,10 @@
text {* Some property holds "sufficiently close" to the limit point. *}
-lemma eventually_at: (* FIXME: this replaces Metric_Spaces.eventually_at *)
+lemma eventually_at2:
"eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
unfolding eventually_at dist_nz by auto
-lemma eventually_within: (* FIXME: this replaces Topological_Spaces.eventually_within *)
- "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)"
- by (rule eventually_within_less)
-
lemma eventually_happens: "eventually P net ==> trivial_limit net \<or> (\<exists>x. P x)"
unfolding trivial_limit_def
by (auto elim: eventually_rev_mp)
@@ -1301,11 +1296,11 @@
lemma Lim_within: "(f ---> l) (at a within S) \<longleftrightarrow>
(\<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)"
- by (auto simp add: tendsto_iff eventually_within)
+ by (auto simp add: tendsto_iff eventually_within_less)
lemma Lim_at: "(f ---> l) (at a) \<longleftrightarrow>
(\<forall>e >0. \<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)"
- by (auto simp add: tendsto_iff eventually_at)
+ by (auto simp add: tendsto_iff eventually_at2)
lemma Lim_at_infinity:
"(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x >= b \<longrightarrow> dist (f x) l < e)"
@@ -1374,34 +1369,33 @@
by (simp add: sequentially_imp_eventually_within)
lemma Lim_right_bound:
- fixes f :: "real \<Rightarrow> real"
+ fixes f :: "'a :: {linorder_topology, conditional_complete_linorder, no_top} \<Rightarrow>
+ 'b::{linorder_topology, conditional_complete_linorder}"
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"
assumes bnd: "\<And>a. a \<in> I \<Longrightarrow> x < a \<Longrightarrow> K \<le> f a"
shows "(f ---> Inf (f ` ({x<..} \<inter> I))) (at x within ({x<..} \<inter> I))"
proof cases
assume "{x<..} \<inter> I = {}" then show ?thesis by (simp add: Lim_within_empty)
next
- assume [simp]: "{x<..} \<inter> I \<noteq> {}"
+ assume e: "{x<..} \<inter> I \<noteq> {}"
show ?thesis
- proof (rule Lim_within_LIMSEQ, safe)
- fix S assume S: "\<forall>n. S n \<noteq> x \<and> S n \<in> {x <..} \<inter> I" "S ----> x"
-
- show "(\<lambda>n. f (S n)) ----> Inf (f ` ({x<..} \<inter> I))"
- proof (rule LIMSEQ_I, rule ccontr)
- fix r :: real assume "0 < r"
- with cInf_close[of "f ` ({x<..} \<inter> I)" r]
- obtain y where y: "x < y" "y \<in> I" "f y < Inf (f ` ({x <..} \<inter> I)) + r" by auto
- from `x < y` have "0 < y - x" by auto
- from S(2)[THEN LIMSEQ_D, OF this]
- obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> \<bar>S n - x\<bar> < y - x" by auto
-
- assume "\<not> (\<exists>N. \<forall>n\<ge>N. norm (f (S n) - Inf (f ` ({x<..} \<inter> I))) < r)"
- moreover have "\<And>n. Inf (f ` ({x<..} \<inter> I)) \<le> f (S n)"
- using S bnd by (intro cInf_lower[where z=K]) auto
- ultimately obtain n where n: "N \<le> n" "r + Inf (f ` ({x<..} \<inter> I)) \<le> f (S n)"
- by (auto simp: not_less field_simps)
- with N[OF n(1)] mono[OF _ `y \<in> I`, of "S n"] S(1)[THEN spec, of n] y
- show False by auto
+ proof (rule order_tendstoI)
+ fix a assume a: "a < Inf (f ` ({x<..} \<inter> I))"
+ { fix y assume "y \<in> {x<..} \<inter> I"
+ with e bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> f y"
+ by (auto intro: cInf_lower)
+ with a have "a < f y" by (blast intro: less_le_trans) }
+ then show "eventually (\<lambda>x. a < f x) (at x within ({x<..} \<inter> I))"
+ by (auto simp: Topological_Spaces.eventually_within intro: exI[of _ 1] zero_less_one)
+ next
+ fix a assume "Inf (f ` ({x<..} \<inter> I)) < a"
+ from cInf_lessD[OF _ this] e obtain y where y: "x < y" "y \<in> I" "f y < a" by auto
+ show "eventually (\<lambda>x. f x < a) (at x within ({x<..} \<inter> I))"
+ unfolding within_within_eq[symmetric]
+ Topological_Spaces.eventually_within[of _ _ I] eventually_at_right
+ proof (safe intro!: exI[of _ y] y)
+ fix z assume "x < z" "z \<in> I" "z < y"
+ with mono[OF `z\<in>I` `y\<in>I`] `f y < a` show "f z < a" by (auto simp: less_imp_le)
qed
qed
qed
@@ -1602,7 +1596,7 @@
shows "(g ---> l) (at x within S)"
proof (rule Lim_transform_eventually)
show "eventually (\<lambda>x. f x = g x) (at x within S)"
- unfolding eventually_within
+ unfolding eventually_within_less
using assms(1,2) by auto
show "(f ---> l) (at x within S)" by fact
qed
@@ -1613,7 +1607,7 @@
shows "(g ---> l) (at x)"
proof (rule Lim_transform_eventually)
show "eventually (\<lambda>x. f x = g x) (at x)"
- unfolding eventually_at
+ unfolding eventually_at2
using assms(1,2) by auto
show "(f ---> l) (at x)" by fact
qed
@@ -1718,12 +1712,11 @@
using cInf_lower_EX[of _ S] assms by metis
fix e :: real assume "0 < e"
- then obtain x where x: "x \<in> S" "x < Inf S + e"
- using cInf_close `S \<noteq> {}` by auto
- moreover then have "x > Inf S - e" using * by auto
- ultimately have "abs (x - Inf S) < e" by (simp add: abs_diff_less_iff)
- then show "\<exists>x\<in>S. dist x (Inf S) < e"
- using x by (auto simp: dist_norm)
+ then have "Inf S < Inf S + e" by simp
+ with assms obtain x where "x \<in> S" "x < Inf S + e"
+ by (subst (asm) cInf_less_iff[of _ B]) auto
+ with * show "\<exists>x\<in>S. dist x (Inf S) < e"
+ by (intro bexI[of _ x]) (auto simp add: dist_real_def)
qed
lemma closed_contains_Inf:
@@ -3790,7 +3783,7 @@
{ 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"
fix T::"'b set" assume "open T" and "f a \<in> T"
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"
- unfolding continuous_within tendsto_def eventually_within by auto
+ unfolding continuous_within tendsto_def eventually_within_less by auto
have "eventually (\<lambda>n. dist (x n) a < d) sequentially"
using x(2) `d>0` by simp
hence "eventually (\<lambda>n. (f \<circ> x) n \<in> T) sequentially"
@@ -4188,8 +4181,7 @@
hence "eventually (\<lambda>y. f y \<noteq> a) (at x within s)"
using `a \<notin> U` by (fast elim: eventually_mono [rotated])
thus ?thesis
- unfolding Limits.eventually_within Metric_Spaces.eventually_at
- by (rule ex_forward, cut_tac `f x \<noteq> a`, auto simp: dist_commute)
+ using `f x \<noteq> a` by (auto simp: dist_commute zero_less_dist_iff eventually_within_less)
qed
lemma continuous_at_avoid:
--- a/src/HOL/NSA/HDeriv.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/NSA/HDeriv.thy Tue Mar 26 15:10:28 2013 +0100
@@ -7,7 +7,7 @@
header{* Differentiation (Nonstandard) *}
theory HDeriv
-imports Deriv HLim
+imports HLim
begin
text{*Nonstandard Definitions*}
--- a/src/HOL/NSA/HLim.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/NSA/HLim.thy Tue Mar 26 15:10:28 2013 +0100
@@ -6,7 +6,7 @@
header{* Limits and Continuity (Nonstandard) *}
theory HLim
-imports Star Lim
+imports Star
begin
text{*Nonstandard Definitions*}
--- a/src/HOL/NSA/HLog.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/NSA/HLog.thy Tue Mar 26 15:10:28 2013 +0100
@@ -6,7 +6,7 @@
header{*Logarithms: Non-Standard Version*}
theory HLog
-imports Log HTranscendental
+imports HTranscendental
begin
--- a/src/HOL/NSA/HSEQ.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/NSA/HSEQ.thy Tue Mar 26 15:10:28 2013 +0100
@@ -9,7 +9,7 @@
header {* Sequences and Convergence (Nonstandard) *}
theory HSEQ
-imports SEQ NatStar
+imports Limits NatStar
begin
definition
--- a/src/HOL/NSA/HSeries.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/NSA/HSeries.thy Tue Mar 26 15:10:28 2013 +0100
@@ -8,7 +8,7 @@
header{*Finite Summation and Infinite Series for Hyperreals*}
theory HSeries
-imports Series HSEQ
+imports HSEQ
begin
definition
--- a/src/HOL/NSA/HyperDef.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/NSA/HyperDef.thy Tue Mar 26 15:10:28 2013 +0100
@@ -7,7 +7,7 @@
header{*Construction of Hyperreals Using Ultrafilters*}
theory HyperDef
-imports HyperNat Real
+imports Complex_Main HyperNat
begin
type_synonym hypreal = "real star"
--- a/src/HOL/NSA/Hyperreal.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/NSA/Hyperreal.thy Tue Mar 26 15:10:28 2013 +0100
@@ -7,7 +7,7 @@
*)
theory Hyperreal
-imports Ln Deriv Taylor HLog
+imports HLog
begin
end
--- a/src/HOL/NSA/NSA.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/NSA/NSA.thy Tue Mar 26 15:10:28 2013 +0100
@@ -6,7 +6,7 @@
header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*}
theory NSA
-imports HyperDef RComplete
+imports HyperDef
begin
definition
--- a/src/HOL/NSA/NSComplex.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/NSA/NSComplex.thy Tue Mar 26 15:10:28 2013 +0100
@@ -6,7 +6,7 @@
header{*Nonstandard Complex Numbers*}
theory NSComplex
-imports Complex NSA
+imports NSA
begin
type_synonym hcomplex = "complex star"
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Mar 26 15:10:28 2013 +0100
@@ -12,7 +12,7 @@
suite. *)
theory Manual_Nits
-imports Main "~~/src/HOL/Library/Quotient_Product" RealDef
+imports Main "~~/src/HOL/Library/Quotient_Product" Real
begin
chapter {* 2. First Steps *}
--- a/src/HOL/Quickcheck_Benchmark/Find_Unused_Assms_Examples.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Quickcheck_Benchmark/Find_Unused_Assms_Examples.thy Tue Mar 26 15:10:28 2013 +0100
@@ -8,8 +8,7 @@
find_unused_assms Divides
find_unused_assms GCD
-find_unused_assms RealDef
-find_unused_assms RComplete
+find_unused_assms Real
section {* Set Theory *}
--- a/src/HOL/RComplete.thy Tue Mar 26 14:38:44 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,544 +0,0 @@
-(* Title: HOL/RComplete.thy
- Author: Jacques D. Fleuriot, University of Edinburgh
- Author: Larry Paulson, University of Cambridge
- Author: Jeremy Avigad, Carnegie Mellon University
- Author: Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen
-*)
-
-header {* Completeness of the Reals; Floor and Ceiling Functions *}
-
-theory RComplete
-imports Lubs RealDef
-begin
-
-lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
- by simp
-
-lemma abs_diff_less_iff:
- "(\<bar>x - a\<bar> < (r::'a::linordered_idom)) = (a - r < x \<and> x < a + r)"
- by auto
-
-subsection {* Completeness of Positive Reals *}
-
-text {*
- Supremum property for the set of positive reals
-
- Let @{text "P"} be a non-empty set of positive reals, with an upper
- bound @{text "y"}. Then @{text "P"} has a least upper bound
- (written @{text "S"}).
-
- FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
-*}
-
-text {* Only used in HOL/Import/HOL4Compat.thy; delete? *}
-
-lemma posreal_complete:
- fixes P :: "real set"
- assumes not_empty_P: "\<exists>x. x \<in> P"
- and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
- shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
-proof -
- from upper_bound_Ex have "\<exists>z. \<forall>x\<in>P. x \<le> z"
- by (auto intro: less_imp_le)
- from complete_real [OF not_empty_P this] obtain S
- 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
- have "\<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
- proof
- fix y show "(\<exists>x\<in>P. y < x) = (y < S)"
- apply (cases "\<exists>x\<in>P. y < x", simp_all)
- apply (clarify, drule S1, simp)
- apply (simp add: not_less S2)
- done
- qed
- thus ?thesis ..
-qed
-
-text {*
- \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
-*}
-
-lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)"
- apply (frule isLub_isUb)
- apply (frule_tac x = y in isLub_isUb)
- apply (blast intro!: order_antisym dest!: isLub_le_isUb)
- done
-
-
-text {*
- \medskip reals Completeness (again!)
-*}
-
-lemma reals_complete:
- assumes notempty_S: "\<exists>X. X \<in> S"
- and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
- shows "\<exists>t. isLub (UNIV :: real set) S t"
-proof -
- from assms have "\<exists>X. X \<in> S" and "\<exists>Y. \<forall>x\<in>S. x \<le> Y"
- unfolding isUb_def setle_def by simp_all
- from complete_real [OF this] show ?thesis
- by (simp add: isLub_def leastP_def isUb_def setle_def setge_def)
-qed
-
-
-subsection {* The Archimedean Property of the Reals *}
-
-theorem reals_Archimedean:
- assumes x_pos: "0 < x"
- shows "\<exists>n. inverse (real (Suc n)) < x"
- unfolding real_of_nat_def using x_pos
- by (rule ex_inverse_of_nat_Suc_less)
-
-lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
- unfolding real_of_nat_def by (rule ex_less_of_nat)
-
-lemma reals_Archimedean3:
- assumes x_greater_zero: "0 < x"
- shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"
- unfolding real_of_nat_def using `0 < x`
- by (auto intro: ex_less_of_nat_mult)
-
-
-subsection{*Density of the Rational Reals in the Reals*}
-
-text{* This density proof is due to Stefan Richter and was ported by TN. The
-original source is \emph{Real Analysis} by H.L. Royden.
-It employs the Archimedean property of the reals. *}
-
-lemma Rats_dense_in_real:
- fixes x :: real
- assumes "x < y" shows "\<exists>r\<in>\<rat>. x < r \<and> r < y"
-proof -
- from `x<y` have "0 < y-x" by simp
- with reals_Archimedean obtain q::nat
- where q: "inverse (real q) < y-x" and "0 < q" by auto
- def p \<equiv> "ceiling (y * real q) - 1"
- def r \<equiv> "of_int p / real q"
- from q have "x < y - inverse (real q)" by simp
- also have "y - inverse (real q) \<le> r"
- unfolding r_def p_def
- by (simp add: le_divide_eq left_diff_distrib le_of_int_ceiling `0 < q`)
- finally have "x < r" .
- moreover have "r < y"
- unfolding r_def p_def
- by (simp add: divide_less_eq diff_less_eq `0 < q`
- less_ceiling_iff [symmetric])
- moreover from r_def have "r \<in> \<rat>" by simp
- ultimately show ?thesis by fast
-qed
-
-
-subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
-
-(* FIXME: theorems for negative numerals *)
-lemma numeral_less_real_of_int_iff [simp]:
- "((numeral n) < real (m::int)) = (numeral n < m)"
-apply auto
-apply (rule real_of_int_less_iff [THEN iffD1])
-apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
-done
-
-lemma numeral_less_real_of_int_iff2 [simp]:
- "(real (m::int) < (numeral n)) = (m < numeral n)"
-apply auto
-apply (rule real_of_int_less_iff [THEN iffD1])
-apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
-done
-
-lemma numeral_le_real_of_int_iff [simp]:
- "((numeral n) \<le> real (m::int)) = (numeral n \<le> m)"
-by (simp add: linorder_not_less [symmetric])
-
-lemma numeral_le_real_of_int_iff2 [simp]:
- "(real (m::int) \<le> (numeral n)) = (m \<le> numeral n)"
-by (simp add: linorder_not_less [symmetric])
-
-lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"
-unfolding real_of_nat_def by simp
-
-lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"
-unfolding real_of_nat_def by (simp add: floor_minus)
-
-lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"
-unfolding real_of_int_def by simp
-
-lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"
-unfolding real_of_int_def by (simp add: floor_minus)
-
-lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
-unfolding real_of_int_def by (rule floor_exists)
-
-lemma lemma_floor:
- assumes a1: "real m \<le> r" and a2: "r < real n + 1"
- shows "m \<le> (n::int)"
-proof -
- have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)
- also have "... = real (n + 1)" by simp
- finally have "m < n + 1" by (simp only: real_of_int_less_iff)
- thus ?thesis by arith
-qed
-
-lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"
-unfolding real_of_int_def by (rule of_int_floor_le)
-
-lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \<le> x"
-by (auto intro: lemma_floor)
-
-lemma real_of_int_floor_cancel [simp]:
- "(real (floor x) = x) = (\<exists>n::int. x = real n)"
- using floor_real_of_int by metis
-
-lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
- unfolding real_of_int_def using floor_unique [of n x] by simp
-
-lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"
- unfolding real_of_int_def by (rule floor_unique)
-
-lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n"
-apply (rule inj_int [THEN injD])
-apply (simp add: real_of_nat_Suc)
-apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"])
-done
-
-lemma floor_eq4: "[| real n \<le> x; x < real (Suc n) |] ==> nat(floor x) = n"
-apply (drule order_le_imp_less_or_eq)
-apply (auto intro: floor_eq3)
-done
-
-lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
- unfolding real_of_int_def using floor_correct [of r] by simp
-
-lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"
- unfolding real_of_int_def using floor_correct [of r] by simp
-
-lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"
- unfolding real_of_int_def using floor_correct [of r] by simp
-
-lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"
- unfolding real_of_int_def using floor_correct [of r] by simp
-
-lemma le_floor: "real a <= x ==> a <= floor x"
- unfolding real_of_int_def by (simp add: le_floor_iff)
-
-lemma real_le_floor: "a <= floor x ==> real a <= x"
- unfolding real_of_int_def by (simp add: le_floor_iff)
-
-lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
- unfolding real_of_int_def by (rule le_floor_iff)
-
-lemma floor_less_eq: "(floor x < a) = (x < real a)"
- unfolding real_of_int_def by (rule floor_less_iff)
-
-lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
- unfolding real_of_int_def by (rule less_floor_iff)
-
-lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
- unfolding real_of_int_def by (rule floor_le_iff)
-
-lemma floor_add [simp]: "floor (x + real a) = floor x + a"
- unfolding real_of_int_def by (rule floor_add_of_int)
-
-lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
- unfolding real_of_int_def by (rule floor_diff_of_int)
-
-lemma le_mult_floor:
- assumes "0 \<le> (a :: real)" and "0 \<le> b"
- shows "floor a * floor b \<le> floor (a * b)"
-proof -
- have "real (floor a) \<le> a"
- and "real (floor b) \<le> b" by auto
- hence "real (floor a * floor b) \<le> a * b"
- using assms by (auto intro!: mult_mono)
- also have "a * b < real (floor (a * b) + 1)" by auto
- finally show ?thesis unfolding real_of_int_less_iff by simp
-qed
-
-lemma floor_divide_eq_div:
- "floor (real a / real b) = a div b"
-proof cases
- assume "b \<noteq> 0 \<or> b dvd a"
- with real_of_int_div3[of a b] show ?thesis
- by (auto simp: real_of_int_div[symmetric] intro!: floor_eq2 real_of_int_div4 neq_le_trans)
- (metis add_left_cancel zero_neq_one real_of_int_div_aux real_of_int_inject
- real_of_int_zero_cancel right_inverse_eq div_self mod_div_trivial)
-qed (auto simp: real_of_int_div)
-
-lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
- unfolding real_of_nat_def by simp
-
-lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"
- unfolding real_of_int_def by (rule le_of_int_ceiling)
-
-lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"
- unfolding real_of_int_def by simp
-
-lemma real_of_int_ceiling_cancel [simp]:
- "(real (ceiling x) = x) = (\<exists>n::int. x = real n)"
- using ceiling_real_of_int by metis
-
-lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"
- unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
-
-lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"
- unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
-
-lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n |] ==> ceiling x = n"
- unfolding real_of_int_def using ceiling_unique [of n x] by simp
-
-lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
- unfolding real_of_int_def using ceiling_correct [of r] by simp
-
-lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"
- unfolding real_of_int_def using ceiling_correct [of r] by simp
-
-lemma ceiling_le: "x <= real a ==> ceiling x <= a"
- unfolding real_of_int_def by (simp add: ceiling_le_iff)
-
-lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"
- unfolding real_of_int_def by (simp add: ceiling_le_iff)
-
-lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
- unfolding real_of_int_def by (rule ceiling_le_iff)
-
-lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
- unfolding real_of_int_def by (rule less_ceiling_iff)
-
-lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
- unfolding real_of_int_def by (rule ceiling_less_iff)
-
-lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
- unfolding real_of_int_def by (rule le_ceiling_iff)
-
-lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
- unfolding real_of_int_def by (rule ceiling_add_of_int)
-
-lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
- unfolding real_of_int_def by (rule ceiling_diff_of_int)
-
-
-subsection {* Versions for the natural numbers *}
-
-definition
- natfloor :: "real => nat" where
- "natfloor x = nat(floor x)"
-
-definition
- natceiling :: "real => nat" where
- "natceiling x = nat(ceiling x)"
-
-lemma natfloor_zero [simp]: "natfloor 0 = 0"
- by (unfold natfloor_def, simp)
-
-lemma natfloor_one [simp]: "natfloor 1 = 1"
- by (unfold natfloor_def, simp)
-
-lemma zero_le_natfloor [simp]: "0 <= natfloor x"
- by (unfold natfloor_def, simp)
-
-lemma natfloor_numeral_eq [simp]: "natfloor (numeral n) = numeral n"
- by (unfold natfloor_def, simp)
-
-lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"
- by (unfold natfloor_def, simp)
-
-lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"
- by (unfold natfloor_def, simp)
-
-lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"
- unfolding natfloor_def by simp
-
-lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
- unfolding natfloor_def by (intro nat_mono floor_mono)
-
-lemma le_natfloor: "real x <= a ==> x <= natfloor a"
- apply (unfold natfloor_def)
- apply (subst nat_int [THEN sym])
- apply (rule nat_mono)
- apply (rule le_floor)
- apply simp
-done
-
-lemma natfloor_less_iff: "0 \<le> x \<Longrightarrow> natfloor x < n \<longleftrightarrow> x < real n"
- unfolding natfloor_def real_of_nat_def
- by (simp add: nat_less_iff floor_less_iff)
-
-lemma less_natfloor:
- assumes "0 \<le> x" and "x < real (n :: nat)"
- shows "natfloor x < n"
- using assms by (simp add: natfloor_less_iff)
-
-lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)"
- apply (rule iffI)
- apply (rule order_trans)
- prefer 2
- apply (erule real_natfloor_le)
- apply (subst real_of_nat_le_iff)
- apply assumption
- apply (erule le_natfloor)
-done
-
-lemma le_natfloor_eq_numeral [simp]:
- "~ neg((numeral n)::int) ==> 0 <= x ==>
- (numeral n <= natfloor x) = (numeral n <= x)"
- apply (subst le_natfloor_eq, assumption)
- apply simp
-done
-
-lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
- apply (case_tac "0 <= x")
- apply (subst le_natfloor_eq, assumption, simp)
- apply (rule iffI)
- apply (subgoal_tac "natfloor x <= natfloor 0")
- apply simp
- apply (rule natfloor_mono)
- apply simp
- apply simp
-done
-
-lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n"
- unfolding natfloor_def by (simp add: floor_eq2 [where n="int n"])
-
-lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"
- apply (case_tac "0 <= x")
- apply (unfold natfloor_def)
- apply simp
- apply simp_all
-done
-
-lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
-using real_natfloor_add_one_gt by (simp add: algebra_simps)
-
-lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"
- apply (subgoal_tac "z < real(natfloor z) + 1")
- apply arith
- apply (rule real_natfloor_add_one_gt)
-done
-
-lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a"
- unfolding natfloor_def
- unfolding real_of_int_of_nat_eq [symmetric] floor_add
- by (simp add: nat_add_distrib)
-
-lemma natfloor_add_numeral [simp]:
- "~neg ((numeral n)::int) ==> 0 <= x ==>
- natfloor (x + numeral n) = natfloor x + numeral n"
- by (simp add: natfloor_add [symmetric])
-
-lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1"
- by (simp add: natfloor_add [symmetric] del: One_nat_def)
-
-lemma natfloor_subtract [simp]:
- "natfloor(x - real a) = natfloor x - a"
- unfolding natfloor_def real_of_int_of_nat_eq [symmetric] floor_subtract
- by simp
-
-lemma natfloor_div_nat:
- assumes "1 <= x" and "y > 0"
- shows "natfloor (x / real y) = natfloor x div y"
-proof (rule natfloor_eq)
- have "(natfloor x) div y * y \<le> natfloor x"
- by (rule add_leD1 [where k="natfloor x mod y"], simp)
- thus "real (natfloor x div y) \<le> x / real y"
- using assms by (simp add: le_divide_eq le_natfloor_eq)
- have "natfloor x < (natfloor x) div y * y + y"
- apply (subst mod_div_equality [symmetric])
- apply (rule add_strict_left_mono)
- apply (rule mod_less_divisor)
- apply fact
- done
- thus "x / real y < real (natfloor x div y) + 1"
- using assms
- by (simp add: divide_less_eq natfloor_less_iff distrib_right)
-qed
-
-lemma le_mult_natfloor:
- shows "natfloor a * natfloor b \<le> natfloor (a * b)"
- by (cases "0 <= a & 0 <= b")
- (auto simp add: le_natfloor_eq mult_nonneg_nonneg mult_mono' real_natfloor_le natfloor_neg)
-
-lemma natceiling_zero [simp]: "natceiling 0 = 0"
- by (unfold natceiling_def, simp)
-
-lemma natceiling_one [simp]: "natceiling 1 = 1"
- by (unfold natceiling_def, simp)
-
-lemma zero_le_natceiling [simp]: "0 <= natceiling x"
- by (unfold natceiling_def, simp)
-
-lemma natceiling_numeral_eq [simp]: "natceiling (numeral n) = numeral n"
- by (unfold natceiling_def, simp)
-
-lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"
- by (unfold natceiling_def, simp)
-
-lemma real_natceiling_ge: "x <= real(natceiling x)"
- unfolding natceiling_def by (cases "x < 0", simp_all)
-
-lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"
- unfolding natceiling_def by simp
-
-lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"
- unfolding natceiling_def by (intro nat_mono ceiling_mono)
-
-lemma natceiling_le: "x <= real a ==> natceiling x <= a"
- unfolding natceiling_def real_of_nat_def
- by (simp add: nat_le_iff ceiling_le_iff)
-
-lemma natceiling_le_eq: "(natceiling x <= a) = (x <= real a)"
- unfolding natceiling_def real_of_nat_def
- by (simp add: nat_le_iff ceiling_le_iff)
-
-lemma natceiling_le_eq_numeral [simp]:
- "~ neg((numeral n)::int) ==>
- (natceiling x <= numeral n) = (x <= numeral n)"
- by (simp add: natceiling_le_eq)
-
-lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"
- unfolding natceiling_def
- by (simp add: nat_le_iff ceiling_le_iff)
-
-lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1"
- unfolding natceiling_def
- by (simp add: ceiling_eq2 [where n="int n"])
-
-lemma natceiling_add [simp]: "0 <= x ==>
- natceiling (x + real a) = natceiling x + a"
- unfolding natceiling_def
- unfolding real_of_int_of_nat_eq [symmetric] ceiling_add
- by (simp add: nat_add_distrib)
-
-lemma natceiling_add_numeral [simp]:
- "~ neg ((numeral n)::int) ==> 0 <= x ==>
- natceiling (x + numeral n) = natceiling x + numeral n"
- by (simp add: natceiling_add [symmetric])
-
-lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1"
- by (simp add: natceiling_add [symmetric] del: One_nat_def)
-
-lemma natceiling_subtract [simp]: "natceiling(x - real a) = natceiling x - a"
- unfolding natceiling_def real_of_int_of_nat_eq [symmetric] ceiling_subtract
- by simp
-
-subsection {* Exponentiation with floor *}
-
-lemma floor_power:
- assumes "x = real (floor x)"
- shows "floor (x ^ n) = floor x ^ n"
-proof -
- have *: "x ^ n = real (floor x ^ n)"
- using assms by (induct n arbitrary: x) simp_all
- show ?thesis unfolding real_of_int_inject[symmetric]
- unfolding * floor_real_of_int ..
-qed
-
-lemma natfloor_power:
- assumes "x = real (natfloor x)"
- shows "natfloor (x ^ n) = natfloor x ^ n"
-proof -
- from assms have "0 \<le> floor x" by auto
- note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \<le> floor x`]]
- from floor_power[OF this]
- show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \<le> floor x`, symmetric]
- by simp
-qed
-
-end
--- a/src/HOL/Real.thy Tue Mar 26 14:38:44 2013 +0100
+++ b/src/HOL/Real.thy Tue Mar 26 15:10:28 2013 +0100
@@ -1,7 +1,2228 @@
+(* Title: HOL/Real.thy
+ Author: Jacques D. Fleuriot, University of Edinburgh, 1998
+ Author: Larry Paulson, University of Cambridge
+ Author: Jeremy Avigad, Carnegie Mellon University
+ Author: Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen
+ Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
+ Construction of Cauchy Reals by Brian Huffman, 2010
+*)
+
+header {* Development of the Reals using Cauchy Sequences *}
+
theory Real
-imports RComplete RealVector
+imports Rat Conditional_Complete_Lattices
+begin
+
+text {*
+ This theory contains a formalization of the real numbers as
+ equivalence classes of Cauchy sequences of rationals. See
+ @{file "~~/src/HOL/ex/Dedekind_Real.thy"} for an alternative
+ construction using Dedekind cuts.
+*}
+
+subsection {* Preliminary lemmas *}
+
+lemma add_diff_add:
+ fixes a b c d :: "'a::ab_group_add"
+ shows "(a + c) - (b + d) = (a - b) + (c - d)"
+ by simp
+
+lemma minus_diff_minus:
+ fixes a b :: "'a::ab_group_add"
+ shows "- a - - b = - (a - b)"
+ by simp
+
+lemma mult_diff_mult:
+ fixes x y a b :: "'a::ring"
+ shows "(x * y - a * b) = x * (y - b) + (x - a) * b"
+ by (simp add: algebra_simps)
+
+lemma inverse_diff_inverse:
+ fixes a b :: "'a::division_ring"
+ assumes "a \<noteq> 0" and "b \<noteq> 0"
+ shows "inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
+ using assms by (simp add: algebra_simps)
+
+lemma obtain_pos_sum:
+ fixes r :: rat assumes r: "0 < r"
+ obtains s t where "0 < s" and "0 < t" and "r = s + t"
+proof
+ from r show "0 < r/2" by simp
+ from r show "0 < r/2" by simp
+ show "r = r/2 + r/2" by simp
+qed
+
+subsection {* Sequences that converge to zero *}
+
+definition
+ vanishes :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
+where
+ "vanishes X = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r)"
+
+lemma vanishesI: "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r) \<Longrightarrow> vanishes X"
+ unfolding vanishes_def by simp
+
+lemma vanishesD: "\<lbrakk>vanishes X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r"
+ unfolding vanishes_def by simp
+
+lemma vanishes_const [simp]: "vanishes (\<lambda>n. c) \<longleftrightarrow> c = 0"
+ unfolding vanishes_def
+ apply (cases "c = 0", auto)
+ apply (rule exI [where x="\<bar>c\<bar>"], auto)
+ done
+
+lemma vanishes_minus: "vanishes X \<Longrightarrow> vanishes (\<lambda>n. - X n)"
+ unfolding vanishes_def by simp
+
+lemma vanishes_add:
+ assumes X: "vanishes X" and Y: "vanishes Y"
+ shows "vanishes (\<lambda>n. X n + Y n)"
+proof (rule vanishesI)
+ fix r :: rat assume "0 < r"
+ then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
+ by (rule obtain_pos_sum)
+ obtain i where i: "\<forall>n\<ge>i. \<bar>X n\<bar> < s"
+ using vanishesD [OF X s] ..
+ obtain j where j: "\<forall>n\<ge>j. \<bar>Y n\<bar> < t"
+ using vanishesD [OF Y t] ..
+ have "\<forall>n\<ge>max i j. \<bar>X n + Y n\<bar> < r"
+ proof (clarsimp)
+ fix n assume n: "i \<le> n" "j \<le> n"
+ have "\<bar>X n + Y n\<bar> \<le> \<bar>X n\<bar> + \<bar>Y n\<bar>" by (rule abs_triangle_ineq)
+ also have "\<dots> < s + t" by (simp add: add_strict_mono i j n)
+ finally show "\<bar>X n + Y n\<bar> < r" unfolding r .
+ qed
+ thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n + Y n\<bar> < r" ..
+qed
+
+lemma vanishes_diff:
+ assumes X: "vanishes X" and Y: "vanishes Y"
+ shows "vanishes (\<lambda>n. X n - Y n)"
+unfolding diff_minus by (intro vanishes_add vanishes_minus X Y)
+
+lemma vanishes_mult_bounded:
+ assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
+ assumes Y: "vanishes (\<lambda>n. Y n)"
+ shows "vanishes (\<lambda>n. X n * Y n)"
+proof (rule vanishesI)
+ fix r :: rat assume r: "0 < r"
+ obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
+ using X by fast
+ obtain b where b: "0 < b" "r = a * b"
+ proof
+ show "0 < r / a" using r a by (simp add: divide_pos_pos)
+ show "r = a * (r / a)" using a by simp
+ qed
+ obtain k where k: "\<forall>n\<ge>k. \<bar>Y n\<bar> < b"
+ using vanishesD [OF Y b(1)] ..
+ have "\<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r"
+ by (simp add: b(2) abs_mult mult_strict_mono' a k)
+ thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r" ..
+qed
+
+subsection {* Cauchy sequences *}
+
+definition
+ cauchy :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
+where
+ "cauchy X \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r)"
+
+lemma cauchyI:
+ "(\<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"
+ unfolding cauchy_def by simp
+
+lemma cauchyD:
+ "\<lbrakk>cauchy X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r"
+ unfolding cauchy_def by simp
+
+lemma cauchy_const [simp]: "cauchy (\<lambda>n. x)"
+ unfolding cauchy_def by simp
+
+lemma cauchy_add [simp]:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "cauchy (\<lambda>n. X n + Y n)"
+proof (rule cauchyI)
+ fix r :: rat assume "0 < r"
+ then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
+ by (rule obtain_pos_sum)
+ obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
+ using cauchyD [OF X s] ..
+ obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
+ using cauchyD [OF Y t] ..
+ have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r"
+ proof (clarsimp)
+ fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
+ have "\<bar>(X m + Y m) - (X n + Y n)\<bar> \<le> \<bar>X m - X n\<bar> + \<bar>Y m - Y n\<bar>"
+ unfolding add_diff_add by (rule abs_triangle_ineq)
+ also have "\<dots> < s + t"
+ by (rule add_strict_mono, simp_all add: i j *)
+ finally show "\<bar>(X m + Y m) - (X n + Y n)\<bar> < r" unfolding r .
+ qed
+ thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r" ..
+qed
+
+lemma cauchy_minus [simp]:
+ assumes X: "cauchy X"
+ shows "cauchy (\<lambda>n. - X n)"
+using assms unfolding cauchy_def
+unfolding minus_diff_minus abs_minus_cancel .
+
+lemma cauchy_diff [simp]:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "cauchy (\<lambda>n. X n - Y n)"
+using assms unfolding diff_minus by simp
+
+lemma cauchy_imp_bounded:
+ assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
+proof -
+ obtain k where k: "\<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < 1"
+ using cauchyD [OF assms zero_less_one] ..
+ show "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
+ proof (intro exI conjI allI)
+ have "0 \<le> \<bar>X 0\<bar>" by simp
+ also have "\<bar>X 0\<bar> \<le> Max (abs ` X ` {..k})" by simp
+ finally have "0 \<le> Max (abs ` X ` {..k})" .
+ thus "0 < Max (abs ` X ` {..k}) + 1" by simp
+ next
+ fix n :: nat
+ show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1"
+ proof (rule linorder_le_cases)
+ assume "n \<le> k"
+ hence "\<bar>X n\<bar> \<le> Max (abs ` X ` {..k})" by simp
+ thus "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" by simp
+ next
+ assume "k \<le> n"
+ have "\<bar>X n\<bar> = \<bar>X k + (X n - X k)\<bar>" by simp
+ also have "\<bar>X k + (X n - X k)\<bar> \<le> \<bar>X k\<bar> + \<bar>X n - X k\<bar>"
+ by (rule abs_triangle_ineq)
+ also have "\<dots> < Max (abs ` X ` {..k}) + 1"
+ by (rule add_le_less_mono, simp, simp add: k `k \<le> n`)
+ finally show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" .
+ qed
+ qed
+qed
+
+lemma cauchy_mult [simp]:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "cauchy (\<lambda>n. X n * Y n)"
+proof (rule cauchyI)
+ fix r :: rat assume "0 < r"
+ then obtain u v where u: "0 < u" and v: "0 < v" and "r = u + v"
+ by (rule obtain_pos_sum)
+ obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
+ using cauchy_imp_bounded [OF X] by fast
+ obtain b where b: "0 < b" "\<forall>n. \<bar>Y n\<bar> < b"
+ using cauchy_imp_bounded [OF Y] by fast
+ obtain s t where s: "0 < s" and t: "0 < t" and r: "r = a * t + s * b"
+ proof
+ show "0 < v/b" using v b(1) by (rule divide_pos_pos)
+ show "0 < u/a" using u a(1) by (rule divide_pos_pos)
+ show "r = a * (u/a) + (v/b) * b"
+ using a(1) b(1) `r = u + v` by simp
+ qed
+ obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
+ using cauchyD [OF X s] ..
+ obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
+ using cauchyD [OF Y t] ..
+ have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>X m * Y m - X n * Y n\<bar> < r"
+ proof (clarsimp)
+ fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
+ 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>"
+ unfolding mult_diff_mult ..
+ also have "\<dots> \<le> \<bar>X m * (Y m - Y n)\<bar> + \<bar>(X m - X n) * Y n\<bar>"
+ by (rule abs_triangle_ineq)
+ also have "\<dots> = \<bar>X m\<bar> * \<bar>Y m - Y n\<bar> + \<bar>X m - X n\<bar> * \<bar>Y n\<bar>"
+ unfolding abs_mult ..
+ also have "\<dots> < a * t + s * b"
+ by (simp_all add: add_strict_mono mult_strict_mono' a b i j *)
+ finally show "\<bar>X m * Y m - X n * Y n\<bar> < r" unfolding r .
+ qed
+ thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m * Y m - X n * Y n\<bar> < r" ..
+qed
+
+lemma cauchy_not_vanishes_cases:
+ assumes X: "cauchy X"
+ assumes nz: "\<not> vanishes X"
+ shows "\<exists>b>0. \<exists>k. (\<forall>n\<ge>k. b < - X n) \<or> (\<forall>n\<ge>k. b < X n)"
+proof -
+ obtain r where "0 < r" and r: "\<forall>k. \<exists>n\<ge>k. r \<le> \<bar>X n\<bar>"
+ using nz unfolding vanishes_def by (auto simp add: not_less)
+ obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t"
+ using `0 < r` by (rule obtain_pos_sum)
+ obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
+ using cauchyD [OF X s] ..
+ obtain k where "i \<le> k" and "r \<le> \<bar>X k\<bar>"
+ using r by fast
+ have k: "\<forall>n\<ge>k. \<bar>X n - X k\<bar> < s"
+ using i `i \<le> k` by auto
+ have "X k \<le> - r \<or> r \<le> X k"
+ using `r \<le> \<bar>X k\<bar>` by auto
+ hence "(\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
+ unfolding `r = s + t` using k by auto
+ hence "\<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)" ..
+ thus "\<exists>t>0. \<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
+ using t by auto
+qed
+
+lemma cauchy_not_vanishes:
+ assumes X: "cauchy X"
+ assumes nz: "\<not> vanishes X"
+ shows "\<exists>b>0. \<exists>k. \<forall>n\<ge>k. b < \<bar>X n\<bar>"
+using cauchy_not_vanishes_cases [OF assms]
+by clarify (rule exI, erule conjI, rule_tac x=k in exI, auto)
+
+lemma cauchy_inverse [simp]:
+ assumes X: "cauchy X"
+ assumes nz: "\<not> vanishes X"
+ shows "cauchy (\<lambda>n. inverse (X n))"
+proof (rule cauchyI)
+ fix r :: rat assume "0 < r"
+ obtain b i where b: "0 < b" and i: "\<forall>n\<ge>i. b < \<bar>X n\<bar>"
+ using cauchy_not_vanishes [OF X nz] by fast
+ from b i have nz: "\<forall>n\<ge>i. X n \<noteq> 0" by auto
+ obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b"
+ proof
+ show "0 < b * r * b"
+ by (simp add: `0 < r` b mult_pos_pos)
+ show "r = inverse b * (b * r * b) * inverse b"
+ using b by simp
+ qed
+ obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>X m - X n\<bar> < s"
+ using cauchyD [OF X s] ..
+ have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>inverse (X m) - inverse (X n)\<bar> < r"
+ proof (clarsimp)
+ fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
+ have "\<bar>inverse (X m) - inverse (X n)\<bar> =
+ inverse \<bar>X m\<bar> * \<bar>X m - X n\<bar> * inverse \<bar>X n\<bar>"
+ by (simp add: inverse_diff_inverse nz * abs_mult)
+ also have "\<dots> < inverse b * s * inverse b"
+ by (simp add: mult_strict_mono less_imp_inverse_less
+ mult_pos_pos i j b * s)
+ finally show "\<bar>inverse (X m) - inverse (X n)\<bar> < r" unfolding r .
+ qed
+ thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>inverse (X m) - inverse (X n)\<bar> < r" ..
+qed
+
+lemma vanishes_diff_inverse:
+ assumes X: "cauchy X" "\<not> vanishes X"
+ assumes Y: "cauchy Y" "\<not> vanishes Y"
+ assumes XY: "vanishes (\<lambda>n. X n - Y n)"
+ shows "vanishes (\<lambda>n. inverse (X n) - inverse (Y n))"
+proof (rule vanishesI)
+ fix r :: rat assume r: "0 < r"
+ obtain a i where a: "0 < a" and i: "\<forall>n\<ge>i. a < \<bar>X n\<bar>"
+ using cauchy_not_vanishes [OF X] by fast
+ obtain b j where b: "0 < b" and j: "\<forall>n\<ge>j. b < \<bar>Y n\<bar>"
+ using cauchy_not_vanishes [OF Y] by fast
+ obtain s where s: "0 < s" and "inverse a * s * inverse b = r"
+ proof
+ show "0 < a * r * b"
+ using a r b by (simp add: mult_pos_pos)
+ show "inverse a * (a * r * b) * inverse b = r"
+ using a r b by simp
+ qed
+ obtain k where k: "\<forall>n\<ge>k. \<bar>X n - Y n\<bar> < s"
+ using vanishesD [OF XY s] ..
+ have "\<forall>n\<ge>max (max i j) k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r"
+ proof (clarsimp)
+ fix n assume n: "i \<le> n" "j \<le> n" "k \<le> n"
+ have "X n \<noteq> 0" and "Y n \<noteq> 0"
+ using i j a b n by auto
+ hence "\<bar>inverse (X n) - inverse (Y n)\<bar> =
+ inverse \<bar>X n\<bar> * \<bar>X n - Y n\<bar> * inverse \<bar>Y n\<bar>"
+ by (simp add: inverse_diff_inverse abs_mult)
+ also have "\<dots> < inverse a * s * inverse b"
+ apply (intro mult_strict_mono' less_imp_inverse_less)
+ apply (simp_all add: a b i j k n mult_nonneg_nonneg)
+ done
+ also note `inverse a * s * inverse b = r`
+ finally show "\<bar>inverse (X n) - inverse (Y n)\<bar> < r" .
+ qed
+ thus "\<exists>k. \<forall>n\<ge>k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r" ..
+qed
+
+subsection {* Equivalence relation on Cauchy sequences *}
+
+definition realrel :: "(nat \<Rightarrow> rat) \<Rightarrow> (nat \<Rightarrow> rat) \<Rightarrow> bool"
+ where "realrel = (\<lambda>X Y. cauchy X \<and> cauchy Y \<and> vanishes (\<lambda>n. X n - Y n))"
+
+lemma realrelI [intro?]:
+ assumes "cauchy X" and "cauchy Y" and "vanishes (\<lambda>n. X n - Y n)"
+ shows "realrel X Y"
+ using assms unfolding realrel_def by simp
+
+lemma realrel_refl: "cauchy X \<Longrightarrow> realrel X X"
+ unfolding realrel_def by simp
+
+lemma symp_realrel: "symp realrel"
+ unfolding realrel_def
+ by (rule sympI, clarify, drule vanishes_minus, simp)
+
+lemma transp_realrel: "transp realrel"
+ unfolding realrel_def
+ apply (rule transpI, clarify)
+ apply (drule (1) vanishes_add)
+ apply (simp add: algebra_simps)
+ done
+
+lemma part_equivp_realrel: "part_equivp realrel"
+ by (fast intro: part_equivpI symp_realrel transp_realrel
+ realrel_refl cauchy_const)
+
+subsection {* The field of real numbers *}
+
+quotient_type real = "nat \<Rightarrow> rat" / partial: realrel
+ morphisms rep_real Real
+ by (rule part_equivp_realrel)
+
+lemma cr_real_eq: "pcr_real = (\<lambda>x y. cauchy x \<and> Real x = y)"
+ unfolding real.pcr_cr_eq cr_real_def realrel_def by auto
+
+lemma Real_induct [induct type: real]: (* TODO: generate automatically *)
+ assumes "\<And>X. cauchy X \<Longrightarrow> P (Real X)" shows "P x"
+proof (induct x)
+ case (1 X)
+ hence "cauchy X" by (simp add: realrel_def)
+ thus "P (Real X)" by (rule assms)
+qed
+
+lemma eq_Real:
+ "cauchy X \<Longrightarrow> cauchy Y \<Longrightarrow> Real X = Real Y \<longleftrightarrow> vanishes (\<lambda>n. X n - Y n)"
+ using real.rel_eq_transfer
+ unfolding real.pcr_cr_eq cr_real_def fun_rel_def realrel_def by simp
+
+declare real.forall_transfer [transfer_rule del]
+
+lemma forall_real_transfer [transfer_rule]: (* TODO: generate automatically *)
+ "(fun_rel (fun_rel pcr_real op =) op =)
+ (transfer_bforall cauchy) transfer_forall"
+ using real.forall_transfer
+ by (simp add: realrel_def)
+
+instantiation real :: field_inverse_zero
+begin
+
+lift_definition zero_real :: "real" is "\<lambda>n. 0"
+ by (simp add: realrel_refl)
+
+lift_definition one_real :: "real" is "\<lambda>n. 1"
+ by (simp add: realrel_refl)
+
+lift_definition plus_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n + Y n"
+ unfolding realrel_def add_diff_add
+ by (simp only: cauchy_add vanishes_add simp_thms)
+
+lift_definition uminus_real :: "real \<Rightarrow> real" is "\<lambda>X n. - X n"
+ unfolding realrel_def minus_diff_minus
+ by (simp only: cauchy_minus vanishes_minus simp_thms)
+
+lift_definition times_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n * Y n"
+ unfolding realrel_def mult_diff_mult
+ by (subst (4) mult_commute, simp only: cauchy_mult vanishes_add
+ vanishes_mult_bounded cauchy_imp_bounded simp_thms)
+
+lift_definition inverse_real :: "real \<Rightarrow> real"
+ is "\<lambda>X. if vanishes X then (\<lambda>n. 0) else (\<lambda>n. inverse (X n))"
+proof -
+ fix X Y assume "realrel X Y"
+ hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\<lambda>n. X n - Y n)"
+ unfolding realrel_def by simp_all
+ have "vanishes X \<longleftrightarrow> vanishes Y"
+ proof
+ assume "vanishes X"
+ from vanishes_diff [OF this XY] show "vanishes Y" by simp
+ next
+ assume "vanishes Y"
+ from vanishes_add [OF this XY] show "vanishes X" by simp
+ qed
+ thus "?thesis X Y"
+ unfolding realrel_def
+ by (simp add: vanishes_diff_inverse X Y XY)
+qed
+
+definition
+ "x - y = (x::real) + - y"
+
+definition
+ "x / y = (x::real) * inverse y"
+
+lemma add_Real:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "Real X + Real Y = Real (\<lambda>n. X n + Y n)"
+ using assms plus_real.transfer
+ unfolding cr_real_eq fun_rel_def by simp
+
+lemma minus_Real:
+ assumes X: "cauchy X"
+ shows "- Real X = Real (\<lambda>n. - X n)"
+ using assms uminus_real.transfer
+ unfolding cr_real_eq fun_rel_def by simp
+
+lemma diff_Real:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
+ unfolding minus_real_def diff_minus
+ by (simp add: minus_Real add_Real X Y)
+
+lemma mult_Real:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "Real X * Real Y = Real (\<lambda>n. X n * Y n)"
+ using assms times_real.transfer
+ unfolding cr_real_eq fun_rel_def by simp
+
+lemma inverse_Real:
+ assumes X: "cauchy X"
+ shows "inverse (Real X) =
+ (if vanishes X then 0 else Real (\<lambda>n. inverse (X n)))"
+ using assms inverse_real.transfer zero_real.transfer
+ unfolding cr_real_eq fun_rel_def by (simp split: split_if_asm, metis)
+
+instance proof
+ fix a b c :: real
+ show "a + b = b + a"
+ by transfer (simp add: add_ac realrel_def)
+ show "(a + b) + c = a + (b + c)"
+ by transfer (simp add: add_ac realrel_def)
+ show "0 + a = a"
+ by transfer (simp add: realrel_def)
+ show "- a + a = 0"
+ by transfer (simp add: realrel_def)
+ show "a - b = a + - b"
+ by (rule minus_real_def)
+ show "(a * b) * c = a * (b * c)"
+ by transfer (simp add: mult_ac realrel_def)
+ show "a * b = b * a"
+ by transfer (simp add: mult_ac realrel_def)
+ show "1 * a = a"
+ by transfer (simp add: mult_ac realrel_def)
+ show "(a + b) * c = a * c + b * c"
+ by transfer (simp add: distrib_right realrel_def)
+ show "(0\<Colon>real) \<noteq> (1\<Colon>real)"
+ by transfer (simp add: realrel_def)
+ show "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
+ apply transfer
+ apply (simp add: realrel_def)
+ apply (rule vanishesI)
+ apply (frule (1) cauchy_not_vanishes, clarify)
+ apply (rule_tac x=k in exI, clarify)
+ apply (drule_tac x=n in spec, simp)
+ done
+ show "a / b = a * inverse b"
+ by (rule divide_real_def)
+ show "inverse (0::real) = 0"
+ by transfer (simp add: realrel_def)
+qed
+
+end
+
+subsection {* Positive reals *}
+
+lift_definition positive :: "real \<Rightarrow> bool"
+ is "\<lambda>X. \<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
+proof -
+ { fix X Y
+ assume "realrel X Y"
+ hence XY: "vanishes (\<lambda>n. X n - Y n)"
+ unfolding realrel_def by simp_all
+ assume "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
+ then obtain r i where "0 < r" and i: "\<forall>n\<ge>i. r < X n"
+ by fast
+ obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
+ using `0 < r` by (rule obtain_pos_sum)
+ obtain j where j: "\<forall>n\<ge>j. \<bar>X n - Y n\<bar> < s"
+ using vanishesD [OF XY s] ..
+ have "\<forall>n\<ge>max i j. t < Y n"
+ proof (clarsimp)
+ fix n assume n: "i \<le> n" "j \<le> n"
+ have "\<bar>X n - Y n\<bar> < s" and "r < X n"
+ using i j n by simp_all
+ thus "t < Y n" unfolding r by simp
+ qed
+ hence "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < Y n" using t by fast
+ } note 1 = this
+ fix X Y assume "realrel X Y"
+ hence "realrel X Y" and "realrel Y X"
+ using symp_realrel unfolding symp_def by auto
+ thus "?thesis X Y"
+ by (safe elim!: 1)
+qed
+
+lemma positive_Real:
+ assumes X: "cauchy X"
+ shows "positive (Real X) \<longleftrightarrow> (\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n)"
+ using assms positive.transfer
+ unfolding cr_real_eq fun_rel_def by simp
+
+lemma positive_zero: "\<not> positive 0"
+ by transfer auto
+
+lemma positive_add:
+ "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x + y)"
+apply transfer
+apply (clarify, rename_tac a b i j)
+apply (rule_tac x="a + b" in exI, simp)
+apply (rule_tac x="max i j" in exI, clarsimp)
+apply (simp add: add_strict_mono)
+done
+
+lemma positive_mult:
+ "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x * y)"
+apply transfer
+apply (clarify, rename_tac a b i j)
+apply (rule_tac x="a * b" in exI, simp add: mult_pos_pos)
+apply (rule_tac x="max i j" in exI, clarsimp)
+apply (rule mult_strict_mono, auto)
+done
+
+lemma positive_minus:
+ "\<not> positive x \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> positive (- x)"
+apply transfer
+apply (simp add: realrel_def)
+apply (drule (1) cauchy_not_vanishes_cases, safe, fast, fast)
+done
+
+instantiation real :: linordered_field_inverse_zero
+begin
+
+definition
+ "x < y \<longleftrightarrow> positive (y - x)"
+
+definition
+ "x \<le> (y::real) \<longleftrightarrow> x < y \<or> x = y"
+
+definition
+ "abs (a::real) = (if a < 0 then - a else a)"
+
+definition
+ "sgn (a::real) = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
+
+instance proof
+ fix a b c :: real
+ show "\<bar>a\<bar> = (if a < 0 then - a else a)"
+ by (rule abs_real_def)
+ show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a"
+ unfolding less_eq_real_def less_real_def
+ by (auto, drule (1) positive_add, simp_all add: positive_zero)
+ show "a \<le> a"
+ unfolding less_eq_real_def by simp
+ show "a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> a \<le> c"
+ unfolding less_eq_real_def less_real_def
+ by (auto, drule (1) positive_add, simp add: algebra_simps)
+ show "a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> a = b"
+ unfolding less_eq_real_def less_real_def
+ by (auto, drule (1) positive_add, simp add: positive_zero)
+ show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
+ unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *)
+ (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \<equiv> b + - a *)
+ (* Should produce c + b - (c + a) \<equiv> b - a *)
+ show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
+ by (rule sgn_real_def)
+ show "a \<le> b \<or> b \<le> a"
+ unfolding less_eq_real_def less_real_def
+ by (auto dest!: positive_minus)
+ show "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
+ unfolding less_real_def
+ by (drule (1) positive_mult, simp add: algebra_simps)
+qed
+
+end
+
+instantiation real :: distrib_lattice
+begin
+
+definition
+ "(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"
+
+definition
+ "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
+
+instance proof
+qed (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
+
+end
+
+lemma of_nat_Real: "of_nat x = Real (\<lambda>n. of_nat x)"
+apply (induct x)
+apply (simp add: zero_real_def)
+apply (simp add: one_real_def add_Real)
+done
+
+lemma of_int_Real: "of_int x = Real (\<lambda>n. of_int x)"
+apply (cases x rule: int_diff_cases)
+apply (simp add: of_nat_Real diff_Real)
+done
+
+lemma of_rat_Real: "of_rat x = Real (\<lambda>n. x)"
+apply (induct x)
+apply (simp add: Fract_of_int_quotient of_rat_divide)
+apply (simp add: of_int_Real divide_inverse)
+apply (simp add: inverse_Real mult_Real)
+done
+
+instance real :: archimedean_field
+proof
+ fix x :: real
+ show "\<exists>z. x \<le> of_int z"
+ apply (induct x)
+ apply (frule cauchy_imp_bounded, clarify)
+ apply (rule_tac x="ceiling b + 1" in exI)
+ apply (rule less_imp_le)
+ apply (simp add: of_int_Real less_real_def diff_Real positive_Real)
+ apply (rule_tac x=1 in exI, simp add: algebra_simps)
+ apply (rule_tac x=0 in exI, clarsimp)
+ apply (rule le_less_trans [OF abs_ge_self])
+ apply (rule less_le_trans [OF _ le_of_int_ceiling])
+ apply simp
+ done
+qed
+
+instantiation real :: floor_ceiling
+begin
+
+definition [code del]:
+ "floor (x::real) = (THE z. of_int z \<le> x \<and> x < of_int (z + 1))"
+
+instance proof
+ fix x :: real
+ show "of_int (floor x) \<le> x \<and> x < of_int (floor x + 1)"
+ unfolding floor_real_def using floor_exists1 by (rule theI')
+qed
+
+end
+
+subsection {* Completeness *}
+
+lemma not_positive_Real:
+ assumes X: "cauchy X"
+ shows "\<not> positive (Real X) \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> r)"
+unfolding positive_Real [OF X]
+apply (auto, unfold not_less)
+apply (erule obtain_pos_sum)
+apply (drule_tac x=s in spec, simp)
+apply (drule_tac r=t in cauchyD [OF X], clarify)
+apply (drule_tac x=k in spec, clarsimp)
+apply (rule_tac x=n in exI, clarify, rename_tac m)
+apply (drule_tac x=m in spec, simp)
+apply (drule_tac x=n in spec, simp)
+apply (drule spec, drule (1) mp, clarify, rename_tac i)
+apply (rule_tac x="max i k" in exI, simp)
+done
+
+lemma le_Real:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "Real X \<le> Real Y = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r)"
+unfolding not_less [symmetric, where 'a=real] less_real_def
+apply (simp add: diff_Real not_positive_Real X Y)
+apply (simp add: diff_le_eq add_ac)
+done
+
+lemma le_RealI:
+ assumes Y: "cauchy Y"
+ shows "\<forall>n. x \<le> of_rat (Y n) \<Longrightarrow> x \<le> Real Y"
+proof (induct x)
+ fix X assume X: "cauchy X" and "\<forall>n. Real X \<le> of_rat (Y n)"
+ hence le: "\<And>m r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. X n \<le> Y m + r"
+ by (simp add: of_rat_Real le_Real)
+ {
+ fix r :: rat assume "0 < r"
+ then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
+ by (rule obtain_pos_sum)
+ obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>Y m - Y n\<bar> < s"
+ using cauchyD [OF Y s] ..
+ obtain j where j: "\<forall>n\<ge>j. X n \<le> Y i + t"
+ using le [OF t] ..
+ have "\<forall>n\<ge>max i j. X n \<le> Y n + r"
+ proof (clarsimp)
+ fix n assume n: "i \<le> n" "j \<le> n"
+ have "X n \<le> Y i + t" using n j by simp
+ moreover have "\<bar>Y i - Y n\<bar> < s" using n i by simp
+ ultimately show "X n \<le> Y n + r" unfolding r by simp
+ qed
+ hence "\<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r" ..
+ }
+ thus "Real X \<le> Real Y"
+ by (simp add: of_rat_Real le_Real X Y)
+qed
+
+lemma Real_leI:
+ assumes X: "cauchy X"
+ assumes le: "\<forall>n. of_rat (X n) \<le> y"
+ shows "Real X \<le> y"
+proof -
+ have "- y \<le> - Real X"
+ by (simp add: minus_Real X le_RealI of_rat_minus le)
+ thus ?thesis by simp
+qed
+
+lemma less_RealD:
+ assumes Y: "cauchy Y"
+ shows "x < Real Y \<Longrightarrow> \<exists>n. x < of_rat (Y n)"
+by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y])
+
+lemma of_nat_less_two_power:
+ "of_nat n < (2::'a::linordered_idom) ^ n"
+apply (induct n)
+apply simp
+apply (subgoal_tac "(1::'a) \<le> 2 ^ n")
+apply (drule (1) add_le_less_mono, simp)
+apply simp
+done
+
+lemma complete_real:
+ fixes S :: "real set"
+ assumes "\<exists>x. x \<in> S" and "\<exists>z. \<forall>x\<in>S. x \<le> z"
+ shows "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
+proof -
+ obtain x where x: "x \<in> S" using assms(1) ..
+ obtain z where z: "\<forall>x\<in>S. x \<le> z" using assms(2) ..
+
+ def P \<equiv> "\<lambda>x. \<forall>y\<in>S. y \<le> of_rat x"
+ obtain a where a: "\<not> P a"
+ proof
+ have "of_int (floor (x - 1)) \<le> x - 1" by (rule of_int_floor_le)
+ also have "x - 1 < x" by simp
+ finally have "of_int (floor (x - 1)) < x" .
+ hence "\<not> x \<le> of_int (floor (x - 1))" by (simp only: not_le)
+ then show "\<not> P (of_int (floor (x - 1)))"
+ unfolding P_def of_rat_of_int_eq using x by fast
+ qed
+ obtain b where b: "P b"
+ proof
+ show "P (of_int (ceiling z))"
+ unfolding P_def of_rat_of_int_eq
+ proof
+ fix y assume "y \<in> S"
+ hence "y \<le> z" using z by simp
+ also have "z \<le> of_int (ceiling z)" by (rule le_of_int_ceiling)
+ finally show "y \<le> of_int (ceiling z)" .
+ qed
+ qed
+
+ def avg \<equiv> "\<lambda>x y :: rat. x/2 + y/2"
+ def bisect \<equiv> "\<lambda>(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)"
+ def A \<equiv> "\<lambda>n. fst ((bisect ^^ n) (a, b))"
+ def B \<equiv> "\<lambda>n. snd ((bisect ^^ n) (a, b))"
+ def C \<equiv> "\<lambda>n. avg (A n) (B n)"
+ have A_0 [simp]: "A 0 = a" unfolding A_def by simp
+ have B_0 [simp]: "B 0 = b" unfolding B_def by simp
+ have A_Suc [simp]: "\<And>n. A (Suc n) = (if P (C n) then A n else C n)"
+ unfolding A_def B_def C_def bisect_def split_def by simp
+ have B_Suc [simp]: "\<And>n. B (Suc n) = (if P (C n) then C n else B n)"
+ unfolding A_def B_def C_def bisect_def split_def by simp
+
+ have width: "\<And>n. B n - A n = (b - a) / 2^n"
+ apply (simp add: eq_divide_eq)
+ apply (induct_tac n, simp)
+ apply (simp add: C_def avg_def algebra_simps)
+ done
+
+ have twos: "\<And>y r :: rat. 0 < r \<Longrightarrow> \<exists>n. y / 2 ^ n < r"
+ apply (simp add: divide_less_eq)
+ apply (subst mult_commute)
+ apply (frule_tac y=y in ex_less_of_nat_mult)
+ apply clarify
+ apply (rule_tac x=n in exI)
+ apply (erule less_trans)
+ apply (rule mult_strict_right_mono)
+ apply (rule le_less_trans [OF _ of_nat_less_two_power])
+ apply simp
+ apply assumption
+ done
+
+ have PA: "\<And>n. \<not> P (A n)"
+ by (induct_tac n, simp_all add: a)
+ have PB: "\<And>n. P (B n)"
+ by (induct_tac n, simp_all add: b)
+ have ab: "a < b"
+ using a b unfolding P_def
+ apply (clarsimp simp add: not_le)
+ apply (drule (1) bspec)
+ apply (drule (1) less_le_trans)
+ apply (simp add: of_rat_less)
+ done
+ have AB: "\<And>n. A n < B n"
+ by (induct_tac n, simp add: ab, simp add: C_def avg_def)
+ have A_mono: "\<And>i j. i \<le> j \<Longrightarrow> A i \<le> A j"
+ apply (auto simp add: le_less [where 'a=nat])
+ apply (erule less_Suc_induct)
+ apply (clarsimp simp add: C_def avg_def)
+ apply (simp add: add_divide_distrib [symmetric])
+ apply (rule AB [THEN less_imp_le])
+ apply simp
+ done
+ have B_mono: "\<And>i j. i \<le> j \<Longrightarrow> B j \<le> B i"
+ apply (auto simp add: le_less [where 'a=nat])
+ apply (erule less_Suc_induct)
+ apply (clarsimp simp add: C_def avg_def)
+ apply (simp add: add_divide_distrib [symmetric])
+ apply (rule AB [THEN less_imp_le])
+ apply simp
+ done
+ have cauchy_lemma:
+ "\<And>X. \<forall>n. \<forall>i\<ge>n. A n \<le> X i \<and> X i \<le> B n \<Longrightarrow> cauchy X"
+ apply (rule cauchyI)
+ apply (drule twos [where y="b - a"])
+ apply (erule exE)
+ apply (rule_tac x=n in exI, clarify, rename_tac i j)
+ apply (rule_tac y="B n - A n" in le_less_trans) defer
+ apply (simp add: width)
+ apply (drule_tac x=n in spec)
+ apply (frule_tac x=i in spec, drule (1) mp)
+ apply (frule_tac x=j in spec, drule (1) mp)
+ apply (frule A_mono, drule B_mono)
+ apply (frule A_mono, drule B_mono)
+ apply arith
+ done
+ have "cauchy A"
+ apply (rule cauchy_lemma [rule_format])
+ apply (simp add: A_mono)
+ apply (erule order_trans [OF less_imp_le [OF AB] B_mono])
+ done
+ have "cauchy B"
+ apply (rule cauchy_lemma [rule_format])
+ apply (simp add: B_mono)
+ apply (erule order_trans [OF A_mono less_imp_le [OF AB]])
+ done
+ have 1: "\<forall>x\<in>S. x \<le> Real B"
+ proof
+ fix x assume "x \<in> S"
+ then show "x \<le> Real B"
+ using PB [unfolded P_def] `cauchy B`
+ by (simp add: le_RealI)
+ qed
+ have 2: "\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> Real A \<le> z"
+ apply clarify
+ apply (erule contrapos_pp)
+ apply (simp add: not_le)
+ apply (drule less_RealD [OF `cauchy A`], clarify)
+ apply (subgoal_tac "\<not> P (A n)")
+ apply (simp add: P_def not_le, clarify)
+ apply (erule rev_bexI)
+ apply (erule (1) less_trans)
+ apply (simp add: PA)
+ done
+ have "vanishes (\<lambda>n. (b - a) / 2 ^ n)"
+ proof (rule vanishesI)
+ fix r :: rat assume "0 < r"
+ then obtain k where k: "\<bar>b - a\<bar> / 2 ^ k < r"
+ using twos by fast
+ have "\<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r"
+ proof (clarify)
+ fix n assume n: "k \<le> n"
+ have "\<bar>(b - a) / 2 ^ n\<bar> = \<bar>b - a\<bar> / 2 ^ n"
+ by simp
+ also have "\<dots> \<le> \<bar>b - a\<bar> / 2 ^ k"
+ using n by (simp add: divide_left_mono mult_pos_pos)
+ also note k
+ finally show "\<bar>(b - a) / 2 ^ n\<bar> < r" .
+ qed
+ thus "\<exists>k. \<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r" ..
+ qed
+ hence 3: "Real B = Real A"
+ by (simp add: eq_Real `cauchy A` `cauchy B` width)
+ show "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
+ using 1 2 3 by (rule_tac x="Real B" in exI, simp)
+qed
+
+
+instantiation real :: conditional_complete_linorder
begin
+subsection{*Supremum of a set of reals*}
+
+definition
+ Sup_real_def: "Sup X \<equiv> LEAST z::real. \<forall>x\<in>X. x\<le>z"
+
+definition
+ Inf_real_def: "Inf (X::real set) \<equiv> - Sup (uminus ` X)"
+
+instance
+proof
+ { fix z x :: real and X :: "real set"
+ assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
+ 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"
+ using complete_real[of X] by blast
+ then show "x \<le> Sup X"
+ unfolding Sup_real_def by (rule LeastI2_order) (auto simp: x) }
+ note Sup_upper = this
+
+ { fix z :: real and X :: "real set"
+ assume x: "X \<noteq> {}" and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
+ 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"
+ using complete_real[of X] by blast
+ then have "Sup X = s"
+ unfolding Sup_real_def by (best intro: Least_equality)
+ also with s z have "... \<le> z"
+ by blast
+ finally show "Sup X \<le> z" . }
+ note Sup_least = this
+
+ { fix x z :: real and X :: "real set"
+ assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
+ have "-x \<le> Sup (uminus ` X)"
+ by (rule Sup_upper[of _ _ "- z"]) (auto simp add: image_iff x z)
+ then show "Inf X \<le> x"
+ by (auto simp add: Inf_real_def) }
+
+ { fix z :: real and X :: "real set"
+ assume x: "X \<noteq> {}" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
+ have "Sup (uminus ` X) \<le> -z"
+ using x z by (force intro: Sup_least)
+ then show "z \<le> Inf X"
+ by (auto simp add: Inf_real_def) }
+qed
+end
+
+text {*
+ \medskip Completeness properties using @{text "isUb"}, @{text "isLub"}:
+*}
+
+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"
+ by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro: cSup_upper)
+
+
+subsection {* Hiding implementation details *}
+
+hide_const (open) vanishes cauchy positive Real
+
+declare Real_induct [induct del]
+declare Abs_real_induct [induct del]
+declare Abs_real_cases [cases del]
+
+lemmas [transfer_rule del] =
+ real.All_transfer real.Ex_transfer real.rel_eq_transfer forall_real_transfer
+ zero_real.transfer one_real.transfer plus_real.transfer uminus_real.transfer
+ times_real.transfer inverse_real.transfer positive.transfer real.right_unique
+ real.right_total
+
+subsection{*More Lemmas*}
+
+text {* BH: These lemmas should not be necessary; they should be
+covered by existing simp rules and simplification procedures. *}
+
+lemma real_mult_left_cancel: "(c::real) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
+by simp (* redundant with mult_cancel_left *)
+
+lemma real_mult_right_cancel: "(c::real) \<noteq> 0 ==> (a*c=b*c) = (a=b)"
+by simp (* redundant with mult_cancel_right *)
+
+lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
+by simp (* solved by linordered_ring_less_cancel_factor simproc *)
+
+lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
+by simp (* solved by linordered_ring_le_cancel_factor simproc *)
+
+lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
+by simp (* solved by linordered_ring_le_cancel_factor simproc *)
+
+
+subsection {* Embedding numbers into the Reals *}
+
+abbreviation
+ real_of_nat :: "nat \<Rightarrow> real"
+where
+ "real_of_nat \<equiv> of_nat"
+
+abbreviation
+ real_of_int :: "int \<Rightarrow> real"
+where
+ "real_of_int \<equiv> of_int"
+
+abbreviation
+ real_of_rat :: "rat \<Rightarrow> real"
+where
+ "real_of_rat \<equiv> of_rat"
+
+consts
+ (*overloaded constant for injecting other types into "real"*)
+ real :: "'a => real"
+
+defs (overloaded)
+ real_of_nat_def [code_unfold]: "real == real_of_nat"
+ real_of_int_def [code_unfold]: "real == real_of_int"
+
+declare [[coercion_enabled]]
+declare [[coercion "real::nat\<Rightarrow>real"]]
+declare [[coercion "real::int\<Rightarrow>real"]]
+declare [[coercion "int"]]
+
+declare [[coercion_map map]]
+declare [[coercion_map "% f g h x. g (h (f x))"]]
+declare [[coercion_map "% f g (x,y) . (f x, g y)"]]
+
+lemma real_eq_of_nat: "real = of_nat"
+ unfolding real_of_nat_def ..
+
+lemma real_eq_of_int: "real = of_int"
+ unfolding real_of_int_def ..
+
+lemma real_of_int_zero [simp]: "real (0::int) = 0"
+by (simp add: real_of_int_def)
+
+lemma real_of_one [simp]: "real (1::int) = (1::real)"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_power [simp]: "real (x ^ n) = real (x::int) ^ n"
+by (simp add: real_of_int_def of_int_power)
+
+lemmas power_real_of_int = real_of_int_power [symmetric]
+
+lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))"
+ apply (subst real_eq_of_int)+
+ apply (rule of_int_setsum)
+done
+
+lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) =
+ (PROD x:A. real(f x))"
+ apply (subst real_eq_of_int)+
+ apply (rule of_int_setprod)
+done
+
+lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \<le> real y) = (x \<le> y)"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)"
+by (simp add: real_of_int_def)
+
+lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)"
+by (simp add: real_of_int_def)
+
+lemma one_less_real_of_int_cancel_iff: "1 < real (i :: int) \<longleftrightarrow> 1 < i"
+ unfolding real_of_one[symmetric] real_of_int_less_iff ..
+
+lemma one_le_real_of_int_cancel_iff: "1 \<le> real (i :: int) \<longleftrightarrow> 1 \<le> i"
+ unfolding real_of_one[symmetric] real_of_int_le_iff ..
+
+lemma real_of_int_less_one_cancel_iff: "real (i :: int) < 1 \<longleftrightarrow> i < 1"
+ unfolding real_of_one[symmetric] real_of_int_less_iff ..
+
+lemma real_of_int_le_one_cancel_iff: "real (i :: int) \<le> 1 \<longleftrightarrow> i \<le> 1"
+ unfolding real_of_one[symmetric] real_of_int_le_iff ..
+
+lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))"
+by (auto simp add: abs_if)
+
+lemma int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)"
+ apply (subgoal_tac "real n + 1 = real (n + 1)")
+ apply (simp del: real_of_int_add)
+ apply auto
+done
+
+lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)"
+ apply (subgoal_tac "real m + 1 = real (m + 1)")
+ apply (simp del: real_of_int_add)
+ apply simp
+done
+
+lemma real_of_int_div_aux: "(real (x::int)) / (real d) =
+ real (x div d) + (real (x mod d)) / (real d)"
+proof -
+ have "x = (x div d) * d + x mod d"
+ by auto
+ then have "real x = real (x div d) * real d + real(x mod d)"
+ by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym])
+ then have "real x / real d = ... / real d"
+ by simp
+ then show ?thesis
+ by (auto simp add: add_divide_distrib algebra_simps)
+qed
+
+lemma real_of_int_div: "(d :: int) dvd n ==>
+ real(n div d) = real n / real d"
+ apply (subst real_of_int_div_aux)
+ apply simp
+ apply (simp add: dvd_eq_mod_eq_0)
+done
+
+lemma real_of_int_div2:
+ "0 <= real (n::int) / real (x) - real (n div x)"
+ apply (case_tac "x = 0")
+ apply simp
+ apply (case_tac "0 < x")
+ apply (simp add: algebra_simps)
+ apply (subst real_of_int_div_aux)
+ apply simp
+ apply (subst zero_le_divide_iff)
+ apply auto
+ apply (simp add: algebra_simps)
+ apply (subst real_of_int_div_aux)
+ apply simp
+ apply (subst zero_le_divide_iff)
+ apply auto
+done
+
+lemma real_of_int_div3:
+ "real (n::int) / real (x) - real (n div x) <= 1"
+ apply (simp add: algebra_simps)
+ apply (subst real_of_int_div_aux)
+ apply (auto simp add: divide_le_eq intro: order_less_imp_le)
+done
+
+lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x"
+by (insert real_of_int_div2 [of n x], simp)
+
+lemma Ints_real_of_int [simp]: "real (x::int) \<in> Ints"
+unfolding real_of_int_def by (rule Ints_of_int)
+
+
+subsection{*Embedding the Naturals into the Reals*}
+
+lemma real_of_nat_zero [simp]: "real (0::nat) = 0"
+by (simp add: real_of_nat_def)
+
+lemma real_of_nat_1 [simp]: "real (1::nat) = 1"
+by (simp add: real_of_nat_def)
+
+lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)"
+by (simp add: real_of_nat_def)
+
+lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n"
+by (simp add: real_of_nat_def)
+
+(*Not for addsimps: often the LHS is used to represent a positive natural*)
+lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)"
+by (simp add: real_of_nat_def)
+
+lemma real_of_nat_less_iff [iff]:
+ "(real (n::nat) < real m) = (n < m)"
+by (simp add: real_of_nat_def)
+
+lemma real_of_nat_le_iff [iff]: "(real (n::nat) \<le> real m) = (n \<le> m)"
+by (simp add: real_of_nat_def)
+
+lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)"
+by (simp add: real_of_nat_def)
+
+lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)"
+by (simp add: real_of_nat_def del: of_nat_Suc)
+
+lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n"
+by (simp add: real_of_nat_def of_nat_mult)
+
+lemma real_of_nat_power [simp]: "real (m ^ n) = real (m::nat) ^ n"
+by (simp add: real_of_nat_def of_nat_power)
+
+lemmas power_real_of_nat = real_of_nat_power [symmetric]
+
+lemma real_of_nat_setsum [simp]: "real ((SUM x:A. f x)::nat) =
+ (SUM x:A. real(f x))"
+ apply (subst real_eq_of_nat)+
+ apply (rule of_nat_setsum)
+done
+
+lemma real_of_nat_setprod [simp]: "real ((PROD x:A. f x)::nat) =
+ (PROD x:A. real(f x))"
+ apply (subst real_eq_of_nat)+
+ apply (rule of_nat_setprod)
+done
+
+lemma real_of_card: "real (card A) = setsum (%x.1) A"
+ apply (subst card_eq_setsum)
+ apply (subst real_of_nat_setsum)
+ apply simp
+done
+
+lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)"
+by (simp add: real_of_nat_def)
+
+lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)"
+by (simp add: real_of_nat_def)
+
+lemma real_of_nat_diff: "n \<le> m ==> real (m - n) = real (m::nat) - real n"
+by (simp add: add: real_of_nat_def of_nat_diff)
+
+lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)"
+by (auto simp: real_of_nat_def)
+
+lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \<le> 0) = (n = 0)"
+by (simp add: add: real_of_nat_def)
+
+lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0"
+by (simp add: add: real_of_nat_def)
+
+lemma nat_less_real_le: "((n::nat) < m) = (real n + 1 <= real m)"
+ apply (subgoal_tac "real n + 1 = real (Suc n)")
+ apply simp
+ apply (auto simp add: real_of_nat_Suc)
+done
+
+lemma nat_le_real_less: "((n::nat) <= m) = (real n < real m + 1)"
+ apply (subgoal_tac "real m + 1 = real (Suc m)")
+ apply (simp add: less_Suc_eq_le)
+ apply (simp add: real_of_nat_Suc)
+done
+
+lemma real_of_nat_div_aux: "(real (x::nat)) / (real d) =
+ real (x div d) + (real (x mod d)) / (real d)"
+proof -
+ have "x = (x div d) * d + x mod d"
+ by auto
+ then have "real x = real (x div d) * real d + real(x mod d)"
+ by (simp only: real_of_nat_mult [THEN sym] real_of_nat_add [THEN sym])
+ then have "real x / real d = \<dots> / real d"
+ by simp
+ then show ?thesis
+ by (auto simp add: add_divide_distrib algebra_simps)
+qed
+
+lemma real_of_nat_div: "(d :: nat) dvd n ==>
+ real(n div d) = real n / real d"
+ by (subst real_of_nat_div_aux)
+ (auto simp add: dvd_eq_mod_eq_0 [symmetric])
+
+lemma real_of_nat_div2:
+ "0 <= real (n::nat) / real (x) - real (n div x)"
+apply (simp add: algebra_simps)
+apply (subst real_of_nat_div_aux)
+apply simp
+apply (subst zero_le_divide_iff)
+apply simp
+done
+
+lemma real_of_nat_div3:
+ "real (n::nat) / real (x) - real (n div x) <= 1"
+apply(case_tac "x = 0")
+apply (simp)
+apply (simp add: algebra_simps)
+apply (subst real_of_nat_div_aux)
+apply simp
+done
+
+lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x"
+by (insert real_of_nat_div2 [of n x], simp)
+
+lemma real_of_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n"
+by (simp add: real_of_int_def real_of_nat_def)
+
+lemma real_nat_eq_real [simp]: "0 <= x ==> real(nat x) = real x"
+ apply (subgoal_tac "real(int(nat x)) = real(nat x)")
+ apply force
+ apply (simp only: real_of_int_of_nat_eq)
+done
+
+lemma Nats_real_of_nat [simp]: "real (n::nat) \<in> Nats"
+unfolding real_of_nat_def by (rule of_nat_in_Nats)
+
+lemma Ints_real_of_nat [simp]: "real (n::nat) \<in> Ints"
+unfolding real_of_nat_def by (rule Ints_of_nat)
+
+subsection {* The Archimedean Property of the Reals *}
+
+theorem reals_Archimedean:
+ assumes x_pos: "0 < x"
+ shows "\<exists>n. inverse (real (Suc n)) < x"
+ unfolding real_of_nat_def using x_pos
+ by (rule ex_inverse_of_nat_Suc_less)
+
+lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
+ unfolding real_of_nat_def by (rule ex_less_of_nat)
+
+lemma reals_Archimedean3:
+ assumes x_greater_zero: "0 < x"
+ shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"
+ unfolding real_of_nat_def using `0 < x`
+ by (auto intro: ex_less_of_nat_mult)
+
+
+subsection{* Rationals *}
+
+lemma Rats_real_nat[simp]: "real(n::nat) \<in> \<rat>"
+by (simp add: real_eq_of_nat)
+
+
+lemma Rats_eq_int_div_int:
+ "\<rat> = { real(i::int)/real(j::int) |i j. j \<noteq> 0}" (is "_ = ?S")
+proof
+ show "\<rat> \<subseteq> ?S"
+ proof
+ fix x::real assume "x : \<rat>"
+ then obtain r where "x = of_rat r" unfolding Rats_def ..
+ have "of_rat r : ?S"
+ by (cases r)(auto simp add:of_rat_rat real_eq_of_int)
+ thus "x : ?S" using `x = of_rat r` by simp
+ qed
+next
+ show "?S \<subseteq> \<rat>"
+ proof(auto simp:Rats_def)
+ fix i j :: int assume "j \<noteq> 0"
+ hence "real i / real j = of_rat(Fract i j)"
+ by (simp add:of_rat_rat real_eq_of_int)
+ thus "real i / real j \<in> range of_rat" by blast
+ qed
+qed
+
+lemma Rats_eq_int_div_nat:
+ "\<rat> = { real(i::int)/real(n::nat) |i n. n \<noteq> 0}"
+proof(auto simp:Rats_eq_int_div_int)
+ fix i j::int assume "j \<noteq> 0"
+ show "EX (i'::int) (n::nat). real i/real j = real i'/real n \<and> 0<n"
+ proof cases
+ assume "j>0"
+ hence "real i/real j = real i/real(nat j) \<and> 0<nat j"
+ by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
+ thus ?thesis by blast
+ next
+ assume "~ j>0"
+ hence "real i/real j = real(-i)/real(nat(-j)) \<and> 0<nat(-j)" using `j\<noteq>0`
+ by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
+ thus ?thesis by blast
+ qed
+next
+ fix i::int and n::nat assume "0 < n"
+ hence "real i/real n = real i/real(int n) \<and> int n \<noteq> 0" by simp
+ thus "\<exists>(i'::int) j::int. real i/real n = real i'/real j \<and> j \<noteq> 0" by blast
+qed
+
+lemma Rats_abs_nat_div_natE:
+ assumes "x \<in> \<rat>"
+ obtains m n :: nat
+ where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd m n = 1"
+proof -
+ from `x \<in> \<rat>` obtain i::int and n::nat where "n \<noteq> 0" and "x = real i / real n"
+ by(auto simp add: Rats_eq_int_div_nat)
+ hence "\<bar>x\<bar> = real(nat(abs i)) / real n" by simp
+ then obtain m :: nat where x_rat: "\<bar>x\<bar> = real m / real n" by blast
+ let ?gcd = "gcd m n"
+ from `n\<noteq>0` have gcd: "?gcd \<noteq> 0" by simp
+ let ?k = "m div ?gcd"
+ let ?l = "n div ?gcd"
+ let ?gcd' = "gcd ?k ?l"
+ have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m"
+ by (rule dvd_mult_div_cancel)
+ have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
+ by (rule dvd_mult_div_cancel)
+ from `n\<noteq>0` and gcd_l have "?l \<noteq> 0" by (auto iff del: neq0_conv)
+ moreover
+ have "\<bar>x\<bar> = real ?k / real ?l"
+ proof -
+ from gcd have "real ?k / real ?l =
+ real (?gcd * ?k) / real (?gcd * ?l)" by simp
+ also from gcd_k and gcd_l have "\<dots> = real m / real n" by simp
+ also from x_rat have "\<dots> = \<bar>x\<bar>" ..
+ finally show ?thesis ..
+ qed
+ moreover
+ have "?gcd' = 1"
+ proof -
+ have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)"
+ by (rule gcd_mult_distrib_nat)
+ with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
+ with gcd show ?thesis by auto
+ qed
+ ultimately show ?thesis ..
+qed
+
+subsection{*Density of the Rational Reals in the Reals*}
+
+text{* This density proof is due to Stefan Richter and was ported by TN. The
+original source is \emph{Real Analysis} by H.L. Royden.
+It employs the Archimedean property of the reals. *}
+
+lemma Rats_dense_in_real:
+ fixes x :: real
+ assumes "x < y" shows "\<exists>r\<in>\<rat>. x < r \<and> r < y"
+proof -
+ from `x<y` have "0 < y-x" by simp
+ with reals_Archimedean obtain q::nat
+ where q: "inverse (real q) < y-x" and "0 < q" by auto
+ def p \<equiv> "ceiling (y * real q) - 1"
+ def r \<equiv> "of_int p / real q"
+ from q have "x < y - inverse (real q)" by simp
+ also have "y - inverse (real q) \<le> r"
+ unfolding r_def p_def
+ by (simp add: le_divide_eq left_diff_distrib le_of_int_ceiling `0 < q`)
+ finally have "x < r" .
+ moreover have "r < y"
+ unfolding r_def p_def
+ by (simp add: divide_less_eq diff_less_eq `0 < q`
+ less_ceiling_iff [symmetric])
+ moreover from r_def have "r \<in> \<rat>" by simp
+ ultimately show ?thesis by fast
+qed
+
+
+
+subsection{*Numerals and Arithmetic*}
+
+lemma [code_abbrev]:
+ "real_of_int (numeral k) = numeral k"
+ "real_of_int (neg_numeral k) = neg_numeral k"
+ by simp_all
+
+text{*Collapse applications of @{term real} to @{term number_of}*}
+lemma real_numeral [simp]:
+ "real (numeral v :: int) = numeral v"
+ "real (neg_numeral v :: int) = neg_numeral v"
+by (simp_all add: real_of_int_def)
+
+lemma real_of_nat_numeral [simp]:
+ "real (numeral v :: nat) = numeral v"
+by (simp add: real_of_nat_def)
+
+declaration {*
+ K (Lin_Arith.add_inj_thms [@{thm real_of_nat_le_iff} RS iffD2, @{thm real_of_nat_inject} RS iffD2]
+ (* not needed because x < (y::nat) can be rewritten as Suc x <= y: real_of_nat_less_iff RS iffD2 *)
+ #> Lin_Arith.add_inj_thms [@{thm real_of_int_le_iff} RS iffD2, @{thm real_of_int_inject} RS iffD2]
+ (* not needed because x < (y::int) can be rewritten as x + 1 <= y: real_of_int_less_iff RS iffD2 *)
+ #> Lin_Arith.add_simps [@{thm real_of_nat_zero}, @{thm real_of_nat_Suc}, @{thm real_of_nat_add},
+ @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one},
+ @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff},
+ @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq},
+ @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)}]
+ #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \<Rightarrow> real"})
+ #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \<Rightarrow> real"}))
+*}
+
+
+subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
+
+lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)"
+by arith
+
+text {* FIXME: redundant with @{text add_eq_0_iff} below *}
+lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)"
+by auto
+
+lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)"
+by auto
+
+lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)"
+by auto
+
+lemma real_add_le_0_iff: "(x+y \<le> (0::real)) = (y \<le> -x)"
+by auto
+
+lemma real_0_le_add_iff: "((0::real) \<le> x+y) = (-x \<le> y)"
+by auto
+
+subsection {* Lemmas about powers *}
+
+text {* FIXME: declare this in Rings.thy or not at all *}
+declare abs_mult_self [simp]
+
+(* used by Import/HOL/real.imp *)
+lemma two_realpow_ge_one: "(1::real) \<le> 2 ^ n"
+by simp
+
+lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n"
+apply (induct "n")
+apply (auto simp add: real_of_nat_Suc)
+apply (subst mult_2)
+apply (erule add_less_le_mono)
+apply (rule two_realpow_ge_one)
+done
+
+text {* TODO: no longer real-specific; rename and move elsewhere *}
+lemma realpow_Suc_le_self:
+ fixes r :: "'a::linordered_semidom"
+ shows "[| 0 \<le> r; r \<le> 1 |] ==> r ^ Suc n \<le> r"
+by (insert power_decreasing [of 1 "Suc n" r], simp)
+
+text {* TODO: no longer real-specific; rename and move elsewhere *}
+lemma realpow_minus_mult:
+ fixes x :: "'a::monoid_mult"
+ shows "0 < n \<Longrightarrow> x ^ (n - 1) * x = x ^ n"
+by (simp add: power_commutes split add: nat_diff_split)
+
+text {* FIXME: declare this [simp] for all types, or not at all *}
+lemma real_two_squares_add_zero_iff [simp]:
+ "(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)"
+by (rule sum_squares_eq_zero_iff)
+
+text {* FIXME: declare this [simp] for all types, or not at all *}
+lemma realpow_two_sum_zero_iff [simp]:
+ "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)"
+by (rule sum_power2_eq_zero_iff)
+
+lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))"
+by (rule_tac y = 0 in order_trans, auto)
+
+lemma realpow_square_minus_le [simp]: "-(u ^ 2) \<le> (x::real) ^ 2"
+by (auto simp add: power2_eq_square)
+
+
+lemma numeral_power_le_real_of_nat_cancel_iff[simp]:
+ "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::nat) ^ n \<le> a"
+ unfolding real_of_nat_le_iff[symmetric] by simp
+
+lemma real_of_nat_le_numeral_power_cancel_iff[simp]:
+ "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::nat) ^ n"
+ unfolding real_of_nat_le_iff[symmetric] by simp
+
+lemma numeral_power_le_real_of_int_cancel_iff[simp]:
+ "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::int) ^ n \<le> a"
+ unfolding real_of_int_le_iff[symmetric] by simp
+
+lemma real_of_int_le_numeral_power_cancel_iff[simp]:
+ "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::int) ^ n"
+ unfolding real_of_int_le_iff[symmetric] by simp
+
+lemma neg_numeral_power_le_real_of_int_cancel_iff[simp]:
+ "(neg_numeral x::real) ^ n \<le> real a \<longleftrightarrow> (neg_numeral x::int) ^ n \<le> a"
+ unfolding real_of_int_le_iff[symmetric] by simp
+
+lemma real_of_int_le_neg_numeral_power_cancel_iff[simp]:
+ "real a \<le> (neg_numeral x::real) ^ n \<longleftrightarrow> a \<le> (neg_numeral x::int) ^ n"
+ unfolding real_of_int_le_iff[symmetric] by simp
+
+subsection{*Density of the Reals*}
+
+lemma real_lbound_gt_zero:
+ "[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2"
+apply (rule_tac x = " (min d1 d2) /2" in exI)
+apply (simp add: min_def)
+done
+
+
+text{*Similar results are proved in @{text Fields}*}
+lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
+ by auto
+
+lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y"
+ by auto
+
+lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
+ by simp
+
+subsection{*Absolute Value Function for the Reals*}
+
+lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))"
+by (simp add: abs_if)
+
+(* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *)
+lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
+by (force simp add: abs_le_iff)
+
+lemma abs_add_one_gt_zero: "(0::real) < 1 + abs(x)"
+by (simp add: abs_if)
+
+lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
+by (rule abs_of_nonneg [OF real_of_nat_ge_zero])
+
+lemma abs_add_one_not_less_self: "~ abs(x) + (1::real) < x"
+by simp
+
+lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)"
+by simp
+
+
+subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
+
+(* FIXME: theorems for negative numerals *)
+lemma numeral_less_real_of_int_iff [simp]:
+ "((numeral n) < real (m::int)) = (numeral n < m)"
+apply auto
+apply (rule real_of_int_less_iff [THEN iffD1])
+apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
+done
+
+lemma numeral_less_real_of_int_iff2 [simp]:
+ "(real (m::int) < (numeral n)) = (m < numeral n)"
+apply auto
+apply (rule real_of_int_less_iff [THEN iffD1])
+apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
+done
+
+lemma numeral_le_real_of_int_iff [simp]:
+ "((numeral n) \<le> real (m::int)) = (numeral n \<le> m)"
+by (simp add: linorder_not_less [symmetric])
+
+lemma numeral_le_real_of_int_iff2 [simp]:
+ "(real (m::int) \<le> (numeral n)) = (m \<le> numeral n)"
+by (simp add: linorder_not_less [symmetric])
+
+lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"
+unfolding real_of_nat_def by simp
+
+lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"
+unfolding real_of_nat_def by (simp add: floor_minus)
+
+lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"
+unfolding real_of_int_def by simp
+
+lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"
+unfolding real_of_int_def by (simp add: floor_minus)
+
+lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
+unfolding real_of_int_def by (rule floor_exists)
+
+lemma lemma_floor:
+ assumes a1: "real m \<le> r" and a2: "r < real n + 1"
+ shows "m \<le> (n::int)"
+proof -
+ have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)
+ also have "... = real (n + 1)" by simp
+ finally have "m < n + 1" by (simp only: real_of_int_less_iff)
+ thus ?thesis by arith
+qed
+
+lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"
+unfolding real_of_int_def by (rule of_int_floor_le)
+
+lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \<le> x"
+by (auto intro: lemma_floor)
+
+lemma real_of_int_floor_cancel [simp]:
+ "(real (floor x) = x) = (\<exists>n::int. x = real n)"
+ using floor_real_of_int by metis
+
+lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
+ unfolding real_of_int_def using floor_unique [of n x] by simp
+
+lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"
+ unfolding real_of_int_def by (rule floor_unique)
+
+lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n"
+apply (rule inj_int [THEN injD])
+apply (simp add: real_of_nat_Suc)
+apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"])
+done
+
+lemma floor_eq4: "[| real n \<le> x; x < real (Suc n) |] ==> nat(floor x) = n"
+apply (drule order_le_imp_less_or_eq)
+apply (auto intro: floor_eq3)
+done
+
+lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
+ unfolding real_of_int_def using floor_correct [of r] by simp
+
+lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"
+ unfolding real_of_int_def using floor_correct [of r] by simp
+
+lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"
+ unfolding real_of_int_def using floor_correct [of r] by simp
+
+lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"
+ unfolding real_of_int_def using floor_correct [of r] by simp
+
+lemma le_floor: "real a <= x ==> a <= floor x"
+ unfolding real_of_int_def by (simp add: le_floor_iff)
+
+lemma real_le_floor: "a <= floor x ==> real a <= x"
+ unfolding real_of_int_def by (simp add: le_floor_iff)
+
+lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
+ unfolding real_of_int_def by (rule le_floor_iff)
+
+lemma floor_less_eq: "(floor x < a) = (x < real a)"
+ unfolding real_of_int_def by (rule floor_less_iff)
+
+lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
+ unfolding real_of_int_def by (rule less_floor_iff)
+
+lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
+ unfolding real_of_int_def by (rule floor_le_iff)
+
+lemma floor_add [simp]: "floor (x + real a) = floor x + a"
+ unfolding real_of_int_def by (rule floor_add_of_int)
+
+lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
+ unfolding real_of_int_def by (rule floor_diff_of_int)
+
+lemma le_mult_floor:
+ assumes "0 \<le> (a :: real)" and "0 \<le> b"
+ shows "floor a * floor b \<le> floor (a * b)"
+proof -
+ have "real (floor a) \<le> a"
+ and "real (floor b) \<le> b" by auto
+ hence "real (floor a * floor b) \<le> a * b"
+ using assms by (auto intro!: mult_mono)
+ also have "a * b < real (floor (a * b) + 1)" by auto
+ finally show ?thesis unfolding real_of_int_less_iff by simp
+qed
+
+lemma floor_divide_eq_div:
+ "floor (real a / real b) = a div b"
+proof cases
+ assume "b \<noteq> 0 \<or> b dvd a"
+ with real_of_int_div3[of a b] show ?thesis
+ by (auto simp: real_of_int_div[symmetric] intro!: floor_eq2 real_of_int_div4 neq_le_trans)
+ (metis add_left_cancel zero_neq_one real_of_int_div_aux real_of_int_inject
+ real_of_int_zero_cancel right_inverse_eq div_self mod_div_trivial)
+qed (auto simp: real_of_int_div)
+
+lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
+ unfolding real_of_nat_def by simp
+
+lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"
+ unfolding real_of_int_def by (rule le_of_int_ceiling)
+
+lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"
+ unfolding real_of_int_def by simp
+
+lemma real_of_int_ceiling_cancel [simp]:
+ "(real (ceiling x) = x) = (\<exists>n::int. x = real n)"
+ using ceiling_real_of_int by metis
+
+lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"
+ unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
+
+lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"
+ unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
+
+lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n |] ==> ceiling x = n"
+ unfolding real_of_int_def using ceiling_unique [of n x] by simp
+
+lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
+ unfolding real_of_int_def using ceiling_correct [of r] by simp
+
+lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"
+ unfolding real_of_int_def using ceiling_correct [of r] by simp
+
+lemma ceiling_le: "x <= real a ==> ceiling x <= a"
+ unfolding real_of_int_def by (simp add: ceiling_le_iff)
+
+lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"
+ unfolding real_of_int_def by (simp add: ceiling_le_iff)
+
+lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
+ unfolding real_of_int_def by (rule ceiling_le_iff)
+
+lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
+ unfolding real_of_int_def by (rule less_ceiling_iff)
+
+lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
+ unfolding real_of_int_def by (rule ceiling_less_iff)
+
+lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
+ unfolding real_of_int_def by (rule le_ceiling_iff)
+
+lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
+ unfolding real_of_int_def by (rule ceiling_add_of_int)
+
+lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
+ unfolding real_of_int_def by (rule ceiling_diff_of_int)
+
+
+subsubsection {* Versions for the natural numbers *}
+
+definition
+ natfloor :: "real => nat" where
+ "natfloor x = nat(floor x)"
+
+definition
+ natceiling :: "real => nat" where
+ "natceiling x = nat(ceiling x)"
+
+lemma natfloor_zero [simp]: "natfloor 0 = 0"
+ by (unfold natfloor_def, simp)
+
+lemma natfloor_one [simp]: "natfloor 1 = 1"
+ by (unfold natfloor_def, simp)
+
+lemma zero_le_natfloor [simp]: "0 <= natfloor x"
+ by (unfold natfloor_def, simp)
+
+lemma natfloor_numeral_eq [simp]: "natfloor (numeral n) = numeral n"
+ by (unfold natfloor_def, simp)
+
+lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"
+ by (unfold natfloor_def, simp)
+
+lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"
+ by (unfold natfloor_def, simp)
+
+lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"
+ unfolding natfloor_def by simp
+
+lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
+ unfolding natfloor_def by (intro nat_mono floor_mono)
+
+lemma le_natfloor: "real x <= a ==> x <= natfloor a"
+ apply (unfold natfloor_def)
+ apply (subst nat_int [THEN sym])
+ apply (rule nat_mono)
+ apply (rule le_floor)
+ apply simp
+done
+
+lemma natfloor_less_iff: "0 \<le> x \<Longrightarrow> natfloor x < n \<longleftrightarrow> x < real n"
+ unfolding natfloor_def real_of_nat_def
+ by (simp add: nat_less_iff floor_less_iff)
+
+lemma less_natfloor:
+ assumes "0 \<le> x" and "x < real (n :: nat)"
+ shows "natfloor x < n"
+ using assms by (simp add: natfloor_less_iff)
+
+lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)"
+ apply (rule iffI)
+ apply (rule order_trans)
+ prefer 2
+ apply (erule real_natfloor_le)
+ apply (subst real_of_nat_le_iff)
+ apply assumption
+ apply (erule le_natfloor)
+done
+
+lemma le_natfloor_eq_numeral [simp]:
+ "~ neg((numeral n)::int) ==> 0 <= x ==>
+ (numeral n <= natfloor x) = (numeral n <= x)"
+ apply (subst le_natfloor_eq, assumption)
+ apply simp
+done
+
+lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
+ apply (case_tac "0 <= x")
+ apply (subst le_natfloor_eq, assumption, simp)
+ apply (rule iffI)
+ apply (subgoal_tac "natfloor x <= natfloor 0")
+ apply simp
+ apply (rule natfloor_mono)
+ apply simp
+ apply simp
+done
+
+lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n"
+ unfolding natfloor_def by (simp add: floor_eq2 [where n="int n"])
+
+lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"
+ apply (case_tac "0 <= x")
+ apply (unfold natfloor_def)
+ apply simp
+ apply simp_all
+done
+
+lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
+using real_natfloor_add_one_gt by (simp add: algebra_simps)
+
+lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"
+ apply (subgoal_tac "z < real(natfloor z) + 1")
+ apply arith
+ apply (rule real_natfloor_add_one_gt)
+done
+
+lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a"
+ unfolding natfloor_def
+ unfolding real_of_int_of_nat_eq [symmetric] floor_add
+ by (simp add: nat_add_distrib)
+
+lemma natfloor_add_numeral [simp]:
+ "~neg ((numeral n)::int) ==> 0 <= x ==>
+ natfloor (x + numeral n) = natfloor x + numeral n"
+ by (simp add: natfloor_add [symmetric])
+
+lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1"
+ by (simp add: natfloor_add [symmetric] del: One_nat_def)
+
+lemma natfloor_subtract [simp]:
+ "natfloor(x - real a) = natfloor x - a"
+ unfolding natfloor_def real_of_int_of_nat_eq [symmetric] floor_subtract
+ by simp
+
+lemma natfloor_div_nat:
+ assumes "1 <= x" and "y > 0"
+ shows "natfloor (x / real y) = natfloor x div y"
+proof (rule natfloor_eq)
+ have "(natfloor x) div y * y \<le> natfloor x"
+ by (rule add_leD1 [where k="natfloor x mod y"], simp)
+ thus "real (natfloor x div y) \<le> x / real y"
+ using assms by (simp add: le_divide_eq le_natfloor_eq)
+ have "natfloor x < (natfloor x) div y * y + y"
+ apply (subst mod_div_equality [symmetric])
+ apply (rule add_strict_left_mono)
+ apply (rule mod_less_divisor)
+ apply fact
+ done
+ thus "x / real y < real (natfloor x div y) + 1"
+ using assms
+ by (simp add: divide_less_eq natfloor_less_iff distrib_right)
+qed
+
+lemma le_mult_natfloor:
+ shows "natfloor a * natfloor b \<le> natfloor (a * b)"
+ by (cases "0 <= a & 0 <= b")
+ (auto simp add: le_natfloor_eq mult_nonneg_nonneg mult_mono' real_natfloor_le natfloor_neg)
+
+lemma natceiling_zero [simp]: "natceiling 0 = 0"
+ by (unfold natceiling_def, simp)
+
+lemma natceiling_one [simp]: "natceiling 1 = 1"
+ by (unfold natceiling_def, simp)
+
+lemma zero_le_natceiling [simp]: "0 <= natceiling x"
+ by (unfold natceiling_def, simp)
+
+lemma natceiling_numeral_eq [simp]: "natceiling (numeral n) = numeral n"
+ by (unfold natceiling_def, simp)
+
+lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"
+ by (unfold natceiling_def, simp)
+
+lemma real_natceiling_ge: "x <= real(natceiling x)"
+ unfolding natceiling_def by (cases "x < 0", simp_all)
+
+lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"
+ unfolding natceiling_def by simp
+
+lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"
+ unfolding natceiling_def by (intro nat_mono ceiling_mono)
+
+lemma natceiling_le: "x <= real a ==> natceiling x <= a"
+ unfolding natceiling_def real_of_nat_def
+ by (simp add: nat_le_iff ceiling_le_iff)
+
+lemma natceiling_le_eq: "(natceiling x <= a) = (x <= real a)"
+ unfolding natceiling_def real_of_nat_def
+ by (simp add: nat_le_iff ceiling_le_iff)
+
+lemma natceiling_le_eq_numeral [simp]:
+ "~ neg((numeral n)::int) ==>
+ (natceiling x <= numeral n) = (x <= numeral n)"
+ by (simp add: natceiling_le_eq)
+
+lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"
+ unfolding natceiling_def
+ by (simp add: nat_le_iff ceiling_le_iff)
+
+lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1"
+ unfolding natceiling_def
+ by (simp add: ceiling_eq2 [where n="int n"])
+
+lemma natceiling_add [simp]: "0 <= x ==>
+ natceiling (x + real a) = natceiling x + a"
+ unfolding natceiling_def
+ unfolding real_of_int_of_nat_eq [symmetric] ceiling_add
+ by (simp add: nat_add_distrib)
+
+lemma natceiling_add_numeral [simp]:
+ "~ neg ((numeral n)::int) ==> 0 <= x ==>
+ natceiling (x + numeral n) = natceiling x + numeral n"
+ by (simp add: natceiling_add [symmetric])
+
+lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1"
+ by (simp add: natceiling_add [symmetric] del: One_nat_def)
+
+lemma natceiling_subtract [simp]: "natceiling(x - real a) = natceiling x - a"
+ unfolding natceiling_def real_of_int_of_nat_eq [symmetric] ceiling_subtract
+ by simp
+
+subsection {* Exponentiation with floor *}
+
+lemma floor_power:
+ assumes "x = real (floor x)"
+ shows "floor (x ^ n) = floor x ^ n"
+proof -
+ have *: "x ^ n = real (floor x ^ n)"
+ using assms by (induct n arbitrary: x) simp_all
+ show ?thesis unfolding real_of_int_inject[symmetric]
+ unfolding * floor_real_of_int ..
+qed
+
+lemma natfloor_power:
+ assumes "x = real (natfloor x)"
+ shows "natfloor (x ^ n) = natfloor x ^ n"
+proof -
+ from assms have "0 \<le> floor x" by auto
+ note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \<le> floor x`]]
+ from floor_power[OF this]
+ show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \<le> floor x`, symmetric]
+ by simp
+qed
+
+
+subsection {* Implementation of rational real numbers *}
+
+text {* Formal constructor *}
+
+definition Ratreal :: "rat \<Rightarrow> real" where
+ [code_abbrev, simp]: "Ratreal = of_rat"
+
+code_datatype Ratreal
+
+
+text {* Numerals *}
+
+lemma [code_abbrev]:
+ "(of_rat (of_int a) :: real) = of_int a"
+ by simp
+
+lemma [code_abbrev]:
+ "(of_rat 0 :: real) = 0"
+ by simp
+
+lemma [code_abbrev]:
+ "(of_rat 1 :: real) = 1"
+ by simp
+
+lemma [code_abbrev]:
+ "(of_rat (numeral k) :: real) = numeral k"
+ by simp
+
+lemma [code_abbrev]:
+ "(of_rat (neg_numeral k) :: real) = neg_numeral k"
+ by simp
+
+lemma [code_post]:
+ "(of_rat (0 / r) :: real) = 0"
+ "(of_rat (r / 0) :: real) = 0"
+ "(of_rat (1 / 1) :: real) = 1"
+ "(of_rat (numeral k / 1) :: real) = numeral k"
+ "(of_rat (neg_numeral k / 1) :: real) = neg_numeral k"
+ "(of_rat (1 / numeral k) :: real) = 1 / numeral k"
+ "(of_rat (1 / neg_numeral k) :: real) = 1 / neg_numeral k"
+ "(of_rat (numeral k / numeral l) :: real) = numeral k / numeral l"
+ "(of_rat (numeral k / neg_numeral l) :: real) = numeral k / neg_numeral l"
+ "(of_rat (neg_numeral k / numeral l) :: real) = neg_numeral k / numeral l"
+ "(of_rat (neg_numeral k / neg_numeral l) :: real) = neg_numeral k / neg_numeral l"
+ by (simp_all add: of_rat_divide)
+
+
+text {* Operations *}
+
+lemma zero_real_code [code]:
+ "0 = Ratreal 0"
+by simp
+
+lemma one_real_code [code]:
+ "1 = Ratreal 1"
+by simp
+
+instantiation real :: equal
+begin
+
+definition "HOL.equal (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
+
+instance proof
+qed (simp add: equal_real_def)
+
+lemma real_equal_code [code]:
+ "HOL.equal (Ratreal x) (Ratreal y) \<longleftrightarrow> HOL.equal x y"
+ by (simp add: equal_real_def equal)
+
+lemma [code nbe]:
+ "HOL.equal (x::real) x \<longleftrightarrow> True"
+ by (rule equal_refl)
+
+end
+
+lemma real_less_eq_code [code]: "Ratreal x \<le> Ratreal y \<longleftrightarrow> x \<le> y"
+ by (simp add: of_rat_less_eq)
+
+lemma real_less_code [code]: "Ratreal x < Ratreal y \<longleftrightarrow> x < y"
+ by (simp add: of_rat_less)
+
+lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)"
+ by (simp add: of_rat_add)
+
+lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)"
+ by (simp add: of_rat_mult)
+
+lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)"
+ by (simp add: of_rat_minus)
+
+lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)"
+ by (simp add: of_rat_diff)
+
+lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)"
+ by (simp add: of_rat_inverse)
+
+lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)"
+ by (simp add: of_rat_divide)
+
+lemma real_floor_code [code]: "floor (Ratreal x) = floor x"
+ 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)
+
+
+text {* Quickcheck *}
+
+definition (in term_syntax)
+ valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+ [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\<cdot>} k"
+
+notation fcomp (infixl "\<circ>>" 60)
+notation scomp (infixl "\<circ>\<rightarrow>" 60)
+
+instantiation real :: random
+begin
+
+definition
+ "Quickcheck_Random.random i = Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>r. Pair (valterm_ratreal r))"
+
+instance ..
+
+end
+
+no_notation fcomp (infixl "\<circ>>" 60)
+no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
+
+instantiation real :: exhaustive
+begin
+
+definition
+ "exhaustive_real f d = Quickcheck_Exhaustive.exhaustive (%r. f (Ratreal r)) d"
+
+instance ..
+
+end
+
+instantiation real :: full_exhaustive
+begin
+
+definition
+ "full_exhaustive_real f d = Quickcheck_Exhaustive.full_exhaustive (%r. f (valterm_ratreal r)) d"
+
+instance ..
+
+end
+
+instantiation real :: narrowing
+begin
+
+definition
+ "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing"
+
+instance ..
+
+end
+
+
+subsection {* Setup for Nitpick *}
+
+declaration {*
+ Nitpick_HOL.register_frac_type @{type_name real}
+ [(@{const_name zero_real_inst.zero_real}, @{const_name Nitpick.zero_frac}),
+ (@{const_name one_real_inst.one_real}, @{const_name Nitpick.one_frac}),
+ (@{const_name plus_real_inst.plus_real}, @{const_name Nitpick.plus_frac}),
+ (@{const_name times_real_inst.times_real}, @{const_name Nitpick.times_frac}),
+ (@{const_name uminus_real_inst.uminus_real}, @{const_name Nitpick.uminus_frac}),
+ (@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}),
+ (@{const_name ord_real_inst.less_real}, @{const_name Nitpick.less_frac}),
+ (@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})]
+*}
+
+lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real
+ ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real
+ times_real_inst.times_real uminus_real_inst.uminus_real
+ zero_real_inst.zero_real
+
ML_file "Tools/SMT/smt_real.ML"
setup SMT_Real.setup
--- a/src/HOL/RealDef.thy Tue Mar 26 14:38:44 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1704 +0,0 @@
-(* Title : HOL/RealDef.thy
- Author : Jacques D. Fleuriot, 1998
- Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
- Additional contributions by Jeremy Avigad
- Construction of Cauchy Reals by Brian Huffman, 2010
-*)
-
-header {* Development of the Reals using Cauchy Sequences *}
-
-theory RealDef
-imports Rat
-begin
-
-text {*
- This theory contains a formalization of the real numbers as
- equivalence classes of Cauchy sequences of rationals. See
- @{file "~~/src/HOL/ex/Dedekind_Real.thy"} for an alternative
- construction using Dedekind cuts.
-*}
-
-subsection {* Preliminary lemmas *}
-
-lemma add_diff_add:
- fixes a b c d :: "'a::ab_group_add"
- shows "(a + c) - (b + d) = (a - b) + (c - d)"
- by simp
-
-lemma minus_diff_minus:
- fixes a b :: "'a::ab_group_add"
- shows "- a - - b = - (a - b)"
- by simp
-
-lemma mult_diff_mult:
- fixes x y a b :: "'a::ring"
- shows "(x * y - a * b) = x * (y - b) + (x - a) * b"
- by (simp add: algebra_simps)
-
-lemma inverse_diff_inverse:
- fixes a b :: "'a::division_ring"
- assumes "a \<noteq> 0" and "b \<noteq> 0"
- shows "inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
- using assms by (simp add: algebra_simps)
-
-lemma obtain_pos_sum:
- fixes r :: rat assumes r: "0 < r"
- obtains s t where "0 < s" and "0 < t" and "r = s + t"
-proof
- from r show "0 < r/2" by simp
- from r show "0 < r/2" by simp
- show "r = r/2 + r/2" by simp
-qed
-
-subsection {* Sequences that converge to zero *}
-
-definition
- vanishes :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
-where
- "vanishes X = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r)"
-
-lemma vanishesI: "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r) \<Longrightarrow> vanishes X"
- unfolding vanishes_def by simp
-
-lemma vanishesD: "\<lbrakk>vanishes X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r"
- unfolding vanishes_def by simp
-
-lemma vanishes_const [simp]: "vanishes (\<lambda>n. c) \<longleftrightarrow> c = 0"
- unfolding vanishes_def
- apply (cases "c = 0", auto)
- apply (rule exI [where x="\<bar>c\<bar>"], auto)
- done
-
-lemma vanishes_minus: "vanishes X \<Longrightarrow> vanishes (\<lambda>n. - X n)"
- unfolding vanishes_def by simp
-
-lemma vanishes_add:
- assumes X: "vanishes X" and Y: "vanishes Y"
- shows "vanishes (\<lambda>n. X n + Y n)"
-proof (rule vanishesI)
- fix r :: rat assume "0 < r"
- then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
- by (rule obtain_pos_sum)
- obtain i where i: "\<forall>n\<ge>i. \<bar>X n\<bar> < s"
- using vanishesD [OF X s] ..
- obtain j where j: "\<forall>n\<ge>j. \<bar>Y n\<bar> < t"
- using vanishesD [OF Y t] ..
- have "\<forall>n\<ge>max i j. \<bar>X n + Y n\<bar> < r"
- proof (clarsimp)
- fix n assume n: "i \<le> n" "j \<le> n"
- have "\<bar>X n + Y n\<bar> \<le> \<bar>X n\<bar> + \<bar>Y n\<bar>" by (rule abs_triangle_ineq)
- also have "\<dots> < s + t" by (simp add: add_strict_mono i j n)
- finally show "\<bar>X n + Y n\<bar> < r" unfolding r .
- qed
- thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n + Y n\<bar> < r" ..
-qed
-
-lemma vanishes_diff:
- assumes X: "vanishes X" and Y: "vanishes Y"
- shows "vanishes (\<lambda>n. X n - Y n)"
-unfolding diff_minus by (intro vanishes_add vanishes_minus X Y)
-
-lemma vanishes_mult_bounded:
- assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
- assumes Y: "vanishes (\<lambda>n. Y n)"
- shows "vanishes (\<lambda>n. X n * Y n)"
-proof (rule vanishesI)
- fix r :: rat assume r: "0 < r"
- obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
- using X by fast
- obtain b where b: "0 < b" "r = a * b"
- proof
- show "0 < r / a" using r a by (simp add: divide_pos_pos)
- show "r = a * (r / a)" using a by simp
- qed
- obtain k where k: "\<forall>n\<ge>k. \<bar>Y n\<bar> < b"
- using vanishesD [OF Y b(1)] ..
- have "\<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r"
- by (simp add: b(2) abs_mult mult_strict_mono' a k)
- thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r" ..
-qed
-
-subsection {* Cauchy sequences *}
-
-definition
- cauchy :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
-where
- "cauchy X \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r)"
-
-lemma cauchyI:
- "(\<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"
- unfolding cauchy_def by simp
-
-lemma cauchyD:
- "\<lbrakk>cauchy X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r"
- unfolding cauchy_def by simp
-
-lemma cauchy_const [simp]: "cauchy (\<lambda>n. x)"
- unfolding cauchy_def by simp
-
-lemma cauchy_add [simp]:
- assumes X: "cauchy X" and Y: "cauchy Y"
- shows "cauchy (\<lambda>n. X n + Y n)"
-proof (rule cauchyI)
- fix r :: rat assume "0 < r"
- then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
- by (rule obtain_pos_sum)
- obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
- using cauchyD [OF X s] ..
- obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
- using cauchyD [OF Y t] ..
- have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r"
- proof (clarsimp)
- fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
- have "\<bar>(X m + Y m) - (X n + Y n)\<bar> \<le> \<bar>X m - X n\<bar> + \<bar>Y m - Y n\<bar>"
- unfolding add_diff_add by (rule abs_triangle_ineq)
- also have "\<dots> < s + t"
- by (rule add_strict_mono, simp_all add: i j *)
- finally show "\<bar>(X m + Y m) - (X n + Y n)\<bar> < r" unfolding r .
- qed
- thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r" ..
-qed
-
-lemma cauchy_minus [simp]:
- assumes X: "cauchy X"
- shows "cauchy (\<lambda>n. - X n)"
-using assms unfolding cauchy_def
-unfolding minus_diff_minus abs_minus_cancel .
-
-lemma cauchy_diff [simp]:
- assumes X: "cauchy X" and Y: "cauchy Y"
- shows "cauchy (\<lambda>n. X n - Y n)"
-using assms unfolding diff_minus by simp
-
-lemma cauchy_imp_bounded:
- assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
-proof -
- obtain k where k: "\<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < 1"
- using cauchyD [OF assms zero_less_one] ..
- show "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
- proof (intro exI conjI allI)
- have "0 \<le> \<bar>X 0\<bar>" by simp
- also have "\<bar>X 0\<bar> \<le> Max (abs ` X ` {..k})" by simp
- finally have "0 \<le> Max (abs ` X ` {..k})" .
- thus "0 < Max (abs ` X ` {..k}) + 1" by simp
- next
- fix n :: nat
- show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1"
- proof (rule linorder_le_cases)
- assume "n \<le> k"
- hence "\<bar>X n\<bar> \<le> Max (abs ` X ` {..k})" by simp
- thus "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" by simp
- next
- assume "k \<le> n"
- have "\<bar>X n\<bar> = \<bar>X k + (X n - X k)\<bar>" by simp
- also have "\<bar>X k + (X n - X k)\<bar> \<le> \<bar>X k\<bar> + \<bar>X n - X k\<bar>"
- by (rule abs_triangle_ineq)
- also have "\<dots> < Max (abs ` X ` {..k}) + 1"
- by (rule add_le_less_mono, simp, simp add: k `k \<le> n`)
- finally show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" .
- qed
- qed
-qed
-
-lemma cauchy_mult [simp]:
- assumes X: "cauchy X" and Y: "cauchy Y"
- shows "cauchy (\<lambda>n. X n * Y n)"
-proof (rule cauchyI)
- fix r :: rat assume "0 < r"
- then obtain u v where u: "0 < u" and v: "0 < v" and "r = u + v"
- by (rule obtain_pos_sum)
- obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
- using cauchy_imp_bounded [OF X] by fast
- obtain b where b: "0 < b" "\<forall>n. \<bar>Y n\<bar> < b"
- using cauchy_imp_bounded [OF Y] by fast
- obtain s t where s: "0 < s" and t: "0 < t" and r: "r = a * t + s * b"
- proof
- show "0 < v/b" using v b(1) by (rule divide_pos_pos)
- show "0 < u/a" using u a(1) by (rule divide_pos_pos)
- show "r = a * (u/a) + (v/b) * b"
- using a(1) b(1) `r = u + v` by simp
- qed
- obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
- using cauchyD [OF X s] ..
- obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
- using cauchyD [OF Y t] ..
- have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>X m * Y m - X n * Y n\<bar> < r"
- proof (clarsimp)
- fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
- 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>"
- unfolding mult_diff_mult ..
- also have "\<dots> \<le> \<bar>X m * (Y m - Y n)\<bar> + \<bar>(X m - X n) * Y n\<bar>"
- by (rule abs_triangle_ineq)
- also have "\<dots> = \<bar>X m\<bar> * \<bar>Y m - Y n\<bar> + \<bar>X m - X n\<bar> * \<bar>Y n\<bar>"
- unfolding abs_mult ..
- also have "\<dots> < a * t + s * b"
- by (simp_all add: add_strict_mono mult_strict_mono' a b i j *)
- finally show "\<bar>X m * Y m - X n * Y n\<bar> < r" unfolding r .
- qed
- thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m * Y m - X n * Y n\<bar> < r" ..
-qed
-
-lemma cauchy_not_vanishes_cases:
- assumes X: "cauchy X"
- assumes nz: "\<not> vanishes X"
- shows "\<exists>b>0. \<exists>k. (\<forall>n\<ge>k. b < - X n) \<or> (\<forall>n\<ge>k. b < X n)"
-proof -
- obtain r where "0 < r" and r: "\<forall>k. \<exists>n\<ge>k. r \<le> \<bar>X n\<bar>"
- using nz unfolding vanishes_def by (auto simp add: not_less)
- obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t"
- using `0 < r` by (rule obtain_pos_sum)
- obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
- using cauchyD [OF X s] ..
- obtain k where "i \<le> k" and "r \<le> \<bar>X k\<bar>"
- using r by fast
- have k: "\<forall>n\<ge>k. \<bar>X n - X k\<bar> < s"
- using i `i \<le> k` by auto
- have "X k \<le> - r \<or> r \<le> X k"
- using `r \<le> \<bar>X k\<bar>` by auto
- hence "(\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
- unfolding `r = s + t` using k by auto
- hence "\<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)" ..
- thus "\<exists>t>0. \<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
- using t by auto
-qed
-
-lemma cauchy_not_vanishes:
- assumes X: "cauchy X"
- assumes nz: "\<not> vanishes X"
- shows "\<exists>b>0. \<exists>k. \<forall>n\<ge>k. b < \<bar>X n\<bar>"
-using cauchy_not_vanishes_cases [OF assms]
-by clarify (rule exI, erule conjI, rule_tac x=k in exI, auto)
-
-lemma cauchy_inverse [simp]:
- assumes X: "cauchy X"
- assumes nz: "\<not> vanishes X"
- shows "cauchy (\<lambda>n. inverse (X n))"
-proof (rule cauchyI)
- fix r :: rat assume "0 < r"
- obtain b i where b: "0 < b" and i: "\<forall>n\<ge>i. b < \<bar>X n\<bar>"
- using cauchy_not_vanishes [OF X nz] by fast
- from b i have nz: "\<forall>n\<ge>i. X n \<noteq> 0" by auto
- obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b"
- proof
- show "0 < b * r * b"
- by (simp add: `0 < r` b mult_pos_pos)
- show "r = inverse b * (b * r * b) * inverse b"
- using b by simp
- qed
- obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>X m - X n\<bar> < s"
- using cauchyD [OF X s] ..
- have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>inverse (X m) - inverse (X n)\<bar> < r"
- proof (clarsimp)
- fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
- have "\<bar>inverse (X m) - inverse (X n)\<bar> =
- inverse \<bar>X m\<bar> * \<bar>X m - X n\<bar> * inverse \<bar>X n\<bar>"
- by (simp add: inverse_diff_inverse nz * abs_mult)
- also have "\<dots> < inverse b * s * inverse b"
- by (simp add: mult_strict_mono less_imp_inverse_less
- mult_pos_pos i j b * s)
- finally show "\<bar>inverse (X m) - inverse (X n)\<bar> < r" unfolding r .
- qed
- thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>inverse (X m) - inverse (X n)\<bar> < r" ..
-qed
-
-lemma vanishes_diff_inverse:
- assumes X: "cauchy X" "\<not> vanishes X"
- assumes Y: "cauchy Y" "\<not> vanishes Y"
- assumes XY: "vanishes (\<lambda>n. X n - Y n)"
- shows "vanishes (\<lambda>n. inverse (X n) - inverse (Y n))"
-proof (rule vanishesI)
- fix r :: rat assume r: "0 < r"
- obtain a i where a: "0 < a" and i: "\<forall>n\<ge>i. a < \<bar>X n\<bar>"
- using cauchy_not_vanishes [OF X] by fast
- obtain b j where b: "0 < b" and j: "\<forall>n\<ge>j. b < \<bar>Y n\<bar>"
- using cauchy_not_vanishes [OF Y] by fast
- obtain s where s: "0 < s" and "inverse a * s * inverse b = r"
- proof
- show "0 < a * r * b"
- using a r b by (simp add: mult_pos_pos)
- show "inverse a * (a * r * b) * inverse b = r"
- using a r b by simp
- qed
- obtain k where k: "\<forall>n\<ge>k. \<bar>X n - Y n\<bar> < s"
- using vanishesD [OF XY s] ..
- have "\<forall>n\<ge>max (max i j) k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r"
- proof (clarsimp)
- fix n assume n: "i \<le> n" "j \<le> n" "k \<le> n"
- have "X n \<noteq> 0" and "Y n \<noteq> 0"
- using i j a b n by auto
- hence "\<bar>inverse (X n) - inverse (Y n)\<bar> =
- inverse \<bar>X n\<bar> * \<bar>X n - Y n\<bar> * inverse \<bar>Y n\<bar>"
- by (simp add: inverse_diff_inverse abs_mult)
- also have "\<dots> < inverse a * s * inverse b"
- apply (intro mult_strict_mono' less_imp_inverse_less)
- apply (simp_all add: a b i j k n mult_nonneg_nonneg)
- done
- also note `inverse a * s * inverse b = r`
- finally show "\<bar>inverse (X n) - inverse (Y n)\<bar> < r" .
- qed
- thus "\<exists>k. \<forall>n\<ge>k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r" ..
-qed
-
-subsection {* Equivalence relation on Cauchy sequences *}
-
-definition realrel :: "(nat \<Rightarrow> rat) \<Rightarrow> (nat \<Rightarrow> rat) \<Rightarrow> bool"
- where "realrel = (\<lambda>X Y. cauchy X \<and> cauchy Y \<and> vanishes (\<lambda>n. X n - Y n))"
-
-lemma realrelI [intro?]:
- assumes "cauchy X" and "cauchy Y" and "vanishes (\<lambda>n. X n - Y n)"
- shows "realrel X Y"
- using assms unfolding realrel_def by simp
-
-lemma realrel_refl: "cauchy X \<Longrightarrow> realrel X X"
- unfolding realrel_def by simp
-
-lemma symp_realrel: "symp realrel"
- unfolding realrel_def
- by (rule sympI, clarify, drule vanishes_minus, simp)
-
-lemma transp_realrel: "transp realrel"
- unfolding realrel_def
- apply (rule transpI, clarify)
- apply (drule (1) vanishes_add)
- apply (simp add: algebra_simps)
- done
-
-lemma part_equivp_realrel: "part_equivp realrel"
- by (fast intro: part_equivpI symp_realrel transp_realrel
- realrel_refl cauchy_const)
-
-subsection {* The field of real numbers *}
-
-quotient_type real = "nat \<Rightarrow> rat" / partial: realrel
- morphisms rep_real Real
- by (rule part_equivp_realrel)
-
-lemma cr_real_eq: "pcr_real = (\<lambda>x y. cauchy x \<and> Real x = y)"
- unfolding real.pcr_cr_eq cr_real_def realrel_def by auto
-
-lemma Real_induct [induct type: real]: (* TODO: generate automatically *)
- assumes "\<And>X. cauchy X \<Longrightarrow> P (Real X)" shows "P x"
-proof (induct x)
- case (1 X)
- hence "cauchy X" by (simp add: realrel_def)
- thus "P (Real X)" by (rule assms)
-qed
-
-lemma eq_Real:
- "cauchy X \<Longrightarrow> cauchy Y \<Longrightarrow> Real X = Real Y \<longleftrightarrow> vanishes (\<lambda>n. X n - Y n)"
- using real.rel_eq_transfer
- unfolding real.pcr_cr_eq cr_real_def fun_rel_def realrel_def by simp
-
-declare real.forall_transfer [transfer_rule del]
-
-lemma forall_real_transfer [transfer_rule]: (* TODO: generate automatically *)
- "(fun_rel (fun_rel pcr_real op =) op =)
- (transfer_bforall cauchy) transfer_forall"
- using real.forall_transfer
- by (simp add: realrel_def)
-
-instantiation real :: field_inverse_zero
-begin
-
-lift_definition zero_real :: "real" is "\<lambda>n. 0"
- by (simp add: realrel_refl)
-
-lift_definition one_real :: "real" is "\<lambda>n. 1"
- by (simp add: realrel_refl)
-
-lift_definition plus_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n + Y n"
- unfolding realrel_def add_diff_add
- by (simp only: cauchy_add vanishes_add simp_thms)
-
-lift_definition uminus_real :: "real \<Rightarrow> real" is "\<lambda>X n. - X n"
- unfolding realrel_def minus_diff_minus
- by (simp only: cauchy_minus vanishes_minus simp_thms)
-
-lift_definition times_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n * Y n"
- unfolding realrel_def mult_diff_mult
- by (subst (4) mult_commute, simp only: cauchy_mult vanishes_add
- vanishes_mult_bounded cauchy_imp_bounded simp_thms)
-
-lift_definition inverse_real :: "real \<Rightarrow> real"
- is "\<lambda>X. if vanishes X then (\<lambda>n. 0) else (\<lambda>n. inverse (X n))"
-proof -
- fix X Y assume "realrel X Y"
- hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\<lambda>n. X n - Y n)"
- unfolding realrel_def by simp_all
- have "vanishes X \<longleftrightarrow> vanishes Y"
- proof
- assume "vanishes X"
- from vanishes_diff [OF this XY] show "vanishes Y" by simp
- next
- assume "vanishes Y"
- from vanishes_add [OF this XY] show "vanishes X" by simp
- qed
- thus "?thesis X Y"
- unfolding realrel_def
- by (simp add: vanishes_diff_inverse X Y XY)
-qed
-
-definition
- "x - y = (x::real) + - y"
-
-definition
- "x / y = (x::real) * inverse y"
-
-lemma add_Real:
- assumes X: "cauchy X" and Y: "cauchy Y"
- shows "Real X + Real Y = Real (\<lambda>n. X n + Y n)"
- using assms plus_real.transfer
- unfolding cr_real_eq fun_rel_def by simp
-
-lemma minus_Real:
- assumes X: "cauchy X"
- shows "- Real X = Real (\<lambda>n. - X n)"
- using assms uminus_real.transfer
- unfolding cr_real_eq fun_rel_def by simp
-
-lemma diff_Real:
- assumes X: "cauchy X" and Y: "cauchy Y"
- shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
- unfolding minus_real_def diff_minus
- by (simp add: minus_Real add_Real X Y)
-
-lemma mult_Real:
- assumes X: "cauchy X" and Y: "cauchy Y"
- shows "Real X * Real Y = Real (\<lambda>n. X n * Y n)"
- using assms times_real.transfer
- unfolding cr_real_eq fun_rel_def by simp
-
-lemma inverse_Real:
- assumes X: "cauchy X"
- shows "inverse (Real X) =
- (if vanishes X then 0 else Real (\<lambda>n. inverse (X n)))"
- using assms inverse_real.transfer zero_real.transfer
- unfolding cr_real_eq fun_rel_def by (simp split: split_if_asm, metis)
-
-instance proof
- fix a b c :: real
- show "a + b = b + a"
- by transfer (simp add: add_ac realrel_def)
- show "(a + b) + c = a + (b + c)"
- by transfer (simp add: add_ac realrel_def)
- show "0 + a = a"
- by transfer (simp add: realrel_def)
- show "- a + a = 0"
- by transfer (simp add: realrel_def)
- show "a - b = a + - b"
- by (rule minus_real_def)
- show "(a * b) * c = a * (b * c)"
- by transfer (simp add: mult_ac realrel_def)
- show "a * b = b * a"
- by transfer (simp add: mult_ac realrel_def)
- show "1 * a = a"
- by transfer (simp add: mult_ac realrel_def)
- show "(a + b) * c = a * c + b * c"
- by transfer (simp add: distrib_right realrel_def)
- show "(0\<Colon>real) \<noteq> (1\<Colon>real)"
- by transfer (simp add: realrel_def)
- show "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
- apply transfer
- apply (simp add: realrel_def)
- apply (rule vanishesI)
- apply (frule (1) cauchy_not_vanishes, clarify)
- apply (rule_tac x=k in exI, clarify)
- apply (drule_tac x=n in spec, simp)
- done
- show "a / b = a * inverse b"
- by (rule divide_real_def)
- show "inverse (0::real) = 0"
- by transfer (simp add: realrel_def)
-qed
-
-end
-
-subsection {* Positive reals *}
-
-lift_definition positive :: "real \<Rightarrow> bool"
- is "\<lambda>X. \<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
-proof -
- { fix X Y
- assume "realrel X Y"
- hence XY: "vanishes (\<lambda>n. X n - Y n)"
- unfolding realrel_def by simp_all
- assume "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
- then obtain r i where "0 < r" and i: "\<forall>n\<ge>i. r < X n"
- by fast
- obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
- using `0 < r` by (rule obtain_pos_sum)
- obtain j where j: "\<forall>n\<ge>j. \<bar>X n - Y n\<bar> < s"
- using vanishesD [OF XY s] ..
- have "\<forall>n\<ge>max i j. t < Y n"
- proof (clarsimp)
- fix n assume n: "i \<le> n" "j \<le> n"
- have "\<bar>X n - Y n\<bar> < s" and "r < X n"
- using i j n by simp_all
- thus "t < Y n" unfolding r by simp
- qed
- hence "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < Y n" using t by fast
- } note 1 = this
- fix X Y assume "realrel X Y"
- hence "realrel X Y" and "realrel Y X"
- using symp_realrel unfolding symp_def by auto
- thus "?thesis X Y"
- by (safe elim!: 1)
-qed
-
-lemma positive_Real:
- assumes X: "cauchy X"
- shows "positive (Real X) \<longleftrightarrow> (\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n)"
- using assms positive.transfer
- unfolding cr_real_eq fun_rel_def by simp
-
-lemma positive_zero: "\<not> positive 0"
- by transfer auto
-
-lemma positive_add:
- "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x + y)"
-apply transfer
-apply (clarify, rename_tac a b i j)
-apply (rule_tac x="a + b" in exI, simp)
-apply (rule_tac x="max i j" in exI, clarsimp)
-apply (simp add: add_strict_mono)
-done
-
-lemma positive_mult:
- "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x * y)"
-apply transfer
-apply (clarify, rename_tac a b i j)
-apply (rule_tac x="a * b" in exI, simp add: mult_pos_pos)
-apply (rule_tac x="max i j" in exI, clarsimp)
-apply (rule mult_strict_mono, auto)
-done
-
-lemma positive_minus:
- "\<not> positive x \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> positive (- x)"
-apply transfer
-apply (simp add: realrel_def)
-apply (drule (1) cauchy_not_vanishes_cases, safe, fast, fast)
-done
-
-instantiation real :: linordered_field_inverse_zero
-begin
-
-definition
- "x < y \<longleftrightarrow> positive (y - x)"
-
-definition
- "x \<le> (y::real) \<longleftrightarrow> x < y \<or> x = y"
-
-definition
- "abs (a::real) = (if a < 0 then - a else a)"
-
-definition
- "sgn (a::real) = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
-
-instance proof
- fix a b c :: real
- show "\<bar>a\<bar> = (if a < 0 then - a else a)"
- by (rule abs_real_def)
- show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a"
- unfolding less_eq_real_def less_real_def
- by (auto, drule (1) positive_add, simp_all add: positive_zero)
- show "a \<le> a"
- unfolding less_eq_real_def by simp
- show "a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> a \<le> c"
- unfolding less_eq_real_def less_real_def
- by (auto, drule (1) positive_add, simp add: algebra_simps)
- show "a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> a = b"
- unfolding less_eq_real_def less_real_def
- by (auto, drule (1) positive_add, simp add: positive_zero)
- show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
- unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *)
- (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \<equiv> b + - a *)
- (* Should produce c + b - (c + a) \<equiv> b - a *)
- show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
- by (rule sgn_real_def)
- show "a \<le> b \<or> b \<le> a"
- unfolding less_eq_real_def less_real_def
- by (auto dest!: positive_minus)
- show "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
- unfolding less_real_def
- by (drule (1) positive_mult, simp add: algebra_simps)
-qed
-
-end
-
-instantiation real :: distrib_lattice
-begin
-
-definition
- "(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"
-
-definition
- "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
-
-instance proof
-qed (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
-
-end
-
-lemma of_nat_Real: "of_nat x = Real (\<lambda>n. of_nat x)"
-apply (induct x)
-apply (simp add: zero_real_def)
-apply (simp add: one_real_def add_Real)
-done
-
-lemma of_int_Real: "of_int x = Real (\<lambda>n. of_int x)"
-apply (cases x rule: int_diff_cases)
-apply (simp add: of_nat_Real diff_Real)
-done
-
-lemma of_rat_Real: "of_rat x = Real (\<lambda>n. x)"
-apply (induct x)
-apply (simp add: Fract_of_int_quotient of_rat_divide)
-apply (simp add: of_int_Real divide_inverse)
-apply (simp add: inverse_Real mult_Real)
-done
-
-instance real :: archimedean_field
-proof
- fix x :: real
- show "\<exists>z. x \<le> of_int z"
- apply (induct x)
- apply (frule cauchy_imp_bounded, clarify)
- apply (rule_tac x="ceiling b + 1" in exI)
- apply (rule less_imp_le)
- apply (simp add: of_int_Real less_real_def diff_Real positive_Real)
- apply (rule_tac x=1 in exI, simp add: algebra_simps)
- apply (rule_tac x=0 in exI, clarsimp)
- apply (rule le_less_trans [OF abs_ge_self])
- apply (rule less_le_trans [OF _ le_of_int_ceiling])
- apply simp
- done
-qed
-
-instantiation real :: floor_ceiling
-begin
-
-definition [code del]:
- "floor (x::real) = (THE z. of_int z \<le> x \<and> x < of_int (z + 1))"
-
-instance proof
- fix x :: real
- show "of_int (floor x) \<le> x \<and> x < of_int (floor x + 1)"
- unfolding floor_real_def using floor_exists1 by (rule theI')
-qed
-
-end
-
-subsection {* Completeness *}
-
-lemma not_positive_Real:
- assumes X: "cauchy X"
- shows "\<not> positive (Real X) \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> r)"
-unfolding positive_Real [OF X]
-apply (auto, unfold not_less)
-apply (erule obtain_pos_sum)
-apply (drule_tac x=s in spec, simp)
-apply (drule_tac r=t in cauchyD [OF X], clarify)
-apply (drule_tac x=k in spec, clarsimp)
-apply (rule_tac x=n in exI, clarify, rename_tac m)
-apply (drule_tac x=m in spec, simp)
-apply (drule_tac x=n in spec, simp)
-apply (drule spec, drule (1) mp, clarify, rename_tac i)
-apply (rule_tac x="max i k" in exI, simp)
-done
-
-lemma le_Real:
- assumes X: "cauchy X" and Y: "cauchy Y"
- shows "Real X \<le> Real Y = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r)"
-unfolding not_less [symmetric, where 'a=real] less_real_def
-apply (simp add: diff_Real not_positive_Real X Y)
-apply (simp add: diff_le_eq add_ac)
-done
-
-lemma le_RealI:
- assumes Y: "cauchy Y"
- shows "\<forall>n. x \<le> of_rat (Y n) \<Longrightarrow> x \<le> Real Y"
-proof (induct x)
- fix X assume X: "cauchy X" and "\<forall>n. Real X \<le> of_rat (Y n)"
- hence le: "\<And>m r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. X n \<le> Y m + r"
- by (simp add: of_rat_Real le_Real)
- {
- fix r :: rat assume "0 < r"
- then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
- by (rule obtain_pos_sum)
- obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>Y m - Y n\<bar> < s"
- using cauchyD [OF Y s] ..
- obtain j where j: "\<forall>n\<ge>j. X n \<le> Y i + t"
- using le [OF t] ..
- have "\<forall>n\<ge>max i j. X n \<le> Y n + r"
- proof (clarsimp)
- fix n assume n: "i \<le> n" "j \<le> n"
- have "X n \<le> Y i + t" using n j by simp
- moreover have "\<bar>Y i - Y n\<bar> < s" using n i by simp
- ultimately show "X n \<le> Y n + r" unfolding r by simp
- qed
- hence "\<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r" ..
- }
- thus "Real X \<le> Real Y"
- by (simp add: of_rat_Real le_Real X Y)
-qed
-
-lemma Real_leI:
- assumes X: "cauchy X"
- assumes le: "\<forall>n. of_rat (X n) \<le> y"
- shows "Real X \<le> y"
-proof -
- have "- y \<le> - Real X"
- by (simp add: minus_Real X le_RealI of_rat_minus le)
- thus ?thesis by simp
-qed
-
-lemma less_RealD:
- assumes Y: "cauchy Y"
- shows "x < Real Y \<Longrightarrow> \<exists>n. x < of_rat (Y n)"
-by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y])
-
-lemma of_nat_less_two_power:
- "of_nat n < (2::'a::linordered_idom) ^ n"
-apply (induct n)
-apply simp
-apply (subgoal_tac "(1::'a) \<le> 2 ^ n")
-apply (drule (1) add_le_less_mono, simp)
-apply simp
-done
-
-lemma complete_real:
- fixes S :: "real set"
- assumes "\<exists>x. x \<in> S" and "\<exists>z. \<forall>x\<in>S. x \<le> z"
- shows "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
-proof -
- obtain x where x: "x \<in> S" using assms(1) ..
- obtain z where z: "\<forall>x\<in>S. x \<le> z" using assms(2) ..
-
- def P \<equiv> "\<lambda>x. \<forall>y\<in>S. y \<le> of_rat x"
- obtain a where a: "\<not> P a"
- proof
- have "of_int (floor (x - 1)) \<le> x - 1" by (rule of_int_floor_le)
- also have "x - 1 < x" by simp
- finally have "of_int (floor (x - 1)) < x" .
- hence "\<not> x \<le> of_int (floor (x - 1))" by (simp only: not_le)
- then show "\<not> P (of_int (floor (x - 1)))"
- unfolding P_def of_rat_of_int_eq using x by fast
- qed
- obtain b where b: "P b"
- proof
- show "P (of_int (ceiling z))"
- unfolding P_def of_rat_of_int_eq
- proof
- fix y assume "y \<in> S"
- hence "y \<le> z" using z by simp
- also have "z \<le> of_int (ceiling z)" by (rule le_of_int_ceiling)
- finally show "y \<le> of_int (ceiling z)" .
- qed
- qed
-
- def avg \<equiv> "\<lambda>x y :: rat. x/2 + y/2"
- def bisect \<equiv> "\<lambda>(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)"
- def A \<equiv> "\<lambda>n. fst ((bisect ^^ n) (a, b))"
- def B \<equiv> "\<lambda>n. snd ((bisect ^^ n) (a, b))"
- def C \<equiv> "\<lambda>n. avg (A n) (B n)"
- have A_0 [simp]: "A 0 = a" unfolding A_def by simp
- have B_0 [simp]: "B 0 = b" unfolding B_def by simp
- have A_Suc [simp]: "\<And>n. A (Suc n) = (if P (C n) then A n else C n)"
- unfolding A_def B_def C_def bisect_def split_def by simp
- have B_Suc [simp]: "\<And>n. B (Suc n) = (if P (C n) then C n else B n)"
- unfolding A_def B_def C_def bisect_def split_def by simp
-
- have width: "\<And>n. B n - A n = (b - a) / 2^n"
- apply (simp add: eq_divide_eq)
- apply (induct_tac n, simp)
- apply (simp add: C_def avg_def algebra_simps)
- done
-
- have twos: "\<And>y r :: rat. 0 < r \<Longrightarrow> \<exists>n. y / 2 ^ n < r"
- apply (simp add: divide_less_eq)
- apply (subst mult_commute)
- apply (frule_tac y=y in ex_less_of_nat_mult)
- apply clarify
- apply (rule_tac x=n in exI)
- apply (erule less_trans)
- apply (rule mult_strict_right_mono)
- apply (rule le_less_trans [OF _ of_nat_less_two_power])
- apply simp
- apply assumption
- done
-
- have PA: "\<And>n. \<not> P (A n)"
- by (induct_tac n, simp_all add: a)
- have PB: "\<And>n. P (B n)"
- by (induct_tac n, simp_all add: b)
- have ab: "a < b"
- using a b unfolding P_def
- apply (clarsimp simp add: not_le)
- apply (drule (1) bspec)
- apply (drule (1) less_le_trans)
- apply (simp add: of_rat_less)
- done
- have AB: "\<And>n. A n < B n"
- by (induct_tac n, simp add: ab, simp add: C_def avg_def)
- have A_mono: "\<And>i j. i \<le> j \<Longrightarrow> A i \<le> A j"
- apply (auto simp add: le_less [where 'a=nat])
- apply (erule less_Suc_induct)
- apply (clarsimp simp add: C_def avg_def)
- apply (simp add: add_divide_distrib [symmetric])
- apply (rule AB [THEN less_imp_le])
- apply simp
- done
- have B_mono: "\<And>i j. i \<le> j \<Longrightarrow> B j \<le> B i"
- apply (auto simp add: le_less [where 'a=nat])
- apply (erule less_Suc_induct)
- apply (clarsimp simp add: C_def avg_def)
- apply (simp add: add_divide_distrib [symmetric])
- apply (rule AB [THEN less_imp_le])
- apply simp
- done
- have cauchy_lemma:
- "\<And>X. \<forall>n. \<forall>i\<ge>n. A n \<le> X i \<and> X i \<le> B n \<Longrightarrow> cauchy X"
- apply (rule cauchyI)
- apply (drule twos [where y="b - a"])
- apply (erule exE)
- apply (rule_tac x=n in exI, clarify, rename_tac i j)
- apply (rule_tac y="B n - A n" in le_less_trans) defer
- apply (simp add: width)
- apply (drule_tac x=n in spec)
- apply (frule_tac x=i in spec, drule (1) mp)
- apply (frule_tac x=j in spec, drule (1) mp)
- apply (frule A_mono, drule B_mono)
- apply (frule A_mono, drule B_mono)
- apply arith
- done
- have "cauchy A"
- apply (rule cauchy_lemma [rule_format])
- apply (simp add: A_mono)
- apply (erule order_trans [OF less_imp_le [OF AB] B_mono])
- done
- have "cauchy B"
- apply (rule cauchy_lemma [rule_format])
- apply (simp add: B_mono)
- apply (erule order_trans [OF A_mono less_imp_le [OF AB]])
- done
- have 1: "\<forall>x\<in>S. x \<le> Real B"
- proof
- fix x assume "x \<in> S"
- then show "x \<le> Real B"
- using PB [unfolded P_def] `cauchy B`
- by (simp add: le_RealI)
- qed
- have 2: "\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> Real A \<le> z"
- apply clarify
- apply (erule contrapos_pp)
- apply (simp add: not_le)
- apply (drule less_RealD [OF `cauchy A`], clarify)
- apply (subgoal_tac "\<not> P (A n)")
- apply (simp add: P_def not_le, clarify)
- apply (erule rev_bexI)
- apply (erule (1) less_trans)
- apply (simp add: PA)
- done
- have "vanishes (\<lambda>n. (b - a) / 2 ^ n)"
- proof (rule vanishesI)
- fix r :: rat assume "0 < r"
- then obtain k where k: "\<bar>b - a\<bar> / 2 ^ k < r"
- using twos by fast
- have "\<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r"
- proof (clarify)
- fix n assume n: "k \<le> n"
- have "\<bar>(b - a) / 2 ^ n\<bar> = \<bar>b - a\<bar> / 2 ^ n"
- by simp
- also have "\<dots> \<le> \<bar>b - a\<bar> / 2 ^ k"
- using n by (simp add: divide_left_mono mult_pos_pos)
- also note k
- finally show "\<bar>(b - a) / 2 ^ n\<bar> < r" .
- qed
- thus "\<exists>k. \<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r" ..
- qed
- hence 3: "Real B = Real A"
- by (simp add: eq_Real `cauchy A` `cauchy B` width)
- show "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
- using 1 2 3 by (rule_tac x="Real B" in exI, simp)
-qed
-
-subsection {* Hiding implementation details *}
-
-hide_const (open) vanishes cauchy positive Real
-
-declare Real_induct [induct del]
-declare Abs_real_induct [induct del]
-declare Abs_real_cases [cases del]
-
-lemmas [transfer_rule del] =
- real.All_transfer real.Ex_transfer real.rel_eq_transfer forall_real_transfer
- zero_real.transfer one_real.transfer plus_real.transfer uminus_real.transfer
- times_real.transfer inverse_real.transfer positive.transfer real.right_unique
- real.right_total
-
-subsection{*More Lemmas*}
-
-text {* BH: These lemmas should not be necessary; they should be
-covered by existing simp rules and simplification procedures. *}
-
-lemma real_mult_left_cancel: "(c::real) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
-by simp (* redundant with mult_cancel_left *)
-
-lemma real_mult_right_cancel: "(c::real) \<noteq> 0 ==> (a*c=b*c) = (a=b)"
-by simp (* redundant with mult_cancel_right *)
-
-lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
-by simp (* solved by linordered_ring_less_cancel_factor simproc *)
-
-lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
-by simp (* solved by linordered_ring_le_cancel_factor simproc *)
-
-lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
-by simp (* solved by linordered_ring_le_cancel_factor simproc *)
-
-
-subsection {* Embedding numbers into the Reals *}
-
-abbreviation
- real_of_nat :: "nat \<Rightarrow> real"
-where
- "real_of_nat \<equiv> of_nat"
-
-abbreviation
- real_of_int :: "int \<Rightarrow> real"
-where
- "real_of_int \<equiv> of_int"
-
-abbreviation
- real_of_rat :: "rat \<Rightarrow> real"
-where
- "real_of_rat \<equiv> of_rat"
-
-consts
- (*overloaded constant for injecting other types into "real"*)
- real :: "'a => real"
-
-defs (overloaded)
- real_of_nat_def [code_unfold]: "real == real_of_nat"
- real_of_int_def [code_unfold]: "real == real_of_int"
-
-declare [[coercion_enabled]]
-declare [[coercion "real::nat\<Rightarrow>real"]]
-declare [[coercion "real::int\<Rightarrow>real"]]
-declare [[coercion "int"]]
-
-declare [[coercion_map map]]
-declare [[coercion_map "% f g h x. g (h (f x))"]]
-declare [[coercion_map "% f g (x,y) . (f x, g y)"]]
-
-lemma real_eq_of_nat: "real = of_nat"
- unfolding real_of_nat_def ..
-
-lemma real_eq_of_int: "real = of_int"
- unfolding real_of_int_def ..
-
-lemma real_of_int_zero [simp]: "real (0::int) = 0"
-by (simp add: real_of_int_def)
-
-lemma real_of_one [simp]: "real (1::int) = (1::real)"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_power [simp]: "real (x ^ n) = real (x::int) ^ n"
-by (simp add: real_of_int_def of_int_power)
-
-lemmas power_real_of_int = real_of_int_power [symmetric]
-
-lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))"
- apply (subst real_eq_of_int)+
- apply (rule of_int_setsum)
-done
-
-lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) =
- (PROD x:A. real(f x))"
- apply (subst real_eq_of_int)+
- apply (rule of_int_setprod)
-done
-
-lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \<le> real y) = (x \<le> y)"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)"
-by (simp add: real_of_int_def)
-
-lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)"
-by (simp add: real_of_int_def)
-
-lemma one_less_real_of_int_cancel_iff: "1 < real (i :: int) \<longleftrightarrow> 1 < i"
- unfolding real_of_one[symmetric] real_of_int_less_iff ..
-
-lemma one_le_real_of_int_cancel_iff: "1 \<le> real (i :: int) \<longleftrightarrow> 1 \<le> i"
- unfolding real_of_one[symmetric] real_of_int_le_iff ..
-
-lemma real_of_int_less_one_cancel_iff: "real (i :: int) < 1 \<longleftrightarrow> i < 1"
- unfolding real_of_one[symmetric] real_of_int_less_iff ..
-
-lemma real_of_int_le_one_cancel_iff: "real (i :: int) \<le> 1 \<longleftrightarrow> i \<le> 1"
- unfolding real_of_one[symmetric] real_of_int_le_iff ..
-
-lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))"
-by (auto simp add: abs_if)
-
-lemma int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)"
- apply (subgoal_tac "real n + 1 = real (n + 1)")
- apply (simp del: real_of_int_add)
- apply auto
-done
-
-lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)"
- apply (subgoal_tac "real m + 1 = real (m + 1)")
- apply (simp del: real_of_int_add)
- apply simp
-done
-
-lemma real_of_int_div_aux: "(real (x::int)) / (real d) =
- real (x div d) + (real (x mod d)) / (real d)"
-proof -
- have "x = (x div d) * d + x mod d"
- by auto
- then have "real x = real (x div d) * real d + real(x mod d)"
- by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym])
- then have "real x / real d = ... / real d"
- by simp
- then show ?thesis
- by (auto simp add: add_divide_distrib algebra_simps)
-qed
-
-lemma real_of_int_div: "(d :: int) dvd n ==>
- real(n div d) = real n / real d"
- apply (subst real_of_int_div_aux)
- apply simp
- apply (simp add: dvd_eq_mod_eq_0)
-done
-
-lemma real_of_int_div2:
- "0 <= real (n::int) / real (x) - real (n div x)"
- apply (case_tac "x = 0")
- apply simp
- apply (case_tac "0 < x")
- apply (simp add: algebra_simps)
- apply (subst real_of_int_div_aux)
- apply simp
- apply (subst zero_le_divide_iff)
- apply auto
- apply (simp add: algebra_simps)
- apply (subst real_of_int_div_aux)
- apply simp
- apply (subst zero_le_divide_iff)
- apply auto
-done
-
-lemma real_of_int_div3:
- "real (n::int) / real (x) - real (n div x) <= 1"
- apply (simp add: algebra_simps)
- apply (subst real_of_int_div_aux)
- apply (auto simp add: divide_le_eq intro: order_less_imp_le)
-done
-
-lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x"
-by (insert real_of_int_div2 [of n x], simp)
-
-lemma Ints_real_of_int [simp]: "real (x::int) \<in> Ints"
-unfolding real_of_int_def by (rule Ints_of_int)
-
-
-subsection{*Embedding the Naturals into the Reals*}
-
-lemma real_of_nat_zero [simp]: "real (0::nat) = 0"
-by (simp add: real_of_nat_def)
-
-lemma real_of_nat_1 [simp]: "real (1::nat) = 1"
-by (simp add: real_of_nat_def)
-
-lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)"
-by (simp add: real_of_nat_def)
-
-lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n"
-by (simp add: real_of_nat_def)
-
-(*Not for addsimps: often the LHS is used to represent a positive natural*)
-lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)"
-by (simp add: real_of_nat_def)
-
-lemma real_of_nat_less_iff [iff]:
- "(real (n::nat) < real m) = (n < m)"
-by (simp add: real_of_nat_def)
-
-lemma real_of_nat_le_iff [iff]: "(real (n::nat) \<le> real m) = (n \<le> m)"
-by (simp add: real_of_nat_def)
-
-lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)"
-by (simp add: real_of_nat_def)
-
-lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)"
-by (simp add: real_of_nat_def del: of_nat_Suc)
-
-lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n"
-by (simp add: real_of_nat_def of_nat_mult)
-
-lemma real_of_nat_power [simp]: "real (m ^ n) = real (m::nat) ^ n"
-by (simp add: real_of_nat_def of_nat_power)
-
-lemmas power_real_of_nat = real_of_nat_power [symmetric]
-
-lemma real_of_nat_setsum [simp]: "real ((SUM x:A. f x)::nat) =
- (SUM x:A. real(f x))"
- apply (subst real_eq_of_nat)+
- apply (rule of_nat_setsum)
-done
-
-lemma real_of_nat_setprod [simp]: "real ((PROD x:A. f x)::nat) =
- (PROD x:A. real(f x))"
- apply (subst real_eq_of_nat)+
- apply (rule of_nat_setprod)
-done
-
-lemma real_of_card: "real (card A) = setsum (%x.1) A"
- apply (subst card_eq_setsum)
- apply (subst real_of_nat_setsum)
- apply simp
-done
-
-lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)"
-by (simp add: real_of_nat_def)
-
-lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)"
-by (simp add: real_of_nat_def)
-
-lemma real_of_nat_diff: "n \<le> m ==> real (m - n) = real (m::nat) - real n"
-by (simp add: add: real_of_nat_def of_nat_diff)
-
-lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)"
-by (auto simp: real_of_nat_def)
-
-lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \<le> 0) = (n = 0)"
-by (simp add: add: real_of_nat_def)
-
-lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0"
-by (simp add: add: real_of_nat_def)
-
-lemma nat_less_real_le: "((n::nat) < m) = (real n + 1 <= real m)"
- apply (subgoal_tac "real n + 1 = real (Suc n)")
- apply simp
- apply (auto simp add: real_of_nat_Suc)
-done
-
-lemma nat_le_real_less: "((n::nat) <= m) = (real n < real m + 1)"
- apply (subgoal_tac "real m + 1 = real (Suc m)")
- apply (simp add: less_Suc_eq_le)
- apply (simp add: real_of_nat_Suc)
-done
-
-lemma real_of_nat_div_aux: "(real (x::nat)) / (real d) =
- real (x div d) + (real (x mod d)) / (real d)"
-proof -
- have "x = (x div d) * d + x mod d"
- by auto
- then have "real x = real (x div d) * real d + real(x mod d)"
- by (simp only: real_of_nat_mult [THEN sym] real_of_nat_add [THEN sym])
- then have "real x / real d = \<dots> / real d"
- by simp
- then show ?thesis
- by (auto simp add: add_divide_distrib algebra_simps)
-qed
-
-lemma real_of_nat_div: "(d :: nat) dvd n ==>
- real(n div d) = real n / real d"
- by (subst real_of_nat_div_aux)
- (auto simp add: dvd_eq_mod_eq_0 [symmetric])
-
-lemma real_of_nat_div2:
- "0 <= real (n::nat) / real (x) - real (n div x)"
-apply (simp add: algebra_simps)
-apply (subst real_of_nat_div_aux)
-apply simp
-apply (subst zero_le_divide_iff)
-apply simp
-done
-
-lemma real_of_nat_div3:
- "real (n::nat) / real (x) - real (n div x) <= 1"
-apply(case_tac "x = 0")
-apply (simp)
-apply (simp add: algebra_simps)
-apply (subst real_of_nat_div_aux)
-apply simp
-done
-
-lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x"
-by (insert real_of_nat_div2 [of n x], simp)
-
-lemma real_of_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n"
-by (simp add: real_of_int_def real_of_nat_def)
-
-lemma real_nat_eq_real [simp]: "0 <= x ==> real(nat x) = real x"
- apply (subgoal_tac "real(int(nat x)) = real(nat x)")
- apply force
- apply (simp only: real_of_int_of_nat_eq)
-done
-
-lemma Nats_real_of_nat [simp]: "real (n::nat) \<in> Nats"
-unfolding real_of_nat_def by (rule of_nat_in_Nats)
-
-lemma Ints_real_of_nat [simp]: "real (n::nat) \<in> Ints"
-unfolding real_of_nat_def by (rule Ints_of_nat)
-
-
-subsection{* Rationals *}
-
-lemma Rats_real_nat[simp]: "real(n::nat) \<in> \<rat>"
-by (simp add: real_eq_of_nat)
-
-
-lemma Rats_eq_int_div_int:
- "\<rat> = { real(i::int)/real(j::int) |i j. j \<noteq> 0}" (is "_ = ?S")
-proof
- show "\<rat> \<subseteq> ?S"
- proof
- fix x::real assume "x : \<rat>"
- then obtain r where "x = of_rat r" unfolding Rats_def ..
- have "of_rat r : ?S"
- by (cases r)(auto simp add:of_rat_rat real_eq_of_int)
- thus "x : ?S" using `x = of_rat r` by simp
- qed
-next
- show "?S \<subseteq> \<rat>"
- proof(auto simp:Rats_def)
- fix i j :: int assume "j \<noteq> 0"
- hence "real i / real j = of_rat(Fract i j)"
- by (simp add:of_rat_rat real_eq_of_int)
- thus "real i / real j \<in> range of_rat" by blast
- qed
-qed
-
-lemma Rats_eq_int_div_nat:
- "\<rat> = { real(i::int)/real(n::nat) |i n. n \<noteq> 0}"
-proof(auto simp:Rats_eq_int_div_int)
- fix i j::int assume "j \<noteq> 0"
- show "EX (i'::int) (n::nat). real i/real j = real i'/real n \<and> 0<n"
- proof cases
- assume "j>0"
- hence "real i/real j = real i/real(nat j) \<and> 0<nat j"
- by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
- thus ?thesis by blast
- next
- assume "~ j>0"
- hence "real i/real j = real(-i)/real(nat(-j)) \<and> 0<nat(-j)" using `j\<noteq>0`
- by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
- thus ?thesis by blast
- qed
-next
- fix i::int and n::nat assume "0 < n"
- hence "real i/real n = real i/real(int n) \<and> int n \<noteq> 0" by simp
- thus "\<exists>(i'::int) j::int. real i/real n = real i'/real j \<and> j \<noteq> 0" by blast
-qed
-
-lemma Rats_abs_nat_div_natE:
- assumes "x \<in> \<rat>"
- obtains m n :: nat
- where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd m n = 1"
-proof -
- from `x \<in> \<rat>` obtain i::int and n::nat where "n \<noteq> 0" and "x = real i / real n"
- by(auto simp add: Rats_eq_int_div_nat)
- hence "\<bar>x\<bar> = real(nat(abs i)) / real n" by simp
- then obtain m :: nat where x_rat: "\<bar>x\<bar> = real m / real n" by blast
- let ?gcd = "gcd m n"
- from `n\<noteq>0` have gcd: "?gcd \<noteq> 0" by simp
- let ?k = "m div ?gcd"
- let ?l = "n div ?gcd"
- let ?gcd' = "gcd ?k ?l"
- have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m"
- by (rule dvd_mult_div_cancel)
- have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
- by (rule dvd_mult_div_cancel)
- from `n\<noteq>0` and gcd_l have "?l \<noteq> 0" by (auto iff del: neq0_conv)
- moreover
- have "\<bar>x\<bar> = real ?k / real ?l"
- proof -
- from gcd have "real ?k / real ?l =
- real (?gcd * ?k) / real (?gcd * ?l)" by simp
- also from gcd_k and gcd_l have "\<dots> = real m / real n" by simp
- also from x_rat have "\<dots> = \<bar>x\<bar>" ..
- finally show ?thesis ..
- qed
- moreover
- have "?gcd' = 1"
- proof -
- have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)"
- by (rule gcd_mult_distrib_nat)
- with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
- with gcd show ?thesis by auto
- qed
- ultimately show ?thesis ..
-qed
-
-
-subsection{*Numerals and Arithmetic*}
-
-lemma [code_abbrev]:
- "real_of_int (numeral k) = numeral k"
- "real_of_int (neg_numeral k) = neg_numeral k"
- by simp_all
-
-text{*Collapse applications of @{term real} to @{term number_of}*}
-lemma real_numeral [simp]:
- "real (numeral v :: int) = numeral v"
- "real (neg_numeral v :: int) = neg_numeral v"
-by (simp_all add: real_of_int_def)
-
-lemma real_of_nat_numeral [simp]:
- "real (numeral v :: nat) = numeral v"
-by (simp add: real_of_nat_def)
-
-declaration {*
- K (Lin_Arith.add_inj_thms [@{thm real_of_nat_le_iff} RS iffD2, @{thm real_of_nat_inject} RS iffD2]
- (* not needed because x < (y::nat) can be rewritten as Suc x <= y: real_of_nat_less_iff RS iffD2 *)
- #> Lin_Arith.add_inj_thms [@{thm real_of_int_le_iff} RS iffD2, @{thm real_of_int_inject} RS iffD2]
- (* not needed because x < (y::int) can be rewritten as x + 1 <= y: real_of_int_less_iff RS iffD2 *)
- #> Lin_Arith.add_simps [@{thm real_of_nat_zero}, @{thm real_of_nat_Suc}, @{thm real_of_nat_add},
- @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one},
- @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff},
- @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq},
- @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)}]
- #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \<Rightarrow> real"})
- #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \<Rightarrow> real"}))
-*}
-
-
-subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
-
-lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)"
-by arith
-
-text {* FIXME: redundant with @{text add_eq_0_iff} below *}
-lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)"
-by auto
-
-lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)"
-by auto
-
-lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)"
-by auto
-
-lemma real_add_le_0_iff: "(x+y \<le> (0::real)) = (y \<le> -x)"
-by auto
-
-lemma real_0_le_add_iff: "((0::real) \<le> x+y) = (-x \<le> y)"
-by auto
-
-subsection {* Lemmas about powers *}
-
-text {* FIXME: declare this in Rings.thy or not at all *}
-declare abs_mult_self [simp]
-
-(* used by Import/HOL/real.imp *)
-lemma two_realpow_ge_one: "(1::real) \<le> 2 ^ n"
-by simp
-
-lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n"
-apply (induct "n")
-apply (auto simp add: real_of_nat_Suc)
-apply (subst mult_2)
-apply (erule add_less_le_mono)
-apply (rule two_realpow_ge_one)
-done
-
-text {* TODO: no longer real-specific; rename and move elsewhere *}
-lemma realpow_Suc_le_self:
- fixes r :: "'a::linordered_semidom"
- shows "[| 0 \<le> r; r \<le> 1 |] ==> r ^ Suc n \<le> r"
-by (insert power_decreasing [of 1 "Suc n" r], simp)
-
-text {* TODO: no longer real-specific; rename and move elsewhere *}
-lemma realpow_minus_mult:
- fixes x :: "'a::monoid_mult"
- shows "0 < n \<Longrightarrow> x ^ (n - 1) * x = x ^ n"
-by (simp add: power_commutes split add: nat_diff_split)
-
-text {* FIXME: declare this [simp] for all types, or not at all *}
-lemma real_two_squares_add_zero_iff [simp]:
- "(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)"
-by (rule sum_squares_eq_zero_iff)
-
-text {* FIXME: declare this [simp] for all types, or not at all *}
-lemma realpow_two_sum_zero_iff [simp]:
- "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)"
-by (rule sum_power2_eq_zero_iff)
-
-lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))"
-by (rule_tac y = 0 in order_trans, auto)
-
-lemma realpow_square_minus_le [simp]: "-(u ^ 2) \<le> (x::real) ^ 2"
-by (auto simp add: power2_eq_square)
-
-
-lemma numeral_power_le_real_of_nat_cancel_iff[simp]:
- "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::nat) ^ n \<le> a"
- unfolding real_of_nat_le_iff[symmetric] by simp
-
-lemma real_of_nat_le_numeral_power_cancel_iff[simp]:
- "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::nat) ^ n"
- unfolding real_of_nat_le_iff[symmetric] by simp
-
-lemma numeral_power_le_real_of_int_cancel_iff[simp]:
- "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::int) ^ n \<le> a"
- unfolding real_of_int_le_iff[symmetric] by simp
-
-lemma real_of_int_le_numeral_power_cancel_iff[simp]:
- "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::int) ^ n"
- unfolding real_of_int_le_iff[symmetric] by simp
-
-lemma neg_numeral_power_le_real_of_int_cancel_iff[simp]:
- "(neg_numeral x::real) ^ n \<le> real a \<longleftrightarrow> (neg_numeral x::int) ^ n \<le> a"
- unfolding real_of_int_le_iff[symmetric] by simp
-
-lemma real_of_int_le_neg_numeral_power_cancel_iff[simp]:
- "real a \<le> (neg_numeral x::real) ^ n \<longleftrightarrow> a \<le> (neg_numeral x::int) ^ n"
- unfolding real_of_int_le_iff[symmetric] by simp
-
-subsection{*Density of the Reals*}
-
-lemma real_lbound_gt_zero:
- "[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2"
-apply (rule_tac x = " (min d1 d2) /2" in exI)
-apply (simp add: min_def)
-done
-
-
-text{*Similar results are proved in @{text Fields}*}
-lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
- by auto
-
-lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y"
- by auto
-
-
-subsection{*Absolute Value Function for the Reals*}
-
-lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))"
-by (simp add: abs_if)
-
-(* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *)
-lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
-by (force simp add: abs_le_iff)
-
-lemma abs_add_one_gt_zero: "(0::real) < 1 + abs(x)"
-by (simp add: abs_if)
-
-lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
-by (rule abs_of_nonneg [OF real_of_nat_ge_zero])
-
-lemma abs_add_one_not_less_self: "~ abs(x) + (1::real) < x"
-by simp
-
-lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)"
-by simp
-
-
-subsection {* Implementation of rational real numbers *}
-
-text {* Formal constructor *}
-
-definition Ratreal :: "rat \<Rightarrow> real" where
- [code_abbrev, simp]: "Ratreal = of_rat"
-
-code_datatype Ratreal
-
-
-text {* Numerals *}
-
-lemma [code_abbrev]:
- "(of_rat (of_int a) :: real) = of_int a"
- by simp
-
-lemma [code_abbrev]:
- "(of_rat 0 :: real) = 0"
- by simp
-
-lemma [code_abbrev]:
- "(of_rat 1 :: real) = 1"
- by simp
-
-lemma [code_abbrev]:
- "(of_rat (numeral k) :: real) = numeral k"
- by simp
-
-lemma [code_abbrev]:
- "(of_rat (neg_numeral k) :: real) = neg_numeral k"
- by simp
-
-lemma [code_post]:
- "(of_rat (0 / r) :: real) = 0"
- "(of_rat (r / 0) :: real) = 0"
- "(of_rat (1 / 1) :: real) = 1"
- "(of_rat (numeral k / 1) :: real) = numeral k"
- "(of_rat (neg_numeral k / 1) :: real) = neg_numeral k"
- "(of_rat (1 / numeral k) :: real) = 1 / numeral k"
- "(of_rat (1 / neg_numeral k) :: real) = 1 / neg_numeral k"
- "(of_rat (numeral k / numeral l) :: real) = numeral k / numeral l"
- "(of_rat (numeral k / neg_numeral l) :: real) = numeral k / neg_numeral l"
- "(of_rat (neg_numeral k / numeral l) :: real) = neg_numeral k / numeral l"
- "(of_rat (neg_numeral k / neg_numeral l) :: real) = neg_numeral k / neg_numeral l"
- by (simp_all add: of_rat_divide)
-
-
-text {* Operations *}
-
-lemma zero_real_code [code]:
- "0 = Ratreal 0"
-by simp
-
-lemma one_real_code [code]:
- "1 = Ratreal 1"
-by simp
-
-instantiation real :: equal
-begin
-
-definition "HOL.equal (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
-
-instance proof
-qed (simp add: equal_real_def)
-
-lemma real_equal_code [code]:
- "HOL.equal (Ratreal x) (Ratreal y) \<longleftrightarrow> HOL.equal x y"
- by (simp add: equal_real_def equal)
-
-lemma [code nbe]:
- "HOL.equal (x::real) x \<longleftrightarrow> True"
- by (rule equal_refl)
-
-end
-
-lemma real_less_eq_code [code]: "Ratreal x \<le> Ratreal y \<longleftrightarrow> x \<le> y"
- by (simp add: of_rat_less_eq)
-
-lemma real_less_code [code]: "Ratreal x < Ratreal y \<longleftrightarrow> x < y"
- by (simp add: of_rat_less)
-
-lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)"
- by (simp add: of_rat_add)
-
-lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)"
- by (simp add: of_rat_mult)
-
-lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)"
- by (simp add: of_rat_minus)
-
-lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)"
- by (simp add: of_rat_diff)
-
-lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)"
- by (simp add: of_rat_inverse)
-
-lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)"
- by (simp add: of_rat_divide)
-
-lemma real_floor_code [code]: "floor (Ratreal x) = floor x"
- 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)
-
-
-text {* Quickcheck *}
-
-definition (in term_syntax)
- valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
- [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\<cdot>} k"
-
-notation fcomp (infixl "\<circ>>" 60)
-notation scomp (infixl "\<circ>\<rightarrow>" 60)
-
-instantiation real :: random
-begin
-
-definition
- "Quickcheck_Random.random i = Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>r. Pair (valterm_ratreal r))"
-
-instance ..
-
-end
-
-no_notation fcomp (infixl "\<circ>>" 60)
-no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
-
-instantiation real :: exhaustive
-begin
-
-definition
- "exhaustive_real f d = Quickcheck_Exhaustive.exhaustive (%r. f (Ratreal r)) d"
-
-instance ..
-
-end
-
-instantiation real :: full_exhaustive
-begin
-
-definition
- "full_exhaustive_real f d = Quickcheck_Exhaustive.full_exhaustive (%r. f (valterm_ratreal r)) d"
-
-instance ..
-
-end
-
-instantiation real :: narrowing
-begin
-
-definition
- "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing"
-
-instance ..
-
-end
-
-
-subsection {* Setup for Nitpick *}
-
-declaration {*
- Nitpick_HOL.register_frac_type @{type_name real}
- [(@{const_name zero_real_inst.zero_real}, @{const_name Nitpick.zero_frac}),
- (@{const_name one_real_inst.one_real}, @{const_name Nitpick.one_frac}),
- (@{const_name plus_real_inst.plus_real}, @{const_name Nitpick.plus_frac}),
- (@{const_name times_real_inst.times_real}, @{const_name Nitpick.times_frac}),
- (@{const_name uminus_real_inst.uminus_real}, @{const_name Nitpick.uminus_frac}),
- (@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}),
- (@{const_name ord_real_inst.less_real}, @{const_name Nitpick.less_frac}),
- (@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})]
-*}
-
-lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real
- ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real
- times_real_inst.times_real uminus_real_inst.uminus_real
- zero_real_inst.zero_real
-
-end
--- a/src/HOL/RealVector.thy Tue Mar 26 14:38:44 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1096 +0,0 @@
-(* Title: HOL/RealVector.thy
- Author: Brian Huffman
-*)
-
-header {* Vector Spaces and Algebras over the Reals *}
-
-theory RealVector
-imports RComplete Metric_Spaces SupInf
-begin
-
-subsection {* Locale for additive functions *}
-
-locale additive =
- fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
- assumes add: "f (x + y) = f x + f y"
-begin
-
-lemma zero: "f 0 = 0"
-proof -
- have "f 0 = f (0 + 0)" by simp
- also have "\<dots> = f 0 + f 0" by (rule add)
- finally show "f 0 = 0" by simp
-qed
-
-lemma minus: "f (- x) = - f x"
-proof -
- have "f (- x) + f x = f (- x + x)" by (rule add [symmetric])
- also have "\<dots> = - f x + f x" by (simp add: zero)
- finally show "f (- x) = - f x" by (rule add_right_imp_eq)
-qed
-
-lemma diff: "f (x - y) = f x - f y"
-by (simp add: add minus diff_minus)
-
-lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
-apply (cases "finite A")
-apply (induct set: finite)
-apply (simp add: zero)
-apply (simp add: add)
-apply (simp add: zero)
-done
-
-end
-
-subsection {* Vector spaces *}
-
-locale vector_space =
- fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
- assumes scale_right_distrib [algebra_simps]:
- "scale a (x + y) = scale a x + scale a y"
- and scale_left_distrib [algebra_simps]:
- "scale (a + b) x = scale a x + scale b x"
- and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
- and scale_one [simp]: "scale 1 x = x"
-begin
-
-lemma scale_left_commute:
- "scale a (scale b x) = scale b (scale a x)"
-by (simp add: mult_commute)
-
-lemma scale_zero_left [simp]: "scale 0 x = 0"
- and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
- and scale_left_diff_distrib [algebra_simps]:
- "scale (a - b) x = scale a x - scale b x"
- and scale_setsum_left: "scale (setsum f A) x = (\<Sum>a\<in>A. scale (f a) x)"
-proof -
- interpret s: additive "\<lambda>a. scale a x"
- proof qed (rule scale_left_distrib)
- show "scale 0 x = 0" by (rule s.zero)
- show "scale (- a) x = - (scale a x)" by (rule s.minus)
- show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
- show "scale (setsum f A) x = (\<Sum>a\<in>A. scale (f a) x)" by (rule s.setsum)
-qed
-
-lemma scale_zero_right [simp]: "scale a 0 = 0"
- and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
- and scale_right_diff_distrib [algebra_simps]:
- "scale a (x - y) = scale a x - scale a y"
- and scale_setsum_right: "scale a (setsum f A) = (\<Sum>x\<in>A. scale a (f x))"
-proof -
- interpret s: additive "\<lambda>x. scale a x"
- proof qed (rule scale_right_distrib)
- show "scale a 0 = 0" by (rule s.zero)
- show "scale a (- x) = - (scale a x)" by (rule s.minus)
- show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
- show "scale a (setsum f A) = (\<Sum>x\<in>A. scale a (f x))" by (rule s.setsum)
-qed
-
-lemma scale_eq_0_iff [simp]:
- "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
-proof cases
- assume "a = 0" thus ?thesis by simp
-next
- assume anz [simp]: "a \<noteq> 0"
- { assume "scale a x = 0"
- hence "scale (inverse a) (scale a x) = 0" by simp
- hence "x = 0" by simp }
- thus ?thesis by force
-qed
-
-lemma scale_left_imp_eq:
- "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
-proof -
- assume nonzero: "a \<noteq> 0"
- assume "scale a x = scale a y"
- hence "scale a (x - y) = 0"
- by (simp add: scale_right_diff_distrib)
- hence "x - y = 0" by (simp add: nonzero)
- thus "x = y" by (simp only: right_minus_eq)
-qed
-
-lemma scale_right_imp_eq:
- "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
-proof -
- assume nonzero: "x \<noteq> 0"
- assume "scale a x = scale b x"
- hence "scale (a - b) x = 0"
- by (simp add: scale_left_diff_distrib)
- hence "a - b = 0" by (simp add: nonzero)
- thus "a = b" by (simp only: right_minus_eq)
-qed
-
-lemma scale_cancel_left [simp]:
- "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
-by (auto intro: scale_left_imp_eq)
-
-lemma scale_cancel_right [simp]:
- "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
-by (auto intro: scale_right_imp_eq)
-
-end
-
-subsection {* Real vector spaces *}
-
-class scaleR =
- fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
-begin
-
-abbreviation
- divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
-where
- "x /\<^sub>R r == scaleR (inverse r) x"
-
-end
-
-class real_vector = scaleR + ab_group_add +
- assumes scaleR_add_right: "scaleR a (x + y) = scaleR a x + scaleR a y"
- and scaleR_add_left: "scaleR (a + b) x = scaleR a x + scaleR b x"
- and scaleR_scaleR: "scaleR a (scaleR b x) = scaleR (a * b) x"
- and scaleR_one: "scaleR 1 x = x"
-
-interpretation real_vector:
- vector_space "scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"
-apply unfold_locales
-apply (rule scaleR_add_right)
-apply (rule scaleR_add_left)
-apply (rule scaleR_scaleR)
-apply (rule scaleR_one)
-done
-
-text {* Recover original theorem names *}
-
-lemmas scaleR_left_commute = real_vector.scale_left_commute
-lemmas scaleR_zero_left = real_vector.scale_zero_left
-lemmas scaleR_minus_left = real_vector.scale_minus_left
-lemmas scaleR_diff_left = real_vector.scale_left_diff_distrib
-lemmas scaleR_setsum_left = real_vector.scale_setsum_left
-lemmas scaleR_zero_right = real_vector.scale_zero_right
-lemmas scaleR_minus_right = real_vector.scale_minus_right
-lemmas scaleR_diff_right = real_vector.scale_right_diff_distrib
-lemmas scaleR_setsum_right = real_vector.scale_setsum_right
-lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
-lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
-lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
-lemmas scaleR_cancel_left = real_vector.scale_cancel_left
-lemmas scaleR_cancel_right = real_vector.scale_cancel_right
-
-text {* Legacy names *}
-
-lemmas scaleR_left_distrib = scaleR_add_left
-lemmas scaleR_right_distrib = scaleR_add_right
-lemmas scaleR_left_diff_distrib = scaleR_diff_left
-lemmas scaleR_right_diff_distrib = scaleR_diff_right
-
-lemma scaleR_minus1_left [simp]:
- fixes x :: "'a::real_vector"
- shows "scaleR (-1) x = - x"
- using scaleR_minus_left [of 1 x] by simp
-
-class real_algebra = real_vector + ring +
- assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
- and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
-
-class real_algebra_1 = real_algebra + ring_1
-
-class real_div_algebra = real_algebra_1 + division_ring
-
-class real_field = real_div_algebra + field
-
-instantiation real :: real_field
-begin
-
-definition
- real_scaleR_def [simp]: "scaleR a x = a * x"
-
-instance proof
-qed (simp_all add: algebra_simps)
-
-end
-
-interpretation scaleR_left: additive "(\<lambda>a. scaleR a x::'a::real_vector)"
-proof qed (rule scaleR_left_distrib)
-
-interpretation scaleR_right: additive "(\<lambda>x. scaleR a x::'a::real_vector)"
-proof qed (rule scaleR_right_distrib)
-
-lemma nonzero_inverse_scaleR_distrib:
- fixes x :: "'a::real_div_algebra" shows
- "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
-by (rule inverse_unique, simp)
-
-lemma inverse_scaleR_distrib:
- fixes x :: "'a::{real_div_algebra, division_ring_inverse_zero}"
- shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
-apply (case_tac "a = 0", simp)
-apply (case_tac "x = 0", simp)
-apply (erule (1) nonzero_inverse_scaleR_distrib)
-done
-
-
-subsection {* Embedding of the Reals into any @{text real_algebra_1}:
-@{term of_real} *}
-
-definition
- of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
- "of_real r = scaleR r 1"
-
-lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
-by (simp add: of_real_def)
-
-lemma of_real_0 [simp]: "of_real 0 = 0"
-by (simp add: of_real_def)
-
-lemma of_real_1 [simp]: "of_real 1 = 1"
-by (simp add: of_real_def)
-
-lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
-by (simp add: of_real_def scaleR_left_distrib)
-
-lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
-by (simp add: of_real_def)
-
-lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
-by (simp add: of_real_def scaleR_left_diff_distrib)
-
-lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
-by (simp add: of_real_def mult_commute)
-
-lemma nonzero_of_real_inverse:
- "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
- inverse (of_real x :: 'a::real_div_algebra)"
-by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
-
-lemma of_real_inverse [simp]:
- "of_real (inverse x) =
- inverse (of_real x :: 'a::{real_div_algebra, division_ring_inverse_zero})"
-by (simp add: of_real_def inverse_scaleR_distrib)
-
-lemma nonzero_of_real_divide:
- "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
- (of_real x / of_real y :: 'a::real_field)"
-by (simp add: divide_inverse nonzero_of_real_inverse)
-
-lemma of_real_divide [simp]:
- "of_real (x / y) =
- (of_real x / of_real y :: 'a::{real_field, field_inverse_zero})"
-by (simp add: divide_inverse)
-
-lemma of_real_power [simp]:
- "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1}) ^ n"
-by (induct n) simp_all
-
-lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
-by (simp add: of_real_def)
-
-lemma inj_of_real:
- "inj of_real"
- by (auto intro: injI)
-
-lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
-
-lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
-proof
- fix r
- show "of_real r = id r"
- by (simp add: of_real_def)
-qed
-
-text{*Collapse nested embeddings*}
-lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
-by (induct n) auto
-
-lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
-by (cases z rule: int_diff_cases, simp)
-
-lemma of_real_numeral: "of_real (numeral w) = numeral w"
-using of_real_of_int_eq [of "numeral w"] by simp
-
-lemma of_real_neg_numeral: "of_real (neg_numeral w) = neg_numeral w"
-using of_real_of_int_eq [of "neg_numeral w"] by simp
-
-text{*Every real algebra has characteristic zero*}
-
-instance real_algebra_1 < ring_char_0
-proof
- from inj_of_real inj_of_nat have "inj (of_real \<circ> of_nat)" by (rule inj_comp)
- then show "inj (of_nat :: nat \<Rightarrow> 'a)" by (simp add: comp_def)
-qed
-
-instance real_field < field_char_0 ..
-
-
-subsection {* The Set of Real Numbers *}
-
-definition Reals :: "'a::real_algebra_1 set" where
- "Reals = range of_real"
-
-notation (xsymbols)
- Reals ("\<real>")
-
-lemma Reals_of_real [simp]: "of_real r \<in> Reals"
-by (simp add: Reals_def)
-
-lemma Reals_of_int [simp]: "of_int z \<in> Reals"
-by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
-
-lemma Reals_of_nat [simp]: "of_nat n \<in> Reals"
-by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
-
-lemma Reals_numeral [simp]: "numeral w \<in> Reals"
-by (subst of_real_numeral [symmetric], rule Reals_of_real)
-
-lemma Reals_neg_numeral [simp]: "neg_numeral w \<in> Reals"
-by (subst of_real_neg_numeral [symmetric], rule Reals_of_real)
-
-lemma Reals_0 [simp]: "0 \<in> Reals"
-apply (unfold Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_0 [symmetric])
-done
-
-lemma Reals_1 [simp]: "1 \<in> Reals"
-apply (unfold Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_1 [symmetric])
-done
-
-lemma Reals_add [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a + b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_add [symmetric])
-done
-
-lemma Reals_minus [simp]: "a \<in> Reals \<Longrightarrow> - a \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_minus [symmetric])
-done
-
-lemma Reals_diff [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a - b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_diff [symmetric])
-done
-
-lemma Reals_mult [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a * b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_mult [symmetric])
-done
-
-lemma nonzero_Reals_inverse:
- fixes a :: "'a::real_div_algebra"
- shows "\<lbrakk>a \<in> Reals; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_real_inverse [symmetric])
-done
-
-lemma Reals_inverse [simp]:
- fixes a :: "'a::{real_div_algebra, division_ring_inverse_zero}"
- shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_inverse [symmetric])
-done
-
-lemma nonzero_Reals_divide:
- fixes a b :: "'a::real_field"
- shows "\<lbrakk>a \<in> Reals; b \<in> Reals; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_real_divide [symmetric])
-done
-
-lemma Reals_divide [simp]:
- fixes a b :: "'a::{real_field, field_inverse_zero}"
- shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_divide [symmetric])
-done
-
-lemma Reals_power [simp]:
- fixes a :: "'a::{real_algebra_1}"
- shows "a \<in> Reals \<Longrightarrow> a ^ n \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_power [symmetric])
-done
-
-lemma Reals_cases [cases set: Reals]:
- assumes "q \<in> \<real>"
- obtains (of_real) r where "q = of_real r"
- unfolding Reals_def
-proof -
- from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
- then obtain r where "q = of_real r" ..
- then show thesis ..
-qed
-
-lemma Reals_induct [case_names of_real, induct set: Reals]:
- "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
- by (rule Reals_cases) auto
-
-
-subsection {* Real normed vector spaces *}
-
-class norm =
- fixes norm :: "'a \<Rightarrow> real"
-
-class sgn_div_norm = scaleR + norm + sgn +
- assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x"
-
-class dist_norm = dist + norm + minus +
- assumes dist_norm: "dist x y = norm (x - y)"
-
-class real_normed_vector = real_vector + sgn_div_norm + dist_norm + open_dist +
- assumes norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
- and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
- and norm_scaleR [simp]: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
-begin
-
-lemma norm_ge_zero [simp]: "0 \<le> norm x"
-proof -
- have "0 = norm (x + -1 *\<^sub>R x)"
- using scaleR_add_left[of 1 "-1" x] norm_scaleR[of 0 x] by (simp add: scaleR_one)
- also have "\<dots> \<le> norm x + norm (-1 *\<^sub>R x)" by (rule norm_triangle_ineq)
- finally show ?thesis by simp
-qed
-
-end
-
-class real_normed_algebra = real_algebra + real_normed_vector +
- assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
-
-class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
- assumes norm_one [simp]: "norm 1 = 1"
-
-class real_normed_div_algebra = real_div_algebra + real_normed_vector +
- assumes norm_mult: "norm (x * y) = norm x * norm y"
-
-class real_normed_field = real_field + real_normed_div_algebra
-
-instance real_normed_div_algebra < real_normed_algebra_1
-proof
- fix x y :: 'a
- show "norm (x * y) \<le> norm x * norm y"
- by (simp add: norm_mult)
-next
- have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
- by (rule norm_mult)
- thus "norm (1::'a) = 1" by simp
-qed
-
-lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
-by simp
-
-lemma zero_less_norm_iff [simp]:
- fixes x :: "'a::real_normed_vector"
- shows "(0 < norm x) = (x \<noteq> 0)"
-by (simp add: order_less_le)
-
-lemma norm_not_less_zero [simp]:
- fixes x :: "'a::real_normed_vector"
- shows "\<not> norm x < 0"
-by (simp add: linorder_not_less)
-
-lemma norm_le_zero_iff [simp]:
- fixes x :: "'a::real_normed_vector"
- shows "(norm x \<le> 0) = (x = 0)"
-by (simp add: order_le_less)
-
-lemma norm_minus_cancel [simp]:
- fixes x :: "'a::real_normed_vector"
- shows "norm (- x) = norm x"
-proof -
- have "norm (- x) = norm (scaleR (- 1) x)"
- by (simp only: scaleR_minus_left scaleR_one)
- also have "\<dots> = \<bar>- 1\<bar> * norm x"
- by (rule norm_scaleR)
- finally show ?thesis by simp
-qed
-
-lemma norm_minus_commute:
- fixes a b :: "'a::real_normed_vector"
- shows "norm (a - b) = norm (b - a)"
-proof -
- have "norm (- (b - a)) = norm (b - a)"
- by (rule norm_minus_cancel)
- thus ?thesis by simp
-qed
-
-lemma norm_triangle_ineq2:
- fixes a b :: "'a::real_normed_vector"
- shows "norm a - norm b \<le> norm (a - b)"
-proof -
- have "norm (a - b + b) \<le> norm (a - b) + norm b"
- by (rule norm_triangle_ineq)
- thus ?thesis by simp
-qed
-
-lemma norm_triangle_ineq3:
- fixes a b :: "'a::real_normed_vector"
- shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
-apply (subst abs_le_iff)
-apply auto
-apply (rule norm_triangle_ineq2)
-apply (subst norm_minus_commute)
-apply (rule norm_triangle_ineq2)
-done
-
-lemma norm_triangle_ineq4:
- fixes a b :: "'a::real_normed_vector"
- shows "norm (a - b) \<le> norm a + norm b"
-proof -
- have "norm (a + - b) \<le> norm a + norm (- b)"
- by (rule norm_triangle_ineq)
- thus ?thesis
- by (simp only: diff_minus norm_minus_cancel)
-qed
-
-lemma norm_diff_ineq:
- fixes a b :: "'a::real_normed_vector"
- shows "norm a - norm b \<le> norm (a + b)"
-proof -
- have "norm a - norm (- b) \<le> norm (a - - b)"
- by (rule norm_triangle_ineq2)
- thus ?thesis by simp
-qed
-
-lemma norm_diff_triangle_ineq:
- fixes a b c d :: "'a::real_normed_vector"
- shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
-proof -
- have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
- by (simp add: diff_minus add_ac)
- also have "\<dots> \<le> norm (a - c) + norm (b - d)"
- by (rule norm_triangle_ineq)
- finally show ?thesis .
-qed
-
-lemma abs_norm_cancel [simp]:
- fixes a :: "'a::real_normed_vector"
- shows "\<bar>norm a\<bar> = norm a"
-by (rule abs_of_nonneg [OF norm_ge_zero])
-
-lemma norm_add_less:
- fixes x y :: "'a::real_normed_vector"
- shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
-by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
-
-lemma norm_mult_less:
- fixes x y :: "'a::real_normed_algebra"
- shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
-apply (rule order_le_less_trans [OF norm_mult_ineq])
-apply (simp add: mult_strict_mono')
-done
-
-lemma norm_of_real [simp]:
- "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
-unfolding of_real_def by simp
-
-lemma norm_numeral [simp]:
- "norm (numeral w::'a::real_normed_algebra_1) = numeral w"
-by (subst of_real_numeral [symmetric], subst norm_of_real, simp)
-
-lemma norm_neg_numeral [simp]:
- "norm (neg_numeral w::'a::real_normed_algebra_1) = numeral w"
-by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp)
-
-lemma norm_of_int [simp]:
- "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
-by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
-
-lemma norm_of_nat [simp]:
- "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
-apply (subst of_real_of_nat_eq [symmetric])
-apply (subst norm_of_real, simp)
-done
-
-lemma nonzero_norm_inverse:
- fixes a :: "'a::real_normed_div_algebra"
- shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
-apply (rule inverse_unique [symmetric])
-apply (simp add: norm_mult [symmetric])
-done
-
-lemma norm_inverse:
- fixes a :: "'a::{real_normed_div_algebra, division_ring_inverse_zero}"
- shows "norm (inverse a) = inverse (norm a)"
-apply (case_tac "a = 0", simp)
-apply (erule nonzero_norm_inverse)
-done
-
-lemma nonzero_norm_divide:
- fixes a b :: "'a::real_normed_field"
- shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
-by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
-
-lemma norm_divide:
- fixes a b :: "'a::{real_normed_field, field_inverse_zero}"
- shows "norm (a / b) = norm a / norm b"
-by (simp add: divide_inverse norm_mult norm_inverse)
-
-lemma norm_power_ineq:
- fixes x :: "'a::{real_normed_algebra_1}"
- shows "norm (x ^ n) \<le> norm x ^ n"
-proof (induct n)
- case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
-next
- case (Suc n)
- have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
- by (rule norm_mult_ineq)
- also from Suc have "\<dots> \<le> norm x * norm x ^ n"
- using norm_ge_zero by (rule mult_left_mono)
- finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n"
- by simp
-qed
-
-lemma norm_power:
- fixes x :: "'a::{real_normed_div_algebra}"
- shows "norm (x ^ n) = norm x ^ n"
-by (induct n) (simp_all add: norm_mult)
-
-text {* Every normed vector space is a metric space. *}
-
-instance real_normed_vector < metric_space
-proof
- fix x y :: 'a show "dist x y = 0 \<longleftrightarrow> x = y"
- unfolding dist_norm by simp
-next
- fix x y z :: 'a show "dist x y \<le> dist x z + dist y z"
- unfolding dist_norm
- using norm_triangle_ineq4 [of "x - z" "y - z"] by simp
-qed
-
-subsection {* Class instances for real numbers *}
-
-instantiation real :: real_normed_field
-begin
-
-definition real_norm_def [simp]:
- "norm r = \<bar>r\<bar>"
-
-instance
-apply (intro_classes, unfold real_norm_def real_scaleR_def)
-apply (rule dist_real_def)
-apply (simp add: sgn_real_def)
-apply (rule abs_eq_0)
-apply (rule abs_triangle_ineq)
-apply (rule abs_mult)
-apply (rule abs_mult)
-done
-
-end
-
-subsection {* Extra type constraints *}
-
-text {* Only allow @{term "open"} in class @{text topological_space}. *}
-
-setup {* Sign.add_const_constraint
- (@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"}) *}
-
-text {* Only allow @{term dist} in class @{text metric_space}. *}
-
-setup {* Sign.add_const_constraint
- (@{const_name dist}, SOME @{typ "'a::metric_space \<Rightarrow> 'a \<Rightarrow> real"}) *}
-
-text {* Only allow @{term norm} in class @{text real_normed_vector}. *}
-
-setup {* Sign.add_const_constraint
- (@{const_name norm}, SOME @{typ "'a::real_normed_vector \<Rightarrow> real"}) *}
-
-subsection {* Sign function *}
-
-lemma norm_sgn:
- "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
-by (simp add: sgn_div_norm)
-
-lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
-by (simp add: sgn_div_norm)
-
-lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
-by (simp add: sgn_div_norm)
-
-lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
-by (simp add: sgn_div_norm)
-
-lemma sgn_scaleR:
- "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
-by (simp add: sgn_div_norm mult_ac)
-
-lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
-by (simp add: sgn_div_norm)
-
-lemma sgn_of_real:
- "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
-unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
-
-lemma sgn_mult:
- fixes x y :: "'a::real_normed_div_algebra"
- shows "sgn (x * y) = sgn x * sgn y"
-by (simp add: sgn_div_norm norm_mult mult_commute)
-
-lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
-by (simp add: sgn_div_norm divide_inverse)
-
-lemma real_sgn_pos: "0 < (x::real) \<Longrightarrow> sgn x = 1"
-unfolding real_sgn_eq by simp
-
-lemma real_sgn_neg: "(x::real) < 0 \<Longrightarrow> sgn x = -1"
-unfolding real_sgn_eq by simp
-
-lemma norm_conv_dist: "norm x = dist x 0"
- unfolding dist_norm by simp
-
-subsection {* Bounded Linear and Bilinear Operators *}
-
-locale bounded_linear = additive f for f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" +
- assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
- assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
-begin
-
-lemma pos_bounded:
- "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
-proof -
- obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
- using bounded by fast
- show ?thesis
- proof (intro exI impI conjI allI)
- show "0 < max 1 K"
- by (rule order_less_le_trans [OF zero_less_one le_maxI1])
- next
- fix x
- have "norm (f x) \<le> norm x * K" using K .
- also have "\<dots> \<le> norm x * max 1 K"
- by (rule mult_left_mono [OF le_maxI2 norm_ge_zero])
- finally show "norm (f x) \<le> norm x * max 1 K" .
- qed
-qed
-
-lemma nonneg_bounded:
- "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
-proof -
- from pos_bounded
- show ?thesis by (auto intro: order_less_imp_le)
-qed
-
-end
-
-lemma bounded_linear_intro:
- assumes "\<And>x y. f (x + y) = f x + f y"
- assumes "\<And>r x. f (scaleR r x) = scaleR r (f x)"
- assumes "\<And>x. norm (f x) \<le> norm x * K"
- shows "bounded_linear f"
- by default (fast intro: assms)+
-
-locale bounded_bilinear =
- fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
- \<Rightarrow> 'c::real_normed_vector"
- (infixl "**" 70)
- assumes add_left: "prod (a + a') b = prod a b + prod a' b"
- assumes add_right: "prod a (b + b') = prod a b + prod a b'"
- assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
- assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
- assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
-begin
-
-lemma pos_bounded:
- "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
-apply (cut_tac bounded, erule exE)
-apply (rule_tac x="max 1 K" in exI, safe)
-apply (rule order_less_le_trans [OF zero_less_one le_maxI1])
-apply (drule spec, drule spec, erule order_trans)
-apply (rule mult_left_mono [OF le_maxI2])
-apply (intro mult_nonneg_nonneg norm_ge_zero)
-done
-
-lemma nonneg_bounded:
- "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
-proof -
- from pos_bounded
- show ?thesis by (auto intro: order_less_imp_le)
-qed
-
-lemma additive_right: "additive (\<lambda>b. prod a b)"
-by (rule additive.intro, rule add_right)
-
-lemma additive_left: "additive (\<lambda>a. prod a b)"
-by (rule additive.intro, rule add_left)
-
-lemma zero_left: "prod 0 b = 0"
-by (rule additive.zero [OF additive_left])
-
-lemma zero_right: "prod a 0 = 0"
-by (rule additive.zero [OF additive_right])
-
-lemma minus_left: "prod (- a) b = - prod a b"
-by (rule additive.minus [OF additive_left])
-
-lemma minus_right: "prod a (- b) = - prod a b"
-by (rule additive.minus [OF additive_right])
-
-lemma diff_left:
- "prod (a - a') b = prod a b - prod a' b"
-by (rule additive.diff [OF additive_left])
-
-lemma diff_right:
- "prod a (b - b') = prod a b - prod a b'"
-by (rule additive.diff [OF additive_right])
-
-lemma bounded_linear_left:
- "bounded_linear (\<lambda>a. a ** b)"
-apply (cut_tac bounded, safe)
-apply (rule_tac K="norm b * K" in bounded_linear_intro)
-apply (rule add_left)
-apply (rule scaleR_left)
-apply (simp add: mult_ac)
-done
-
-lemma bounded_linear_right:
- "bounded_linear (\<lambda>b. a ** b)"
-apply (cut_tac bounded, safe)
-apply (rule_tac K="norm a * K" in bounded_linear_intro)
-apply (rule add_right)
-apply (rule scaleR_right)
-apply (simp add: mult_ac)
-done
-
-lemma prod_diff_prod:
- "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
-by (simp add: diff_left diff_right)
-
-end
-
-lemma bounded_bilinear_mult:
- "bounded_bilinear (op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra)"
-apply (rule bounded_bilinear.intro)
-apply (rule distrib_right)
-apply (rule distrib_left)
-apply (rule mult_scaleR_left)
-apply (rule mult_scaleR_right)
-apply (rule_tac x="1" in exI)
-apply (simp add: norm_mult_ineq)
-done
-
-lemma bounded_linear_mult_left:
- "bounded_linear (\<lambda>x::'a::real_normed_algebra. x * y)"
- using bounded_bilinear_mult
- by (rule bounded_bilinear.bounded_linear_left)
-
-lemma bounded_linear_mult_right:
- "bounded_linear (\<lambda>y::'a::real_normed_algebra. x * y)"
- using bounded_bilinear_mult
- by (rule bounded_bilinear.bounded_linear_right)
-
-lemma bounded_linear_divide:
- "bounded_linear (\<lambda>x::'a::real_normed_field. x / y)"
- unfolding divide_inverse by (rule bounded_linear_mult_left)
-
-lemma bounded_bilinear_scaleR: "bounded_bilinear scaleR"
-apply (rule bounded_bilinear.intro)
-apply (rule scaleR_left_distrib)
-apply (rule scaleR_right_distrib)
-apply simp
-apply (rule scaleR_left_commute)
-apply (rule_tac x="1" in exI, simp)
-done
-
-lemma bounded_linear_scaleR_left: "bounded_linear (\<lambda>r. scaleR r x)"
- using bounded_bilinear_scaleR
- by (rule bounded_bilinear.bounded_linear_left)
-
-lemma bounded_linear_scaleR_right: "bounded_linear (\<lambda>x. scaleR r x)"
- using bounded_bilinear_scaleR
- by (rule bounded_bilinear.bounded_linear_right)
-
-lemma bounded_linear_of_real: "bounded_linear (\<lambda>r. of_real r)"
- unfolding of_real_def by (rule bounded_linear_scaleR_left)
-
-instance real_normed_algebra_1 \<subseteq> perfect_space
-proof
- fix x::'a
- show "\<not> open {x}"
- unfolding open_dist dist_norm
- by (clarsimp, rule_tac x="x + of_real (e/2)" in exI, simp)
-qed
-
-section {* Connectedness *}
-
-class linear_continuum_topology = linorder_topology + conditional_complete_linorder + inner_dense_linorder
-begin
-
-lemma Inf_notin_open:
- assumes A: "open A" and bnd: "\<forall>a\<in>A. x < a"
- shows "Inf A \<notin> A"
-proof
- assume "Inf A \<in> A"
- then obtain b where "b < Inf A" "{b <.. Inf A} \<subseteq> A"
- using open_left[of A "Inf A" x] assms by auto
- with dense[of b "Inf A"] obtain c where "c < Inf A" "c \<in> A"
- by (auto simp: subset_eq)
- then show False
- using cInf_lower[OF `c \<in> A`, of x] bnd by (metis less_imp_le not_le)
-qed
-
-lemma Sup_notin_open:
- assumes A: "open A" and bnd: "\<forall>a\<in>A. a < x"
- shows "Sup A \<notin> A"
-proof
- assume "Sup A \<in> A"
- then obtain b where "Sup A < b" "{Sup A ..< b} \<subseteq> A"
- using open_right[of A "Sup A" x] assms by auto
- with dense[of "Sup A" b] obtain c where "Sup A < c" "c \<in> A"
- by (auto simp: subset_eq)
- then show False
- using cSup_upper[OF `c \<in> A`, of x] bnd by (metis less_imp_le not_le)
-qed
-
-end
-
-instance real :: linear_continuum_topology ..
-
-lemma connectedI_interval:
- fixes U :: "'a :: linear_continuum_topology set"
- assumes *: "\<And>x y z. x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x \<le> z \<Longrightarrow> z \<le> y \<Longrightarrow> z \<in> U"
- shows "connected U"
-proof (rule connectedI)
- { fix A B assume "open A" "open B" "A \<inter> B \<inter> U = {}" "U \<subseteq> A \<union> B"
- fix x y assume "x < y" "x \<in> A" "y \<in> B" "x \<in> U" "y \<in> U"
-
- let ?z = "Inf (B \<inter> {x <..})"
-
- have "x \<le> ?z" "?z \<le> y"
- using `y \<in> B` `x < y` by (auto intro: cInf_lower[where z=x] cInf_greatest)
- with `x \<in> U` `y \<in> U` have "?z \<in> U"
- by (rule *)
- moreover have "?z \<notin> B \<inter> {x <..}"
- using `open B` by (intro Inf_notin_open) auto
- ultimately have "?z \<in> A"
- using `x \<le> ?z` `A \<inter> B \<inter> U = {}` `x \<in> A` `U \<subseteq> A \<union> B` by auto
-
- { assume "?z < y"
- obtain a where "?z < a" "{?z ..< a} \<subseteq> A"
- using open_right[OF `open A` `?z \<in> A` `?z < y`] by auto
- moreover obtain b where "b \<in> B" "x < b" "b < min a y"
- using cInf_less_iff[of "B \<inter> {x <..}" x "min a y"] `?z < a` `?z < y` `x < y` `y \<in> B`
- by (auto intro: less_imp_le)
- moreover then have "?z \<le> b"
- by (intro cInf_lower[where z=x]) auto
- moreover have "b \<in> U"
- using `x \<le> ?z` `?z \<le> b` `b < min a y`
- by (intro *[OF `x \<in> U` `y \<in> U`]) (auto simp: less_imp_le)
- ultimately have "\<exists>b\<in>B. b \<in> A \<and> b \<in> U"
- by (intro bexI[of _ b]) auto }
- then have False
- using `?z \<le> y` `?z \<in> A` `y \<in> B` `y \<in> U` `A \<inter> B \<inter> U = {}` unfolding le_less by blast }
- note not_disjoint = this
-
- fix A B assume AB: "open A" "open B" "U \<subseteq> A \<union> B" "A \<inter> B \<inter> U = {}"
- moreover assume "A \<inter> U \<noteq> {}" then obtain x where x: "x \<in> U" "x \<in> A" by auto
- moreover assume "B \<inter> U \<noteq> {}" then obtain y where y: "y \<in> U" "y \<in> B" by auto
- moreover note not_disjoint[of B A y x] not_disjoint[of A B x y]
- ultimately show False by (cases x y rule: linorder_cases) auto
-qed
-
-lemma connected_iff_interval:
- fixes U :: "'a :: linear_continuum_topology set"
- shows "connected U \<longleftrightarrow> (\<forall>x\<in>U. \<forall>y\<in>U. \<forall>z. x \<le> z \<longrightarrow> z \<le> y \<longrightarrow> z \<in> U)"
- by (auto intro: connectedI_interval dest: connectedD_interval)
-
-lemma connected_UNIV[simp]: "connected (UNIV::'a::linear_continuum_topology set)"
- unfolding connected_iff_interval by auto
-
-lemma connected_Ioi[simp]: "connected {a::'a::linear_continuum_topology <..}"
- unfolding connected_iff_interval by auto
-
-lemma connected_Ici[simp]: "connected {a::'a::linear_continuum_topology ..}"
- unfolding connected_iff_interval by auto
-
-lemma connected_Iio[simp]: "connected {..< a::'a::linear_continuum_topology}"
- unfolding connected_iff_interval by auto
-
-lemma connected_Iic[simp]: "connected {.. a::'a::linear_continuum_topology}"
- unfolding connected_iff_interval by auto
-
-lemma connected_Ioo[simp]: "connected {a <..< b::'a::linear_continuum_topology}"
- unfolding connected_iff_interval by auto
-
-lemma connected_Ioc[simp]: "connected {a <.. b::'a::linear_continuum_topology}"
- unfolding connected_iff_interval by auto
-
-lemma connected_Ico[simp]: "connected {a ..< b::'a::linear_continuum_topology}"
- unfolding connected_iff_interval by auto
-
-lemma connected_Icc[simp]: "connected {a .. b::'a::linear_continuum_topology}"
- unfolding connected_iff_interval by auto
-
-lemma connected_contains_Ioo:
- fixes A :: "'a :: linorder_topology set"
- assumes A: "connected A" "a \<in> A" "b \<in> A" shows "{a <..< b} \<subseteq> A"
- using connectedD_interval[OF A] by (simp add: subset_eq Ball_def less_imp_le)
-
-subsection {* Intermediate Value Theorem *}
-
-lemma IVT':
- fixes f :: "'a :: linear_continuum_topology \<Rightarrow> 'b :: linorder_topology"
- assumes y: "f a \<le> y" "y \<le> f b" "a \<le> b"
- assumes *: "continuous_on {a .. b} f"
- shows "\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y"
-proof -
- have "connected {a..b}"
- unfolding connected_iff_interval by auto
- from connected_continuous_image[OF * this, THEN connectedD_interval, of "f a" "f b" y] y
- show ?thesis
- by (auto simp add: atLeastAtMost_def atLeast_def atMost_def)
-qed
-
-lemma IVT2':
- fixes f :: "'a :: linear_continuum_topology \<Rightarrow> 'b :: linorder_topology"
- assumes y: "f b \<le> y" "y \<le> f a" "a \<le> b"
- assumes *: "continuous_on {a .. b} f"
- shows "\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y"
-proof -
- have "connected {a..b}"
- unfolding connected_iff_interval by auto
- from connected_continuous_image[OF * this, THEN connectedD_interval, of "f b" "f a" y] y
- show ?thesis
- by (auto simp add: atLeastAtMost_def atLeast_def atMost_def)
-qed
-
-lemma IVT:
- fixes f :: "'a :: linear_continuum_topology \<Rightarrow> 'b :: linorder_topology"
- shows "f a \<le> y \<Longrightarrow> y \<le> f b \<Longrightarrow> a \<le> b \<Longrightarrow> (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x) \<Longrightarrow> \<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y"
- by (rule IVT') (auto intro: continuous_at_imp_continuous_on)
-
-lemma IVT2:
- fixes f :: "'a :: linear_continuum_topology \<Rightarrow> 'b :: linorder_topology"
- shows "f b \<le> y \<Longrightarrow> y \<le> f a \<Longrightarrow> a \<le> b \<Longrightarrow> (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x) \<Longrightarrow> \<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y"
- by (rule IVT2') (auto intro: continuous_at_imp_continuous_on)
-
-lemma continuous_inj_imp_mono: