| author | blanchet | 
| Wed, 16 Nov 2011 17:59:58 +0100 | |
| changeset 45524 | 43ca06e6c168 | 
| parent 44708 | 37ce74ff4203 | 
| child 45966 | 03ce2b2a29a2 | 
| permissions | -rw-r--r-- | 
| 30122 | 1  | 
(* Title: HOL/RComplete.thy  | 
2  | 
Author: Jacques D. Fleuriot, University of Edinburgh  | 
|
3  | 
Author: Larry Paulson, University of Cambridge  | 
|
4  | 
Author: Jeremy Avigad, Carnegie Mellon University  | 
|
5  | 
Author: Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen  | 
|
| 16893 | 6  | 
*)  | 
| 5078 | 7  | 
|
| 16893 | 8  | 
header {* Completeness of the Reals; Floor and Ceiling Functions *}
 | 
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
9  | 
|
| 15131 | 10  | 
theory RComplete  | 
| 15140 | 11  | 
imports Lubs RealDef  | 
| 15131 | 12  | 
begin  | 
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
13  | 
|
| 
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
14  | 
lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"  | 
| 16893 | 15  | 
by simp  | 
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
16  | 
|
| 
32707
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
30242 
diff
changeset
 | 
17  | 
lemma abs_diff_less_iff:  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
32960 
diff
changeset
 | 
18  | 
"(\<bar>x - a\<bar> < (r::'a::linordered_idom)) = (a - r < x \<and> x < a + r)"  | 
| 
32707
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
30242 
diff
changeset
 | 
19  | 
by auto  | 
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
20  | 
|
| 16893 | 21  | 
subsection {* Completeness of Positive Reals *}
 | 
22  | 
||
23  | 
text {*
 | 
|
24  | 
Supremum property for the set of positive reals  | 
|
25  | 
||
26  | 
  Let @{text "P"} be a non-empty set of positive reals, with an upper
 | 
|
27  | 
  bound @{text "y"}.  Then @{text "P"} has a least upper bound
 | 
|
28  | 
  (written @{text "S"}).
 | 
|
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
29  | 
|
| 16893 | 30  | 
  FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
 | 
31  | 
*}  | 
|
32  | 
||
| 
36795
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
33  | 
text {* Only used in HOL/Import/HOL4Compat.thy; delete? *}
 | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
34  | 
|
| 16893 | 35  | 
lemma posreal_complete:  | 
| 
44690
 
b6d8b11ed399
remove unused assumption from lemma posreal_complete
 
huffman 
parents: 
44679 
diff
changeset
 | 
36  | 
fixes P :: "real set"  | 
| 
 
b6d8b11ed399
remove unused assumption from lemma posreal_complete
 
huffman 
parents: 
44679 
diff
changeset
 | 
