| author | wenzelm | 
| Tue, 02 May 2017 18:27:51 +0200 | |
| changeset 65686 | 4a762cad298f | 
| parent 64317 | 029e6247210e | 
| child 66154 | bc5e6461f759 | 
| 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: 
43733diff
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: 
62348diff
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: 
62348diff
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: 
62348diff
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: 
62348diff
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: 
62348diff
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: 
43704diff
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 | ||
| 63489 | 227 | lemma floor_unique_iff: "\<lfloor>x\<rfloor> = a \<longleftrightarrow> of_int a \<le> x \<and> x < of_int a + 1" | 
| 228 | for x :: "'a::floor_ceiling" | |
| 229 | 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 | 230 | |
| 61942 | 231 | lemma of_int_floor_le [simp]: "of_int \<lfloor>x\<rfloor> \<le> x" | 
| 30096 | 232 | using floor_correct .. | 
| 233 | ||
| 61942 | 234 | lemma le_floor_iff: "z \<le> \<lfloor>x\<rfloor> \<longleftrightarrow> of_int z \<le> x" | 
| 30096 | 235 | proof | 
| 61942 | 236 | assume "z \<le> \<lfloor>x\<rfloor>" | 
| 237 | then have "(of_int z :: 'a) \<le> of_int \<lfloor>x\<rfloor>" by simp | |
| 238 | also have "of_int \<lfloor>x\<rfloor> \<le> x" by (rule of_int_floor_le) | |
| 30096 | 239 | finally show "of_int z \<le> x" . | 
| 240 | next | |
| 241 | assume "of_int z \<le> x" | |
| 61942 | 242 | also have "x < of_int (\<lfloor>x\<rfloor> + 1)" using floor_correct .. | 
| 243 | finally show "z \<le> \<lfloor>x\<rfloor>" by (simp del: of_int_add) | |
| 30096 | 244 | qed | 
| 245 | ||
| 61942 | 246 | lemma floor_less_iff: "\<lfloor>x\<rfloor> < z \<longleftrightarrow> x < of_int z" | 
| 30096 | 247 | by (simp add: not_le [symmetric] le_floor_iff) | 
| 248 | ||
| 61942 | 249 | lemma less_floor_iff: "z < \<lfloor>x\<rfloor> \<longleftrightarrow> of_int z + 1 \<le> x" | 
| 30096 | 250 | using le_floor_iff [of "z + 1" x] by auto | 
| 251 | ||
| 61942 | 252 | lemma floor_le_iff: "\<lfloor>x\<rfloor> \<le> z \<longleftrightarrow> x < of_int z + 1" | 
| 30096 | 253 | by (simp add: not_less [symmetric] less_floor_iff) | 
| 254 | ||
| 61942 | 255 | 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 | 256 | 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 | 257 | |
| 61942 | 258 | lemma floor_mono: | 
| 259 | assumes "x \<le> y" | |
| 260 | shows "\<lfloor>x\<rfloor> \<le> \<lfloor>y\<rfloor>" | |
| 30096 | 261 | proof - | 
| 61942 | 262 | have "of_int \<lfloor>x\<rfloor> \<le> x" by (rule of_int_floor_le) | 
| 60758 | 263 | also note \<open>x \<le> y\<close> | 
| 30096 | 264 | finally show ?thesis by (simp add: le_floor_iff) | 
| 265 | qed | |
| 266 | ||
| 61942 | 267 | lemma floor_less_cancel: "\<lfloor>x\<rfloor> < \<lfloor>y\<rfloor> \<Longrightarrow> x < y" | 
| 30096 | 268 | by (auto simp add: not_le [symmetric] floor_mono) | 
| 269 | ||
| 61942 | 270 | lemma floor_of_int [simp]: "\<lfloor>of_int z\<rfloor> = z" | 
| 30096 | 271 | by (rule floor_unique) simp_all | 
| 272 | ||
| 61942 | 273 | lemma floor_of_nat [simp]: "\<lfloor>of_nat n\<rfloor> = int n" | 
| 30096 | 274 | using floor_of_int [of "of_nat n"] by simp | 
| 275 | ||
| 61942 | 276 | 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 | 277 | 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 | 278 | |
| 63489 | 279 | |
| 280 | text \<open>Floor with numerals.\<close> | |
| 30096 | 281 | |
| 61942 | 282 | lemma floor_zero [simp]: "\<lfloor>0\<rfloor> = 0" | 
| 30096 | 283 | using floor_of_int [of 0] by simp | 
| 284 | ||
| 61942 | 285 | lemma floor_one [simp]: "\<lfloor>1\<rfloor> = 1" | 
| 30096 | 286 | using floor_of_int [of 1] by simp | 
| 287 | ||
| 61942 | 288 | lemma floor_numeral [simp]: "\<lfloor>numeral v\<rfloor> = numeral v" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 289 | using floor_of_int [of "numeral v"] by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 290 | |
| 61942 | 291 | lemma floor_neg_numeral [simp]: "\<lfloor>- numeral v\<rfloor> = - numeral v" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54281diff
changeset | 292 | using floor_of_int [of "- numeral v"] by simp | 
| 30096 | 293 | |
| 61942 | 294 | lemma zero_le_floor [simp]: "0 \<le> \<lfloor>x\<rfloor> \<longleftrightarrow> 0 \<le> x" | 
| 30096 | 295 | by (simp add: le_floor_iff) | 
| 296 | ||
| 61942 | 297 | lemma one_le_floor [simp]: "1 \<le> \<lfloor>x\<rfloor> \<longleftrightarrow> 1 \<le> x" | 
| 30096 | 298 | by (simp add: le_floor_iff) | 
| 299 | ||
| 63489 | 300 | 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 | 301 | by (simp add: le_floor_iff) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 302 | |
| 63489 | 303 | lemma neg_numeral_le_floor [simp]: "- numeral v \<le> \<lfloor>x\<rfloor> \<longleftrightarrow> - numeral v \<le> x" | 
| 30096 | 304 | by (simp add: le_floor_iff) | 
| 305 | ||
| 61942 | 306 | lemma zero_less_floor [simp]: "0 < \<lfloor>x\<rfloor> \<longleftrightarrow> 1 \<le> x" | 
| 30096 | 307 | by (simp add: less_floor_iff) | 
| 308 | ||
| 61942 | 309 | lemma one_less_floor [simp]: "1 < \<lfloor>x\<rfloor> \<longleftrightarrow> 2 \<le> x" | 
| 30096 | 310 | by (simp add: less_floor_iff) | 
| 311 | ||
| 63489 | 312 | 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 | 313 | by (simp add: less_floor_iff) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 314 | |
| 63489 | 315 | lemma neg_numeral_less_floor [simp]: "- numeral v < \<lfloor>x\<rfloor> \<longleftrightarrow> - numeral v + 1 \<le> x" | 
| 30096 | 316 | by (simp add: less_floor_iff) | 
| 317 | ||
| 61942 | 318 | lemma floor_le_zero [simp]: "\<lfloor>x\<rfloor> \<le> 0 \<longleftrightarrow> x < 1" | 
| 30096 | 319 | by (simp add: floor_le_iff) | 
| 320 | ||
| 61942 | 321 | lemma floor_le_one [simp]: "\<lfloor>x\<rfloor> \<le> 1 \<longleftrightarrow> x < 2" | 
| 30096 | 322 | by (simp add: floor_le_iff) | 
| 323 | ||
| 63489 | 324 | 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 | 325 | by (simp add: floor_le_iff) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 326 | |
| 63489 | 327 | lemma floor_le_neg_numeral [simp]: "\<lfloor>x\<rfloor> \<le> - numeral v \<longleftrightarrow> x < - numeral v + 1" | 
| 30096 | 328 | by (simp add: floor_le_iff) | 
| 329 | ||
| 61942 | 330 | lemma floor_less_zero [simp]: "\<lfloor>x\<rfloor> < 0 \<longleftrightarrow> x < 0" | 
| 30096 | 331 | by (simp add: floor_less_iff) | 
| 332 | ||
| 61942 | 333 | lemma floor_less_one [simp]: "\<lfloor>x\<rfloor> < 1 \<longleftrightarrow> x < 1" | 
| 30096 | 334 | by (simp add: floor_less_iff) | 
| 335 | ||
| 63489 | 336 | 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 | 337 | by (simp add: floor_less_iff) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 338 | |
| 63489 | 339 | lemma floor_less_neg_numeral [simp]: "\<lfloor>x\<rfloor> < - numeral v \<longleftrightarrow> x < - numeral v" | 
| 30096 | 340 | by (simp add: floor_less_iff) | 
| 341 | ||
| 63489 | 342 | |
| 343 | text \<open>Addition and subtraction of integers.\<close> | |
| 30096 | 344 | |
| 63599 | 345 | lemma floor_add_int: "\<lfloor>x\<rfloor> + z = \<lfloor>x + of_int z\<rfloor>" | 
| 346 | using floor_correct [of x] by (simp add: floor_unique[symmetric]) | |
| 30096 | 347 | |
| 63599 | 348 | lemma int_add_floor: "z + \<lfloor>x\<rfloor> = \<lfloor>of_int z + x\<rfloor>" | 
| 349 | 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 | 350 | |
| 63599 | 351 | lemma one_add_floor: "\<lfloor>x\<rfloor> + 1 = \<lfloor>x + 1\<rfloor>" | 
| 352 | using floor_add_int [of x 1] by simp | |
| 30096 | 353 | |
| 61942 | 354 | lemma floor_diff_of_int [simp]: "\<lfloor>x - of_int z\<rfloor> = \<lfloor>x\<rfloor> - z" | 
| 63599 | 355 | using floor_add_int [of x "- z"] by (simp add: algebra_simps) | 
| 30096 | 356 | |
| 61942 | 357 | 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 | 358 | 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 | 359 | |
| 63489 | 360 | 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 | 361 | using floor_diff_of_int [of x "numeral v"] by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 362 | |
| 61942 | 363 | lemma floor_diff_one [simp]: "\<lfloor>x - 1\<rfloor> = \<lfloor>x\<rfloor> - 1" | 
| 30096 | 364 | using floor_diff_of_int [of x 1] by simp | 
| 365 | ||
| 58097 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 hoelzl parents: 
58040diff
changeset | 366 | lemma le_mult_floor: | 
| 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 hoelzl parents: 
58040diff
changeset | 367 | assumes "0 \<le> a" and "0 \<le> b" | 
| 61942 | 368 | 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 | 369 | proof - | 
| 63489 | 370 | have "of_int \<lfloor>a\<rfloor> \<le> a" and "of_int \<lfloor>b\<rfloor> \<le> b" | 
| 371 | by (auto intro: of_int_floor_le) | |
| 372 | 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 | 373 | using assms by (auto intro!: mult_mono) | 
| 61942 | 374 | 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 | 375 | using floor_correct[of "a * b"] by auto | 
| 63489 | 376 | finally show ?thesis | 
| 377 | unfolding of_int_less_iff by simp | |
| 58097 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 hoelzl parents: 
58040diff
changeset | 378 | qed | 
| 
cfd3cff9387b
add simp rules for divisions of numerals in floor and ceiling.
 hoelzl parents: 
