| author | hoelzl | 
| Wed, 09 Apr 2014 09:37:47 +0200 | |
| changeset 56479 | 91958d4b30f7 | 
| parent 53015 | a1119cf551e8 | 
| child 57514 | bdc2c6b40bf2 | 
| permissions | -rw-r--r-- | 
| 51174 | 1 | (* Author: Florian Haftmann, TU Muenchen *) | 
| 2 | ||
| 3 | header {* Common discrete functions *}
 | |
| 4 | ||
| 5 | theory Discrete | |
| 6 | imports Main | |
| 7 | begin | |
| 8 | ||
| 9 | subsection {* Discrete logarithm *}
 | |
| 10 | ||
| 11 | fun log :: "nat \<Rightarrow> nat" where | |
| 12 | [simp del]: "log n = (if n < 2 then 0 else Suc (log (n div 2)))" | |
| 13 | ||
| 14 | lemma log_zero [simp]: | |
| 15 | "log 0 = 0" | |
| 16 | by (simp add: log.simps) | |
| 17 | ||
| 18 | lemma log_one [simp]: | |
| 19 | "log 1 = 0" | |
| 20 | by (simp add: log.simps) | |
| 21 | ||
| 22 | lemma log_Suc_zero [simp]: | |
| 23 | "log (Suc 0) = 0" | |
| 24 | using log_one by simp | |
| 25 | ||
| 26 | lemma log_rec: | |
| 27 | "n \<ge> 2 \<Longrightarrow> log n = Suc (log (n div 2))" | |
| 28 | by (simp add: log.simps) | |
| 29 | ||
| 30 | lemma log_twice [simp]: | |
| 31 | "n \<noteq> 0 \<Longrightarrow> log (2 * n) = Suc (log n)" | |
| 32 | by (simp add: log_rec) | |
| 33 | ||
| 34 | lemma log_half [simp]: | |
| 35 | "log (n div 2) = log n - 1" | |
| 36 | proof (cases "n < 2") | |
| 37 | case True | |
| 38 | then have "n = 0 \<or> n = 1" by arith | |
| 39 | then show ?thesis by (auto simp del: One_nat_def) | |
| 40 | next | |
| 41 | case False then show ?thesis by (simp add: log_rec) | |
| 42 | qed | |
| 43 | ||
| 44 | lemma log_exp [simp]: | |
| 45 | "log (2 ^ n) = n" | |
| 46 | by (induct n) simp_all | |
| 47 | ||
| 48 | lemma log_mono: | |
| 49 | "mono log" | |
| 50 | proof | |
| 51 | fix m n :: nat | |
| 52 | assume "m \<le> n" | |
| 53 | then show "log m \<le> log n" | |
| 54 | proof (induct m arbitrary: n rule: log.induct) | |
| 55 | case (1 m) | |
| 56 | then have mn2: "m div 2 \<le> n div 2" by arith | |
| 57 | show "log m \<le> log n" | |
| 58 | proof (cases "m < 2") | |
| 59 | case True | |
| 60 | then have "m = 0 \<or> m = 1" by arith | |
| 61 | then show ?thesis by (auto simp del: One_nat_def) | |
| 62 | next | |
| 63 | case False | |
| 64 | with mn2 have "m \<ge> 2" and "n \<ge> 2" by auto arith | |
| 65 | from False have m2_0: "m div 2 \<noteq> 0" by arith | |
| 66 | with mn2 have n2_0: "n div 2 \<noteq> 0" by arith | |
| 67 | from False "1.hyps" mn2 have "log (m div 2) \<le> log (n div 2)" by blast | |
| 68 | with m2_0 n2_0 have "log (2 * (m div 2)) \<le> log (2 * (n div 2))" by simp | |
| 69 | with m2_0 n2_0 `m \<ge> 2` `n \<ge> 2` show ?thesis by (simp only: log_rec [of m] log_rec [of n]) simp | |
| 70 | qed | |
| 71 | qed | |
| 72 | qed | |
| 73 | ||
| 74 | ||
| 75 | subsection {* Discrete square root *}
 | |
| 76 | ||
| 77 | definition sqrt :: "nat \<Rightarrow> nat" | |
| 78 | where | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 79 |   "sqrt n = Max {m. m\<^sup>2 \<le> n}"
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 80 | |
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 81 | lemma sqrt_aux: | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 82 | fixes n :: nat | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 83 |   shows "finite {m. m\<^sup>2 \<le> n}" and "{m. m\<^sup>2 \<le> n} \<noteq> {}"
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 84 | proof - | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 85 |   { fix m
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 86 | assume "m\<^sup>2 \<le> n" | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 87 | then have "m \<le> n" | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 88 | by (cases m) (simp_all add: power2_eq_square) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 89 | } note ** = this | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 90 |   then have "{m. m\<^sup>2 \<le> n} \<subseteq> {m. m \<le> n}" by auto
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 91 |   then show "finite {m. m\<^sup>2 \<le> n}" by (rule finite_subset) rule
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 92 | have "0\<^sup>2 \<le> n" by simp | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 93 |   then show *: "{m. m\<^sup>2 \<le> n} \<noteq> {}" by blast
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 94 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 95 | |
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 96 | lemma [code]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 97 |   "sqrt n = Max (Set.filter (\<lambda>m. m\<^sup>2 \<le> n) {0..n})"
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 98 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 99 |   from power2_nat_le_imp_le [of _ n] have "{m. m \<le> n \<and> m\<^sup>2 \<le> n} = {m. m\<^sup>2 \<le> n}" by auto
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 100 | then show ?thesis by (simp add: sqrt_def Set.filter_def) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 101 | qed | 
| 51174 | 102 | |
| 103 | lemma sqrt_inverse_power2 [simp]: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 104 | "sqrt (n\<^sup>2) = n" | 
| 51174 | 105 | proof - | 
| 106 |   have "{m. m \<le> n} \<noteq> {}" by auto
 | |