37  | 
assumes not_empty_P: "\<exists>x. x \<in> P"  | 
| 16893 | 38  | 
and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"  | 
39  | 
shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"  | 
|
| 
36795
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
40  | 
proof -  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
41  | 
from upper_bound_Ex have "\<exists>z. \<forall>x\<in>P. x \<le> z"  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
42  | 
by (auto intro: less_imp_le)  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
43  | 
from complete_real [OF not_empty_P this] obtain S  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
44  | 
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  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
45  | 
have "\<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
46  | 
proof  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
47  | 
fix y show "(\<exists>x\<in>P. y < x) = (y < S)"  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
48  | 
apply (cases "\<exists>x\<in>P. y < x", simp_all)  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
49  | 
apply (clarify, drule S1, simp)  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
50  | 
apply (simp add: not_less S2)  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
51  | 
done  | 
| 16893 | 52  | 
qed  | 
| 
36795
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
53  | 
thus ?thesis ..  | 
| 16893 | 54  | 
qed  | 
55  | 
||
56  | 
text {*
 | 
|
57  | 
  \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
 | 
|
58  | 
*}  | 
|
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
59  | 
|
| 
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
60  | 
lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)"  | 
| 16893 | 61  | 
apply (frule isLub_isUb)  | 
62  | 
apply (frule_tac x = y in isLub_isUb)  | 
|
63  | 
apply (blast intro!: order_antisym dest!: isLub_le_isUb)  | 
|
64  | 
done  | 
|
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
65  | 
|
| 5078 | 66  | 
|
| 16893 | 67  | 
text {*
 | 
68  | 
\medskip reals Completeness (again!)  | 
|
69  | 
*}  | 
|
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
70  | 
|
| 16893 | 71  | 
lemma reals_complete:  | 
72  | 
assumes notempty_S: "\<exists>X. X \<in> S"  | 
|
73  | 
and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"  | 
|
74  | 
shows "\<exists>t. isLub (UNIV :: real set) S t"  | 
|
75  | 
proof -  | 
|
| 
36795
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
76  | 
from assms have "\<exists>X. X \<in> S" and "\<exists>Y. \<forall>x\<in>S. x \<le> Y"  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
77  | 
unfolding isUb_def setle_def by simp_all  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
78  | 
from complete_real [OF this] show ?thesis  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
79  | 
unfolding isLub_def leastP_def setle_def setge_def Ball_def  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
80  | 
Collect_def mem_def isUb_def UNIV_def by simp  | 
| 16893 | 81  | 
qed  | 
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
82  | 
|
| 
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
83  | 
|
| 16893 | 84  | 
subsection {* The Archimedean Property of the Reals *}
 | 
85  | 
||
86  | 
theorem reals_Archimedean:  | 
|
87  | 
assumes x_pos: "0 < x"  | 
|
88  | 
shows "\<exists>n. inverse (real (Suc n)) < x"  | 
|
| 
36795
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
89  | 
unfolding real_of_nat_def using x_pos  | 
| 
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
90  | 
by (rule ex_inverse_of_nat_Suc_less)  | 
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
91  | 
|
| 
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
92  | 
lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"  | 
| 
36795
 
e05e1283c550
new construction of real numbers using Cauchy sequences
 
huffman 
parents: 
35578 
diff
changeset
 | 
93  | 
unfolding real_of_nat_def by (rule ex_less_of_nat)  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
94  | 
|
| 16893 | 95  | 
lemma reals_Archimedean3:  | 
96  | 
assumes x_greater_zero: "0 < x"  | 
|
97  | 
shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
98  | 
unfolding real_of_nat_def using `0 < x`  | 
| 
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
99  | 
by (auto intro: ex_less_of_nat_mult)  | 
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
100  | 
|
| 16819 | 101  | 
|
| 
28091
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
102  | 
subsection{*Density of the Rational Reals in the Reals*}
 | 
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
103  | 
|
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
104  | 
text{* This density proof is due to Stefan Richter and was ported by TN.  The
 | 
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
105  | 
original source is \emph{Real Analysis} by H.L. Royden.
 | 
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
106  | 
It employs the Archimedean property of the reals. *}  | 
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
107  | 
|
| 44668 | 108  | 
lemma Rats_dense_in_real:  | 
109  | 
fixes x :: real  | 
|
110  | 
assumes "x < y" shows "\<exists>r\<in>\<rat>. x < r \<and> r < y"  | 
|
| 
28091
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
111  | 
proof -  | 
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
112  | 
from `x<y` have "0 < y-x" by simp  | 
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
113  | 
with reals_Archimedean obtain q::nat  | 
| 44668 | 114  | 
where q: "inverse (real q) < y-x" and "0 < q" by auto  | 
115  | 
def p \<equiv> "ceiling (y * real q) - 1"  | 
|
116  | 
def r \<equiv> "of_int p / real q"  | 
|
117  | 
from q have "x < y - inverse (real q)" by simp  | 
|
118  | 
also have "y - inverse (real q) \<le> r"  | 
|
119  | 
unfolding r_def p_def  | 
|
120  | 
by (simp add: le_divide_eq left_diff_distrib le_of_int_ceiling `0 < q`)  | 
|
121  | 
finally have "x < r" .  | 
|
122  | 
moreover have "r < y"  | 
|
123  | 
unfolding r_def p_def  | 
|
124  | 
by (simp add: divide_less_eq diff_less_eq `0 < q`  | 
|
125  | 
less_ceiling_iff [symmetric])  | 
|
126  | 
moreover from r_def have "r \<in> \<rat>" by simp  | 
|
| 
28091
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
127  | 
ultimately show ?thesis by fast  | 
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
128  | 
qed  | 
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
129  | 
|
| 
 
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
 
