| author | wenzelm | 
| Wed, 16 Jan 2019 18:54:18 +0100 | |
| changeset 69673 | cc47e7e06f38 | 
| parent 68721 | 53ad5c01be3f | 
| child 70365 | 4df0628e8545 | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: HOL/Archimedean_Field.thy  | 
2  | 
Author: Brian Huffman  | 
|
| 30096 | 3  | 
*)  | 
4  | 
||
| 60758 | 5  | 
section \<open>Archimedean Fields, Floor and Ceiling Functions\<close>  | 
| 30096 | 6  | 
|
7  | 
theory Archimedean_Field  | 
|
8  | 
imports Main  | 
|
9  | 
begin  | 
|
10  | 
||
| 63331 | 11  | 
lemma cInf_abs_ge:  | 
| 63489 | 12  | 
  fixes S :: "'a::{linordered_idom,conditionally_complete_linorder} set"
 | 
13  | 
  assumes "S \<noteq> {}"
 | 
|
14  | 
and bdd: "\<And>x. x\<in>S \<Longrightarrow> \<bar>x\<bar> \<le> a"  | 
|
| 63331 | 15  | 
shows "\<bar>Inf S\<bar> \<le> a"  | 
16  | 
proof -  | 
|
17  | 
have "Sup (uminus ` S) = - (Inf S)"  | 
|
18  | 
proof (rule antisym)  | 
|
| 63489 | 19  | 
show "- (Inf S) \<le> Sup (uminus ` S)"  | 
| 63331 | 20  | 
apply (subst minus_le_iff)  | 
21  | 
      apply (rule cInf_greatest [OF \<open>S \<noteq> {}\<close>])
 | 
|
22  | 
apply (subst minus_le_iff)  | 
|
| 63489 | 23  | 
apply (rule cSup_upper)  | 
24  | 
apply force  | 
|
25  | 
using bdd  | 
|
26  | 
apply (force simp: abs_le_iff bdd_above_def)  | 
|
| 63331 | 27  | 
done  | 
28  | 
next  | 
|
29  | 
show "Sup (uminus ` S) \<le> - Inf S"  | 
|
30  | 
apply (rule cSup_least)  | 
|
| 63489 | 31  | 
      using \<open>S \<noteq> {}\<close>
 | 
32  | 
apply force  | 
|
| 63331 | 33  | 
apply clarsimp  | 
34  | 
apply (rule cInf_lower)  | 
|
| 63489 | 35  | 
apply assumption  | 
36  | 
using bdd  | 
|
37  | 
apply (simp add: bdd_below_def)  | 
|
38  | 
apply (rule_tac x = "- a" in exI)  | 
|
| 63331 | 39  | 
apply force  | 
40  | 
done  | 
|
41  | 
qed  | 
|
| 63489 | 42  | 
with cSup_abs_le [of "uminus ` S"] assms show ?thesis  | 
43  | 
by fastforce  | 
|
| 63331 | 44  | 
qed  | 
45  | 
||
46  | 
lemma cSup_asclose:  | 
|
| 63489 | 47  | 
  fixes S :: "'a::{linordered_idom,conditionally_complete_linorder} set"
 | 
| 63331 | 48  | 
  assumes S: "S \<noteq> {}"
 | 
49  | 
and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e"  | 
|
50  | 
shows "\<bar>Sup S - l\<bar> \<le> e"  | 
|
51  | 
proof -  | 
|
| 63489 | 52  | 
have *: "\<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e" for x l e :: 'a  | 
| 63331 | 53  | 
by arith  | 
54  | 
have "bdd_above S"  | 
|
55  | 
using b by (auto intro!: bdd_aboveI[of _ "l + e"])  | 
|
56  | 
with S b show ?thesis  | 
|
| 63489 | 57  | 
unfolding * by (auto intro!: cSup_upper2 cSup_least)  | 
| 63331 | 58  | 
qed  | 
59  | 
||
60  | 
lemma cInf_asclose:  | 
|
| 63489 | 61  | 
  fixes S :: "'a::{linordered_idom,conditionally_complete_linorder} set"
 | 
| 63331 | 62  | 
  assumes S: "S \<noteq> {}"
 | 
63  | 
and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e"  | 
|
64  | 
shows "\<bar>Inf S - l\<bar> \<le> e"  | 
|
65  | 
proof -  | 
|
| 63489 | 66  | 
have *: "\<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e" for x l e :: 'a  | 
| 63331 | 67  | 
by arith  | 
68  | 
have "bdd_below S"  | 
|
69  | 
using b by (auto intro!: bdd_belowI[of _ "l - e"])  | 
|
70  | 
with S b show ?thesis  | 
|
| 63489 | 71  | 
unfolding * by (auto intro!: cInf_lower2 cInf_greatest)  | 
| 63331 | 72  | 
qed  | 
73  | 
||
| 63489 | 74  | 
|
| 60758 | 75  | 
subsection \<open>Class of Archimedean fields\<close>  | 
| 30096 | 76  | 
|
| 60758 | 77  | 
text \<open>Archimedean fields have no infinite elements.\<close>  | 
| 30096 | 78  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
79  | 
class archimedean_field = linordered_field +  | 
| 30096 | 80  | 
assumes ex_le_of_int: "\<exists>z. x \<le> of_int z"  | 
81  | 
||
| 63489 | 82  | 
lemma ex_less_of_int: "\<exists>z. x < of_int z"  | 
83  | 
for x :: "'a::archimedean_field"  | 
|
| 30096 | 84  | 
proof -  | 
85  | 
from ex_le_of_int obtain z where "x \<le> of_int z" ..  | 
|
86  | 
then have "x < of_int (z + 1)" by simp  | 
|
87  | 
then show ?thesis ..  | 
|
88  | 
qed  | 
|
89  | 
||
| 63489 | 90  | 
lemma ex_of_int_less: "\<exists>z. of_int z < x"  | 
91  | 
for x :: "'a::archimedean_field"  | 
|
| 30096 | 92  | 
proof -  | 
93  | 
from ex_less_of_int obtain z where "- x < of_int z" ..  | 
|
94  | 
then have "of_int (- z) < x" by simp  | 
|
95  | 
then show ?thesis ..  | 
|
96  | 
qed  | 
|
97  | 
||
| 63489 | 98  | 
lemma reals_Archimedean2: "\<exists>n. x < of_nat n"  | 
99  | 
for x :: "'a::archimedean_field"  | 
|
| 30096 | 100  | 
proof -  | 
| 63489 | 101  | 
obtain z where "x < of_int z"  | 
102  | 
using ex_less_of_int ..  | 
|
103  | 
also have "\<dots> \<le> of_int (int (nat z))"  | 
|
104  | 
by simp  | 
|
105  | 
also have "\<dots> = of_nat (nat z)"  | 
|
106  | 
by (simp only: of_int_of_nat_eq)  | 
|
| 30096 | 107  | 
finally show ?thesis ..  | 
108  | 
qed  | 
|
109  | 
||
| 63489 | 110  | 
lemma real_arch_simple: "\<exists>n. x \<le> of_nat n"  | 
111  | 
for x :: "'a::archimedean_field"  | 
|
| 30096 | 112  | 
proof -  | 
| 63489 | 113  | 
obtain n where "x < of_nat n"  | 
114  | 
using reals_Archimedean2 ..  | 
|
115  | 
then have "x \<le> of_nat n"  | 
|
116  | 
by simp  | 
|
| 30096 | 117  | 
then show ?thesis ..  | 
118  | 
qed  | 
|
119  | 
||
| 60758 | 120  | 
text \<open>Archimedean fields have no infinitesimal elements.\<close>  | 
| 30096 | 121  | 
|
| 
62623
 
dbc62f86a1a9
rationalisation of theorem names esp about "real Archimedian" etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
62348 
diff
changeset
 | 
122  | 
lemma reals_Archimedean:  | 
| 30096 | 123  | 
fixes x :: "'a::archimedean_field"  | 
| 63489 | 124  | 
assumes "0 < x"  | 
125  | 
shows "\<exists>n. inverse (of_nat (Suc n)) < x"  | 
|
| 30096 | 126  | 
proof -  | 
| 60758 | 127  | 
from \<open>0 < x\<close> have "0 < inverse x"  | 
| 30096 | 128  | 
by (rule positive_imp_inverse_positive)  | 
129  | 
obtain n where "inverse x < of_nat n"  | 
|
| 
62623
 
dbc62f86a1a9
rationalisation of theorem names esp about "real Archimedian" etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
62348 
diff
changeset
 | 
