author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 81410 | da3bf4a755b3 |
permissions | -rw-r--r-- |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
1 |
(* Title: HOL/Library/Log_Nat.thy |
81410
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
2 |
Author: Johannes Hölzl, Fabian Immler, Manuel Eberl |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
3 |
Copyright 2012 TU München |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
4 |
*) |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
5 |
|
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
6 |
section \<open>Logarithm of Natural Numbers\<close> |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
7 |
|
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
8 |
theory Log_Nat |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
9 |
imports Complex_Main |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
10 |
begin |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
11 |
|
70349 | 12 |
subsection \<open>Preliminaries\<close> |
13 |
||
14 |
lemma divide_nat_diff_div_nat_less_one: |
|
15 |
"real x / real b - real (x div b) < 1" for x b :: nat |
|
16 |
proof (cases "b = 0") |
|
17 |
case True |
|
18 |
then show ?thesis |
|
19 |
by simp |
|
20 |
next |
|
21 |
case False |
|
22 |
then have "real (x div b) + real (x mod b) / real b - real (x div b) < 1" |
|
23 |
by (simp add: field_simps) |
|
24 |
then show ?thesis |
|
80732 | 25 |
by (metis of_nat_of_nat_div_aux) |
70349 | 26 |
qed |
27 |
||
28 |
||
29 |
subsection \<open>Floorlog\<close> |
|
30 |
||
31 |
definition floorlog :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
32 |
where "floorlog b a = (if a > 0 \<and> b > 1 then nat \<lfloor>log b a\<rfloor> + 1 else 0)" |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
33 |
|
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
34 |
lemma floorlog_mono: "x \<le> y \<Longrightarrow> floorlog b x \<le> floorlog b y" |
70349 | 35 |
by (auto simp: floorlog_def floor_mono nat_mono) |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
36 |
|
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
37 |
lemma floorlog_bounds: |
70349 | 38 |
"b ^ (floorlog b x - 1) \<le> x \<and> x < b ^ (floorlog b x)" if "x > 0" "b > 1" |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
39 |
proof |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
40 |
show "b ^ (floorlog b x - 1) \<le> x" |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
41 |
proof - |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
42 |
have "b ^ nat \<lfloor>log b x\<rfloor> = b powr \<lfloor>log b x\<rfloor>" |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
43 |
using powr_realpow[symmetric, of b "nat \<lfloor>log b x\<rfloor>"] \<open>x > 0\<close> \<open>b > 1\<close> |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
44 |
by simp |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
45 |
also have "\<dots> \<le> b powr log b x" using \<open>b > 1\<close> by simp |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
46 |
also have "\<dots> = real_of_int x" using \<open>0 < x\<close> \<open>b > 1\<close> by simp |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
47 |
finally have "b ^ nat \<lfloor>log b x\<rfloor> \<le> real_of_int x" by simp |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
48 |
then show ?thesis |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
49 |
using \<open>0 < x\<close> \<open>b > 1\<close> of_nat_le_iff |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
50 |
by (fastforce simp add: floorlog_def) |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
51 |
qed |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
52 |
show "x < b ^ (floorlog b x)" |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
53 |
proof - |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
54 |
have "x \<le> b powr (log b x)" using \<open>x > 0\<close> \<open>b > 1\<close> by simp |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
55 |
also have "\<dots> < b powr (\<lfloor>log b x\<rfloor> + 1)" |
70349 | 56 |
using that by (intro powr_less_mono) auto |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
57 |
also have "\<dots> = b ^ nat (\<lfloor>log b (real_of_int x)\<rfloor> + 1)" |
70349 | 58 |
using that by (simp flip: powr_realpow) |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
59 |
finally |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
60 |
have "x < b ^ nat (\<lfloor>log b (int x)\<rfloor> + 1)" |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
61 |
by (rule of_nat_less_imp_less) |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
62 |
then show ?thesis |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
63 |
using \<open>x > 0\<close> \<open>b > 1\<close> by (simp add: floorlog_def nat_add_distrib) |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
64 |
qed |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
65 |
qed |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
66 |
|
70349 | 67 |
lemma floorlog_power [simp]: |
68 |
"floorlog b (a * b ^ c) = floorlog b a + c" if "a > 0" "b > 1" |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
69 |
proof - |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
70 |
have "\<lfloor>log b a + real c\<rfloor> = \<lfloor>log b a\<rfloor> + c" by arith |
70349 | 71 |
then show ?thesis using that |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
72 |
by (auto simp: floorlog_def log_mult powr_realpow[symmetric] nat_add_distrib) |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
73 |
qed |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
74 |
|
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
75 |
lemma floor_log_add_eqI: |
70349 | 76 |
"\<lfloor>log b (a + r)\<rfloor> = \<lfloor>log b a\<rfloor>" if "b > 1" "a \<ge> 1" "0 \<le> r" "r < 1" |
77 |
for a b :: nat and r :: real |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
78 |
proof (rule floor_eq2) |
70349 | 79 |
have "log b a \<le> log b (a + r)" using that by force |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
80 |
then show "\<lfloor>log b a\<rfloor> \<le> log b (a + r)" by arith |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
81 |
next |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
82 |
define l::int where "l = int b ^ (nat \<lfloor>log b a\<rfloor> + 1)" |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
83 |
have l_def_real: "l = b powr (\<lfloor>log b a\<rfloor> + 1)" |
70349 | 84 |
using that by (simp add: l_def powr_add powr_real_of_int) |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
85 |
have "a < l" |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
86 |
proof - |
70349 | 87 |
have "a = b powr (log b a)" using that by simp |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
88 |
also have "\<dots> < b powr floor ((log b a) + 1)" |
70349 | 89 |
using that(1) by auto |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
90 |
also have "\<dots> = l" |
70349 | 91 |
using that by (simp add: l_def powr_real_of_int powr_add) |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
92 |
finally show ?thesis by simp |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
93 |
qed |
70349 | 94 |
then have "a + r < l" using that by simp |
95 |
then have "log b (a + r) < log b l" using that by simp |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
96 |
also have "\<dots> = real_of_int \<lfloor>log b a\<rfloor> + 1" |
70349 | 97 |
using that by (simp add: l_def_real) |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
98 |
finally show "log b (a + r) < real_of_int \<lfloor>log b a\<rfloor> + 1" . |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
99 |
qed |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
100 |
|
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
101 |
lemma floor_log_div: |
70349 | 102 |
"\<lfloor>log b x\<rfloor> = \<lfloor>log b (x div b)\<rfloor> + 1" if "b > 1" "x > 0" "x div b > 0" |
103 |
for b x :: nat |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
104 |
proof- |
70349 | 105 |
have "\<lfloor>log b x\<rfloor> = \<lfloor>log b (x / b * b)\<rfloor>" using that by simp |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
106 |
also have "\<dots> = \<lfloor>log b (x / b) + log b b\<rfloor>" |
70349 | 107 |
using that by (subst log_mult) auto |
108 |
also have "\<dots> = \<lfloor>log b (x / b)\<rfloor> + 1" using that by simp |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
109 |
also have "\<lfloor>log b (x / b)\<rfloor> = \<lfloor>log b (x div b + (x / b - x div b))\<rfloor>" by simp |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
110 |
also have "\<dots> = \<lfloor>log b (x div b)\<rfloor>" |
80732 | 111 |
using that of_nat_div_le_of_nat divide_nat_diff_div_nat_less_one |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
112 |
by (intro floor_log_add_eqI) auto |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
113 |
finally show ?thesis . |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
114 |
qed |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
115 |
|
70349 | 116 |
lemma compute_floorlog [code]: |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
117 |
"floorlog b x = (if x > 0 \<and> b > 1 then floorlog b (x div b) + 1 else 0)" |
70349 | 118 |
by (auto simp: floorlog_def floor_log_div[of b x] div_eq_0_iff nat_add_distrib |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
119 |
intro!: floor_eq2) |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
120 |
|
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
121 |
lemma floor_log_eq_if: |
70349 | 122 |
"\<lfloor>log b x\<rfloor> = \<lfloor>log b y\<rfloor>" if "x div b = y div b" "b > 1" "x > 0" "x div b \<ge> 1" |
123 |
for b x y :: nat |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
124 |
proof - |
70349 | 125 |
have "y > 0" using that by (auto intro: ccontr) |
126 |
thus ?thesis using that by (simp add: floor_log_div) |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
127 |
qed |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
128 |
|
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
129 |
lemma floorlog_eq_if: |
70349 | 130 |
"floorlog b x = floorlog b y" if "x div b = y div b" "b > 1" "x > 0" "x div b \<ge> 1" |
131 |
for b x y :: nat |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
132 |
proof - |
70349 | 133 |
have "y > 0" using that by (auto intro: ccontr) |
134 |
then show ?thesis using that |
|
135 |
by (auto simp add: floorlog_def eq_nat_nat_iff intro: floor_log_eq_if) |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
136 |
qed |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
137 |
|
70349 | 138 |
lemma floorlog_leD: |
139 |
"floorlog b x \<le> w \<Longrightarrow> b > 1 \<Longrightarrow> x < b ^ w" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
140 |
by (metis floorlog_bounds leD linorder_neqE_nat order.strict_trans power_strict_increasing_iff |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
141 |
zero_less_one zero_less_power) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
142 |
|
70349 | 143 |
lemma floorlog_leI: |
144 |
"x < b ^ w \<Longrightarrow> 0 \<le> w \<Longrightarrow> b > 1 \<Longrightarrow> floorlog b x \<le> w" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
145 |
by (drule less_imp_of_nat_less[where 'a=real]) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
146 |
(auto simp: floorlog_def Suc_le_eq nat_less_iff floor_less_iff log_of_power_less) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
147 |
|
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
148 |
lemma floorlog_eq_zero_iff: |
70349 | 149 |
"floorlog b x = 0 \<longleftrightarrow> b \<le> 1 \<or> x \<le> 0" |
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
150 |
by (auto simp: floorlog_def) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
151 |
|
70349 | 152 |
lemma floorlog_le_iff: |
153 |
"floorlog b x \<le> w \<longleftrightarrow> b \<le> 1 \<or> b > 1 \<and> 0 \<le> w \<and> x < b ^ w" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
154 |
using floorlog_leD[of b x w] floorlog_leI[of x b w] |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
155 |
by (auto simp: floorlog_eq_zero_iff[THEN iffD2]) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
156 |
|
70349 | 157 |
lemma floorlog_ge_SucI: |
158 |
"Suc w \<le> floorlog b x" if "b ^ w \<le> x" "b > 1" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
159 |
using that le_log_of_power[of b w x] power_not_zero |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
160 |
by (force simp: floorlog_def Suc_le_eq powr_realpow not_less Suc_nat_eq_nat_zadd1 |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
161 |
zless_nat_eq_int_zless int_add_floor less_floor_iff |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
162 |
simp del: floor_add2) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
163 |
|
70349 | 164 |
lemma floorlog_geI: |
165 |
"w \<le> floorlog b x" if "b ^ (w - 1) \<le> x" "b > 1" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
166 |
using floorlog_ge_SucI[of b "w - 1" x] that |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
167 |
by auto |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
168 |
|
70349 | 169 |
lemma floorlog_geD: |
170 |
"b ^ (w - 1) \<le> x" if "w \<le> floorlog b x" "w > 0" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
171 |
proof - |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
172 |
have "b > 1" "0 < x" |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
173 |
using that by (auto simp: floorlog_def split: if_splits) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
174 |
have "b ^ (w - 1) \<le> x" if "b ^ w \<le> x" |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
175 |
proof - |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
176 |
have "b ^ (w - 1) \<le> b ^ w" |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
177 |
using \<open>b > 1\<close> |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
178 |
by (auto intro!: power_increasing) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
179 |
also note that |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
180 |
finally show ?thesis . |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
181 |
qed |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
182 |
moreover have "b ^ nat \<lfloor>log (real b) (real x)\<rfloor> \<le> x" (is "?l \<le> _") |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
183 |
proof - |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
184 |
have "0 \<le> log (real b) (real x)" |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
185 |
using \<open>b > 1\<close> \<open>0 < x\<close> |
75455
91c16c5ad3e9
tidied auto / simp with null arguments
paulson <lp15@cam.ac.uk>
parents:
70350
diff
changeset
|
186 |
by auto |
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
187 |
then have "?l \<le> b powr log (real b) (real x)" |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
188 |
using \<open>b > 1\<close> |
68406 | 189 |
by (auto simp flip: powr_realpow intro!: powr_mono of_nat_floor) |
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
190 |
also have "\<dots> = x" using \<open>b > 1\<close> \<open>0 < x\<close> |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
191 |
by auto |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
192 |
finally show ?thesis |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
193 |
unfolding of_nat_le_iff . |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
194 |
qed |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
195 |
ultimately show ?thesis |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
196 |
using that |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
197 |
by (auto simp: floorlog_def le_nat_iff le_floor_iff le_log_iff powr_realpow |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
198 |
split: if_splits elim!: le_SucE) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
199 |
qed |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
200 |
|
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
201 |
|
81410
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
202 |
subsection \<open>\<close> |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
203 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
204 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
205 |
definition ceillog2 :: "nat \<Rightarrow> nat" where |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
206 |
"ceillog2 n = (if n = 0 then 0 else nat \<lceil>log 2 (real n)\<rceil>)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
207 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
208 |
lemma ceillog2_0 [simp]: "ceillog2 0 = 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
209 |
and ceillog2_Suc_0 [simp]: "ceillog2 (Suc 0) = 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
210 |
and ceillog2_2 [simp]: "ceillog2 2 = 1" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
211 |
by (auto simp: ceillog2_def) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
212 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
213 |
lemma ceillog2_le1_eq_0 [simp]: "n \<le> 1 \<Longrightarrow> ceillog2 n = 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
214 |
by (cases n) auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
215 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
216 |
lemma ceillog2_2_power [simp]: "ceillog2 (2 ^ n) = n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
217 |
by (auto simp: ceillog2_def) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
218 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
219 |
lemma ceillog2_ge_log: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
220 |
assumes "n > 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
221 |
shows "real (ceillog2 n) \<ge> log 2 (real n)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
222 |
proof - |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
223 |
have "real_of_int \<lceil>log 2 (real n)\<rceil> \<ge> log 2 (real n)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
224 |
by linarith |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
225 |
thus ?thesis |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
226 |
using assms unfolding ceillog2_def by auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
227 |
qed |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
228 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
229 |
lemma ceillog2_less_log: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
230 |
assumes "n > 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
231 |
shows "real (ceillog2 n) < log 2 (real n) + 1" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
232 |
proof - |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
233 |
have "real_of_int \<lceil>log 2 (real n)\<rceil> < log 2 (real n) + 1" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
234 |
by linarith |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
235 |
thus ?thesis |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
236 |
using assms unfolding ceillog2_def by auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
237 |
qed |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
238 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
239 |
lemma ceillog2_le_iff: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
240 |
assumes "n > 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
241 |
shows "ceillog2 n \<le> l \<longleftrightarrow> n \<le> 2 ^ l" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
242 |
proof - |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
243 |
have "ceillog2 n \<le> l \<longleftrightarrow> real n \<le> 2 ^ l" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
244 |
unfolding ceillog2_def using assms by (auto simp: log_le_iff powr_realpow) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
245 |
also have "2 ^ l = real (2 ^ l)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
246 |
by simp |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
247 |
also have "real n \<le> real (2 ^ l) \<longleftrightarrow> n \<le> 2 ^ l" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
248 |
by linarith |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
249 |
finally show ?thesis . |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
250 |
qed |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
251 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
252 |
lemma ceillog2_ge_iff: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
253 |
assumes "n > 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
254 |
shows "ceillog2 n \<ge> l \<longleftrightarrow> 2 ^ l < 2 * n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
255 |
proof - |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
256 |
have "-1 < (0 :: real)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
257 |
by auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
258 |
also have "\<dots> \<le> log 2 (real n)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
259 |
using assms by auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
260 |
finally have "ceillog2 n \<ge> l \<longleftrightarrow> real l - 1 < log 2 (real n)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
261 |
unfolding ceillog2_def using assms by (auto simp: le_nat_iff le_ceiling_iff) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
262 |
also have "\<dots> \<longleftrightarrow> real l < log 2 (real (2 * n))" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
263 |
using assms by (auto simp: log_mult) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
264 |
also have "\<dots> \<longleftrightarrow> 2 ^ l < real (2 * n)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
265 |
using assms by (subst less_log_iff) (auto simp: powr_realpow) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
266 |
also have "2 ^ l = real (2 ^ l)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
267 |
by simp |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
268 |
also have "real (2 ^ l) < real (2 * n) \<longleftrightarrow> 2 ^ l < 2 * n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
269 |
by linarith |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
270 |
finally show ?thesis . |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
271 |
qed |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
272 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
273 |
lemma le_two_power_ceillog2: "n \<le> 2 ^ ceillog2 n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
274 |
using neq0_conv ceillog2_le_iff by blast |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
275 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
276 |
lemma two_power_ceillog2_gt: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
277 |
assumes "n > 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
278 |
shows "2 * n > 2 ^ ceillog2 n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
279 |
using ceillog2_ge_iff[of n "ceillog2 n"] assms by simp |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
280 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
281 |
lemma ceillog2_eqI: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
282 |
assumes "n \<le> 2 ^ l" "2 ^ l < 2 * n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
283 |
shows "ceillog2 n = l" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
284 |
by (metis Suc_leI assms bot_nat_0.not_eq_extremum ceillog2_ge_iff ceillog2_le_iff le_antisym mult_is_0 |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
285 |
not_less_eq_eq) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
286 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
287 |
lemma ceillog2_rec_even: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
288 |
assumes "k > 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
289 |
shows "ceillog2 (2 * k) = Suc (ceillog2 k)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
290 |
by (rule ceillog2_eqI) (auto simp: le_two_power_ceillog2 two_power_ceillog2_gt assms) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
291 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
292 |
lemma ceillog2_mono: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
293 |
assumes "m \<le> n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
294 |
shows "ceillog2 m \<le> ceillog2 n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
295 |
proof (cases "m = 0") |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
296 |
case False |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
297 |
have "\<lceil>log 2 (real m)\<rceil> \<le> \<lceil>log 2 (real n)\<rceil>" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
298 |
by (intro ceiling_mono) (use False assms in auto) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
299 |
hence "nat \<lceil>log 2 (real m)\<rceil> \<le> nat \<lceil>log 2 (real n)\<rceil>" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
300 |
by linarith |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
301 |
thus ?thesis using False assms |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
302 |
unfolding ceillog2_def by simp |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
303 |
qed auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
304 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
305 |
lemma ceillog2_rec_odd: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
306 |
assumes "k > 0" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
307 |
shows "ceillog2 (Suc (2 * k)) = Suc (ceillog2 (Suc k))" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
308 |
proof - |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
309 |
have "2 ^ ceillog2 (Suc (2 * k)) > Suc (2 * k)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
310 |
by (metis assms diff_Suc_1 dvd_triv_left le_two_power_ceillog2 mult_pos_pos nat_power_eq_Suc_0_iff |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
311 |
order_less_le pos2 semiring_parity_class.even_mask_iff) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
312 |
then have "ceillog2 (2 * k + 2) \<le> ceillog2 (2 * k + 1)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
313 |
by (simp add: ceillog2_le_iff) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
314 |
moreover have "ceillog2 (2 * k + 2) \<ge> ceillog2 (2 * k + 1)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
315 |
by (rule ceillog2_mono) auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
316 |
ultimately have "ceillog2 (2 * k + 2) = ceillog2 (2 * k + 1)" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
317 |
by (rule antisym) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
318 |
also have "2 * k + 2 = 2 * Suc k" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
319 |
by simp |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
320 |
also have "ceillog2 (2 * Suc k) = Suc (ceillog2 (Suc k))" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
321 |
by (rule ceillog2_rec_even) auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
322 |
finally show ?thesis |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
323 |
by simp |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
324 |
qed |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
325 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
326 |
(* TODO: better code is possible using bitlen and "count trailing 0 bits" *) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
327 |
lemma ceillog2_rec: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
328 |
"ceillog2 n = (if n \<le> 1 then 0 else 1 + ceillog2 ((n + 1) div 2))" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
329 |
proof (cases "n \<le> 1") |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
330 |
case True |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
331 |
thus ?thesis |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
332 |
by (cases n) auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
333 |
next |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
334 |
case False |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
335 |
thus ?thesis |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
336 |
by (cases "even n") (auto elim!: evenE oddE simp: ceillog2_rec_even ceillog2_rec_odd) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
337 |
qed |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
338 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
339 |
lemma funpow_div2_ceillog2_le_1: |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
340 |
"((\<lambda>n. (n + 1) div 2) ^^ ceillog2 n) n \<le> 1" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
341 |
proof (induction n rule: less_induct) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
342 |
case (less n) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
343 |
show ?case |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
344 |
proof (cases "n \<le> 1") |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
345 |
case True |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
346 |
thus ?thesis by (cases n) auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
347 |
next |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
348 |
case False |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
349 |
have "((\<lambda>n. (n + 1) div 2) ^^ Suc (ceillog2 ((n + 1) div 2))) n \<le> 1" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
350 |
using less.IH[of "(n+1) div 2"] False by (subst funpow_Suc_right) auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
351 |
also have "Suc (ceillog2 ((n + 1) div 2)) = ceillog2 n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
352 |
using False by (subst ceillog2_rec[of n]) auto |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
353 |
finally show ?thesis . |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
354 |
qed |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
355 |
qed |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
356 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
357 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
358 |
fun ceillog2_aux :: "nat \<Rightarrow> nat \<Rightarrow> nat" where |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
359 |
"ceillog2_aux acc n = (if n \<le> 1 then acc else ceillog2_aux (acc + 1) ((n + 1) div 2))" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
360 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
361 |
lemmas [simp del] = ceillog2_aux.simps |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
362 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
363 |
lemma ceillog2_aux_correct: "ceillog2_aux acc n = ceillog2 n + acc" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
364 |
proof (induction acc n rule: ceillog2_aux.induct) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
365 |
case (1 acc n) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
366 |
show ?case |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
367 |
proof (cases "n \<le> 1") |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
368 |
case False |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
369 |
thus ?thesis using ceillog2_rec[of n] "1.IH" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
370 |
by (auto simp: ceillog2_aux.simps[of acc n]) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
371 |
qed (auto simp: ceillog2_aux.simps[of acc n]) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
372 |
qed |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
373 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
374 |
(* TODO: better code equation using bit operations *) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
375 |
lemma ceillog2_code [code]: "ceillog2 n = ceillog2_aux 0 n" |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
376 |
by (simp add: ceillog2_aux_correct) |
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
377 |
|
da3bf4a755b3
The ceillog2 function (thanks to Manuel Eberl), replacing two copies of Ceil_Log2
paulson <lp15@cam.ac.uk>
parents:
80732
diff
changeset
|
378 |
|
70349 | 379 |
subsection \<open>Bitlen\<close> |
380 |
||
381 |
definition bitlen :: "int \<Rightarrow> int" |
|
382 |
where "bitlen a = floorlog 2 (nat a)" |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
383 |
|
70349 | 384 |
lemma bitlen_alt_def: |
385 |
"bitlen a = (if a > 0 then \<lfloor>log 2 a\<rfloor> + 1 else 0)" |
|
386 |
by (simp add: bitlen_def floorlog_def) |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
387 |
|
70349 | 388 |
lemma bitlen_zero [simp]: |
389 |
"bitlen 0 = 0" |
|
67573 | 390 |
by (auto simp: bitlen_def floorlog_def) |
391 |
||
70349 | 392 |
lemma bitlen_nonneg: |
393 |
"0 \<le> bitlen x" |
|
67573 | 394 |
by (simp add: bitlen_def) |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
395 |
|
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
396 |
lemma bitlen_bounds: |
70349 | 397 |
"2 ^ nat (bitlen x - 1) \<le> x \<and> x < 2 ^ nat (bitlen x)" if "x > 0" |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
398 |
proof - |
70349 | 399 |
from that have "bitlen x \<ge> 1" by (auto simp: bitlen_alt_def) |
400 |
with that floorlog_bounds[of "nat x" 2] show ?thesis |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
401 |
by (auto simp add: bitlen_def le_nat_iff nat_less_iff nat_diff_distrib) |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
402 |
qed |
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
403 |
|
70349 | 404 |
lemma bitlen_pow2 [simp]: |
405 |
"bitlen (b * 2 ^ c) = bitlen b + c" if "b > 0" |
|
406 |
using that by (simp add: bitlen_def nat_mult_distrib nat_power_eq) |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
407 |
|
70349 | 408 |
lemma compute_bitlen [code]: |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
409 |
"bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)" |
70349 | 410 |
by (simp add: bitlen_def nat_div_distrib compute_floorlog) |
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
411 |
|
70349 | 412 |
lemma bitlen_eq_zero_iff: |
413 |
"bitlen x = 0 \<longleftrightarrow> x \<le> 0" |
|
414 |
by (auto simp add: bitlen_alt_def) |
|
63664 | 415 |
(metis compute_bitlen add.commute bitlen_alt_def bitlen_nonneg less_add_same_cancel2 |
416 |
not_less zero_less_one) |
|
417 |
||
418 |
lemma bitlen_div: |
|
70349 | 419 |
"1 \<le> real_of_int m / 2^nat (bitlen m - 1)" |
420 |
and "real_of_int m / 2^nat (bitlen m - 1) < 2" if "0 < m" |
|
63664 | 421 |
proof - |
422 |
let ?B = "2^nat (bitlen m - 1)" |
|
423 |
||
424 |
have "?B \<le> m" using bitlen_bounds[OF \<open>0 <m\<close>] .. |
|
425 |
then have "1 * ?B \<le> real_of_int m" |
|
426 |
unfolding of_int_le_iff[symmetric] by auto |
|
427 |
then show "1 \<le> real_of_int m / ?B" by auto |
|
428 |
||
70349 | 429 |
from that have "0 \<le> bitlen m - 1" by (auto simp: bitlen_alt_def) |
63664 | 430 |
|
70349 | 431 |
have "m < 2^nat(bitlen m)" using bitlen_bounds[OF that] .. |
432 |
also from that have "\<dots> = 2^nat(bitlen m - 1 + 1)" |
|
63664 | 433 |
by (auto simp: bitlen_def) |
434 |
also have "\<dots> = ?B * 2" |
|
435 |
unfolding nat_add_distrib[OF \<open>0 \<le> bitlen m - 1\<close> zero_le_one] by auto |
|
436 |
finally have "real_of_int m < 2 * ?B" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
437 |
by (metis (full_types) mult.commute power.simps(2) of_int_less_numeral_power_cancel_iff) |
63664 | 438 |
then have "real_of_int m / ?B < 2 * ?B / ?B" |
439 |
by (rule divide_strict_right_mono) auto |
|
440 |
then show "real_of_int m / ?B < 2" by auto |
|
441 |
qed |
|
442 |
||
70349 | 443 |
lemma bitlen_le_iff_floorlog: |
444 |
"bitlen x \<le> w \<longleftrightarrow> w \<ge> 0 \<and> floorlog 2 (nat x) \<le> nat w" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
445 |
by (auto simp: bitlen_def) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
446 |
|
70349 | 447 |
lemma bitlen_le_iff_power: |
448 |
"bitlen x \<le> w \<longleftrightarrow> w \<ge> 0 \<and> x < 2 ^ nat w" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
449 |
by (auto simp: bitlen_le_iff_floorlog floorlog_le_iff) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
450 |
|
70349 | 451 |
lemma less_power_nat_iff_bitlen: |
452 |
"x < 2 ^ w \<longleftrightarrow> bitlen (int x) \<le> w" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
453 |
using bitlen_le_iff_power[of x w] |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
454 |
by auto |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
455 |
|
70349 | 456 |
lemma bitlen_ge_iff_power: |
457 |
"w \<le> bitlen x \<longleftrightarrow> w \<le> 0 \<or> 2 ^ (nat w - 1) \<le> x" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
458 |
unfolding bitlen_def |
68406 | 459 |
by (auto simp flip: nat_le_iff intro: floorlog_geI dest: floorlog_geD) |
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
460 |
|
70349 | 461 |
lemma bitlen_twopow_add_eq: |
462 |
"bitlen (2 ^ w + b) = w + 1" if "0 \<le> b" "b < 2 ^ w" |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
463 |
by (auto simp: that nat_add_distrib bitlen_le_iff_power bitlen_ge_iff_power intro!: antisym) |
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
63664
diff
changeset
|
464 |
|
63663
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents:
diff
changeset
|
465 |
end |