nipkow 
parents: 
27435 
diff
changeset
 | 
130  | 
|
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
131  | 
subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
 | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
132  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
133  | 
lemma number_of_less_real_of_int_iff [simp]:  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
134  | 
"((number_of n) < real (m::int)) = (number_of n < m)"  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
135  | 
apply auto  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
136  | 
apply (rule real_of_int_less_iff [THEN iffD1])  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
137  | 
apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
138  | 
done  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
139  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
140  | 
lemma number_of_less_real_of_int_iff2 [simp]:  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
141  | 
"(real (m::int) < (number_of n)) = (m < number_of n)"  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
142  | 
apply auto  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
143  | 
apply (rule real_of_int_less_iff [THEN iffD1])  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
144  | 
apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
145  | 
done  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
146  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
147  | 
lemma number_of_le_real_of_int_iff [simp]:  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
148  | 
"((number_of n) \<le> real (m::int)) = (number_of n \<le> m)"  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
149  | 
by (simp add: linorder_not_less [symmetric])  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
150  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
151  | 
lemma number_of_le_real_of_int_iff2 [simp]:  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
152  | 
"(real (m::int) \<le> (number_of n)) = (m \<le> number_of n)"  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
153  | 
by (simp add: linorder_not_less [symmetric])  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
154  | 
|
| 24355 | 155  | 
lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
156  | 
unfolding real_of_nat_def by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
157  | 
|
| 24355 | 158  | 
lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"  | 
| 30102 | 159  | 
unfolding real_of_nat_def by (simp add: floor_minus)  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
160  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
161  | 
lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
162  | 
unfolding real_of_int_def by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
163  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
164  | 
lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"  | 
| 30102 | 165  | 
unfolding real_of_int_def by (simp add: floor_minus)  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
166  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
167  | 
lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
168  | 
unfolding real_of_int_def by (rule floor_exists)  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
169  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
170  | 
lemma lemma_floor:  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
171  | 
assumes a1: "real m \<le> r" and a2: "r < real n + 1"  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
172  | 
shows "m \<le> (n::int)"  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
173  | 
proof -  | 
| 23389 | 174  | 
have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)  | 
175  | 
also have "... = real (n + 1)" by simp  | 
|
176  | 
finally have "m < n + 1" by (simp only: real_of_int_less_iff)  | 
|
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
177  | 
thus ?thesis by arith  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
178  | 
qed  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
179  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
180  | 
lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
181  | 
unfolding real_of_int_def by (rule of_int_floor_le)  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
182  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
183  | 
lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \<le> x"  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
184  | 
by (auto intro: lemma_floor)  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
185  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
186  | 
lemma real_of_int_floor_cancel [simp]:  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
187  | 
"(real (floor x) = x) = (\<exists>n::int. x = real n)"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
188  | 
using floor_real_of_int by metis  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
189  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
190  | 
lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
191  | 
unfolding real_of_int_def using floor_unique [of n x] by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
192  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
193  | 
lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
194  | 
unfolding real_of_int_def by (rule floor_unique)  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
195  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
196  | 
lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n"  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
197  | 
apply (rule inj_int [THEN injD])  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
198  | 
apply (simp add: real_of_nat_Suc)  | 
| 15539 | 199  | 
apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"])  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
200  | 
done  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
201  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
202  | 
lemma floor_eq4: "[| real n \<le> x; x < real (Suc n) |] ==> nat(floor x) = n"  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
203  | 
apply (drule order_le_imp_less_or_eq)  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
204  | 
apply (auto intro: floor_eq3)  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
205  | 
done  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
206  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
207  | 
lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
208  | 
unfolding real_of_int_def using floor_correct [of r] by simp  | 
| 16819 | 209  | 
|
210  | 
lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
211  | 
unfolding real_of_int_def using floor_correct [of r] by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
212  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
213  | 
lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
214  | 
unfolding real_of_int_def using floor_correct [of r] by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
215  | 
|
| 16819 | 216  | 
lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
217  | 
unfolding real_of_int_def using floor_correct [of r] by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
218  | 
|
| 16819 | 219  | 
lemma le_floor: "real a <= x ==> a <= floor x"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
220  | 
unfolding real_of_int_def by (simp add: le_floor_iff)  | 
| 16819 | 221  | 
|
222  | 
lemma real_le_floor: "a <= floor x ==> real a <= x"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
223  | 
unfolding real_of_int_def by (simp add: le_floor_iff)  | 
| 16819 | 224  | 
|
225  | 
lemma le_floor_eq: "(a <= floor x) = (real a <= x)"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
226  | 
unfolding real_of_int_def by (rule le_floor_iff)  | 
| 16819 | 227  | 
|
228  | 
lemma floor_less_eq: "(floor x < a) = (x < real a)"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
229  | 
unfolding real_of_int_def by (rule floor_less_iff)  | 
| 16819 | 230  | 
|
231  | 
lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
232  | 
unfolding real_of_int_def by (rule less_floor_iff)  | 
| 16819 | 233  | 
|
234  | 
lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
235  | 
unfolding real_of_int_def by (rule floor_le_iff)  | 
| 16819 | 236  | 
|
237  | 
lemma floor_add [simp]: "floor (x + real a) = floor x + a"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
238  | 
unfolding real_of_int_def by (rule floor_add_of_int)  | 
| 16819 | 239  | 
|
240  | 
lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
241  | 
unfolding real_of_int_def by (rule floor_diff_of_int)  | 
| 16819 | 242  | 
|
| 
35578
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
243  | 
lemma le_mult_floor:  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
244  | 
assumes "0 \<le> (a :: real)" and "0 \<le> b"  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
245  | 
shows "floor a * floor b \<le> floor (a * b)"  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
246  | 
proof -  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
247  | 
have "real (floor a) \<le> a"  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
248  | 
and "real (floor b) \<le> b" by auto  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
249  | 
hence "real (floor a * floor b) \<le> a * b"  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
250  | 
using assms by (auto intro!: mult_mono)  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
251  | 
also have "a * b < real (floor (a * b) + 1)" by auto  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
252  | 
finally show ?thesis unfolding real_of_int_less_iff by simp  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
253  | 
qed  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
254  | 
|
| 24355 | 255  | 
lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
256  | 
unfolding real_of_nat_def by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
257  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
258  | 
lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
259  | 
unfolding real_of_int_def by (rule le_of_int_ceiling)  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
260  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
261  | 
lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"  | 
| 
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
262  | 
unfolding real_of_int_def by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
263  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
264  | 
lemma real_of_int_ceiling_cancel [simp]:  | 
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
265  | 
"(real (ceiling x) = x) = (\<exists>n::int. x = real n)"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
266  | 
using ceiling_real_of_int by metis  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
267  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
268  | 
lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
269  | 
unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
270  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
271  | 
lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
272  | 
unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
273  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
274  | 
lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n |] ==> ceiling x = n"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
275  | 
unfolding real_of_int_def using ceiling_unique [of n x] by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
276  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
277  | 
lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
278  | 
unfolding real_of_int_def using ceiling_correct [of r] by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
279  | 
|
| 
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
280  | 
lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
281  | 
unfolding real_of_int_def using ceiling_correct [of r] by simp  | 
| 
14641
 
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
 