130  | 
using reals_Archimedean2 ..  | 
| 30096 | 131  | 
then obtain m where "inverse x < of_nat (Suc m)"  | 
| 60758 | 132  | 
using \<open>0 < inverse x\<close> by (cases n) (simp_all del: of_nat_Suc)  | 
| 30096 | 133  | 
then have "inverse (of_nat (Suc m)) < inverse (inverse x)"  | 
| 60758 | 134  | 
using \<open>0 < inverse x\<close> by (rule less_imp_inverse_less)  | 
| 30096 | 135  | 
then have "inverse (of_nat (Suc m)) < x"  | 
| 60758 | 136  | 
using \<open>0 < x\<close> by (simp add: nonzero_inverse_inverse_eq)  | 
| 30096 | 137  | 
then show ?thesis ..  | 
138  | 
qed  | 
|
139  | 
||
140  | 
lemma ex_inverse_of_nat_less:  | 
|
141  | 
fixes x :: "'a::archimedean_field"  | 
|
| 63489 | 142  | 
assumes "0 < x"  | 
143  | 
shows "\<exists>n>0. inverse (of_nat n) < x"  | 
|
| 
62623
 
dbc62f86a1a9
rationalisation of theorem names esp about "real Archimedian" etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
62348 
diff
changeset
 | 
144  | 
using reals_Archimedean [OF \<open>0 < x\<close>] by auto  | 
| 30096 | 145  | 
|
146  | 
lemma ex_less_of_nat_mult:  | 
|
147  | 
fixes x :: "'a::archimedean_field"  | 
|
| 63489 | 148  | 
assumes "0 < x"  | 
149  | 
shows "\<exists>n. y < of_nat n * x"  | 
|
| 30096 | 150  | 
proof -  | 
| 63489 | 151  | 
obtain n where "y / x < of_nat n"  | 
152  | 
using reals_Archimedean2 ..  | 
|
153  | 
with \<open>0 < x\<close> have "y < of_nat n * x"  | 
|
154  | 
by (simp add: pos_divide_less_eq)  | 
|
| 30096 | 155  | 
then show ?thesis ..  | 
156  | 
qed  | 
|
157  | 
||
158  | 
||
| 60758 | 159  | 
subsection \<open>Existence and uniqueness of floor function\<close>  | 
| 30096 | 160  | 
|
161  | 
lemma exists_least_lemma:  | 
|
162  | 
assumes "\<not> P 0" and "\<exists>n. P n"  | 
|
163  | 
shows "\<exists>n. \<not> P n \<and> P (Suc n)"  | 
|
164  | 
proof -  | 
|
| 63489 | 165  | 
from \<open>\<exists>n. P n\<close> have "P (Least P)"  | 
166  | 
by (rule LeastI_ex)  | 
|
| 60758 | 167  | 
with \<open>\<not> P 0\<close> obtain n where "Least P = Suc n"  | 
| 30096 | 168  | 
by (cases "Least P") auto  | 
| 63489 | 169  | 
then have "n < Least P"  | 
170  | 
by simp  | 
|
171  | 
then have "\<not> P n"  | 
|
172  | 
by (rule not_less_Least)  | 
|
| 30096 | 173  | 
then have "\<not> P n \<and> P (Suc n)"  | 
| 60758 | 174  | 
using \<open>P (Least P)\<close> \<open>Least P = Suc n\<close> by simp  | 
| 30096 | 175  | 
then show ?thesis ..  | 
176  | 
qed  | 
|
177  | 
||
178  | 
lemma floor_exists:  | 
|
179  | 
fixes x :: "'a::archimedean_field"  | 
|
180  | 
shows "\<exists>z. of_int z \<le> x \<and> x < of_int (z + 1)"  | 
|
| 63489 | 181  | 
proof (cases "0 \<le> x")  | 
182  | 
case True  | 
|
183  | 
then have "\<not> x < of_nat 0"  | 
|
184  | 
by simp  | 
|
| 30096 | 185  | 
then have "\<exists>n. \<not> x < of_nat n \<and> x < of_nat (Suc n)"  | 
| 
62623
 
dbc62f86a1a9
rationalisation of theorem names esp about "real Archimedian" etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
62348 
diff
changeset
 | 
186  | 
using reals_Archimedean2 by (rule exists_least_lemma)  | 
| 30096 | 187  | 
then obtain n where "\<not> x < of_nat n \<and> x < of_nat (Suc n)" ..  | 
| 63489 | 188  | 
then have "of_int (int n) \<le> x \<and> x < of_int (int n + 1)"  | 
189  | 
by simp  | 
|
| 30096 | 190  | 
then show ?thesis ..  | 
191  | 
next  | 
|
| 63489 | 192  | 
case False  | 
193  | 
then have "\<not> - x \<le> of_nat 0"  | 
|
194  | 
by simp  | 
|
| 30096 | 195  | 
then have "\<exists>n. \<not> - x \<le> of_nat n \<and> - x \<le> of_nat (Suc n)"  | 
| 
62623
 
dbc62f86a1a9
rationalisation of theorem names esp about "real Archimedian" etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
62348 
diff
changeset
 | 
196  | 
using real_arch_simple by (rule exists_least_lemma)  | 
| 30096 | 197  | 
then obtain n where "\<not> - x \<le> of_nat n \<and> - x \<le> of_nat (Suc n)" ..  | 
| 63489 | 198  | 
then have "of_int (- int n - 1) \<le> x \<and> x < of_int (- int n - 1 + 1)"  | 
199  | 
by simp  | 
|
| 30096 | 200  | 
then show ?thesis ..  | 
201  | 
qed  | 
|
202  | 
||
| 63489 | 203  | 
lemma floor_exists1: "\<exists>!z. of_int z \<le> x \<and> x < of_int (z + 1)"  | 
204  | 
for x :: "'a::archimedean_field"  | 
|
| 30096 | 205  | 
proof (rule ex_ex1I)  | 
206  | 
show "\<exists>z. of_int z \<le> x \<and> x < of_int (z + 1)"  | 
|
207  | 
by (rule floor_exists)  | 
|
208  | 
next  | 
|
| 63489 | 209  | 
fix y z  | 
210  | 
assume "of_int y \<le> x \<and> x < of_int (y + 1)"  | 
|
211  | 
and "of_int z \<le> x \<and> x < of_int (z + 1)"  | 
|
| 54281 | 212  | 
with le_less_trans [of "of_int y" "x" "of_int (z + 1)"]  | 
| 63489 | 213  | 
le_less_trans [of "of_int z" "x" "of_int (y + 1)"] show "y = z"  | 
214  | 
by (simp del: of_int_add)  | 
|
| 30096 | 215  | 
qed  | 
216  | 
||
217  | 
||
| 60758 | 218  | 
subsection \<open>Floor function\<close>  | 
| 30096 | 219  | 
|
| 
43732
 
6b2bdc57155b
adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
 
bulwahn 
parents: 
43704 
diff
changeset
 | 
220  | 
class floor_ceiling = archimedean_field +  | 
| 61942 | 221  | 
  fixes floor :: "'a \<Rightarrow> int"  ("\<lfloor>_\<rfloor>")
 | 
222  | 
assumes floor_correct: "of_int \<lfloor>x\<rfloor> \<le> x \<and> x < of_int (\<lfloor>x\<rfloor> + 1)"  | 
|
| 30096 | 223  | 
|
| 63489 | 224  | 
lemma floor_unique: "of_int z \<le> x \<Longrightarrow> x < of_int z + 1 \<Longrightarrow> \<lfloor>x\<rfloor> = z"  | 
| 30096 | 225  | 
using floor_correct [of x] floor_exists1 [of x] by auto  | 
226  | 
||
| 66515 | 227  | 
lemma floor_eq_iff: "\<lfloor>x\<rfloor> = a \<longleftrightarrow> of_int a \<le> x \<and> x < of_int a + 1"  | 
228  | 
using floor_correct floor_unique by auto  | 
|
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
229  | 
|
| 61942 | 230  | 
lemma of_int_floor_le [simp]: "of_int \<lfloor>x\<rfloor> \<le> x"  | 
| 30096 | 231  | 
using floor_correct ..  | 
232  | 
||
| 61942 | 233  | 
lemma le_floor_iff: "z \<le> \<lfloor>x\<rfloor> \<longleftrightarrow> of_int z \<le> x"  | 
| 30096 | 234  | 
proof  | 
| 61942 | 235  | 
assume "z \<le> \<lfloor>x\<rfloor>"  | 
236  | 
then have "(of_int z :: 'a) \<le> of_int \<lfloor>x\<rfloor>" by simp  | 
|
237  | 
also have "of_int \<lfloor>x\<rfloor> \<le> x" by (rule of_int_floor_le)  | 
|
| 30096 | 238  | 
finally show "of_int z \<le> x" .  | 
239  | 
next  | 
|
240  | 
assume "of_int z \<le> x"  | 
|
| 61942 | 241  | 
also have "x < of_int (\<lfloor>x\<rfloor> + 1)" using floor_correct ..  | 
242  | 
finally show "z \<le> \<lfloor>x\<rfloor>" by (simp del: of_int_add)  | 
|
| 30096 | 243  | 
qed  | 
244  | 
||
| 61942 | 245  | 
lemma floor_less_iff: "\<lfloor>x\<rfloor> < z \<longleftrightarrow> x < of_int z"  | 
| 30096 | 246  | 
by (simp add: not_le [symmetric] le_floor_iff)  | 
247  | 
||
| 61942 | 248  | 
lemma less_floor_iff: "z < \<lfloor>x\<rfloor> \<longleftrightarrow> of_int z + 1 \<le> x"  | 
| 30096 | 249  | 
using le_floor_iff [of "z + 1" x] by auto  | 
250  | 
||
| 61942 | 251  | 
lemma floor_le_iff: "\<lfloor>x\<rfloor> \<le> z \<longleftrightarrow> x < of_int z + 1"  | 
| 30096 | 252  | 
by (simp add: not_less [symmetric] less_floor_iff)  | 
253  | 
||
| 61942 | 254  | 
lemma floor_split[arith_split]: "P \<lfloor>t\<rfloor> \<longleftrightarrow> (\<forall>i. of_int i \<le> t \<and> t < of_int i + 1 \<longrightarrow> P i)"  | 
| 
58040
 