58040diff
changeset | 379 | |
| 63489 | 380 | lemma floor_divide_of_int_eq: "\<lfloor>of_int k / of_int l\<rfloor> = k div l" | 
| 381 | for k l :: int | |
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 382 | proof (cases "l = 0") | 
| 63489 | 383 | case True | 
| 384 | then show ?thesis by simp | |
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 385 | next | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 386 | case False | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 387 | 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 | 388 | proof (cases "l > 0") | 
| 63489 | 389 | case True | 
| 390 | then show ?thesis | |
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 391 | by (auto intro: floor_unique) | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 392 | next | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 393 | case False | 
| 63489 | 394 | obtain r where "r = - l" | 
| 395 | by blast | |
| 396 | then have l: "l = - r" | |
| 397 | by simp | |
| 63540 | 398 | with \<open>l \<noteq> 0\<close> False have "r > 0" | 
| 63489 | 399 | by simp | 
| 63540 | 400 | with l show ?thesis | 
| 63489 | 401 | using pos_mod_bound [of r] | 
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 402 | 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 | 403 | qed | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 404 | 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 | 405 | by simp | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 406 | 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 | 407 | 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 | 408 | finally have "(of_int k / of_int l :: 'a) = \<dots> / of_int l" | 
| 63331 | 409 | by simp | 
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 410 | 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 | 411 | using False by (simp only:) (simp add: field_simps) | 
| 63331 | 412 | 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 | 413 | by simp | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 414 | 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 | 415 | by (simp add: ac_simps) | 
| 60128 | 416 | 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 | 417 | by (simp add: floor_add_int) | 
| 63489 | 418 | with * show ?thesis | 
| 419 | by simp | |
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 420 | qed | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 421 | |
| 63489 | 422 | lemma floor_divide_of_nat_eq: "\<lfloor>of_nat m / of_nat n\<rfloor> = of_nat (m div n)" | 
| 423 | for m n :: nat | |
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 424 | proof (cases "n = 0") | 
| 63489 | 425 | case True | 
| 426 | then show ?thesis by simp | |
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 427 | next | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 428 | case False | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 429 | 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 | 430 | by (auto intro: floor_unique) | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 431 | 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 | 432 | by simp | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 433 | also have "\<dots> = (of_nat (m div n) + of_nat (m mod n) / of_nat n) * of_nat n" | 
| 63489 | 434 | 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 | 435 | finally have "(of_nat m / of_nat n :: 'a) = \<dots> / of_nat n" | 
| 63331 | 436 | by simp | 
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 437 | 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 | 438 | using False by (simp only:) simp | 
| 63331 | 439 | 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 | 440 | by simp | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 441 | 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 | 442 | by (simp add: ac_simps) | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 443 | 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 | 444 | by simp | 
| 63489 | 445 | ultimately have "\<lfloor>of_nat m / of_nat n :: 'a\<rfloor> = | 
| 446 | \<lfloor>of_nat (m mod n) / of_nat n :: 'a\<rfloor> + of_nat (m div n)" | |
| 63599 | 447 | by (simp only: floor_add_int) | 
| 63489 | 448 | with * show ?thesis | 
| 449 | by simp | |
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 450 | qed | 
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 451 | |
| 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59613diff
changeset | 452 | |
| 60758 | 453 | subsection \<open>Ceiling function\<close> | 
| 30096 | 454 | |
| 61942 | 455 | definition ceiling :: "'a::floor_ceiling \<Rightarrow> int"  ("\<lceil>_\<rceil>")
 | 
