wenzelm@41959: (* Title: HOL/Archimedean_Field.thy wenzelm@41959: Author: Brian Huffman huffman@30096: *) huffman@30096: huffman@30096: header {* Archimedean Fields, Floor and Ceiling Functions *} huffman@30096: huffman@30096: theory Archimedean_Field huffman@30096: imports Main huffman@30096: begin huffman@30096: huffman@30096: subsection {* Class of Archimedean fields *} huffman@30096: huffman@30096: text {* Archimedean fields have no infinite elements. *} huffman@30096: huffman@47108: class archimedean_field = linordered_field + huffman@30096: assumes ex_le_of_int: "\z. x \ of_int z" huffman@30096: huffman@30096: lemma ex_less_of_int: huffman@30096: fixes x :: "'a::archimedean_field" shows "\z. x < of_int z" huffman@30096: proof - huffman@30096: from ex_le_of_int obtain z where "x \ of_int z" .. huffman@30096: then have "x < of_int (z + 1)" by simp huffman@30096: then show ?thesis .. huffman@30096: qed huffman@30096: huffman@30096: lemma ex_of_int_less: huffman@30096: fixes x :: "'a::archimedean_field" shows "\z. of_int z < x" huffman@30096: proof - huffman@30096: from ex_less_of_int obtain z where "- x < of_int z" .. huffman@30096: then have "of_int (- z) < x" by simp huffman@30096: then show ?thesis .. huffman@30096: qed huffman@30096: huffman@30096: lemma ex_less_of_nat: huffman@30096: fixes x :: "'a::archimedean_field" shows "\n. x < of_nat n" huffman@30096: proof - huffman@30096: obtain z where "x < of_int z" using ex_less_of_int .. huffman@30096: also have "\ \ of_int (int (nat z))" by simp huffman@30096: also have "\ = of_nat (nat z)" by (simp only: of_int_of_nat_eq) huffman@30096: finally show ?thesis .. huffman@30096: qed huffman@30096: huffman@30096: lemma ex_le_of_nat: huffman@30096: fixes x :: "'a::archimedean_field" shows "\n. x \ of_nat n" huffman@30096: proof - huffman@30096: obtain n where "x < of_nat n" using ex_less_of_nat .. huffman@30096: then have "x \ of_nat n" by simp huffman@30096: then show ?thesis .. huffman@30096: qed huffman@30096: huffman@30096: text {* Archimedean fields have no infinitesimal elements. *} huffman@30096: huffman@30096: lemma ex_inverse_of_nat_Suc_less: huffman@30096: fixes x :: "'a::archimedean_field" huffman@30096: assumes "0 < x" shows "\n. inverse (of_nat (Suc n)) < x" huffman@30096: proof - huffman@30096: from `0 < x` have "0 < inverse x" huffman@30096: by (rule positive_imp_inverse_positive) huffman@30096: obtain n where "inverse x < of_nat n" huffman@30096: using ex_less_of_nat .. huffman@30096: then obtain m where "inverse x < of_nat (Suc m)" huffman@30096: using `0 < inverse x` by (cases n) (simp_all del: of_nat_Suc) huffman@30096: then have "inverse (of_nat (Suc m)) < inverse (inverse x)" huffman@30096: using `0 < inverse x` by (rule less_imp_inverse_less) huffman@30096: then have "inverse (of_nat (Suc m)) < x" huffman@30096: using `0 < x` by (simp add: nonzero_inverse_inverse_eq) huffman@30096: then show ?thesis .. huffman@30096: qed huffman@30096: huffman@30096: lemma ex_inverse_of_nat_less: huffman@30096: fixes x :: "'a::archimedean_field" huffman@30096: assumes "0 < x" shows "\n>0. inverse (of_nat n) < x" huffman@30096: using ex_inverse_of_nat_Suc_less [OF `0 < x`] by auto huffman@30096: huffman@30096: lemma ex_less_of_nat_mult: huffman@30096: fixes x :: "'a::archimedean_field" huffman@30096: assumes "0 < x" shows "\n. y < of_nat n * x" huffman@30096: proof - huffman@30096: obtain n where "y / x < of_nat n" using ex_less_of_nat .. huffman@30096: with `0 < x` have "y < of_nat n * x" by (simp add: pos_divide_less_eq) huffman@30096: then show ?thesis .. huffman@30096: qed huffman@30096: huffman@30096: huffman@30096: subsection {* Existence and uniqueness of floor function *} huffman@30096: huffman@30096: lemma exists_least_lemma: huffman@30096: assumes "\ P 0" and "\n. P n" huffman@30096: shows "\n. \ P n \ P (Suc n)" huffman@30096: proof - huffman@30096: from `\n. P n` have "P (Least P)" by (rule LeastI_ex) huffman@30096: with `\ P 0` obtain n where "Least P = Suc n" huffman@30096: by (cases "Least P") auto huffman@30096: then have "n < Least P" by simp huffman@30096: then have "\ P n" by (rule not_less_Least) huffman@30096: then have "\ P n \ P (Suc n)" huffman@30096: using `P (Least P)` `Least P = Suc n` by simp huffman@30096: then show ?thesis .. huffman@30096: qed huffman@30096: huffman@30096: lemma floor_exists: huffman@30096: fixes x :: "'a::archimedean_field" huffman@30096: shows "\z. of_int z \ x \ x < of_int (z + 1)" huffman@30096: proof (cases) huffman@30096: assume "0 \ x" huffman@30096: then have "\ x < of_nat 0" by simp huffman@30096: then have "\n. \ x < of_nat n \ x < of_nat (Suc n)" huffman@30096: using ex_less_of_nat by (rule exists_least_lemma) huffman@30096: then obtain n where "\ x < of_nat n \ x < of_nat (Suc n)" .. huffman@30096: then have "of_int (int n) \ x \ x < of_int (int n + 1)" by simp huffman@30096: then show ?thesis .. huffman@30096: next huffman@30096: assume "\ 0 \ x" huffman@30096: then have "\ - x \ of_nat 0" by simp huffman@30096: then have "\n. \ - x \ of_nat n \ - x \ of_nat (Suc n)" huffman@30096: using ex_le_of_nat by (rule exists_least_lemma) huffman@30096: then obtain n where "\ - x \ of_nat n \ - x \ of_nat (Suc n)" .. huffman@30096: then have "of_int (- int n - 1) \ x \ x < of_int (- int n - 1 + 1)" by simp huffman@30096: then show ?thesis .. huffman@30096: qed huffman@30096: huffman@30096: lemma floor_exists1: huffman@30096: fixes x :: "'a::archimedean_field" huffman@30096: shows "\!z. of_int z \ x \ x < of_int (z + 1)" huffman@30096: proof (rule ex_ex1I) huffman@30096: show "\z. of_int z \ x \ x < of_int (z + 1)" huffman@30096: by (rule floor_exists) huffman@30096: next huffman@30096: fix y z assume huffman@30096: "of_int y \ x \ x < of_int (y + 1)" huffman@30096: "of_int z \ x \ x < of_int (z + 1)" huffman@30096: then have huffman@30096: "of_int y \ x" "x < of_int (y + 1)" huffman@30096: "of_int z \ x" "x < of_int (z + 1)" huffman@30096: by simp_all huffman@30096: from le_less_trans [OF `of_int y \ x` `x < of_int (z + 1)`] huffman@30096: le_less_trans [OF `of_int z \ x` `x < of_int (y + 1)`] huffman@30096: show "y = z" by (simp del: of_int_add) huffman@30096: qed huffman@30096: huffman@30096: huffman@30096: subsection {* Floor function *} huffman@30096: bulwahn@43732: class floor_ceiling = archimedean_field + bulwahn@43732: fixes floor :: "'a \ int" bulwahn@43732: assumes floor_correct: "of_int (floor x) \ x \ x < of_int (floor x + 1)" huffman@30096: huffman@30096: notation (xsymbols) huffman@30096: floor ("\_\") huffman@30096: huffman@30096: notation (HTML output) huffman@30096: floor ("\_\") huffman@30096: huffman@30096: lemma floor_unique: "\of_int z \ x; x < of_int z + 1\ \ floor x = z" huffman@30096: using floor_correct [of x] floor_exists1 [of x] by auto huffman@30096: huffman@30096: lemma of_int_floor_le: "of_int (floor x) \ x" huffman@30096: using floor_correct .. huffman@30096: huffman@30096: lemma le_floor_iff: "z \ floor x \ of_int z \ x" huffman@30096: proof huffman@30096: assume "z \ floor x" huffman@30096: then have "(of_int z :: 'a) \ of_int (floor x)" by simp huffman@30096: also have "of_int (floor x) \ x" by (rule of_int_floor_le) huffman@30096: finally show "of_int z \ x" . huffman@30096: next huffman@30096: assume "of_int z \ x" huffman@30096: also have "x < of_int (floor x + 1)" using floor_correct .. huffman@30096: finally show "z \ floor x" by (simp del: of_int_add) huffman@30096: qed huffman@30096: huffman@30096: lemma floor_less_iff: "floor x < z \ x < of_int z" huffman@30096: by (simp add: not_le [symmetric] le_floor_iff) huffman@30096: huffman@30096: lemma less_floor_iff: "z < floor x \ of_int z + 1 \ x" huffman@30096: using le_floor_iff [of "z + 1" x] by auto huffman@30096: huffman@30096: lemma floor_le_iff: "floor x \ z \ x < of_int z + 1" huffman@30096: by (simp add: not_less [symmetric] less_floor_iff) huffman@30096: huffman@30096: lemma floor_mono: assumes "x \ y" shows "floor x \ floor y" huffman@30096: proof - huffman@30096: have "of_int (floor x) \ x" by (rule of_int_floor_le) huffman@30096: also note `x \ y` huffman@30096: finally show ?thesis by (simp add: le_floor_iff) huffman@30096: qed huffman@30096: huffman@30096: lemma floor_less_cancel: "floor x < floor y \ x < y" huffman@30096: by (auto simp add: not_le [symmetric] floor_mono) huffman@30096: huffman@30096: lemma floor_of_int [simp]: "floor (of_int z) = z" huffman@30096: by (rule floor_unique) simp_all huffman@30096: huffman@30096: lemma floor_of_nat [simp]: "floor (of_nat n) = int n" huffman@30096: using floor_of_int [of "of_nat n"] by simp huffman@30096: huffman@47307: lemma le_floor_add: "floor x + floor y \ floor (x + y)" huffman@47307: by (simp only: le_floor_iff of_int_add add_mono of_int_floor_le) huffman@47307: huffman@30096: text {* Floor with numerals *} huffman@30096: huffman@30096: lemma floor_zero [simp]: "floor 0 = 0" huffman@30096: using floor_of_int [of 0] by simp huffman@30096: huffman@30096: lemma floor_one [simp]: "floor 1 = 1" huffman@30096: using floor_of_int [of 1] by simp huffman@30096: huffman@47108: lemma floor_numeral [simp]: "floor (numeral v) = numeral v" huffman@47108: using floor_of_int [of "numeral v"] by simp huffman@47108: huffman@47108: lemma floor_neg_numeral [simp]: "floor (neg_numeral v) = neg_numeral v" huffman@47108: using floor_of_int [of "neg_numeral v"] by simp huffman@30096: huffman@30096: lemma zero_le_floor [simp]: "0 \ floor x \ 0 \ x" huffman@30096: by (simp add: le_floor_iff) huffman@30096: huffman@30096: lemma one_le_floor [simp]: "1 \ floor x \ 1 \ x" huffman@30096: by (simp add: le_floor_iff) huffman@30096: huffman@47108: lemma numeral_le_floor [simp]: huffman@47108: "numeral v \ floor x \ numeral v \ x" huffman@47108: by (simp add: le_floor_iff) huffman@47108: huffman@47108: lemma neg_numeral_le_floor [simp]: huffman@47108: "neg_numeral v \ floor x \ neg_numeral v \ x" huffman@30096: by (simp add: le_floor_iff) huffman@30096: huffman@30096: lemma zero_less_floor [simp]: "0 < floor x \ 1 \ x" huffman@30096: by (simp add: less_floor_iff) huffman@30096: huffman@30096: lemma one_less_floor [simp]: "1 < floor x \ 2 \ x" huffman@30096: by (simp add: less_floor_iff) huffman@30096: huffman@47108: lemma numeral_less_floor [simp]: huffman@47108: "numeral v < floor x \ numeral v + 1 \ x" huffman@47108: by (simp add: less_floor_iff) huffman@47108: huffman@47108: lemma neg_numeral_less_floor [simp]: huffman@47108: "neg_numeral v < floor x \ neg_numeral v + 1 \ x" huffman@30096: by (simp add: less_floor_iff) huffman@30096: huffman@30096: lemma floor_le_zero [simp]: "floor x \ 0 \ x < 1" huffman@30096: by (simp add: floor_le_iff) huffman@30096: huffman@30096: lemma floor_le_one [simp]: "floor x \ 1 \ x < 2" huffman@30096: by (simp add: floor_le_iff) huffman@30096: huffman@47108: lemma floor_le_numeral [simp]: huffman@47108: "floor x \ numeral v \ x < numeral v + 1" huffman@47108: by (simp add: floor_le_iff) huffman@47108: huffman@47108: lemma floor_le_neg_numeral [simp]: huffman@47108: "floor x \ neg_numeral v \ x < neg_numeral v + 1" huffman@30096: by (simp add: floor_le_iff) huffman@30096: huffman@30096: lemma floor_less_zero [simp]: "floor x < 0 \ x < 0" huffman@30096: by (simp add: floor_less_iff) huffman@30096: huffman@30096: lemma floor_less_one [simp]: "floor x < 1 \ x < 1" huffman@30096: by (simp add: floor_less_iff) huffman@30096: huffman@47108: lemma floor_less_numeral [simp]: huffman@47108: "floor x < numeral v \ x < numeral v" huffman@47108: by (simp add: floor_less_iff) huffman@47108: huffman@47108: lemma floor_less_neg_numeral [simp]: huffman@47108: "floor x < neg_numeral v \ x < neg_numeral v" huffman@30096: by (simp add: floor_less_iff) huffman@30096: huffman@30096: text {* Addition and subtraction of integers *} huffman@30096: huffman@30096: lemma floor_add_of_int [simp]: "floor (x + of_int z) = floor x + z" huffman@30096: using floor_correct [of x] by (simp add: floor_unique) huffman@30096: huffman@47108: lemma floor_add_numeral [simp]: huffman@47108: "floor (x + numeral v) = floor x + numeral v" huffman@47108: using floor_add_of_int [of x "numeral v"] by simp huffman@47108: huffman@47108: lemma floor_add_neg_numeral [simp]: huffman@47108: "floor (x + neg_numeral v) = floor x + neg_numeral v" huffman@47108: using floor_add_of_int [of x "neg_numeral v"] by simp huffman@30096: huffman@30096: lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1" huffman@30096: using floor_add_of_int [of x 1] by simp huffman@30096: huffman@30096: lemma floor_diff_of_int [simp]: "floor (x - of_int z) = floor x - z" huffman@30096: using floor_add_of_int [of x "- z"] by (simp add: algebra_simps) huffman@30096: huffman@47108: lemma floor_diff_numeral [simp]: huffman@47108: "floor (x - numeral v) = floor x - numeral v" huffman@47108: using floor_diff_of_int [of x "numeral v"] by simp huffman@47108: huffman@47108: lemma floor_diff_neg_numeral [simp]: huffman@47108: "floor (x - neg_numeral v) = floor x - neg_numeral v" huffman@47108: using floor_diff_of_int [of x "neg_numeral v"] by simp huffman@30096: huffman@30096: lemma floor_diff_one [simp]: "floor (x - 1) = floor x - 1" huffman@30096: using floor_diff_of_int [of x 1] by simp huffman@30096: huffman@30096: huffman@30096: subsection {* Ceiling function *} huffman@30096: huffman@30096: definition bulwahn@43732: ceiling :: "'a::floor_ceiling \ int" where bulwahn@43733: "ceiling x = - floor (- x)" huffman@30096: huffman@30096: notation (xsymbols) huffman@30096: ceiling ("\_\") huffman@30096: huffman@30096: notation (HTML output) huffman@30096: ceiling ("\_\") huffman@30096: huffman@30096: lemma ceiling_correct: "of_int (ceiling x) - 1 < x \ x \ of_int (ceiling x)" huffman@30096: unfolding ceiling_def using floor_correct [of "- x"] by simp huffman@30096: huffman@30096: lemma ceiling_unique: "\of_int z - 1 < x; x \ of_int z\ \ ceiling x = z" huffman@30096: unfolding ceiling_def using floor_unique [of "- z" "- x"] by simp huffman@30096: huffman@30096: lemma le_of_int_ceiling: "x \ of_int (ceiling x)" huffman@30096: using ceiling_correct .. huffman@30096: huffman@30096: lemma ceiling_le_iff: "ceiling x \ z \ x \ of_int z" huffman@30096: unfolding ceiling_def using le_floor_iff [of "- z" "- x"] by auto huffman@30096: huffman@30096: lemma less_ceiling_iff: "z < ceiling x \ of_int z < x" huffman@30096: by (simp add: not_le [symmetric] ceiling_le_iff) huffman@30096: huffman@30096: lemma ceiling_less_iff: "ceiling x < z \ x \ of_int z - 1" huffman@30096: using ceiling_le_iff [of x "z - 1"] by simp huffman@30096: huffman@30096: lemma le_ceiling_iff: "z \ ceiling x \ of_int z - 1 < x" huffman@30096: by (simp add: not_less [symmetric] ceiling_less_iff) huffman@30096: huffman@30096: lemma ceiling_mono: "x \ y \ ceiling x \ ceiling y" huffman@30096: unfolding ceiling_def by (simp add: floor_mono) huffman@30096: huffman@30096: lemma ceiling_less_cancel: "ceiling x < ceiling y \ x < y" huffman@30096: by (auto simp add: not_le [symmetric] ceiling_mono) huffman@30096: huffman@30096: lemma ceiling_of_int [simp]: "ceiling (of_int z) = z" huffman@30096: by (rule ceiling_unique) simp_all huffman@30096: huffman@30096: lemma ceiling_of_nat [simp]: "ceiling (of_nat n) = int n" huffman@30096: using ceiling_of_int [of "of_nat n"] by simp huffman@30096: huffman@47307: lemma ceiling_add_le: "ceiling (x + y) \ ceiling x + ceiling y" huffman@47307: by (simp only: ceiling_le_iff of_int_add add_mono le_of_int_ceiling) huffman@47307: huffman@30096: text {* Ceiling with numerals *} huffman@30096: huffman@30096: lemma ceiling_zero [simp]: "ceiling 0 = 0" huffman@30096: using ceiling_of_int [of 0] by simp huffman@30096: huffman@30096: lemma ceiling_one [simp]: "ceiling 1 = 1" huffman@30096: using ceiling_of_int [of 1] by simp huffman@30096: huffman@47108: lemma ceiling_numeral [simp]: "ceiling (numeral v) = numeral v" huffman@47108: using ceiling_of_int [of "numeral v"] by simp huffman@47108: huffman@47108: lemma ceiling_neg_numeral [simp]: "ceiling (neg_numeral v) = neg_numeral v" huffman@47108: using ceiling_of_int [of "neg_numeral v"] by simp huffman@30096: huffman@30096: lemma ceiling_le_zero [simp]: "ceiling x \ 0 \ x \ 0" huffman@30096: by (simp add: ceiling_le_iff) huffman@30096: huffman@30096: lemma ceiling_le_one [simp]: "ceiling x \ 1 \ x \ 1" huffman@30096: by (simp add: ceiling_le_iff) huffman@30096: huffman@47108: lemma ceiling_le_numeral [simp]: huffman@47108: "ceiling x \ numeral v \ x \ numeral v" huffman@47108: by (simp add: ceiling_le_iff) huffman@47108: huffman@47108: lemma ceiling_le_neg_numeral [simp]: huffman@47108: "ceiling x \ neg_numeral v \ x \ neg_numeral v" huffman@30096: by (simp add: ceiling_le_iff) huffman@30096: huffman@30096: lemma ceiling_less_zero [simp]: "ceiling x < 0 \ x \ -1" huffman@30096: by (simp add: ceiling_less_iff) huffman@30096: huffman@30096: lemma ceiling_less_one [simp]: "ceiling x < 1 \ x \ 0" huffman@30096: by (simp add: ceiling_less_iff) huffman@30096: huffman@47108: lemma ceiling_less_numeral [simp]: huffman@47108: "ceiling x < numeral v \ x \ numeral v - 1" huffman@47108: by (simp add: ceiling_less_iff) huffman@47108: huffman@47108: lemma ceiling_less_neg_numeral [simp]: huffman@47108: "ceiling x < neg_numeral v \ x \ neg_numeral v - 1" huffman@30096: by (simp add: ceiling_less_iff) huffman@30096: huffman@30096: lemma zero_le_ceiling [simp]: "0 \ ceiling x \ -1 < x" huffman@30096: by (simp add: le_ceiling_iff) huffman@30096: huffman@30096: lemma one_le_ceiling [simp]: "1 \ ceiling x \ 0 < x" huffman@30096: by (simp add: le_ceiling_iff) huffman@30096: huffman@47108: lemma numeral_le_ceiling [simp]: huffman@47108: "numeral v \ ceiling x \ numeral v - 1 < x" huffman@47108: by (simp add: le_ceiling_iff) huffman@47108: huffman@47108: lemma neg_numeral_le_ceiling [simp]: huffman@47108: "neg_numeral v \ ceiling x \ neg_numeral v - 1 < x" huffman@30096: by (simp add: le_ceiling_iff) huffman@30096: huffman@30096: lemma zero_less_ceiling [simp]: "0 < ceiling x \ 0 < x" huffman@30096: by (simp add: less_ceiling_iff) huffman@30096: huffman@30096: lemma one_less_ceiling [simp]: "1 < ceiling x \ 1 < x" huffman@30096: by (simp add: less_ceiling_iff) huffman@30096: huffman@47108: lemma numeral_less_ceiling [simp]: huffman@47108: "numeral v < ceiling x \ numeral v < x" huffman@47108: by (simp add: less_ceiling_iff) huffman@47108: huffman@47108: lemma neg_numeral_less_ceiling [simp]: huffman@47108: "neg_numeral v < ceiling x \ neg_numeral v < x" huffman@30096: by (simp add: less_ceiling_iff) huffman@30096: huffman@30096: text {* Addition and subtraction of integers *} huffman@30096: huffman@30096: lemma ceiling_add_of_int [simp]: "ceiling (x + of_int z) = ceiling x + z" huffman@30096: using ceiling_correct [of x] by (simp add: ceiling_unique) huffman@30096: huffman@47108: lemma ceiling_add_numeral [simp]: huffman@47108: "ceiling (x + numeral v) = ceiling x + numeral v" huffman@47108: using ceiling_add_of_int [of x "numeral v"] by simp huffman@47108: huffman@47108: lemma ceiling_add_neg_numeral [simp]: huffman@47108: "ceiling (x + neg_numeral v) = ceiling x + neg_numeral v" huffman@47108: using ceiling_add_of_int [of x "neg_numeral v"] by simp huffman@30096: huffman@30096: lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1" huffman@30096: using ceiling_add_of_int [of x 1] by simp huffman@30096: huffman@30096: lemma ceiling_diff_of_int [simp]: "ceiling (x - of_int z) = ceiling x - z" huffman@30096: using ceiling_add_of_int [of x "- z"] by (simp add: algebra_simps) huffman@30096: huffman@47108: lemma ceiling_diff_numeral [simp]: huffman@47108: "ceiling (x - numeral v) = ceiling x - numeral v" huffman@47108: using ceiling_diff_of_int [of x "numeral v"] by simp huffman@47108: huffman@47108: lemma ceiling_diff_neg_numeral [simp]: huffman@47108: "ceiling (x - neg_numeral v) = ceiling x - neg_numeral v" huffman@47108: using ceiling_diff_of_int [of x "neg_numeral v"] by simp huffman@30096: huffman@30096: lemma ceiling_diff_one [simp]: "ceiling (x - 1) = ceiling x - 1" huffman@30096: using ceiling_diff_of_int [of x 1] by simp huffman@30096: hoelzl@47592: lemma ceiling_diff_floor_le_1: "ceiling x - floor x \ 1" hoelzl@47592: proof - hoelzl@47592: have "of_int \x\ - 1 < x" hoelzl@47592: using ceiling_correct[of x] by simp hoelzl@47592: also have "x < of_int \x\ + 1" hoelzl@47592: using floor_correct[of x] by simp_all hoelzl@47592: finally have "of_int (\x\ - \x\) < (of_int 2::'a)" hoelzl@47592: by simp hoelzl@47592: then show ?thesis hoelzl@47592: unfolding of_int_less_iff by simp hoelzl@47592: qed huffman@30096: huffman@30096: subsection {* Negation *} huffman@30096: huffman@30102: lemma floor_minus: "floor (- x) = - ceiling x" huffman@30096: unfolding ceiling_def by simp huffman@30096: huffman@30102: lemma ceiling_minus: "ceiling (- x) = - floor x" huffman@30096: unfolding ceiling_def by simp huffman@30096: huffman@30096: end