9a867afaab5a
better linarith support for floor, ceiling, natfloor, and natceiling
 
hoelzl 
parents: 
54489 
diff
changeset
 | 
255  | 
by (metis floor_correct floor_unique less_floor_iff not_le order_refl)  | 
| 
 
9a867afaab5a
better linarith support for floor, ceiling, natfloor, and natceiling
 
hoelzl 
parents: 
54489 
diff
changeset
 | 
256  | 
|
| 61942 | 257  | 
lemma floor_mono:  | 
258  | 
assumes "x \<le> y"  | 
|
259  | 
shows "\<lfloor>x\<rfloor> \<le> \<lfloor>y\<rfloor>"  | 
|
| 30096 | 260  | 
proof -  | 
| 61942 | 261  | 
have "of_int \<lfloor>x\<rfloor> \<le> x" by (rule of_int_floor_le)  | 
| 60758 | 262  | 
also note \<open>x \<le> y\<close>  | 
| 30096 | 263  | 
finally show ?thesis by (simp add: le_floor_iff)  | 
264  | 
qed  | 
|
265  | 
||
| 61942 | 266  | 
lemma floor_less_cancel: "\<lfloor>x\<rfloor> < \<lfloor>y\<rfloor> \<Longrightarrow> x < y"  | 
| 30096 | 267  | 
by (auto simp add: not_le [symmetric] floor_mono)  | 
268  | 
||
| 61942 | 269  | 
lemma floor_of_int [simp]: "\<lfloor>of_int z\<rfloor> = z"  | 
| 30096 | 270  | 
by (rule floor_unique) simp_all  | 
271  | 
||
| 61942 | 272  | 
lemma floor_of_nat [simp]: "\<lfloor>of_nat n\<rfloor> = int n"  | 
| 30096 | 273  | 
using floor_of_int [of "of_nat n"] by simp  | 
274  | 
||
| 61942 | 275  | 
lemma le_floor_add: "\<lfloor>x\<rfloor> + \<lfloor>y\<rfloor> \<le> \<lfloor>x + y\<rfloor>"  | 
| 
47307
 
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
 
huffman 
parents: 
47108 
diff
changeset
 | 
276  | 
by (simp only: le_floor_iff of_int_add add_mono of_int_floor_le)  | 
| 
 
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
 
huffman 
parents: 
47108 
diff
changeset
 | 
277  | 
|
| 63489 | 278  | 
|
279  | 
text \<open>Floor with numerals.\<close>  | 
|
| 30096 | 280  | 
|
| 61942 | 281  | 
lemma floor_zero [simp]: "\<lfloor>0\<rfloor> = 0"  | 
| 30096 | 282  | 
using floor_of_int [of 0] by simp  | 
283  | 
||
| 61942 | 284  | 
lemma floor_one [simp]: "\<lfloor>1\<rfloor> = 1"  | 
| 30096 | 285  | 
using floor_of_int [of 1] by simp  | 
286  | 
||
| 61942 | 287  | 
lemma floor_numeral [simp]: "\<lfloor>numeral v\<rfloor> = numeral v"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
288  | 
using floor_of_int [of "numeral v"] by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
289  | 
|
| 61942 | 290  | 
lemma floor_neg_numeral [simp]: "\<lfloor>- numeral v\<rfloor> = - numeral v"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54281 
diff
changeset
 | 
291  | 
using floor_of_int [of "- numeral v"] by simp  | 
| 30096 | 292  | 
|
| 61942 | 293  | 
lemma zero_le_floor [simp]: "0 \<le> \<lfloor>x\<rfloor> \<longleftrightarrow> 0 \<le> x"  | 
| 30096 | 294  | 
by (simp add: le_floor_iff)  | 
295  | 
||
| 61942 | 296  | 
lemma one_le_floor [simp]: "1 \<le> \<lfloor>x\<rfloor> \<longleftrightarrow> 1 \<le> x"  | 
| 30096 | 297  | 
by (simp add: le_floor_iff)  | 
298  | 
||
| 63489 | 299  | 
lemma numeral_le_floor [simp]: "numeral v \<le> \<lfloor>x\<rfloor> \<longleftrightarrow> numeral v \<le> x"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
300  | 
by (simp add: le_floor_iff)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
301  | 
|
| 63489 | 302  | 
lemma neg_numeral_le_floor [simp]: "- numeral v \<le> \<lfloor>x\<rfloor> \<longleftrightarrow> - numeral v \<le> x"  | 
| 30096 | 303  | 
by (simp add: le_floor_iff)  | 
304  | 
||
| 61942 | 305  | 
lemma zero_less_floor [simp]: "0 < \<lfloor>x\<rfloor> \<longleftrightarrow> 1 \<le> x"  | 
| 30096 | 306  | 
by (simp add: less_floor_iff)  | 
307  | 
||
| 61942 | 308  | 
lemma one_less_floor [simp]: "1 < \<lfloor>x\<rfloor> \<longleftrightarrow> 2 \<le> x"  | 
| 30096 | 309  | 
by (simp add: less_floor_iff)  | 
310  | 
||
| 63489 | 311  | 
lemma numeral_less_floor [simp]: "numeral v < \<lfloor>x\<rfloor> \<longleftrightarrow> numeral v + 1 \<le> x"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
312  | 
by (simp add: less_floor_iff)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
313  | 
|
| 63489 | 314  | 
lemma neg_numeral_less_floor [simp]: "- numeral v < \<lfloor>x\<rfloor> \<longleftrightarrow> - numeral v + 1 \<le> x"  | 
| 30096 | 315  | 
by (simp add: less_floor_iff)  | 
316  | 
||
| 61942 | 317  | 
lemma floor_le_zero [simp]: "\<lfloor>x\<rfloor> \<le> 0 \<longleftrightarrow> x < 1"  | 
| 30096 | 318  | 
by (simp add: floor_le_iff)  | 
319  | 
||
| 61942 | 320  | 
lemma floor_le_one [simp]: "\<lfloor>x\<rfloor> \<le> 1 \<longleftrightarrow> x < 2"  | 
| 30096 | 321  | 
by (simp add: floor_le_iff)  | 
322  | 
||
| 63489 | 323  | 
lemma floor_le_numeral [simp]: "\<lfloor>x\<rfloor> \<le> numeral v \<longleftrightarrow> x < numeral v + 1"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
324  | 
by (simp add: floor_le_iff)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
325  | 
|
| 63489 | 326  | 
lemma floor_le_neg_numeral [simp]: "\<lfloor>x\<rfloor> \<le> - numeral v \<longleftrightarrow> x < - numeral v + 1"  | 
| 30096 | 327  | 
by (simp add: floor_le_iff)  | 
328  | 
||
| 61942 | 329  | 
lemma floor_less_zero [simp]: "\<lfloor>x\<rfloor> < 0 \<longleftrightarrow> x < 0"  | 
| 30096 | 330  | 
by (simp add: floor_less_iff)  | 
331  | 
||
| 61942 | 332  | 
lemma floor_less_one [simp]: "\<lfloor>x\<rfloor> < 1 \<longleftrightarrow> x < 1"  | 
| 30096 | 333  | 
by (simp add: floor_less_iff)  | 
334  | 
||
| 63489 | 335  | 
lemma floor_less_numeral [simp]: "\<lfloor>x\<rfloor> < numeral v \<longleftrightarrow> x < numeral v"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
336  | 
by (simp add: floor_less_iff)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
337  | 
|
| 63489 | 338  | 
lemma floor_less_neg_numeral [simp]: "\<lfloor>x\<rfloor> < - numeral v \<longleftrightarrow> x < - numeral v"  | 
| 30096 | 339  | 
by (simp add: floor_less_iff)  | 
340  | 
||
| 
66154
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
341  | 
lemma le_mult_floor_Ints:  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
342  | 
assumes "0 \<le> a" "a \<in> Ints"  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
343  | 
shows "of_int (\<lfloor>a\<rfloor> * \<lfloor>b\<rfloor>) \<le> (of_int\<lfloor>a * b\<rfloor> :: 'a :: linordered_idom)"  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
344  | 
by (metis Ints_cases assms floor_less_iff floor_of_int linorder_not_less mult_left_mono of_int_floor_le of_int_less_iff of_int_mult)  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
345  | 
|
| 63489 | 346  | 
|
347  | 
text \<open>Addition and subtraction of integers.\<close>  | 
|
| 30096 | 348  | 
|
| 63599 | 349  | 
lemma floor_add_int: "\<lfloor>x\<rfloor> + z = \<lfloor>x + of_int z\<rfloor>"  | 
350  | 
using floor_correct [of x] by (simp add: floor_unique[symmetric])  | 
|
| 30096 | 351  | 
|
| 63599 | 352  | 
lemma int_add_floor: "z + \<lfloor>x\<rfloor> = \<lfloor>of_int z + x\<rfloor>"  | 
353  | 
using floor_correct [of x] by (simp add: floor_unique[symmetric])  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
354  | 
|
| 63599 | 355  | 
lemma one_add_floor: "\<lfloor>x\<rfloor> + 1 = \<lfloor>x + 1\<rfloor>"  | 
356  | 
using floor_add_int [of x 1] by simp  | 
|
| 30096 | 357  | 
|
| 61942 | 358  | 
lemma floor_diff_of_int [simp]: "\<lfloor>x - of_int z\<rfloor> = \<lfloor>x\<rfloor> - z"  | 
| 63599 | 359  | 
using floor_add_int [of x "- z"] by (simp add: algebra_simps)  | 
| 30096 | 360  | 
|
| 61942 | 361  | 
lemma floor_uminus_of_int [simp]: "\<lfloor>- (of_int z)\<rfloor> = - z"  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
362  | 
by (metis floor_diff_of_int [of 0] diff_0 floor_zero)  | 
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
363  | 
|
| 63489 | 364  | 
lemma floor_diff_numeral [simp]: "\<lfloor>x - numeral v\<rfloor> = \<lfloor>x\<rfloor> - numeral v"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
365  | 
using floor_diff_of_int [of x "numeral v"] by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
366  | 
|
| 61942 | 367  | 
lemma floor_diff_one [simp]: "\<lfloor>x - 1\<rfloor> = \<lfloor>x\<rfloor> - 1"  | 
| 30096 | 368  | 
using floor_diff_of_int [of x 1] by simp  | 
369  | 
||
| 
58097
 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 