| 456 | where "\<lceil>x\<rceil> = - \<lfloor>- x\<rfloor>" | |
| 30096 | 457 | |
| 61942 | 458 | lemma ceiling_correct: "of_int \<lceil>x\<rceil> - 1 < x \<and> x \<le> of_int \<lceil>x\<rceil>" | 
| 63489 | 459 | unfolding ceiling_def using floor_correct [of "- x"] | 
| 460 | by (simp add: le_minus_iff) | |
| 30096 | 461 | |
| 63489 | 462 | lemma ceiling_unique: "of_int z - 1 < x \<Longrightarrow> x \<le> of_int z \<Longrightarrow> \<lceil>x\<rceil> = z" | 
| 30096 | 463 | unfolding ceiling_def using floor_unique [of "- z" "- x"] by simp | 
| 464 | ||
| 61942 | 465 | lemma le_of_int_ceiling [simp]: "x \<le> of_int \<lceil>x\<rceil>" | 
| 30096 | 466 | using ceiling_correct .. | 
| 467 | ||
| 61942 | 468 | lemma ceiling_le_iff: "\<lceil>x\<rceil> \<le> z \<longleftrightarrow> x \<le> of_int z" | 
| 30096 | 469 | unfolding ceiling_def using le_floor_iff [of "- z" "- x"] by auto | 
| 470 | ||
| 61942 | 471 | lemma less_ceiling_iff: "z < \<lceil>x\<rceil> \<longleftrightarrow> of_int z < x" | 
| 30096 | 472 | by (simp add: not_le [symmetric] ceiling_le_iff) | 
| 473 | ||
| 61942 | 474 | lemma ceiling_less_iff: "\<lceil>x\<rceil> < z \<longleftrightarrow> x \<le> of_int z - 1" | 
| 30096 | 475 | using ceiling_le_iff [of x "z - 1"] by simp | 
| 476 | ||
| 61942 | 477 | lemma le_ceiling_iff: "z \<le> \<lceil>x\<rceil> \<longleftrightarrow> of_int z - 1 < x" | 
| 30096 | 478 | by (simp add: not_less [symmetric] ceiling_less_iff) | 
| 479 | ||
| 61942 | 480 | lemma ceiling_mono: "x \<ge> y \<Longrightarrow> \<lceil>x\<rceil> \<ge> \<lceil>y\<rceil>" | 
| 30096 | 481 | unfolding ceiling_def by (simp add: floor_mono) | 
| 482 | ||
| 61942 | 483 | lemma ceiling_less_cancel: "\<lceil>x\<rceil> < \<lceil>y\<rceil> \<Longrightarrow> x < y" | 
| 30096 | 484 | by (auto simp add: not_le [symmetric] ceiling_mono) | 
| 485 | ||
| 61942 | 486 | lemma ceiling_of_int [simp]: "\<lceil>of_int z\<rceil> = z" | 
| 30096 | 487 | by (rule ceiling_unique) simp_all | 
| 488 | ||
| 61942 | 489 | lemma ceiling_of_nat [simp]: "\<lceil>of_nat n\<rceil> = int n" | 
| 30096 | 490 | using ceiling_of_int [of "of_nat n"] by simp | 
| 491 | ||
| 61942 | 492 | 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 | 493 | 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 | 494 | |
| 63879 
15bbf6360339
simple new lemmas, mostly about sets
 paulson <lp15@cam.ac.uk> parents: 