| 107 |   then have "Max {m. m \<le> n} \<le> n" by auto
 | |
| 108 | then show ?thesis | |
| 109 | by (auto simp add: sqrt_def power2_nat_le_eq_le intro: antisym) | |
| 110 | qed | |
| 111 | ||
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 112 | lemma mono_sqrt: "mono sqrt" | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 113 | proof | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 114 | fix m n :: nat | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 115 | have *: "0 * 0 \<le> m" by simp | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 116 | assume "m \<le> n" | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 117 | then show "sqrt m \<le> sqrt n" | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 118 | by (auto intro!: Max_mono `0 * 0 \<le> m` finite_less_ub simp add: power2_eq_square sqrt_def) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 119 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 120 | |
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 121 | lemma sqrt_greater_zero_iff [simp]: | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 122 | "sqrt n > 0 \<longleftrightarrow> n > 0" | 
| 51174 | 123 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 124 |   have *: "0 < Max {m. m\<^sup>2 \<le> n} \<longleftrightarrow> (\<exists>a\<in>{m. m\<^sup>2 \<le> n}. 0 < a)"
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 125 | by (rule Max_gr_iff) (fact sqrt_aux)+ | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 126 | show ?thesis | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 127 | proof | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 128 | assume "0 < sqrt n" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 129 |     then have "0 < Max {m. m\<^sup>2 \<le> n}" by (simp add: sqrt_def)
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 130 | with * show "0 < n" by (auto dest: power2_nat_le_imp_le) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 131 | next | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 132 | assume "0 < n" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 133 | then have "1\<^sup>2 \<le> n \<and> 0 < (1::nat)" by simp | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 134 | then have "\<exists>q. q\<^sup>2 \<le> n \<and> 0 < q" .. | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 135 |     with * have "0 < Max {m. m\<^sup>2 \<le> n}" by blast
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 136 | then show "0 < sqrt n" by (simp add: sqrt_def) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 137 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 138 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 139 | |
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 140 | lemma sqrt_power2_le [simp]: (* FIXME tune proof *) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 141 | "(sqrt n)\<^sup>2 \<le> n" | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 142 | proof (cases "n > 0") | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 143 | case False then show ?thesis by (simp add: sqrt_def) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 144 | next | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 145 | case True then have "sqrt n > 0" by simp | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 146 |   then have "mono (times (Max {m. m\<^sup>2 \<le> n}))" by (auto intro: mono_times_nat simp add: sqrt_def)
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 147 |   then have *: "Max {m. m\<^sup>2 \<le> n} * Max {m. m\<^sup>2 \<le> n} = Max (times (Max {m. m\<^sup>2 \<le> n}) ` {m. m\<^sup>2 \<le> n})"
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 148 | using sqrt_aux [of n] by (rule mono_Max_commute) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 149 |   have "Max (op * (Max {m. m * m \<le> n}) ` {m. m * m \<le> n}) \<le> n"
 | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 150 | apply (subst Max_le_iff) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 151 | apply (metis (mono_tags) finite_imageI finite_less_ub le_square) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 152 | apply simp | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 153 | apply (metis le0 mult_0_right) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 154 | apply auto | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 155 | proof - | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 156 | fix q | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 157 | assume "q * q \<le> n" | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 158 |       show "Max {m. m * m \<le> n} * q \<le> n"
 | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 159 | proof (cases "q > 0") | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 160 | case False then show ?thesis by simp | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 161 | next | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 162 | case True then have "mono (times q)" by (rule mono_times_nat) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 163 |         then have "q * Max {m. m * m \<le> n} = Max (times q ` {m. m * m \<le> n})"
 | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 164 | using sqrt_aux [of n] by (auto simp add: power2_eq_square intro: mono_Max_commute) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 165 |         then have "Max {m. m * m \<le> n} * q = Max (times q ` {m. m * m \<le> n})" by (simp add: mult_ac)
 | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 166 | then show ?thesis apply simp | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 167 | apply (subst Max_le_iff) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 168 | apply auto | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 169 | apply (metis (mono_tags) finite_imageI finite_less_ub le_square) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 170 | apply (metis `q * q \<le> n`) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 171 | using `q * q \<le> n` by (metis le_cases mult_le_mono1 mult_le_mono2 order_trans) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 172 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 173 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 174 | with * show ?thesis by (simp add: sqrt_def power2_eq_square) | 
| 51174 | 175 | qed | 
| 176 | ||
| 177 | lemma sqrt_le: | |
| 178 | "sqrt n \<le> n" | |
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 179 | using sqrt_aux [of n] by (auto simp add: sqrt_def intro: power2_nat_le_imp_le) | 
| 51174 | 180 | |
| 181 | hide_const (open) log sqrt | |
| 182 | ||
| 183 | end | |
| 184 |