hoelzl 
parents: 
58040 
diff
changeset
 | 
370  | 
lemma le_mult_floor:  | 
| 
 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 
hoelzl 
parents: 
58040 
diff
changeset
 | 
371  | 
assumes "0 \<le> a" and "0 \<le> b"  | 
| 61942 | 372  | 
shows "\<lfloor>a\<rfloor> * \<lfloor>b\<rfloor> \<le> \<lfloor>a * b\<rfloor>"  | 
| 
58097
 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 
hoelzl 
parents: 
58040 
diff
changeset
 | 
373  | 
proof -  | 
| 63489 | 374  | 
have "of_int \<lfloor>a\<rfloor> \<le> a" and "of_int \<lfloor>b\<rfloor> \<le> b"  | 
375  | 
by (auto intro: of_int_floor_le)  | 
|
376  | 
then have "of_int (\<lfloor>a\<rfloor> * \<lfloor>b\<rfloor>) \<le> a * b"  | 
|
| 
58097
 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 
hoelzl 
parents: 
58040 
diff
changeset
 | 
377  | 
using assms by (auto intro!: mult_mono)  | 
| 61942 | 378  | 
also have "a * b < of_int (\<lfloor>a * b\<rfloor> + 1)"  | 
| 
58097
 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 
hoelzl 
parents: 
58040 
diff
changeset
 | 
379  | 
using floor_correct[of "a * b"] by auto  | 
| 63489 | 380  | 
finally show ?thesis  | 
381  | 
unfolding of_int_less_iff by simp  | 
|
| 
58097
 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 
hoelzl 
parents: 
58040 
diff
changeset
 | 
382  | 
qed  | 
| 
 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 
hoelzl 
parents: 
58040 
diff
changeset
 | 
383  | 
|
| 63489 | 384  | 
lemma floor_divide_of_int_eq: "\<lfloor>of_int k / of_int l\<rfloor> = k div l"  | 
385  | 
for k l :: int  | 
|
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
386  | 
proof (cases "l = 0")  | 
| 63489 | 387  | 
case True  | 
388  | 
then show ?thesis by simp  | 
|
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
389  | 
next  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
390  | 
case False  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
391  | 
have *: "\<lfloor>of_int (k mod l) / of_int l :: 'a\<rfloor> = 0"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
392  | 
proof (cases "l > 0")  | 
| 63489 | 393  | 
case True  | 
394  | 
then show ?thesis  | 
|
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
395  | 
by (auto intro: floor_unique)  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
396  | 
next  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
397  | 
case False  | 
| 63489 | 398  | 
obtain r where "r = - l"  | 
399  | 
by blast  | 
|
400  | 
then have l: "l = - r"  | 
|
401  | 
by simp  | 
|
| 63540 | 402  | 
with \<open>l \<noteq> 0\<close> False have "r > 0"  | 
| 63489 | 403  | 
by simp  | 
| 63540 | 404  | 
with l show ?thesis  | 
| 63489 | 405  | 
using pos_mod_bound [of r]  | 
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
406  | 
by (auto simp add: zmod_zminus2_eq_if less_le field_simps intro: floor_unique)  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
407  | 
qed  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
408  | 
have "(of_int k :: 'a) = of_int (k div l * l + k mod l)"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
409  | 
by simp  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
410  | 
also have "\<dots> = (of_int (k div l) + of_int (k mod l) / of_int l) * of_int l"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
411  | 
using False by (simp only: of_int_add) (simp add: field_simps)  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
412  | 
finally have "(of_int k / of_int l :: 'a) = \<dots> / of_int l"  | 
| 63331 | 413  | 
by simp  | 
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
414  | 
then have "(of_int k / of_int l :: 'a) = of_int (k div l) + of_int (k mod l) / of_int l"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
415  | 
using False by (simp only:) (simp add: field_simps)  | 
| 63331 | 416  | 
then have "\<lfloor>of_int k / of_int l :: 'a\<rfloor> = \<lfloor>of_int (k div l) + of_int (k mod l) / of_int l :: 'a\<rfloor>"  | 
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
417  | 
by simp  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
418  | 
then have "\<lfloor>of_int k / of_int l :: 'a\<rfloor> = \<lfloor>of_int (k mod l) / of_int l + of_int (k div l) :: 'a\<rfloor>"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
419  | 
by (simp add: ac_simps)  | 
| 60128 | 420  | 
then have "\<lfloor>of_int k / of_int l :: 'a\<rfloor> = \<lfloor>of_int (k mod l) / of_int l :: 'a\<rfloor> + k div l"  | 
| 63599 | 421  | 
by (simp add: floor_add_int)  | 
| 63489 | 422  | 
with * show ?thesis  | 
423  | 
by simp  | 
|
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
424  | 
qed  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
425  | 
|
| 63489 | 426  | 
lemma floor_divide_of_nat_eq: "\<lfloor>of_nat m / of_nat n\<rfloor> = of_nat (m div n)"  | 
427  | 
for m n :: nat  | 
|
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
428  | 
proof (cases "n = 0")  | 
| 63489 | 429  | 
case True  | 
430  | 
then show ?thesis by simp  | 
|
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
431  | 
next  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
432  | 
case False  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
433  | 
then have *: "\<lfloor>of_nat (m mod n) / of_nat n :: 'a\<rfloor> = 0"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
434  | 
by (auto intro: floor_unique)  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
435  | 
have "(of_nat m :: 'a) = of_nat (m div n * n + m mod n)"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
436  | 
by simp  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
437  | 
also have "\<dots> = (of_nat (m div n) + of_nat (m mod n) / of_nat n) * of_nat n"  | 
| 63489 | 438  | 
using False by (simp only: of_nat_add) (simp add: field_simps)  | 
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
439  | 
finally have "(of_nat m / of_nat n :: 'a) = \<dots> / of_nat n"  | 
| 63331 | 440  | 
by simp  | 
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
441  | 
then have "(of_nat m / of_nat n :: 'a) = of_nat (m div n) + of_nat (m mod n) / of_nat n"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
442  | 
using False by (simp only:) simp  | 
| 63331 | 443  | 
then have "\<lfloor>of_nat m / of_nat n :: 'a\<rfloor> = \<lfloor>of_nat (m div n) + of_nat (m mod n) / of_nat n :: 'a\<rfloor>"  | 
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
444  | 
by simp  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
445  | 
then have "\<lfloor>of_nat m / of_nat n :: 'a\<rfloor> = \<lfloor>of_nat (m mod n) / of_nat n + of_nat (m div n) :: 'a\<rfloor>"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
446  | 
by (simp add: ac_simps)  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
447  | 
moreover have "(of_nat (m div n) :: 'a) = of_int (of_nat (m div n))"  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
448  | 
by simp  | 
| 63489 | 449  | 
ultimately have "\<lfloor>of_nat m / of_nat n :: 'a\<rfloor> =  | 
450  | 
\<lfloor>of_nat (m mod n) / of_nat n :: 'a\<rfloor> + of_nat (m div n)"  | 
|
| 63599 | 451  | 
by (simp only: floor_add_int)  | 
| 63489 | 452  | 
with * show ?thesis  | 
453  | 
by simp  | 
|
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
454  | 
qed  | 
| 
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
455  | 
|
| 
68499
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
456  | 
lemma floor_divide_lower:  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
457  | 
fixes q :: "'a::floor_ceiling"  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
458  | 
shows "q > 0 \<Longrightarrow> of_int \<lfloor>p / q\<rfloor> * q \<le> p"  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
459  | 
using of_int_floor_le pos_le_divide_eq by blast  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
460  | 
|
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
461  | 
lemma floor_divide_upper:  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
462  | 
fixes q :: "'a::floor_ceiling"  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
463  | 
shows "q > 0 \<Longrightarrow> p < (of_int \<lfloor>p / q\<rfloor> + 1) * q"  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
464  | 
by (meson floor_eq_iff pos_divide_less_eq)  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
465  | 
|
| 
59984
 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 