63621diff
changeset | 495 | lemma finite_int_segment: | 
| 
15bbf6360339
simple new lemmas, mostly about sets
 paulson <lp15@cam.ac.uk> parents: 
63621diff
changeset | 496 | fixes a :: "'a::floor_ceiling" | 
| 
15bbf6360339
simple new lemmas, mostly about sets
 paulson <lp15@cam.ac.uk> parents: 
63621diff
changeset | 497 |   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 | 498 | proof - | 
| 
15bbf6360339
simple new lemmas, mostly about sets
 paulson <lp15@cam.ac.uk> parents: 
63621diff
changeset | 499 |   have "finite {ceiling a..floor b}"
 | 
| 
15bbf6360339
simple new lemmas, mostly about sets
 paulson <lp15@cam.ac.uk> parents: 
63621diff
changeset | 500 | by simp | 
| 
15bbf6360339
simple new lemmas, mostly about sets
 paulson <lp15@cam.ac.uk> parents: 
63621diff
changeset | 501 |   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 | 502 | 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 | 503 | ultimately show ?thesis | 
| 
15bbf6360339
simple new lemmas, mostly about sets
 paulson <lp15@cam.ac.uk> parents: 
63621diff
changeset | 504 | by simp | 
| 
15bbf6360339
simple new lemmas, mostly about sets
 paulson <lp15@cam.ac.uk> parents: 