paulson 
parents: 
14476 
diff
changeset
 | 
282  | 
|
| 16819 | 283  | 
lemma ceiling_le: "x <= real a ==> ceiling x <= a"  | 
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
284  | 
unfolding real_of_int_def by (simp add: ceiling_le_iff)  | 
| 16819 | 285  | 
|
286  | 
lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
287  | 
unfolding real_of_int_def by (simp add: ceiling_le_iff)  | 
| 16819 | 288  | 
|
289  | 
lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
290  | 
unfolding real_of_int_def by (rule ceiling_le_iff)  | 
| 16819 | 291  | 
|
292  | 
lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
293  | 
unfolding real_of_int_def by (rule less_ceiling_iff)  | 
| 16819 | 294  | 
|
295  | 
lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
296  | 
unfolding real_of_int_def by (rule ceiling_less_iff)  | 
| 16819 | 297  | 
|
298  | 
lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
299  | 
unfolding real_of_int_def by (rule le_ceiling_iff)  | 
| 16819 | 300  | 
|
301  | 
lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
302  | 
unfolding real_of_int_def by (rule ceiling_add_of_int)  | 
| 16819 | 303  | 
|
304  | 
lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"  | 
|
| 
30097
 
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
 
huffman 
parents: 
29667 
diff
changeset
 | 