haftmann 
parents: 
59613 
diff
changeset
 | 
466  | 
|
| 60758 | 467  | 
subsection \<open>Ceiling function\<close>  | 
| 30096 | 468  | 
|
| 61942 | 469  | 
definition ceiling :: "'a::floor_ceiling \<Rightarrow> int"  ("\<lceil>_\<rceil>")
 | 
470  | 
where "\<lceil>x\<rceil> = - \<lfloor>- x\<rfloor>"  | 
|
| 30096 | 471  | 
|
| 61942 | 472  | 
lemma ceiling_correct: "of_int \<lceil>x\<rceil> - 1 < x \<and> x \<le> of_int \<lceil>x\<rceil>"  | 
| 63489 | 473  | 
unfolding ceiling_def using floor_correct [of "- x"]  | 
474  | 
by (simp add: le_minus_iff)  | 
|
| 30096 | 475  | 
|
| 63489 | 476  | 
lemma ceiling_unique: "of_int z - 1 < x \<Longrightarrow> x \<le> of_int z \<Longrightarrow> \<lceil>x\<rceil> = z"  | 
| 30096 | 477  | 
unfolding ceiling_def using floor_unique [of "- z" "- x"] by simp  | 
478  | 
||
| 66515 | 479  | 
lemma ceiling_eq_iff: "\<lceil>x\<rceil> = a \<longleftrightarrow> of_int a - 1 < x \<and> x \<le> of_int a"  | 
480  | 
using ceiling_correct ceiling_unique by auto  | 
|
481  | 
||
| 61942 | 482  | 
lemma le_of_int_ceiling [simp]: "x \<le> of_int \<lceil>x\<rceil>"  | 
| 30096 | 483  | 
using ceiling_correct ..  | 
484  | 
||
| 61942 | 485  | 
lemma ceiling_le_iff: "\<lceil>x\<rceil> \<le> z \<longleftrightarrow> x \<le> of_int z"  | 
| 30096 | 486  | 
unfolding ceiling_def using le_floor_iff [of "- z" "- x"] by auto  | 
487  | 
||
| 61942 | 488  | 
lemma less_ceiling_iff: "z < \<lceil>x\<rceil> \<longleftrightarrow> of_int z < x"  | 
| 30096 | 489  | 
by (simp add: not_le [symmetric] ceiling_le_iff)  | 
490  | 
||
| 61942 | 491  | 
lemma ceiling_less_iff: "\<lceil>x\<rceil> < z \<longleftrightarrow> x \<le> of_int z - 1"  | 
| 30096 | 492  | 
using ceiling_le_iff [of x "z - 1"] by simp  | 
493  | 
||
| 61942 | 494  | 
lemma le_ceiling_iff: "z \<le> \<lceil>x\<rceil> \<longleftrightarrow> of_int z - 1 < x"  | 
| 30096 | 495  | 
by (simp add: not_less [symmetric] ceiling_less_iff)  | 
496  | 
||
| 61942 | 497  | 
lemma ceiling_mono: "x \<ge> y \<Longrightarrow> \<lceil>x\<rceil> \<ge> \<lceil>y\<rceil>"  | 
| 30096 | 498  | 
unfolding ceiling_def by (simp add: floor_mono)  | 
499  | 
||
| 61942 | 500  | 
lemma ceiling_less_cancel: "\<lceil>x\<rceil> < \<lceil>y\<rceil> \<Longrightarrow> x < y"  | 
| 30096 | 501  | 
by (auto simp add: not_le [symmetric] ceiling_mono)  | 
502  | 
||
| 61942 | 503  | 
lemma ceiling_of_int [simp]: "\<lceil>of_int z\<rceil> = z"  | 
| 30096 | 504  | 
by (rule ceiling_unique) simp_all  | 
505  | 
||
| 61942 | 506  | 
lemma ceiling_of_nat [simp]: "\<lceil>of_nat n\<rceil> = int n"  | 
| 30096 | 507  | 
using ceiling_of_int [of "of_nat n"] by simp  | 
508  | 
||
| 61942 | 509  | 
lemma ceiling_add_le: "\<lceil>x + y\<rceil> \<le> \<lceil>x\<rceil> + \<lceil>y\<rceil>"  | 
| 
47307
 
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
 
huffman 
parents: 
47108 
diff
changeset
 | 
510  | 
by (simp only: ceiling_le_iff of_int_add add_mono le_of_int_ceiling)  | 
| 
 
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
 
huffman 
parents: 
47108 
diff
changeset
 | 
511  | 
|
| 
66154
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
512  | 
lemma mult_ceiling_le:  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
513  | 
assumes "0 \<le> a" and "0 \<le> b"  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
514  | 
shows "\<lceil>a * b\<rceil> \<le> \<lceil>a\<rceil> * \<lceil>b\<rceil>"  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
515  | 
by (metis assms ceiling_le_iff ceiling_mono le_of_int_ceiling mult_mono of_int_mult)  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
516  | 
|
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
517  | 
lemma mult_ceiling_le_Ints:  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
518  | 
assumes "0 \<le> a" "a \<in> Ints"  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
519  | 
shows "(of_int \<lceil>a * b\<rceil> :: 'a :: linordered_idom) \<le> of_int(\<lceil>a\<rceil> * \<lceil>b\<rceil>)"  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
520  | 
by (metis Ints_cases assms ceiling_le_iff ceiling_of_int le_of_int_ceiling mult_left_mono of_int_le_iff of_int_mult)  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
521  | 
|
| 
63879
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
522  | 
lemma finite_int_segment:  | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
523  | 
fixes a :: "'a::floor_ceiling"  | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
524  | 
  shows "finite {x \<in> \<int>. a \<le> x \<and> x \<le> b}"
 | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
525  | 
proof -  | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
526  | 
  have "finite {ceiling a..floor b}"
 | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
527  | 
by simp  | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
528  | 
  moreover have "{x \<in> \<int>. a \<le> x \<and> x \<le> b} = of_int ` {ceiling a..floor b}"
 | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
529  | 
by (auto simp: le_floor_iff ceiling_le_iff elim!: Ints_cases)  | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
530  | 
ultimately show ?thesis  | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
531  | 
by simp  | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
532  | 
qed  | 
| 
 
15bbf6360339
simple new lemmas, mostly about sets
 
paulson <lp15@cam.ac.uk> 
parents: 
63621 
diff
changeset
 | 
533  | 
|
| 
66154
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
534  | 
corollary finite_abs_int_segment:  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
535  | 
fixes a :: "'a::floor_ceiling"  | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
536  | 
  shows "finite {k \<in> \<int>. \<bar>k\<bar> \<le> a}" 
 | 
| 
 
bc5e6461f759
Tidying up integration theory and some new theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
64317 
diff
changeset
 | 