63621diff
changeset | 505 | qed | 
| 
15bbf6360339
simple new lemmas, mostly about sets
 paulson <lp15@cam.ac.uk> parents: 
63621diff
changeset | 506 | |
| 63489 | 507 | |
| 508 | text \<open>Ceiling with numerals.\<close> | |
| 30096 | 509 | |
| 61942 | 510 | lemma ceiling_zero [simp]: "\<lceil>0\<rceil> = 0" | 
| 30096 | 511 | using ceiling_of_int [of 0] by simp | 
| 512 | ||
| 61942 | 513 | lemma ceiling_one [simp]: "\<lceil>1\<rceil> = 1" | 
| 30096 | 514 | using ceiling_of_int [of 1] by simp | 
| 515 | ||
| 61942 | 516 | lemma ceiling_numeral [simp]: "\<lceil>numeral v\<rceil> = numeral v" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 517 | using ceiling_of_int [of "numeral v"] by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 518 | |
| 61942 | 519 | lemma ceiling_neg_numeral [simp]: "\<lceil>- numeral v\<rceil> = - numeral v" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54281diff
changeset | 520 | using ceiling_of_int [of "- numeral v"] by simp | 
| 30096 | 521 | |
| 61942 | 522 | lemma ceiling_le_zero [simp]: "\<lceil>x\<rceil> \<le> 0 \<longleftrightarrow> x \<le> 0" | 
| 30096 | 523 | by (simp add: ceiling_le_iff) | 
| 524 | ||
| 61942 | 525 | lemma ceiling_le_one [simp]: "\<lceil>x\<rceil> \<le> 1 \<longleftrightarrow> x \<le> 1" | 
| 30096 | 526 | by (simp add: ceiling_le_iff) | 
| 527 | ||
| 63489 | 528 | 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 | 529 | by (simp add: ceiling_le_iff) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 530 | |
| 63489 | 531 | lemma ceiling_le_neg_numeral [simp]: "\<lceil>x\<rceil> \<le> - numeral v \<longleftrightarrow> x \<le> - numeral v" | 
| 30096 | 532 | by (simp add: ceiling_le_iff) | 
| 533 | ||
| 61942 | 534 | lemma ceiling_less_zero [simp]: "\<lceil>x\<rceil> < 0 \<longleftrightarrow> x \<le> -1" | 
| 30096 | 535 | by (simp add: ceiling_less_iff) | 
| 536 | ||
| 61942 | 537 | lemma ceiling_less_one [simp]: "\<lceil>x\<rceil> < 1 \<longleftrightarrow> x \<le> 0" | 
| 30096 | 538 | by (simp add: ceiling_less_iff) | 
| 539 | ||
| 63489 | 540 | 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 | 541 | by (simp add: ceiling_less_iff) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 542 | |
| 63489 | 543 | lemma ceiling_less_neg_numeral [simp]: "\<lceil>x\<rceil> < - numeral v \<longleftrightarrow> x \<le> - numeral v - 1" | 
| 30096 | 544 | by (simp add: ceiling_less_iff) | 
| 545 | ||
| 61942 | 546 | lemma zero_le_ceiling [simp]: "0 \<le> \<lceil>x\<rceil> \<longleftrightarrow> -1 < x" | 
| 30096 | 547 | by (simp add: le_ceiling_iff) | 
| 548 | ||
| 61942 | 549 | lemma one_le_ceiling [simp]: "1 \<le> \<lceil>x\<rceil> \<longleftrightarrow> 0 < x" | 
| 30096 | 550 | by (simp add: le_ceiling_iff) | 
| 551 | ||
| 63489 | 552 | 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 | 553 | by (simp add: le_ceiling_iff) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 554 | |
| 63489 | 555 | lemma neg_numeral_le_ceiling [simp]: "- numeral v \<le> \<lceil>x\<rceil> \<longleftrightarrow> - numeral v - 1 < x" | 
| 30096 | 556 | by (simp add: le_ceiling_iff) | 
| 557 | ||
| 61942 | 558 | lemma zero_less_ceiling [simp]: "0 < \<lceil>x\<rceil> \<longleftrightarrow> 0 < x" | 
| 30096 | 559 | by (simp add: less_ceiling_iff) | 
| 560 | ||
| 61942 | 561 | lemma one_less_ceiling [simp]: "1 < \<lceil>x\<rceil> \<longleftrightarrow> 1 < x" | 
| 30096 | 562 | by (simp add: less_ceiling_iff) | 
| 563 | ||
| 63489 | 564 | 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 | 565 | by (simp add: less_ceiling_iff) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 566 | |
| 63489 | 567 | lemma neg_numeral_less_ceiling [simp]: "- numeral v < \<lceil>x\<rceil> \<longleftrightarrow> - numeral v < x" | 
| 30096 | 568 | by (simp add: less_ceiling_iff) | 
| 569 | ||
| 61942 | 570 | lemma ceiling_altdef: "\<lceil>x\<rceil> = (if x = of_int \<lfloor>x\<rfloor> then \<lfloor>x\<rfloor> else \<lfloor>x\<rfloor> + 1)" | 
| 63489 | 571 | by (intro ceiling_unique; simp, linarith?) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 572 | |
| 61942 | 573 | lemma floor_le_ceiling [simp]: "\<lfloor>x\<rfloor> \<le> \<lceil>x\<rceil>" | 
| 574 | by (simp add: ceiling_altdef) | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 575 | |
| 63489 | 576 | |
| 577 | text \<open>Addition and subtraction of integers.\<close> | |
| 30096 | 578 | |
| 61942 | 579 | 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 | 580 | using ceiling_correct [of x] by (simp add: ceiling_def) | 
| 30096 | 581 | |
| 61942 | 582 | 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 | 583 | using ceiling_add_of_int [of x "numeral v"] by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 584 | |
| 61942 | 585 | lemma ceiling_add_one [simp]: "\<lceil>x + 1\<rceil> = \<lceil>x\<rceil> + 1" | 
| 30096 | 586 | using ceiling_add_of_int [of x 1] by simp | 
| 587 | ||
| 61942 | 588 | lemma ceiling_diff_of_int [simp]: "\<lceil>x - of_int z\<rceil> = \<lceil>x\<rceil> - z" | 
| 30096 | 589 | using ceiling_add_of_int [of x "- z"] by (simp add: algebra_simps) | 
| 590 | ||
| 61942 | 591 | 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 | 592 | using ceiling_diff_of_int [of x "numeral v"] by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43733diff
changeset | 593 | |
| 61942 | 594 | lemma ceiling_diff_one [simp]: "\<lceil>x - 1\<rceil> = \<lceil>x\<rceil> - 1" | 
| 30096 | 595 | using ceiling_diff_of_int [of x 1] by simp | 
| 596 | ||
| 61942 | 597 | 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 | 598 | by (auto simp add: ceiling_unique ceiling_correct) | 
| 
9a867afaab5a
better linarith support for floor, ceiling, natfloor, and natceiling
 hoelzl parents: 
