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