537  | 
using finite_int_segment [of "-a" a] by (auto simp add: abs_le_iff conj_commute minus_le_iff)  | 
| 63489 | 538  | 
|
| 
66793
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
539  | 
|
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
540  | 
subsubsection \<open>Ceiling with numerals.\<close>  | 
| 30096 | 541  | 
|
| 61942 | 542  | 
lemma ceiling_zero [simp]: "\<lceil>0\<rceil> = 0"  | 
| 30096 | 543  | 
using ceiling_of_int [of 0] by simp  | 
544  | 
||
| 61942 | 545  | 
lemma ceiling_one [simp]: "\<lceil>1\<rceil> = 1"  | 
| 30096 | 546  | 
using ceiling_of_int [of 1] by simp  | 
547  | 
||
| 61942 | 548  | 
lemma ceiling_numeral [simp]: "\<lceil>numeral v\<rceil> = numeral v"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
549  | 
using ceiling_of_int [of "numeral v"] by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
550  | 
|
| 61942 | 551  | 
lemma ceiling_neg_numeral [simp]: "\<lceil>- numeral v\<rceil> = - numeral v"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54281 
diff
changeset
 | 
552  | 
using ceiling_of_int [of "- numeral v"] by simp  | 
| 30096 | 553  | 
|
| 61942 | 554  | 
lemma ceiling_le_zero [simp]: "\<lceil>x\<rceil> \<le> 0 \<longleftrightarrow> x \<le> 0"  | 
| 30096 | 555  | 
by (simp add: ceiling_le_iff)  | 
556  | 
||
| 61942 | 557  | 
lemma ceiling_le_one [simp]: "\<lceil>x\<rceil> \<le> 1 \<longleftrightarrow> x \<le> 1"  | 
| 30096 | 558  | 
by (simp add: ceiling_le_iff)  | 
559  | 
||
| 63489 | 560  | 
lemma ceiling_le_numeral [simp]: "\<lceil>x\<rceil> \<le> numeral v \<longleftrightarrow> x \<le> numeral v"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
561  | 
by (simp add: ceiling_le_iff)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
562  | 
|
| 63489 | 563  | 
lemma ceiling_le_neg_numeral [simp]: "\<lceil>x\<rceil> \<le> - numeral v \<longleftrightarrow> x \<le> - numeral v"  | 
| 30096 | 564  | 
by (simp add: ceiling_le_iff)  | 
565  | 
||
| 61942 | 566  | 
lemma ceiling_less_zero [simp]: "\<lceil>x\<rceil> < 0 \<longleftrightarrow> x \<le> -1"  | 
| 30096 | 567  | 
by (simp add: ceiling_less_iff)  | 
568  | 
||
| 61942 | 569  | 
lemma ceiling_less_one [simp]: "\<lceil>x\<rceil> < 1 \<longleftrightarrow> x \<le> 0"  | 
| 30096 | 570  | 
by (simp add: ceiling_less_iff)  | 
571  | 
||
| 63489 | 572  | 
lemma ceiling_less_numeral [simp]: "\<lceil>x\<rceil> < numeral v \<longleftrightarrow> x \<le> numeral v - 1"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
573  | 
by (simp add: ceiling_less_iff)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
574  | 
|
| 63489 | 575  | 
lemma ceiling_less_neg_numeral [simp]: "\<lceil>x\<rceil> < - numeral v \<longleftrightarrow> x \<le> - numeral v - 1"  | 
| 30096 | 576  | 
by (simp add: ceiling_less_iff)  | 
577  | 
||
| 61942 | 578  | 
lemma zero_le_ceiling [simp]: "0 \<le> \<lceil>x\<rceil> \<longleftrightarrow> -1 < x"  | 
| 30096 | 579  | 
by (simp add: le_ceiling_iff)  | 
580  | 
||
| 61942 | 581  | 
lemma one_le_ceiling [simp]: "1 \<le> \<lceil>x\<rceil> \<longleftrightarrow> 0 < x"  | 
| 30096 | 582  | 
by (simp add: le_ceiling_iff)  | 
583  | 
||
| 63489 | 584  | 
lemma numeral_le_ceiling [simp]: "numeral v \<le> \<lceil>x\<rceil> \<longleftrightarrow> numeral v - 1 < x"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
585  | 
by (simp add: le_ceiling_iff)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
586  | 
|
| 63489 | 587  | 
lemma neg_numeral_le_ceiling [simp]: "- numeral v \<le> \<lceil>x\<rceil> \<longleftrightarrow> - numeral v - 1 < x"  | 
| 30096 | 588  | 
by (simp add: le_ceiling_iff)  | 
589  | 
||
| 61942 | 590  | 
lemma zero_less_ceiling [simp]: "0 < \<lceil>x\<rceil> \<longleftrightarrow> 0 < x"  | 
| 30096 | 591  | 
by (simp add: less_ceiling_iff)  | 
592  | 
||
| 61942 | 593  | 
lemma one_less_ceiling [simp]: "1 < \<lceil>x\<rceil> \<longleftrightarrow> 1 < x"  | 
| 30096 | 594  | 
by (simp add: less_ceiling_iff)  | 
595  | 
||
| 63489 | 596  | 
lemma numeral_less_ceiling [simp]: "numeral v < \<lceil>x\<rceil> \<longleftrightarrow> numeral v < x"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
597  | 
by (simp add: less_ceiling_iff)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
598  | 
|
| 63489 | 599  | 
lemma neg_numeral_less_ceiling [simp]: "- numeral v < \<lceil>x\<rceil> \<longleftrightarrow> - numeral v < x"  | 
| 30096 | 600  | 
by (simp add: less_ceiling_iff)  | 
601  | 
||
| 61942 | 602  | 
lemma ceiling_altdef: "\<lceil>x\<rceil> = (if x = of_int \<lfloor>x\<rfloor> then \<lfloor>x\<rfloor> else \<lfloor>x\<rfloor> + 1)"  | 
| 63489 | 603  | 
by (intro ceiling_unique; simp, linarith?)  | 
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
604  | 
|
| 61942 | 605  | 
lemma floor_le_ceiling [simp]: "\<lfloor>x\<rfloor> \<le> \<lceil>x\<rceil>"  | 
606  | 
by (simp add: ceiling_altdef)  | 
|
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
607  | 
|
| 63489 | 608  | 
|
| 
66793
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
609  | 
subsubsection \<open>Addition and subtraction of integers.\<close>  | 
| 30096 | 610  | 
|
| 61942 | 611  | 
lemma ceiling_add_of_int [simp]: "\<lceil>x + of_int z\<rceil> = \<lceil>x\<rceil> + z"  | 
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61531 
diff
changeset
 | 
612  | 
using ceiling_correct [of x] by (simp add: ceiling_def)  | 
| 30096 | 613  | 
|
| 61942 | 614  | 
lemma ceiling_add_numeral [simp]: "\<lceil>x + numeral v\<rceil> = \<lceil>x\<rceil> + numeral v"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
615  | 
using ceiling_add_of_int [of x "numeral v"] by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
616  | 
|
| 61942 | 617  | 
lemma ceiling_add_one [simp]: "\<lceil>x + 1\<rceil> = \<lceil>x\<rceil> + 1"  | 
| 30096 | 618  | 
using ceiling_add_of_int [of x 1] by simp  | 
619  | 
||
| 61942 | 620  | 
lemma ceiling_diff_of_int [simp]: "\<lceil>x - of_int z\<rceil> = \<lceil>x\<rceil> - z"  | 
| 30096 | 621  | 
using ceiling_add_of_int [of x "- z"] by (simp add: algebra_simps)  | 
622  | 
||
| 61942 | 623  | 
lemma ceiling_diff_numeral [simp]: "\<lceil>x - numeral v\<rceil> = \<lceil>x\<rceil> - numeral v"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
624  | 
using ceiling_diff_of_int [of x "numeral v"] by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43733 
diff
changeset
 | 
625  | 
|
| 61942 | 626  | 
lemma ceiling_diff_one [simp]: "\<lceil>x - 1\<rceil> = \<lceil>x\<rceil> - 1"  | 
| 30096 | 627  | 
using ceiling_diff_of_int [of x 1] by simp  | 
628  | 
||
| 61942 | 629  | 
lemma ceiling_split[arith_split]: "P \<lceil>t\<rceil> \<longleftrightarrow> (\<forall>i. of_int i - 1 < t \<and> t \<le> of_int i \<longrightarrow> P i)"  | 
| 
58040
 
9a867afaab5a
better linarith support for floor, ceiling, natfloor, and natceiling
 
hoelzl 
parents: 
54489 
diff
changeset
 | 
630  | 
by (auto simp add: ceiling_unique ceiling_correct)  | 
| 
 
9a867afaab5a
better linarith support for floor, ceiling, natfloor, and natceiling
 