54489diff
changeset | 599 | |
| 61942 | 600 | lemma ceiling_diff_floor_le_1: "\<lceil>x\<rceil> - \<lfloor>x\<rfloor> \<le> 1" | 
| 47592 | 601 | proof - | 
| 63331 | 602 | have "of_int \<lceil>x\<rceil> - 1 < x" | 
| 47592 | 603 | using ceiling_correct[of x] by simp | 
| 604 | also have "x < of_int \<lfloor>x\<rfloor> + 1" | |
| 605 | using floor_correct[of x] by simp_all | |
| 606 | finally have "of_int (\<lceil>x\<rceil> - \<lfloor>x\<rfloor>) < (of_int 2::'a)" | |
| 607 | by simp | |
| 608 | then show ?thesis | |
| 609 | unfolding of_int_less_iff by simp | |
| 610 | qed | |
| 30096 | 611 | |
| 63489 | 612 | |
| 60758 | 613 | subsection \<open>Negation\<close> | 
| 30096 | 614 | |
| 61942 | 615 | lemma floor_minus: "\<lfloor>- x\<rfloor> = - \<lceil>x\<rceil>" | 
| 30096 | 616 | unfolding ceiling_def by simp | 
| 617 | ||
| 61942 | 618 | lemma ceiling_minus: "\<lceil>- x\<rceil> = - \<lfloor>x\<rfloor>" | 
| 30096 | 619 | unfolding ceiling_def by simp | 
| 620 | ||
| 61942 | 621 | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63879diff
changeset | 622 | subsection \<open>Natural numbers\<close> | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63879diff
changeset | 623 | |
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63879diff
changeset | 624 | 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 | 625 | by simp | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63879diff
changeset | 626 | |
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63879diff
changeset | 627 | 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 | 628 | by (cases "r\<ge>0") auto | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63879diff
changeset | 629 | |
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63879diff
changeset | 630 | |
| 60758 | 631 | 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 | 632 | |
| 63489 | 633 | definition frac :: "'a \<Rightarrow> 'a::floor_ceiling" | 
| 634 | 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 | 635 | |
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 636 | lemma frac_lt_1: "frac x < 1" | 
| 63489 | 637 | 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 | 638 | |
| 61070 | 639 | 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 | 640 | 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 | 641 | |
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 642 | lemma frac_ge_0 [simp]: "frac x \<ge> 0" | 
| 63489 | 643 | 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 | 644 | |
| 61070 | 645 | 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 | 646 | 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 | 647 | |
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 648 | 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 | 649 | 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 | 650 | |
| 63331 | 651 | 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 | 652 | proof - | 
| 63599 | 653 | 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>" | 
| 654 | 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 | 655 | moreover | 
| 63599 | 656 | 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>)" | 
| 63489 | 657 | apply (simp add: floor_unique_iff) | 
| 658 | apply (auto simp add: algebra_simps) | |
| 659 | apply linarith | |
| 660 | done | |
| 63599 | 661 | 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 | 662 | qed | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 663 | |
| 63621 | 664 | lemma floor_add2[simp]: "x \<in> \<int> \<or> y \<in> \<int> \<Longrightarrow> \<lfloor>x + y\<rfloor> = \<lfloor>x\<rfloor> + \<lfloor>y\<rfloor>" | 
| 665 | by (metis add.commute add.left_neutral frac_lt_1 floor_add frac_eq_0_iff) | |
| 63597 | 666 | |
| 63489 | 667 | lemma frac_add: | 
| 668 | "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 | 669 | 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 | 670 | |
| 63489 | 671 | lemma frac_unique_iff: "frac x = a \<longleftrightarrow> x - a \<in> \<int> \<and> 0 \<le> a \<and> a < 1" | 
| 672 | for x :: "'a::floor_ceiling" | |
| 62348 | 673 | apply (auto simp: Ints_def frac_def algebra_simps floor_unique) | 
| 63489 | 674 | apply linarith+ | 
| 62348 | 675 | done | 
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 676 | |
| 63489 | 677 | 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 | 678 | by (simp add: frac_unique_iff) | 
| 63331 | 679 | |
| 63489 | 680 | lemma frac_neg: "frac (- x) = (if x \<in> \<int> then 0 else 1 - frac x)" | 
| 681 | 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 | 682 | apply (auto simp add: frac_unique_iff) | 
| 63489 | 683 | apply (simp add: frac_def) | 
| 684 | apply (meson frac_lt_1 less_iff_diff_less_0 not_le not_less_iff_gr_or_eq) | |
| 685 | done | |
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 686 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 687 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 688 | subsection \<open>Rounding to the nearest integer\<close> | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 689 | |
| 63489 | 690 | definition round :: "'a::floor_ceiling \<Rightarrow> int" | 
| 691 | where "round x = \<lfloor>x + 1/2\<rfloor>" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 692 | |
| 63489 | 693 | lemma of_int_round_ge: "of_int (round x) \<ge> x - 1/2" | 
| 694 | 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 | 695 | and of_int_round_abs_le: "\<bar>of_int (round x) - x\<bar> \<le> 1/2" | 
| 63489 | 696 | 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 | 697 | proof - | 
| 63489 | 698 | from floor_correct[of "x + 1/2"] have "x + 1/2 < of_int (round x) + 1" | 
| 699 | by (simp add: round_def) | |
| 700 | from add_strict_right_mono[OF this, of "-1"] show A: "of_int (round x) > x - 1/2" | |
| 701 | by simp | |
| 702 | then show "of_int (round x) \<ge> x - 1/2" | |
| 703 | by simp | |
| 704 | from floor_correct[of "x + 1/2"] show "of_int (round x) \<le> x + 1/2" | |
| 705 | by (simp add: round_def) | |
| 706 | with A show "\<bar>of_int (round x) - x\<bar> \<le> 1/2" | |
| 707 | by linarith | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 708 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 709 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 710 | lemma round_of_int [simp]: "round (of_int n) = n" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 711 | unfolding round_def by (subst floor_unique_iff) force | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 712 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 713 | lemma round_0 [simp]: "round 0 = 0" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 714 | using round_of_int[of 0] by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 715 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 716 | lemma round_1 [simp]: "round 1 = 1" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 717 | using round_of_int[of 1] by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 718 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 719 | lemma round_numeral [simp]: "round (numeral n) = numeral n" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 720 | using round_of_int[of "numeral n"] by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 721 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 722 | lemma round_neg_numeral [simp]: "round (-numeral n) = -numeral n" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 723 | using round_of_int[of "-numeral n"] by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 724 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 725 | lemma round_of_nat [simp]: "round (of_nat n) = of_nat n" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 726 | using round_of_int[of "int n"] by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 727 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 728 | lemma round_mono: "x \<le> y \<Longrightarrow> round x \<le> round y" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 729 | unfolding round_def by (intro floor_mono) simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 730 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 731 | lemma round_unique: "of_int y > x - 1/2 \<Longrightarrow> of_int y \<le> x + 1/2 \<Longrightarrow> round x = y" | 
| 63489 | 732 | unfolding round_def | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 733 | proof (rule floor_unique) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 734 | assume "x - 1 / 2 < of_int y" | 
| 63489 | 735 | from add_strict_left_mono[OF this, of 1] show "x + 1 / 2 < of_int y + 1" | 
| 736 | by simp | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 737 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 738 | |
| 64317 | 739 | lemma round_unique': "\<bar>x - of_int n\<bar> < 1/2 \<Longrightarrow> round x = n" | 
| 740 | by (subst (asm) abs_less_iff, rule round_unique) (simp_all add: field_simps) | |
| 741 | ||
| 61942 | 742 | 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 | 743 | by (cases "frac x \<ge> 1/2") | 
| 63489 | 744 | (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 | 745 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 746 | lemma floor_le_round: "\<lfloor>x\<rfloor> \<le> round x" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 747 | unfolding round_def by (intro floor_mono) simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 748 | |
| 63489 | 749 | lemma ceiling_ge_round: "\<lceil>x\<rceil> \<ge> round x" | 
| 750 | unfolding round_altdef by simp | |
| 63331 | 751 | |
| 63489 | 752 | lemma round_diff_minimal: "\<bar>z - of_int (round z)\<bar> \<le> \<bar>z - of_int m\<bar>" | 
| 753 | for z :: "'a::floor_ceiling" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 754 | proof (cases "of_int m \<ge> z") | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 755 | case True | 
| 63489 | 756 | then have "\<bar>z - of_int (round z)\<bar> \<le> \<bar>of_int \<lceil>z\<rceil> - z\<bar>" | 
| 757 | unfolding round_altdef by (simp add: field_simps ceiling_altdef frac_def) linarith | |
| 758 | also have "of_int \<lceil>z\<rceil> - z \<ge> 0" | |
| 759 | by linarith | |
| 61942 | 760 | 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 | 761 | by (simp add: ceiling_le_iff) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 762 | finally show ?thesis . | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 763 | next | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 764 | case False | 
| 63489 | 765 | then have "\<bar>z - of_int (round z)\<bar> \<le> \<bar>of_int \<lfloor>z\<rfloor> - z\<bar>" | 
| 766 | unfolding round_altdef by (simp add: field_simps ceiling_altdef frac_def) linarith | |
| 767 | also have "z - of_int \<lfloor>z\<rfloor> \<ge> 0" | |
| 768 | by linarith | |
| 61942 | 769 | 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 | 770 | by (simp add: le_floor_iff) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 771 | finally show ?thesis . | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 772 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61378diff
changeset | 773 | |
| 30096 | 774 | end |