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