hoelzl 
parents: 
54489 
diff
changeset
 | 
631  | 
|
| 61942 | 632  | 
lemma ceiling_diff_floor_le_1: "\<lceil>x\<rceil> - \<lfloor>x\<rfloor> \<le> 1"  | 
| 47592 | 633  | 
proof -  | 
| 63331 | 634  | 
have "of_int \<lceil>x\<rceil> - 1 < x"  | 
| 47592 | 635  | 
using ceiling_correct[of x] by simp  | 
636  | 
also have "x < of_int \<lfloor>x\<rfloor> + 1"  | 
|
637  | 
using floor_correct[of x] by simp_all  | 
|
638  | 
finally have "of_int (\<lceil>x\<rceil> - \<lfloor>x\<rfloor>) < (of_int 2::'a)"  | 
|
639  | 
by simp  | 
|
640  | 
then show ?thesis  | 
|
641  | 
unfolding of_int_less_iff by simp  | 
|
642  | 
qed  | 
|
| 30096 | 643  | 
|
| 
66793
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
644  | 
lemma nat_approx_posE:  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
645  | 
  fixes e:: "'a::{archimedean_field,floor_ceiling}"
 | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
646  | 
assumes "0 < e"  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
647  | 
obtains n :: nat where "1 / of_nat(Suc n) < e"  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
648  | 
proof  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
649  | 
have "(1::'a) / of_nat (Suc (nat \<lceil>1/e\<rceil>)) < 1 / of_int (\<lceil>1/e\<rceil>)"  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
650  | 
proof (rule divide_strict_left_mono)  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
651  | 
show "(of_int \<lceil>1 / e\<rceil>::'a) < of_nat (Suc (nat \<lceil>1 / e\<rceil>))"  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
652  | 
using assms by (simp add: field_simps)  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
653  | 
show "(0::'a) < of_nat (Suc (nat \<lceil>1 / e\<rceil>)) * of_int \<lceil>1 / e\<rceil>"  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
654  | 
using assms by (auto simp: zero_less_mult_iff pos_add_strict)  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
655  | 
qed auto  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
656  | 
also have "1 / of_int (\<lceil>1/e\<rceil>) \<le> 1 / (1/e)"  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
657  | 
by (rule divide_left_mono) (auto simp: \<open>0 < e\<close> ceiling_correct)  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
658  | 
also have "\<dots> = e" by simp  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
659  | 
finally show "1 / of_nat (Suc (nat \<lceil>1 / e\<rceil>)) < e"  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
660  | 
by metis  | 
| 
 
deabce3ccf1f
new material about connectedness, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
66515 
diff
changeset
 | 
661  | 
qed  | 
| 63489 | 662  | 
|
| 
68499
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
663  | 
lemma ceiling_divide_upper:  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
664  | 
fixes q :: "'a::floor_ceiling"  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
665  | 
shows "q > 0 \<Longrightarrow> p \<le> of_int (ceiling (p / q)) * q"  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
666  | 
by (meson divide_le_eq le_of_int_ceiling)  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
667  | 
|
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
668  | 
lemma ceiling_divide_lower:  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
669  | 
fixes q :: "'a::floor_ceiling"  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
670  | 
shows "q > 0 \<Longrightarrow> (of_int \<lceil>p / q\<rceil> - 1) * q < p"  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
671  | 
by (meson ceiling_eq_iff pos_less_divide_eq)  | 
| 
 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 
paulson <lp15@cam.ac.uk> 
parents: 
66793 
diff
changeset
 | 
672  | 
|
| 60758 | 673  | 
subsection \<open>Negation\<close>  | 
| 30096 | 674  | 
|
| 61942 | 675  | 
lemma floor_minus: "\<lfloor>- x\<rfloor> = - \<lceil>x\<rceil>"  | 
| 30096 | 676  | 
unfolding ceiling_def by simp  | 
677  | 
||
| 61942 | 678  | 
lemma ceiling_minus: "\<lceil>- x\<rceil> = - \<lfloor>x\<rfloor>"  | 
| 30096 | 679  | 
unfolding ceiling_def by simp  | 
680  | 
||
| 
63945
 
444eafb6e864
a few new theorems and a renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
63879 
diff
changeset
 | 
681  | 
subsection \<open>Natural numbers\<close>  | 
| 
 
444eafb6e864
a few new theorems and a renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
63879 
diff
changeset
 | 
682  | 
|
| 
 
444eafb6e864
a few new theorems and a renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
63879 
diff
changeset
 | 
683  | 
lemma of_nat_floor: "r\<ge>0 \<Longrightarrow> of_nat (nat \<lfloor>r\<rfloor>) \<le> r"  | 
| 
 
444eafb6e864
a few new theorems and a renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
63879 
diff
changeset
 | 
684  | 
by simp  | 
| 
 
444eafb6e864
a few new theorems and a renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
63879 
diff
changeset
 | 
685  | 
|
| 
 
444eafb6e864
a few new theorems and a renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
63879 
diff
changeset
 | 
686  | 
lemma of_nat_ceiling: "of_nat (nat \<lceil>r\<rceil>) \<ge> r"  | 
| 
 
444eafb6e864
a few new theorems and a renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
63879 
diff
changeset
 | 
687  | 
by (cases "r\<ge>0") auto  | 
| 
 
444eafb6e864
a few new theorems and a renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
63879 
diff
changeset
 | 
688  | 
|
| 
 
444eafb6e864
a few new theorems and a renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
63879 
diff
changeset
 | 
689  | 
|
| 60758 | 690  | 
subsection \<open>Frac Function\<close>  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
691  | 
|
| 63489 | 692  | 
definition frac :: "'a \<Rightarrow> 'a::floor_ceiling"  | 
693  | 
where "frac x \<equiv> x - of_int \<lfloor>x\<rfloor>"  | 
|
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
694  | 
|
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
695  | 
lemma frac_lt_1: "frac x < 1"  | 
| 63489 | 696  | 
by (simp add: frac_def) linarith  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
697  | 
|
| 61070 | 698  | 
lemma frac_eq_0_iff [simp]: "frac x = 0 \<longleftrightarrow> x \<in> \<int>"  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
699  | 
by (simp add: frac_def) (metis Ints_cases Ints_of_int floor_of_int )  | 
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
700  | 
|
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
701  | 
lemma frac_ge_0 [simp]: "frac x \<ge> 0"  | 
| 63489 | 702  | 
unfolding frac_def by linarith  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
703  | 
|
| 61070 | 704  | 
lemma frac_gt_0_iff [simp]: "frac x > 0 \<longleftrightarrow> x \<notin> \<int>"  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
705  | 
by (metis frac_eq_0_iff frac_ge_0 le_less less_irrefl)  | 
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
706  | 
|
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
707  | 
lemma frac_of_int [simp]: "frac (of_int z) = 0"  | 
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
708  | 
by (simp add: frac_def)  | 
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
709  | 
|
| 68721 | 710  | 
lemma frac_frac [simp]: "frac (frac x) = frac x"  | 
711  | 
by (simp add: frac_def)  | 
|
712  | 
||
| 63331 | 713  | 
lemma floor_add: "\<lfloor>x + y\<rfloor> = (if frac x + frac y < 1 then \<lfloor>x\<rfloor> + \<lfloor>y\<rfloor> else (\<lfloor>x\<rfloor> + \<lfloor>y\<rfloor>) + 1)"  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
714  | 
proof -  | 
| 63599 | 715  | 
have "x + y < 1 + (of_int \<lfloor>x\<rfloor> + of_int \<lfloor>y\<rfloor>) \<Longrightarrow> \<lfloor>x + y\<rfloor> = \<lfloor>x\<rfloor> + \<lfloor>y\<rfloor>"  | 
716  | 
by (metis add.commute floor_unique le_floor_add le_floor_iff of_int_add)  | 
|
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
717  | 
moreover  | 
| 63599 | 718  | 
have "\<not> x + y < 1 + (of_int \<lfloor>x\<rfloor> + of_int \<lfloor>y\<rfloor>) \<Longrightarrow> \<lfloor>x + y\<rfloor> = 1 + (\<lfloor>x\<rfloor> + \<lfloor>y\<rfloor>)"  | 
| 66515 | 719  | 
apply (simp add: floor_eq_iff)  | 
| 63489 | 720  | 
apply (auto simp add: algebra_simps)  | 
721  | 
apply linarith  | 
|
722  | 
done  | 
|
| 63599 | 723  | 
ultimately show ?thesis by (auto simp add: frac_def algebra_simps)  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
724  | 
qed  | 
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
725  | 
|
| 63621 | 726  | 
lemma floor_add2[simp]: "x \<in> \<int> \<or> y \<in> \<int> \<Longrightarrow> \<lfloor>x + y\<rfloor> = \<lfloor>x\<rfloor> + \<lfloor>y\<rfloor>"  | 
727  | 
by (metis add.commute add.left_neutral frac_lt_1 floor_add frac_eq_0_iff)  | 
|
| 63597 | 728  | 
|
| 63489 | 729  | 
lemma frac_add:  | 
730  | 
"frac (x + y) = (if frac x + frac y < 1 then frac x + frac y else (frac x + frac y) - 1)"  | 
|
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
731  | 
by (simp add: frac_def floor_add)  | 
| 
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
732  | 
|
| 63489 | 733  | 
lemma frac_unique_iff: "frac x = a \<longleftrightarrow> x - a \<in> \<int> \<and> 0 \<le> a \<and> a < 1"  | 
734  | 
for x :: "'a::floor_ceiling"  | 
|
| 62348 | 735  | 
apply (auto simp: Ints_def frac_def algebra_simps floor_unique)  | 
| 63489 | 736  | 
apply linarith+  | 
| 62348 | 737  | 
done  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
738  | 
|
| 63489 | 739  | 
lemma frac_eq: "frac x = x \<longleftrightarrow> 0 \<le> x \<and> x < 1"  | 
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
740  | 
by (simp add: frac_unique_iff)  | 
| 63331 | 741  | 
|
| 63489 | 742  | 
lemma frac_neg: "frac (- x) = (if x \<in> \<int> then 0 else 1 - frac x)"  | 
743  | 
for x :: "'a::floor_ceiling"  | 
|
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
744  | 
apply (auto simp add: frac_unique_iff)  | 
| 63489 | 745  | 
apply (simp add: frac_def)  | 
746  | 
apply (meson frac_lt_1 less_iff_diff_less_0 not_le not_less_iff_gr_or_eq)  | 
|
747  | 
done  | 
|
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
58889 
diff
changeset
 | 