305  | 
unfolding real_of_int_def by (rule ceiling_diff_of_int)  | 
| 16819 | 306  | 
|
307  | 
||
308  | 
subsection {* Versions for the natural numbers *}
 | 
|
309  | 
||
| 19765 | 310  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
311  | 
natfloor :: "real => nat" where  | 
| 19765 | 312  | 
"natfloor x = nat(floor x)"  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
313  | 
|
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
314  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
315  | 
natceiling :: "real => nat" where  | 
| 19765 | 316  | 
"natceiling x = nat(ceiling x)"  | 
| 16819 | 317  | 
|
318  | 
lemma natfloor_zero [simp]: "natfloor 0 = 0"  | 
|
319  | 
by (unfold natfloor_def, simp)  | 
|
320  | 
||
321  | 
lemma natfloor_one [simp]: "natfloor 1 = 1"  | 
|
322  | 
by (unfold natfloor_def, simp)  | 
|
323  | 
||
324  | 
lemma zero_le_natfloor [simp]: "0 <= natfloor x"  | 
|
325  | 
by (unfold natfloor_def, simp)  | 
|
326  | 
||
327  | 
lemma natfloor_number_of_eq [simp]: "natfloor (number_of n) = number_of n"  | 
|
328  | 
by (unfold natfloor_def, simp)  | 
|
329  | 
||
330  | 
lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"  | 
|
331  | 
by (unfold natfloor_def, simp)  | 
|
332  | 
||
333  | 
lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"  | 
|
334  | 
by (unfold natfloor_def, simp)  | 
|
335  | 
||
336  | 
lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"  | 
|
| 44679 | 337  | 
unfolding natfloor_def by simp  | 
338  | 
||
| 16819 | 339  | 
lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"  | 
| 44679 | 340  | 
unfolding natfloor_def by (intro nat_mono floor_mono)  | 
| 16819 | 341  | 
|
342  | 
lemma le_natfloor: "real x <= a ==> x <= natfloor a"  | 
|
343  | 
apply (unfold natfloor_def)  | 
|
344  | 
apply (subst nat_int [THEN sym])  | 
|
| 44679 | 345  | 
apply (rule nat_mono)  | 
| 16819 | 346  | 
apply (rule le_floor)  | 
347  | 
apply simp  | 
|
348  | 
done  | 
|
349  | 
||
| 44679 | 350  | 
lemma natfloor_less_iff: "0 \<le> x \<Longrightarrow> natfloor x < n \<longleftrightarrow> x < real n"  | 
351  | 
unfolding natfloor_def real_of_nat_def  | 
|
352  | 
by (simp add: nat_less_iff floor_less_iff)  | 
|
353  | 
||
| 
35578
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
354  | 
lemma less_natfloor:  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
355  | 
assumes "0 \<le> x" and "x < real (n :: nat)"  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
356  | 
shows "natfloor x < n"  | 
| 44679 | 357  | 
using assms by (simp add: natfloor_less_iff)  | 
| 
35578
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
358  | 
|
| 16819 | 359  | 
lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)"  | 
360  | 
apply (rule iffI)  | 
|
361  | 
apply (rule order_trans)  | 
|
362  | 
prefer 2  | 
|
363  | 
apply (erule real_natfloor_le)  | 
|
364  | 
apply (subst real_of_nat_le_iff)  | 
|
365  | 
apply assumption  | 
|
366  | 
apply (erule le_natfloor)  | 
|
367  | 
done  | 
|
368  | 
||
| 16893 | 369  | 
lemma le_natfloor_eq_number_of [simp]:  | 
| 16819 | 370  | 
"~ neg((number_of n)::int) ==> 0 <= x ==>  | 
371  | 
(number_of n <= natfloor x) = (number_of n <= x)"  | 
|
372  | 
apply (subst le_natfloor_eq, assumption)  | 
|
373  | 
apply simp  | 
|
374  | 
done  | 
|
375  | 
||
| 16820 | 376  | 
lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"  | 
| 16819 | 377  | 
apply (case_tac "0 <= x")  | 
378  | 
apply (subst le_natfloor_eq, assumption, simp)  | 
|
379  | 
apply (rule iffI)  | 
|
| 16893 | 380  | 
apply (subgoal_tac "natfloor x <= natfloor 0")  | 
| 16819 | 381  | 
apply simp  | 
382  | 
apply (rule natfloor_mono)  | 
|
383  | 
apply simp  | 
|
384  | 
apply simp  | 
|
385  | 
done  | 
|
386  | 
||
387  | 
lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n"  | 
|
| 44679 | 388  | 
unfolding natfloor_def by (simp add: floor_eq2 [where n="int n"])  | 
| 16819 | 389  | 
|
390  | 
lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"  | 
|
391  | 
apply (case_tac "0 <= x")  | 
|
392  | 
apply (unfold natfloor_def)  | 
|
393  | 
apply simp  | 
|
394  | 
apply simp_all  | 
|
395  | 
done  | 
|
396  | 
||
397  | 
lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"  | 
|
| 29667 | 398  | 
using real_natfloor_add_one_gt by (simp add: algebra_simps)  | 
| 16819 | 399  | 
|
400  | 
lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"  | 
|
401  | 
apply (subgoal_tac "z < real(natfloor z) + 1")  | 
|
402  | 
apply arith  | 
|
403  | 
apply (rule real_natfloor_add_one_gt)  | 
|
404  | 
done  | 
|
405  | 
||
406  | 
lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a"  | 
|
| 44679 | 407  | 
unfolding natfloor_def  | 
408  | 
unfolding real_of_int_of_nat_eq [symmetric] floor_add  | 
|
409  | 
by (simp add: nat_add_distrib)  | 
|
| 16819 | 410  | 
|
| 16893 | 411  | 
lemma natfloor_add_number_of [simp]:  | 
412  | 
"~neg ((number_of n)::int) ==> 0 <= x ==>  | 
|
| 16819 | 413  | 
natfloor (x + number_of n) = natfloor x + number_of n"  | 
| 44679 | 414  | 
by (simp add: natfloor_add [symmetric])  | 
| 16819 | 415  | 
|
416  | 
lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1"  | 
|
| 44679 | 417  | 
by (simp add: natfloor_add [symmetric] del: One_nat_def)  | 
| 16819 | 418  | 
|
| 16893 | 419  | 
lemma natfloor_subtract [simp]: "real a <= x ==>  | 
| 16819 | 420  | 
natfloor(x - real a) = natfloor x - a"  | 
| 44679 | 421  | 
unfolding natfloor_def  | 
422  | 
unfolding real_of_int_of_nat_eq [symmetric] floor_subtract  | 
|
423  | 
by simp  | 
|
| 16819 | 424  | 
|
| 41550 | 425  | 
lemma natfloor_div_nat:  | 
426  | 
assumes "1 <= x" and "y > 0"  | 
|
427  | 
shows "natfloor (x / real y) = natfloor x div y"  | 
|
| 44679 | 428  | 
proof (rule natfloor_eq)  | 
429  | 
have "(natfloor x) div y * y \<le> natfloor x"  | 
|
430  | 
by (rule add_leD1 [where k="natfloor x mod y"], simp)  | 
|
431  | 
thus "real (natfloor x div y) \<le> x / real y"  | 
|
432  | 
using assms by (simp add: le_divide_eq le_natfloor_eq)  | 
|
433  | 
have "natfloor x < (natfloor x) div y * y + y"  | 
|
434  | 
apply (subst mod_div_equality [symmetric])  | 
|
435  | 
apply (rule add_strict_left_mono)  | 
|
436  | 
apply (rule mod_less_divisor)  | 
|
437  | 
apply fact  | 
|
| 
35578
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
438  | 
done  | 
| 44679 | 439  | 
thus "x / real y < real (natfloor x div y) + 1"  | 
440  | 
using assms  | 
|
441  | 
by (simp add: divide_less_eq natfloor_less_iff left_distrib)  | 
|
| 
35578
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
442  | 
qed  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
443  | 
|
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
444  | 
lemma le_mult_natfloor:  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
445  | 
assumes "0 \<le> (a :: real)" and "0 \<le> b"  | 
| 
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
446  | 
shows "natfloor a * natfloor b \<le> natfloor (a * b)"  | 
| 44679 | 447  | 
using assms  | 
448  | 
by (simp add: le_natfloor_eq mult_nonneg_nonneg mult_mono' real_natfloor_le)  | 
|
| 
35578
 
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
 
hoelzl 
parents: 
35028 
diff
changeset
 | 
449  | 
|
| 16819 | 450  | 
lemma natceiling_zero [simp]: "natceiling 0 = 0"  | 
451  | 
by (unfold natceiling_def, simp)  | 
|
452  | 
||
453  | 
lemma natceiling_one [simp]: "natceiling 1 = 1"  | 
|
454  | 
by (unfold natceiling_def, simp)  | 
|
455  | 
||
456  | 
lemma zero_le_natceiling [simp]: "0 <= natceiling x"  | 
|
457  | 
by (unfold natceiling_def, simp)  | 
|
458  | 
||
459  | 
lemma natceiling_number_of_eq [simp]: "natceiling (number_of n) = number_of n"  | 
|
460  | 
by (unfold natceiling_def, simp)  | 
|
461  | 
||
462  | 
lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"  | 
|
463  | 
by (unfold natceiling_def, simp)  | 
|
464  | 
||
465  | 
lemma real_natceiling_ge: "x <= real(natceiling x)"  | 
|
| 44679 | 466  | 
unfolding natceiling_def by (cases "x < 0", simp_all)  | 
| 16819 | 467  | 
|
468  | 
lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"  | 
|
| 44679 | 469  | 
unfolding natceiling_def by simp  | 
| 16819 | 470  | 
|
471  | 
lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"  | 
|
| 44679 | 472  | 
unfolding natceiling_def by (intro nat_mono ceiling_mono)  | 
473  | 
||
| 16819 | 474  | 
lemma natceiling_le: "x <= real a ==> natceiling x <= a"  | 
| 44679 | 475  | 
unfolding natceiling_def real_of_nat_def  | 
476  | 
by (simp add: nat_le_iff ceiling_le_iff)  | 
|
| 16819 | 477  | 
|
| 44708 | 478  | 
lemma natceiling_le_eq: "(natceiling x <= a) = (x <= real a)"  | 
479  | 
unfolding natceiling_def real_of_nat_def  | 
|
| 44679 | 480  | 
by (simp add: nat_le_iff ceiling_le_iff)  | 
| 16819 | 481  | 
|
| 16893 | 482  | 
lemma natceiling_le_eq_number_of [simp]:  | 
| 44708 | 483  | 
"~ neg((number_of n)::int) ==>  | 
| 16820 | 484  | 
(natceiling x <= number_of n) = (x <= number_of n)"  | 
| 44679 | 485  | 
by (simp add: natceiling_le_eq)  | 
| 16819 | 486  | 
|
| 16820 | 487  | 
lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"  | 
| 44679 | 488  | 
unfolding natceiling_def  | 
489  | 
by (simp add: nat_le_iff ceiling_le_iff)  | 
|
| 16819 | 490  | 
|
491  | 
lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1"  | 
|
| 44679 | 492  | 
unfolding natceiling_def  | 
493  | 
by (simp add: ceiling_eq2 [where n="int n"])  | 
|
| 16819 | 494  | 
|
| 16893 | 495  | 
lemma natceiling_add [simp]: "0 <= x ==>  | 
| 16819 | 496  | 
natceiling (x + real a) = natceiling x + a"  | 
| 44679 | 497  | 
unfolding natceiling_def  | 
498  | 
unfolding real_of_int_of_nat_eq [symmetric] ceiling_add  | 
|
499  | 
by (simp add: nat_add_distrib)  | 
|
| 16819 | 500  | 
|
| 16893 | 501  | 
lemma natceiling_add_number_of [simp]:  | 
502  | 
"~ neg ((number_of n)::int) ==> 0 <= x ==>  | 
|
| 16820 | 503  | 
natceiling (x + number_of n) = natceiling x + number_of n"  | 
| 44679 | 504  | 
by (simp add: natceiling_add [symmetric])  | 
| 16819 | 505  | 
|
506  | 
lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1"  | 
|
| 44679 | 507  | 
by (simp add: natceiling_add [symmetric] del: One_nat_def)  | 
| 16819 | 508  | 
|
| 16893 | 509  | 
lemma natceiling_subtract [simp]: "real a <= x ==>  | 
| 16819 | 510  | 
natceiling(x - real a) = natceiling x - a"  | 
| 44679 | 511  | 
unfolding natceiling_def  | 
512  | 
unfolding real_of_int_of_nat_eq [symmetric] ceiling_subtract  | 
|
513  | 
by simp  | 
|
| 16819 | 514  | 
|
| 
36826
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
515  | 
subsection {* Exponentiation with floor *}
 | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
516  | 
|
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
517  | 
lemma floor_power:  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
518  | 
assumes "x = real (floor x)"  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
519  | 
shows "floor (x ^ n) = floor x ^ n"  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
520  | 
proof -  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
521  | 
have *: "x ^ n = real (floor x ^ n)"  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
522  | 
using assms by (induct n arbitrary: x) simp_all  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
523  | 
show ?thesis unfolding real_of_int_inject[symmetric]  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
524  | 
unfolding * floor_real_of_int ..  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
525  | 
qed  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
526  | 
|
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
527  | 
lemma natfloor_power:  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
528  | 
assumes "x = real (natfloor x)"  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
529  | 
shows "natfloor (x ^ n) = natfloor x ^ n"  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
530  | 
proof -  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
531  | 
from assms have "0 \<le> floor x" by auto  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
532  | 
note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \<le> floor x`]]  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
533  | 
from floor_power[OF this]  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
534  | 
show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \<le> floor x`, symmetric]  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
535  | 
by simp  | 
| 
 
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
 
huffman 
parents: 
36795 
diff
changeset
 | 
536  | 
qed  | 
| 16819 | 537  | 
|
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
9429 
diff
changeset
 | 
538  | 
end  |