author | wenzelm |
Wed, 04 May 2011 15:37:39 +0200 | |
changeset 42676 | 8724f20bf69c |
parent 41550 | efa734d9b221 |
child 44667 | ee5772ca7d16 |
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: |
36 |
assumes positive_P: "\<forall>x \<in> P. (0::real) < x" |
|
37 |
and not_empty_P: "\<exists>x. x \<in> P" |
|
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 |
|
32707
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
paulson
parents:
30242
diff
changeset
|
83 |
text{*A version of the same theorem without all those predicates!*} |
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
paulson
parents:
30242
diff
changeset
|
84 |
lemma reals_complete2: |
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
paulson
parents:
30242
diff
changeset
|
85 |
fixes S :: "(real set)" |
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
paulson
parents:
30242
diff
changeset
|
86 |
assumes "\<exists>y. y\<in>S" and "\<exists>(x::real). \<forall>y\<in>S. y \<le> x" |
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
paulson
parents:
30242
diff
changeset
|
87 |
shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) & |
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
paulson
parents:
30242
diff
changeset
|
88 |
(\<forall>z. ((\<forall>y\<in>S. y \<le> z) --> x \<le> z))" |
36795
e05e1283c550
new construction of real numbers using Cauchy sequences
huffman
parents:
35578
diff
changeset
|
89 |
using assms by (rule complete_real) |
32707
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
paulson
parents:
30242
diff
changeset
|
90 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
9429
diff
changeset
|
91 |
|
16893 | 92 |
subsection {* The Archimedean Property of the Reals *} |
93 |
||
94 |
theorem reals_Archimedean: |
|
95 |
assumes x_pos: "0 < x" |
|
96 |
shows "\<exists>n. inverse (real (Suc n)) < x" |
|
36795
e05e1283c550
new construction of real numbers using Cauchy sequences
huffman
parents:
35578
diff
changeset
|
97 |
unfolding real_of_nat_def using x_pos |
e05e1283c550
new construction of real numbers using Cauchy sequences
huffman
parents:
35578
diff
changeset
|
98 |
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
|
99 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
9429
diff
changeset
|
100 |
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
|
101 |
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
|
102 |
|
16893 | 103 |
lemma reals_Archimedean3: |
104 |
assumes x_greater_zero: "0 < x" |
|
105 |
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
|
106 |
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
|
107 |
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
|
108 |
|
16819 | 109 |
lemma reals_Archimedean6: |
110 |
"0 \<le> r ==> \<exists>(n::nat). real (n - 1) \<le> r & r < real (n)" |
|
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
111 |
unfolding real_of_nat_def |
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
112 |
apply (rule exI [where x="nat (floor r + 1)"]) |
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
113 |
apply (insert floor_correct [of r]) |
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
114 |
apply (simp add: nat_add_distrib of_nat_nat) |
16819 | 115 |
done |
116 |
||
117 |
lemma reals_Archimedean6a: "0 \<le> r ==> \<exists>n. real (n) \<le> r & r < real (Suc n)" |
|
16893 | 118 |
by (drule reals_Archimedean6) auto |
16819 | 119 |
|
36979
da7c06ab3169
remove several redundant lemmas about floor and ceiling
huffman
parents:
36826
diff
changeset
|
120 |
text {* TODO: delete *} |
16819 | 121 |
lemma reals_Archimedean_6b_int: |
122 |
"0 \<le> r ==> \<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
|
123 |
unfolding real_of_int_def by (rule floor_exists) |
16819 | 124 |
|
36979
da7c06ab3169
remove several redundant lemmas about floor and ceiling
huffman
parents:
36826
diff
changeset
|
125 |
text {* TODO: delete *} |
16819 | 126 |
lemma reals_Archimedean_6c_int: |
127 |
"r < 0 ==> \<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
|
128 |
unfolding real_of_int_def by (rule floor_exists) |
16819 | 129 |
|
130 |
||
28091
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
131 |
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
|
132 |
|
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
133 |
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
|
134 |
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
|
135 |
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
|
136 |
|
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
137 |
lemma Rats_dense_in_nn_real: fixes x::real |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
138 |
assumes "0\<le>x" and "x<y" shows "\<exists>r \<in> \<rat>. x<r \<and> r<y" |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
139 |
proof - |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
140 |
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
|
141 |
with reals_Archimedean obtain q::nat |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
142 |
where q: "inverse (real q) < y-x" and "0 < real q" by auto |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
143 |
def p \<equiv> "LEAST n. y \<le> real (Suc n)/real q" |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
144 |
from reals_Archimedean2 obtain n::nat where "y * real q < real n" by auto |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
145 |
with `0 < real q` have ex: "y \<le> real n/real q" (is "?P n") |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
146 |
by (simp add: pos_less_divide_eq[THEN sym]) |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
147 |
also from assms have "\<not> y \<le> real (0::nat) / real q" by simp |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
148 |
ultimately have main: "(LEAST n. y \<le> real n/real q) = Suc p" |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
149 |
by (unfold p_def) (rule Least_Suc) |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
150 |
also from ex have "?P (LEAST x. ?P x)" by (rule LeastI) |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
151 |
ultimately have suc: "y \<le> real (Suc p) / real q" by simp |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
152 |
def r \<equiv> "real p/real q" |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
153 |
have "x = y-(y-x)" by simp |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
154 |
also from suc q have "\<dots> < real (Suc p)/real q - inverse (real q)" by arith |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
155 |
also have "\<dots> = real p / real q" |
37887 | 156 |
by (simp only: inverse_eq_divide diff_minus real_of_nat_Suc |
28091
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
157 |
minus_divide_left add_divide_distrib[THEN sym]) simp |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
158 |
finally have "x<r" by (unfold r_def) |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
159 |
have "p<Suc p" .. also note main[THEN sym] |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
160 |
finally have "\<not> ?P p" by (rule not_less_Least) |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
161 |
hence "r<y" by (simp add: r_def) |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
162 |
from r_def have "r \<in> \<rat>" by simp |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
163 |
with `x<r` `r<y` show ?thesis by fast |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
164 |
qed |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
165 |
|
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
166 |
theorem Rats_dense_in_real: fixes x y :: real |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
167 |
assumes "x<y" shows "\<exists>r \<in> \<rat>. x<r \<and> r<y" |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
168 |
proof - |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
169 |
from reals_Archimedean2 obtain n::nat where "-x < real n" by auto |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
170 |
hence "0 \<le> x + real n" by arith |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
171 |
also from `x<y` have "x + real n < y + real n" by arith |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
172 |
ultimately have "\<exists>r \<in> \<rat>. x + real n < r \<and> r < y + real n" |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
173 |
by(rule Rats_dense_in_nn_real) |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
174 |
then obtain r where "r \<in> \<rat>" and r2: "x + real n < r" |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
175 |
and r3: "r < y + real n" |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
176 |
by blast |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
177 |
have "r - real n = r + real (int n)/real (-1::int)" by simp |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
178 |
also from `r\<in>\<rat>` have "r + real (int n)/real (-1::int) \<in> \<rat>" by simp |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
179 |
also from r2 have "x < r - real n" by arith |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
180 |
moreover from r3 have "r - real n < y" by arith |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
181 |
ultimately show ?thesis by fast |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
182 |
qed |
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
183 |
|
50f2d6ba024c
Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents:
27435
diff
changeset
|
184 |
|
14641
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
185 |
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
|
186 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
187 |
lemma number_of_less_real_of_int_iff [simp]: |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
188 |
"((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
|
189 |
apply auto |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
190 |
apply (rule real_of_int_less_iff [THEN iffD1]) |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
191 |
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
|
192 |
done |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
193 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
194 |
lemma number_of_less_real_of_int_iff2 [simp]: |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
195 |
"(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
|
196 |
apply auto |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
197 |
apply (rule real_of_int_less_iff [THEN iffD1]) |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
198 |
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
|
199 |
done |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
200 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
201 |
lemma number_of_le_real_of_int_iff [simp]: |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
202 |
"((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
|
203 |
by (simp add: linorder_not_less [symmetric]) |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
204 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
205 |
lemma number_of_le_real_of_int_iff2 [simp]: |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
206 |
"(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
|
207 |
by (simp add: linorder_not_less [symmetric]) |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
208 |
|
24355 | 209 |
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
|
210 |
unfolding real_of_nat_def by simp |
14641
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
211 |
|
24355 | 212 |
lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n" |
30102 | 213 |
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
|
214 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
215 |
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
|
216 |
unfolding real_of_int_def by simp |
14641
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
217 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
218 |
lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n" |
30102 | 219 |
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
|
220 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
221 |
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
|
222 |
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
|
223 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
224 |
lemma lemma_floor: |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
225 |
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
|
226 |
shows "m \<le> (n::int)" |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
227 |
proof - |
23389 | 228 |
have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans) |
229 |
also have "... = real (n + 1)" by simp |
|
230 |
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
|
231 |
thus ?thesis by arith |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
232 |
qed |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
233 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
234 |
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
|
235 |
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
|
236 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
237 |
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
|
238 |
by (auto intro: lemma_floor) |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
239 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
240 |
lemma real_of_int_floor_cancel [simp]: |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
241 |
"(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
|
242 |
using floor_real_of_int by metis |
14641
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
243 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
244 |
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
|
245 |
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
|
246 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
247 |
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
|
248 |
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
|
249 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
250 |
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
|
251 |
apply (rule inj_int [THEN injD]) |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
252 |
apply (simp add: real_of_nat_Suc) |
15539 | 253 |
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
|
254 |
done |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
255 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
256 |
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
|
257 |
apply (drule order_le_imp_less_or_eq) |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
258 |
apply (auto intro: floor_eq3) |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
259 |
done |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
260 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
261 |
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
|
262 |
unfolding real_of_int_def using floor_correct [of r] by simp |
16819 | 263 |
|
264 |
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
|
265 |
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
|
266 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
267 |
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
|
268 |
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
|
269 |
|
16819 | 270 |
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
|
271 |
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
|
272 |
|
16819 | 273 |
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
|
274 |
unfolding real_of_int_def by (simp add: le_floor_iff) |
16819 | 275 |
|
276 |
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
|
277 |
unfolding real_of_int_def by (simp add: le_floor_iff) |
16819 | 278 |
|
279 |
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
|
280 |
unfolding real_of_int_def by (rule le_floor_iff) |
16819 | 281 |
|
282 |
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
|
283 |
unfolding real_of_int_def by (rule floor_less_iff) |
16819 | 284 |
|
285 |
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
|
286 |
unfolding real_of_int_def by (rule less_floor_iff) |
16819 | 287 |
|
288 |
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
|
289 |
unfolding real_of_int_def by (rule floor_le_iff) |
16819 | 290 |
|
291 |
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
|
292 |
unfolding real_of_int_def by (rule floor_add_of_int) |
16819 | 293 |
|
294 |
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
|
295 |
unfolding real_of_int_def by (rule floor_diff_of_int) |
16819 | 296 |
|
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
297 |
lemma le_mult_floor: |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
298 |
assumes "0 \<le> (a :: real)" and "0 \<le> b" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
299 |
shows "floor a * floor b \<le> floor (a * b)" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
300 |
proof - |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
301 |
have "real (floor a) \<le> a" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
302 |
and "real (floor b) \<le> b" by auto |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
303 |
hence "real (floor a * floor b) \<le> a * b" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
304 |
using assms by (auto intro!: mult_mono) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
305 |
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
|
306 |
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
|
307 |
qed |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
308 |
|
24355 | 309 |
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
|
310 |
unfolding real_of_nat_def by simp |
14641
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
311 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
312 |
lemma ceiling_floor [simp]: "ceiling (real (floor r)) = floor r" |
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
313 |
unfolding real_of_int_def by simp |
14641
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
314 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
315 |
lemma floor_ceiling [simp]: "floor (real (ceiling r)) = ceiling r" |
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
316 |
unfolding real_of_int_def by simp |
14641
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
317 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
318 |
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
|
319 |
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
|
320 |
|
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
321 |
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
|
322 |
unfolding real_of_int_def by simp |
14641
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
323 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
324 |
lemma real_of_int_ceiling_cancel [simp]: |
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
325 |
"(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
|
326 |
using ceiling_real_of_int by metis |
14641
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
327 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
328 |
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
|
329 |
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
|
330 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
331 |
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
|
332 |
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
|
333 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
334 |
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
|
335 |
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
|
336 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
337 |
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
|
338 |
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
|
339 |
|
79b7bd936264
moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents:
14476
diff
changeset
|
340 |
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
|
341 |
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
|
342 |
|
16819 | 343 |
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
|
344 |
unfolding real_of_int_def by (simp add: ceiling_le_iff) |
16819 | 345 |
|
346 |
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
|
347 |
unfolding real_of_int_def by (simp add: ceiling_le_iff) |
16819 | 348 |
|
349 |
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
|
350 |
unfolding real_of_int_def by (rule ceiling_le_iff) |
16819 | 351 |
|
352 |
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
|
353 |
unfolding real_of_int_def by (rule less_ceiling_iff) |
16819 | 354 |
|
355 |
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
|
356 |
unfolding real_of_int_def by (rule ceiling_less_iff) |
16819 | 357 |
|
358 |
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
|
359 |
unfolding real_of_int_def by (rule le_ceiling_iff) |
16819 | 360 |
|
361 |
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
|
362 |
unfolding real_of_int_def by (rule ceiling_add_of_int) |
16819 | 363 |
|
364 |
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
|
365 |
unfolding real_of_int_def by (rule ceiling_diff_of_int) |
16819 | 366 |
|
367 |
||
368 |
subsection {* Versions for the natural numbers *} |
|
369 |
||
19765 | 370 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
371 |
natfloor :: "real => nat" where |
19765 | 372 |
"natfloor x = nat(floor x)" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
373 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
374 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
375 |
natceiling :: "real => nat" where |
19765 | 376 |
"natceiling x = nat(ceiling x)" |
16819 | 377 |
|
378 |
lemma natfloor_zero [simp]: "natfloor 0 = 0" |
|
379 |
by (unfold natfloor_def, simp) |
|
380 |
||
381 |
lemma natfloor_one [simp]: "natfloor 1 = 1" |
|
382 |
by (unfold natfloor_def, simp) |
|
383 |
||
384 |
lemma zero_le_natfloor [simp]: "0 <= natfloor x" |
|
385 |
by (unfold natfloor_def, simp) |
|
386 |
||
387 |
lemma natfloor_number_of_eq [simp]: "natfloor (number_of n) = number_of n" |
|
388 |
by (unfold natfloor_def, simp) |
|
389 |
||
390 |
lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n" |
|
391 |
by (unfold natfloor_def, simp) |
|
392 |
||
393 |
lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x" |
|
394 |
by (unfold natfloor_def, simp) |
|
395 |
||
396 |
lemma natfloor_neg: "x <= 0 ==> natfloor x = 0" |
|
397 |
apply (unfold natfloor_def) |
|
398 |
apply (subgoal_tac "floor x <= floor 0") |
|
399 |
apply simp |
|
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
400 |
apply (erule floor_mono) |
16819 | 401 |
done |
402 |
||
403 |
lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y" |
|
404 |
apply (case_tac "0 <= x") |
|
405 |
apply (subst natfloor_def)+ |
|
406 |
apply (subst nat_le_eq_zle) |
|
407 |
apply force |
|
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
408 |
apply (erule floor_mono) |
16819 | 409 |
apply (subst natfloor_neg) |
410 |
apply simp |
|
411 |
apply simp |
|
412 |
done |
|
413 |
||
414 |
lemma le_natfloor: "real x <= a ==> x <= natfloor a" |
|
415 |
apply (unfold natfloor_def) |
|
416 |
apply (subst nat_int [THEN sym]) |
|
417 |
apply (subst nat_le_eq_zle) |
|
418 |
apply simp |
|
419 |
apply (rule le_floor) |
|
420 |
apply simp |
|
421 |
done |
|
422 |
||
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
423 |
lemma less_natfloor: |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
424 |
assumes "0 \<le> x" and "x < real (n :: nat)" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
425 |
shows "natfloor x < n" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
426 |
proof (rule ccontr) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
427 |
assume "\<not> ?thesis" hence *: "n \<le> natfloor x" by simp |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
428 |
note assms(2) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
429 |
also have "real n \<le> real (natfloor x)" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
430 |
using * unfolding real_of_nat_le_iff . |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
431 |
finally have "x < real (natfloor x)" . |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
432 |
with real_natfloor_le[OF assms(1)] |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
433 |
show False by auto |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
434 |
qed |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
435 |
|
16819 | 436 |
lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)" |
437 |
apply (rule iffI) |
|
438 |
apply (rule order_trans) |
|
439 |
prefer 2 |
|
440 |
apply (erule real_natfloor_le) |
|
441 |
apply (subst real_of_nat_le_iff) |
|
442 |
apply assumption |
|
443 |
apply (erule le_natfloor) |
|
444 |
done |
|
445 |
||
16893 | 446 |
lemma le_natfloor_eq_number_of [simp]: |
16819 | 447 |
"~ neg((number_of n)::int) ==> 0 <= x ==> |
448 |
(number_of n <= natfloor x) = (number_of n <= x)" |
|
449 |
apply (subst le_natfloor_eq, assumption) |
|
450 |
apply simp |
|
451 |
done |
|
452 |
||
16820 | 453 |
lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)" |
16819 | 454 |
apply (case_tac "0 <= x") |
455 |
apply (subst le_natfloor_eq, assumption, simp) |
|
456 |
apply (rule iffI) |
|
16893 | 457 |
apply (subgoal_tac "natfloor x <= natfloor 0") |
16819 | 458 |
apply simp |
459 |
apply (rule natfloor_mono) |
|
460 |
apply simp |
|
461 |
apply simp |
|
462 |
done |
|
463 |
||
464 |
lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n" |
|
465 |
apply (unfold natfloor_def) |
|
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
466 |
apply (subst (2) nat_int [THEN sym]) |
16819 | 467 |
apply (subst eq_nat_nat_iff) |
468 |
apply simp |
|
469 |
apply simp |
|
470 |
apply (rule floor_eq2) |
|
471 |
apply auto |
|
472 |
done |
|
473 |
||
474 |
lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1" |
|
475 |
apply (case_tac "0 <= x") |
|
476 |
apply (unfold natfloor_def) |
|
477 |
apply simp |
|
478 |
apply simp_all |
|
479 |
done |
|
480 |
||
481 |
lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)" |
|
29667 | 482 |
using real_natfloor_add_one_gt by (simp add: algebra_simps) |
16819 | 483 |
|
484 |
lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n" |
|
485 |
apply (subgoal_tac "z < real(natfloor z) + 1") |
|
486 |
apply arith |
|
487 |
apply (rule real_natfloor_add_one_gt) |
|
488 |
done |
|
489 |
||
490 |
lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a" |
|
491 |
apply (unfold natfloor_def) |
|
24355 | 492 |
apply (subgoal_tac "real a = real (int a)") |
16819 | 493 |
apply (erule ssubst) |
23309 | 494 |
apply (simp add: nat_add_distrib del: real_of_int_of_nat_eq) |
16819 | 495 |
apply simp |
496 |
done |
|
497 |
||
16893 | 498 |
lemma natfloor_add_number_of [simp]: |
499 |
"~neg ((number_of n)::int) ==> 0 <= x ==> |
|
16819 | 500 |
natfloor (x + number_of n) = natfloor x + number_of n" |
501 |
apply (subst natfloor_add [THEN sym]) |
|
502 |
apply simp_all |
|
503 |
done |
|
504 |
||
505 |
lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1" |
|
506 |
apply (subst natfloor_add [THEN sym]) |
|
507 |
apply assumption |
|
508 |
apply simp |
|
509 |
done |
|
510 |
||
16893 | 511 |
lemma natfloor_subtract [simp]: "real a <= x ==> |
16819 | 512 |
natfloor(x - real a) = natfloor x - a" |
513 |
apply (unfold natfloor_def) |
|
24355 | 514 |
apply (subgoal_tac "real a = real (int a)") |
16819 | 515 |
apply (erule ssubst) |
23309 | 516 |
apply (simp del: real_of_int_of_nat_eq) |
16819 | 517 |
apply simp |
518 |
done |
|
519 |
||
41550 | 520 |
lemma natfloor_div_nat: |
521 |
assumes "1 <= x" and "y > 0" |
|
522 |
shows "natfloor (x / real y) = natfloor x div y" |
|
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
523 |
proof - |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
524 |
have "natfloor x = (natfloor x) div y * y + (natfloor x) mod y" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
525 |
by simp |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
526 |
then have a: "real(natfloor x) = real ((natfloor x) div y) * real y + |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
527 |
real((natfloor x) mod y)" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
528 |
by (simp only: real_of_nat_add [THEN sym] real_of_nat_mult [THEN sym]) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
529 |
have "x = real(natfloor x) + (x - real(natfloor x))" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
530 |
by simp |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
531 |
then have "x = real ((natfloor x) div y) * real y + |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
532 |
real((natfloor x) mod y) + (x - real(natfloor x))" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
533 |
by (simp add: a) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
534 |
then have "x / real y = ... / real y" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
535 |
by simp |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
536 |
also have "... = real((natfloor x) div y) + real((natfloor x) mod y) / |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
537 |
real y + (x - real(natfloor x)) / real y" |
41550 | 538 |
by (auto simp add: algebra_simps add_divide_distrib diff_divide_distrib) |
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
539 |
finally have "natfloor (x / real y) = natfloor(...)" by simp |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
540 |
also have "... = natfloor(real((natfloor x) mod y) / |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
541 |
real y + (x - real(natfloor x)) / real y + real((natfloor x) div y))" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
542 |
by (simp add: add_ac) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
543 |
also have "... = natfloor(real((natfloor x) mod y) / |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
544 |
real y + (x - real(natfloor x)) / real y) + (natfloor x) div y" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
545 |
apply (rule natfloor_add) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
546 |
apply (rule add_nonneg_nonneg) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
547 |
apply (rule divide_nonneg_pos) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
548 |
apply simp |
41550 | 549 |
apply (simp add: assms) |
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
550 |
apply (rule divide_nonneg_pos) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
551 |
apply (simp add: algebra_simps) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
552 |
apply (rule real_natfloor_le) |
41550 | 553 |
using assms apply auto |
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
554 |
done |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
555 |
also have "natfloor(real((natfloor x) mod y) / |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
556 |
real y + (x - real(natfloor x)) / real y) = 0" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
557 |
apply (rule natfloor_eq) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
558 |
apply simp |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
559 |
apply (rule add_nonneg_nonneg) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
560 |
apply (rule divide_nonneg_pos) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
561 |
apply force |
41550 | 562 |
apply (force simp add: assms) |
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
563 |
apply (rule divide_nonneg_pos) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
564 |
apply (simp add: algebra_simps) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
565 |
apply (rule real_natfloor_le) |
41550 | 566 |
apply (auto simp add: assms) |
567 |
using assms apply arith |
|
568 |
using assms apply (simp add: add_divide_distrib [THEN sym]) |
|
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
569 |
apply (subgoal_tac "real y = real y - 1 + 1") |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
570 |
apply (erule ssubst) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
571 |
apply (rule add_le_less_mono) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
572 |
apply (simp add: algebra_simps) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
573 |
apply (subgoal_tac "1 + real(natfloor x mod y) = |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
574 |
real(natfloor x mod y + 1)") |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
575 |
apply (erule ssubst) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
576 |
apply (subst real_of_nat_le_iff) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
577 |
apply (subgoal_tac "natfloor x mod y < y") |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
578 |
apply arith |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
579 |
apply (rule mod_less_divisor) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
580 |
apply auto |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
581 |
using real_natfloor_add_one_gt |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
582 |
apply (simp add: algebra_simps) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
583 |
done |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
584 |
finally show ?thesis by simp |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
585 |
qed |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
586 |
|
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
587 |
lemma le_mult_natfloor: |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
588 |
assumes "0 \<le> (a :: real)" and "0 \<le> b" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
589 |
shows "natfloor a * natfloor b \<le> natfloor (a * b)" |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
590 |
unfolding natfloor_def |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
591 |
apply (subst nat_mult_distrib[symmetric]) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
592 |
using assms apply simp |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
593 |
apply (subst nat_le_eq_zle) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
594 |
using assms le_mult_floor by (auto intro!: mult_nonneg_nonneg) |
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
595 |
|
16819 | 596 |
lemma natceiling_zero [simp]: "natceiling 0 = 0" |
597 |
by (unfold natceiling_def, simp) |
|
598 |
||
599 |
lemma natceiling_one [simp]: "natceiling 1 = 1" |
|
600 |
by (unfold natceiling_def, simp) |
|
601 |
||
602 |
lemma zero_le_natceiling [simp]: "0 <= natceiling x" |
|
603 |
by (unfold natceiling_def, simp) |
|
604 |
||
605 |
lemma natceiling_number_of_eq [simp]: "natceiling (number_of n) = number_of n" |
|
606 |
by (unfold natceiling_def, simp) |
|
607 |
||
608 |
lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n" |
|
609 |
by (unfold natceiling_def, simp) |
|
610 |
||
611 |
lemma real_natceiling_ge: "x <= real(natceiling x)" |
|
612 |
apply (unfold natceiling_def) |
|
613 |
apply (case_tac "x < 0") |
|
614 |
apply simp |
|
615 |
apply (subst real_nat_eq_real) |
|
616 |
apply (subgoal_tac "ceiling 0 <= ceiling x") |
|
617 |
apply simp |
|
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
618 |
apply (rule ceiling_mono) |
16819 | 619 |
apply simp |
620 |
apply simp |
|
621 |
done |
|
622 |
||
623 |
lemma natceiling_neg: "x <= 0 ==> natceiling x = 0" |
|
624 |
apply (unfold natceiling_def) |
|
625 |
apply simp |
|
626 |
done |
|
627 |
||
628 |
lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y" |
|
629 |
apply (case_tac "0 <= x") |
|
630 |
apply (subst natceiling_def)+ |
|
631 |
apply (subst nat_le_eq_zle) |
|
632 |
apply (rule disjI2) |
|
633 |
apply (subgoal_tac "real (0::int) <= real(ceiling y)") |
|
634 |
apply simp |
|
635 |
apply (rule order_trans) |
|
636 |
apply simp |
|
637 |
apply (erule order_trans) |
|
638 |
apply simp |
|
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
639 |
apply (erule ceiling_mono) |
16819 | 640 |
apply (subst natceiling_neg) |
641 |
apply simp_all |
|
642 |
done |
|
643 |
||
644 |
lemma natceiling_le: "x <= real a ==> natceiling x <= a" |
|
645 |
apply (unfold natceiling_def) |
|
646 |
apply (case_tac "x < 0") |
|
647 |
apply simp |
|
35578
384ad08a1d1b
Added natfloor and floor rules for multiplication and power.
hoelzl
parents:
35028
diff
changeset
|
648 |
apply (subst (2) nat_int [THEN sym]) |
16819 | 649 |
apply (subst nat_le_eq_zle) |
650 |
apply simp |
|
651 |
apply (rule ceiling_le) |
|
652 |
apply simp |
|
653 |
done |
|
654 |
||
655 |
lemma natceiling_le_eq: "0 <= x ==> (natceiling x <= a) = (x <= real a)" |
|
656 |
apply (rule iffI) |
|
657 |
apply (rule order_trans) |
|
658 |
apply (rule real_natceiling_ge) |
|
659 |
apply (subst real_of_nat_le_iff) |
|
660 |
apply assumption |
|
661 |
apply (erule natceiling_le) |
|
662 |
done |
|
663 |
||
16893 | 664 |
lemma natceiling_le_eq_number_of [simp]: |
16820 | 665 |
"~ neg((number_of n)::int) ==> 0 <= x ==> |
666 |
(natceiling x <= number_of n) = (x <= number_of n)" |
|
16819 | 667 |
apply (subst natceiling_le_eq, assumption) |
668 |
apply simp |
|
669 |
done |
|
670 |
||
16820 | 671 |
lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)" |
16819 | 672 |
apply (case_tac "0 <= x") |
673 |
apply (subst natceiling_le_eq) |
|
674 |
apply assumption |
|
675 |
apply simp |
|
676 |
apply (subst natceiling_neg) |
|
677 |
apply simp |
|
678 |
apply simp |
|
679 |
done |
|
680 |
||
681 |
lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1" |
|
682 |
apply (unfold natceiling_def) |
|
19850 | 683 |
apply (simplesubst nat_int [THEN sym]) back back |
16819 | 684 |
apply (subgoal_tac "nat(int n) + 1 = nat(int n + 1)") |
685 |
apply (erule ssubst) |
|
686 |
apply (subst eq_nat_nat_iff) |
|
687 |
apply (subgoal_tac "ceiling 0 <= ceiling x") |
|
688 |
apply simp |
|
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
689 |
apply (rule ceiling_mono) |
16819 | 690 |
apply force |
691 |
apply force |
|
692 |
apply (rule ceiling_eq2) |
|
693 |
apply (simp, simp) |
|
694 |
apply (subst nat_add_distrib) |
|
695 |
apply auto |
|
696 |
done |
|
697 |
||
16893 | 698 |
lemma natceiling_add [simp]: "0 <= x ==> |
16819 | 699 |
natceiling (x + real a) = natceiling x + a" |
700 |
apply (unfold natceiling_def) |
|
24355 | 701 |
apply (subgoal_tac "real a = real (int a)") |
16819 | 702 |
apply (erule ssubst) |
23309 | 703 |
apply (simp del: real_of_int_of_nat_eq) |
16819 | 704 |
apply (subst nat_add_distrib) |
705 |
apply (subgoal_tac "0 = ceiling 0") |
|
706 |
apply (erule ssubst) |
|
30097
57df8626c23b
generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents:
29667
diff
changeset
|
707 |
apply (erule ceiling_mono) |
16819 | 708 |
apply simp_all |
709 |
done |
|
710 |
||
16893 | 711 |
lemma natceiling_add_number_of [simp]: |
712 |
"~ neg ((number_of n)::int) ==> 0 <= x ==> |
|
16820 | 713 |
natceiling (x + number_of n) = natceiling x + number_of n" |
16819 | 714 |
apply (subst natceiling_add [THEN sym]) |
715 |
apply simp_all |
|
716 |
done |
|
717 |
||
718 |
lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1" |
|
719 |
apply (subst natceiling_add [THEN sym]) |
|
720 |
apply assumption |
|
721 |
apply simp |
|
722 |
done |
|
723 |
||
16893 | 724 |
lemma natceiling_subtract [simp]: "real a <= x ==> |
16819 | 725 |
natceiling(x - real a) = natceiling x - a" |
726 |
apply (unfold natceiling_def) |
|
24355 | 727 |
apply (subgoal_tac "real a = real (int a)") |
16819 | 728 |
apply (erule ssubst) |
23309 | 729 |
apply (simp del: real_of_int_of_nat_eq) |
16819 | 730 |
apply simp |
731 |
done |
|
732 |
||
36826
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
733 |
subsection {* Exponentiation with floor *} |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
734 |
|
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
735 |
lemma floor_power: |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
736 |
assumes "x = real (floor x)" |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
737 |
shows "floor (x ^ n) = floor x ^ n" |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
738 |
proof - |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
739 |
have *: "x ^ n = real (floor x ^ n)" |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
740 |
using assms by (induct n arbitrary: x) simp_all |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
741 |
show ?thesis unfolding real_of_int_inject[symmetric] |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
742 |
unfolding * floor_real_of_int .. |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
743 |
qed |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
744 |
|
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
745 |
lemma natfloor_power: |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
746 |
assumes "x = real (natfloor x)" |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
747 |
shows "natfloor (x ^ n) = natfloor x ^ n" |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
748 |
proof - |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
749 |
from assms have "0 \<le> floor x" by auto |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
750 |
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
|
751 |
from floor_power[OF this] |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
752 |
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
|
753 |
by simp |
4d4462d644ae
move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents:
36795
diff
changeset
|
754 |
qed |
16819 | 755 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
9429
diff
changeset
|
756 |
end |