748  | 
|
| 68721 | 749  | 
lemma frac_in_Ints_iff [simp]: "frac x \<in> \<int> \<longleftrightarrow> x \<in> \<int>"  | 
750  | 
proof safe  | 
|
751  | 
assume "frac x \<in> \<int>"  | 
|
752  | 
hence "of_int \<lfloor>x\<rfloor> + frac x \<in> \<int>" by auto  | 
|
753  | 
also have "of_int \<lfloor>x\<rfloor> + frac x = x" by (simp add: frac_def)  | 
|
754  | 
finally show "x \<in> \<int>" .  | 
|
755  | 
qed (auto simp: frac_def)  | 
|
756  | 
||
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
757  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
758  | 
subsection \<open>Rounding to the nearest integer\<close>  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
759  | 
|
| 63489 | 760  | 
definition round :: "'a::floor_ceiling \<Rightarrow> int"  | 
761  | 
where "round x = \<lfloor>x + 1/2\<rfloor>"  | 
|
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
762  | 
|
| 63489 | 763  | 
lemma of_int_round_ge: "of_int (round x) \<ge> x - 1/2"  | 
764  | 
and of_int_round_le: "of_int (round x) \<le> x + 1/2"  | 
|
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
765  | 
and of_int_round_abs_le: "\<bar>of_int (round x) - x\<bar> \<le> 1/2"  | 
| 63489 | 766  | 
and of_int_round_gt: "of_int (round x) > x - 1/2"  | 
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
767  | 
proof -  | 
| 63489 | 768  | 
from floor_correct[of "x + 1/2"] have "x + 1/2 < of_int (round x) + 1"  | 
769  | 
by (simp add: round_def)  | 
|
770  | 
from add_strict_right_mono[OF this, of "-1"] show A: "of_int (round x) > x - 1/2"  | 
|
771  | 
by simp  | 
|
772  | 
then show "of_int (round x) \<ge> x - 1/2"  | 
|
773  | 
by simp  | 
|
774  | 
from floor_correct[of "x + 1/2"] show "of_int (round x) \<le> x + 1/2"  | 
|
775  | 
by (simp add: round_def)  | 
|
776  | 
with A show "\<bar>of_int (round x) - x\<bar> \<le> 1/2"  | 
|
777  | 
by linarith  | 
|
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
778  | 
qed  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
779  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
780  | 
lemma round_of_int [simp]: "round (of_int n) = n"  | 
| 66515 | 781  | 
unfolding round_def by (subst floor_eq_iff) force  | 
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
782  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
783  | 
lemma round_0 [simp]: "round 0 = 0"  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
784  | 
using round_of_int[of 0] by simp  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
785  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
786  | 
lemma round_1 [simp]: "round 1 = 1"  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
787  | 
using round_of_int[of 1] by simp  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
788  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
789  | 
lemma round_numeral [simp]: "round (numeral n) = numeral n"  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
790  | 
using round_of_int[of "numeral n"] by simp  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
791  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
792  | 
lemma round_neg_numeral [simp]: "round (-numeral n) = -numeral n"  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
793  | 
using round_of_int[of "-numeral n"] by simp  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
794  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
795  | 
lemma round_of_nat [simp]: "round (of_nat n) = of_nat n"  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
796  | 
using round_of_int[of "int n"] by simp  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
797  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
798  | 
lemma round_mono: "x \<le> y \<Longrightarrow> round x \<le> round y"  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
799  | 
unfolding round_def by (intro floor_mono) simp  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
800  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
801  | 
lemma round_unique: "of_int y > x - 1/2 \<Longrightarrow> of_int y \<le> x + 1/2 \<Longrightarrow> round x = y"  | 
| 63489 | 802  | 
unfolding round_def  | 
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
803  | 
proof (rule floor_unique)  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
804  | 
assume "x - 1 / 2 < of_int y"  | 
| 63489 | 805  | 
from add_strict_left_mono[OF this, of 1] show "x + 1 / 2 < of_int y + 1"  | 
806  | 
by simp  | 
|
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
807  | 
qed  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
808  | 
|
| 64317 | 809  | 
lemma round_unique': "\<bar>x - of_int n\<bar> < 1/2 \<Longrightarrow> round x = n"  | 
810  | 
by (subst (asm) abs_less_iff, rule round_unique) (simp_all add: field_simps)  | 
|
811  | 
||
| 61942 | 812  | 
lemma round_altdef: "round x = (if frac x \<ge> 1/2 then \<lceil>x\<rceil> else \<lfloor>x\<rfloor>)"  | 
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
813  | 
by (cases "frac x \<ge> 1/2")  | 
| 63489 | 814  | 
(rule round_unique, ((simp add: frac_def field_simps ceiling_altdef; linarith)+)[2])+  | 
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
815  | 
|
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
816  | 
lemma floor_le_round: "\<lfloor>x\<rfloor> \<le> round x"  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
817  | 
unfolding round_def by (intro floor_mono) simp  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
818  | 
|
| 63489 | 819  | 
lemma ceiling_ge_round: "\<lceil>x\<rceil> \<ge> round x"  | 
820  | 
unfolding round_altdef by simp  | 
|
| 63331 | 821  | 
|
| 63489 | 822  | 
lemma round_diff_minimal: "\<bar>z - of_int (round z)\<bar> \<le> \<bar>z - of_int m\<bar>"  | 
823  | 
for z :: "'a::floor_ceiling"  | 
|
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
824  | 
proof (cases "of_int m \<ge> z")  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
825  | 
case True  | 
| 63489 | 826  | 
then have "\<bar>z - of_int (round z)\<bar> \<le> \<bar>of_int \<lceil>z\<rceil> - z\<bar>"  | 
827  | 
unfolding round_altdef by (simp add: field_simps ceiling_altdef frac_def) linarith  | 
|
828  | 
also have "of_int \<lceil>z\<rceil> - z \<ge> 0"  | 
|
829  | 
by linarith  | 
|
| 61942 | 830  | 
with True have "\<bar>of_int \<lceil>z\<rceil> - z\<bar> \<le> \<bar>z - of_int m\<bar>"  | 
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
831  | 
by (simp add: ceiling_le_iff)  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
832  | 
finally show ?thesis .  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
833  | 
next  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
834  | 
case False  | 
| 63489 | 835  | 
then have "\<bar>z - of_int (round z)\<bar> \<le> \<bar>of_int \<lfloor>z\<rfloor> - z\<bar>"  | 
836  | 
unfolding round_altdef by (simp add: field_simps ceiling_altdef frac_def) linarith  | 
|
837  | 
also have "z - of_int \<lfloor>z\<rfloor> \<ge> 0"  | 
|
838  | 
by linarith  | 
|
| 61942 | 839  | 
with False have "\<bar>of_int \<lfloor>z\<rfloor> - z\<bar> \<le> \<bar>z - of_int m\<bar>"  | 
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
840  | 
by (simp add: le_floor_iff)  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
841  | 
finally show ?thesis .  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
842  | 
qed  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
843  | 
|
| 30096 | 844  | 
end  |