| author | wenzelm | 
| Sun, 25 Aug 2024 22:54:56 +0200 | |
| changeset 80771 | fd01ef524169 | 
| parent 80621 | 6c369fec315a | 
| child 81332 | f94b30fa2b6c | 
| permissions | -rw-r--r-- | 
| 59349 | 1 | (* Author: Florian Haftmann, TU Muenchen *) | 
| 51174 | 2 | |
| 60500 | 3 | section \<open>Common discrete functions\<close> | 
| 51174 | 4 | |
| 5 | theory Discrete | |
| 64449 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 6 | imports Complex_Main | 
| 51174 | 7 | begin | 
| 8 | ||
| 60500 | 9 | subsection \<open>Discrete logarithm\<close> | 
| 51174 | 10 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 11 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 12 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 13 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 14 | qualified fun log :: "nat \<Rightarrow> nat" | 
| 59349 | 15 | where [simp del]: "log n = (if n < 2 then 0 else Suc (log (n div 2)))" | 
| 51174 | 16 | |
| 61831 | 17 | lemma log_induct [consumes 1, case_names one double]: | 
| 18 | fixes n :: nat | |
| 19 | assumes "n > 0" | |
| 20 | assumes one: "P 1" | |
| 21 | assumes double: "\<And>n. n \<ge> 2 \<Longrightarrow> P (n div 2) \<Longrightarrow> P n" | |
| 22 | shows "P n" | |
| 61975 | 23 | using \<open>n > 0\<close> proof (induct n rule: log.induct) | 
| 61831 | 24 | fix n | 
| 25 | assume "\<not> n < 2 \<Longrightarrow> | |
| 26 | 0 < n div 2 \<Longrightarrow> P (n div 2)" | |
| 27 | then have *: "n \<ge> 2 \<Longrightarrow> P (n div 2)" by simp | |
| 28 | assume "n > 0" | |
| 29 | show "P n" | |
| 30 | proof (cases "n = 1") | |
| 63540 | 31 | case True | 
| 32 | with one show ?thesis by simp | |
| 61831 | 33 | next | 
| 63540 | 34 | case False | 
| 35 | with \<open>n > 0\<close> have "n \<ge> 2" by auto | |
| 36 | with * have "P (n div 2)" . | |
| 37 | with \<open>n \<ge> 2\<close> show ?thesis by (rule double) | |
| 61831 | 38 | qed | 
| 39 | qed | |
| 40 | ||
| 59349 | 41 | lemma log_zero [simp]: "log 0 = 0" | 
| 51174 | 42 | by (simp add: log.simps) | 
| 43 | ||
| 59349 | 44 | lemma log_one [simp]: "log 1 = 0" | 
| 51174 | 45 | by (simp add: log.simps) | 
| 46 | ||
| 59349 | 47 | lemma log_Suc_zero [simp]: "log (Suc 0) = 0" | 
| 51174 | 48 | using log_one by simp | 
| 49 | ||
| 59349 | 50 | lemma log_rec: "n \<ge> 2 \<Longrightarrow> log n = Suc (log (n div 2))" | 
| 51174 | 51 | by (simp add: log.simps) | 
| 52 | ||
| 59349 | 53 | lemma log_twice [simp]: "n \<noteq> 0 \<Longrightarrow> log (2 * n) = Suc (log n)" | 
| 51174 | 54 | by (simp add: log_rec) | 
| 55 | ||
| 59349 | 56 | lemma log_half [simp]: "log (n div 2) = log n - 1" | 
| 51174 | 57 | proof (cases "n < 2") | 
| 58 | case True | |
| 59 | then have "n = 0 \<or> n = 1" by arith | |
| 60 | then show ?thesis by (auto simp del: One_nat_def) | |
| 61 | next | |
| 59349 | 62 | case False | 
| 63 | then show ?thesis by (simp add: log_rec) | |
| 51174 | 64 | qed | 
| 65 | ||
| 80621 
6c369fec315a
Migration of new material mostly about exp, ln
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 66 | lemma log_power [simp]: "log (2 ^ n) = n" | 
| 51174 | 67 | by (induct n) simp_all | 
| 68 | ||
| 59349 | 69 | lemma log_mono: "mono log" | 
| 51174 | 70 | proof | 
| 71 | fix m n :: nat | |
| 72 | assume "m \<le> n" | |
| 73 | then show "log m \<le> log n" | |
| 74 | proof (induct m arbitrary: n rule: log.induct) | |
| 75 | case (1 m) | |
| 76 | then have mn2: "m div 2 \<le> n div 2" by arith | |
| 77 | show "log m \<le> log n" | |
| 61831 | 78 | proof (cases "m \<ge> 2") | 
| 79 | case False | |
| 51174 | 80 | then have "m = 0 \<or> m = 1" by arith | 
| 81 | then show ?thesis by (auto simp del: One_nat_def) | |
| 82 | next | |
| 61831 | 83 | case True then have "\<not> m < 2" by simp | 
| 84 | with mn2 have "n \<ge> 2" by arith | |
| 85 | from True have m2_0: "m div 2 \<noteq> 0" by arith | |
| 51174 | 86 | with mn2 have n2_0: "n div 2 \<noteq> 0" by arith | 
| 61975 | 87 | from \<open>\<not> m < 2\<close> "1.hyps" mn2 have "log (m div 2) \<le> log (n div 2)" by blast | 
| 51174 | 88 | with m2_0 n2_0 have "log (2 * (m div 2)) \<le> log (2 * (n div 2))" by simp | 
| 60500 | 89 | with m2_0 n2_0 \<open>m \<ge> 2\<close> \<open>n \<ge> 2\<close> show ?thesis by (simp only: log_rec [of m] log_rec [of n]) simp | 
| 51174 | 90 | qed | 
| 91 | qed | |
| 92 | qed | |
| 93 | ||
| 61831 | 94 | lemma log_exp2_le: | 
| 95 | assumes "n > 0" | |
| 96 | shows "2 ^ log n \<le> n" | |
| 63516 | 97 | using assms | 
| 98 | proof (induct n rule: log_induct) | |
| 99 | case one | |
| 100 | then show ?case by simp | |
| 61831 | 101 | next | 
| 63516 | 102 | case (double n) | 
| 61831 | 103 | with log_mono have "log n \<ge> Suc 0" | 
| 104 | by (simp add: log.simps) | |
| 105 | assume "2 ^ log (n div 2) \<le> n div 2" | |
| 61975 | 106 | with \<open>n \<ge> 2\<close> have "2 ^ (log n - Suc 0) \<le> n div 2" by simp | 
| 61831 | 107 | then have "2 ^ (log n - Suc 0) * 2 ^ 1 \<le> n div 2 * 2" by simp | 
| 61975 | 108 | with \<open>log n \<ge> Suc 0\<close> have "2 ^ log n \<le> n div 2 * 2" | 
| 61831 | 109 | unfolding power_add [symmetric] by simp | 
| 110 | also have "n div 2 * 2 \<le> n" by (cases "even n") simp_all | |
| 63516 | 111 | finally show ?case . | 
| 61831 | 112 | qed | 
| 113 | ||
| 64449 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 114 | lemma log_exp2_gt: "2 * 2 ^ log n > n" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 115 | proof (cases "n > 0") | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 116 | case True | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 117 | thus ?thesis | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 118 | proof (induct n rule: log_induct) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 119 | case (double n) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 120 | thus ?case | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 121 | by (cases "even n") (auto elim!: evenE oddE simp: field_simps log.simps) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 122 | qed simp_all | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 123 | qed simp_all | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 124 | |
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 125 | lemma log_exp2_ge: "2 * 2 ^ log n \<ge> n" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 126 | using log_exp2_gt[of n] by simp | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 127 | |
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 128 | lemma log_le_iff: "m \<le> n \<Longrightarrow> log m \<le> log n" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 129 | by (rule monoD [OF log_mono]) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 130 | |
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 131 | lemma log_eqI: | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 132 | assumes "n > 0" "2^k \<le> n" "n < 2 * 2^k" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 133 | shows "log n = k" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 134 | proof (rule antisym) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 135 | from \<open>n > 0\<close> have "2 ^ log n \<le> n" by (rule log_exp2_le) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 136 | also have "\<dots> < 2 ^ Suc k" using assms by simp | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 137 | finally have "log n < Suc k" by (subst (asm) power_strict_increasing_iff) simp_all | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 138 | thus "log n \<le> k" by simp | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 139 | next | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 140 | have "2^k \<le> n" by fact | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 141 | also have "\<dots> < 2^(Suc (log n))" by (simp add: log_exp2_gt) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 142 | finally have "k < Suc (log n)" by (subst (asm) power_strict_increasing_iff) simp_all | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 143 | thus "k \<le> log n" by simp | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 144 | qed | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 145 | |
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 146 | lemma log_altdef: "log n = (if n = 0 then 0 else nat \<lfloor>Transcendental.log 2 (real_of_nat n)\<rfloor>)" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 147 | proof (cases "n = 0") | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 148 | case False | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 149 | have "\<lfloor>Transcendental.log 2 (real_of_nat n)\<rfloor> = int (log n)" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 150 | proof (rule floor_unique) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 151 | from False have "2 powr (real (log n)) \<le> real n" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 152 | by (simp add: powr_realpow log_exp2_le) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 153 | hence "Transcendental.log 2 (2 powr (real (log n))) \<le> Transcendental.log 2 (real n)" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 154 | using False by (subst Transcendental.log_le_cancel_iff) simp_all | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 155 | also have "Transcendental.log 2 (2 powr (real (log n))) = real (log n)" by simp | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 156 | finally show "real_of_int (int (log n)) \<le> Transcendental.log 2 (real n)" by simp | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 157 | next | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 158 | have "real n < real (2 * 2 ^ log n)" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 159 | by (subst of_nat_less_iff) (rule log_exp2_gt) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 160 | also have "\<dots> = 2 powr (real (log n) + 1)" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 161 | by (simp add: powr_add powr_realpow) | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 162 | finally have "Transcendental.log 2 (real n) < Transcendental.log 2 \<dots>" | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 163 | using False by (subst Transcendental.log_less_cancel_iff) simp_all | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 164 | also have "\<dots> = real (log n) + 1" by simp | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 165 | finally show "Transcendental.log 2 (real n) < real_of_int (int (log n)) + 1" by simp | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 166 | qed | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 167 | thus ?thesis by simp | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 168 | qed simp_all | 
| 
8c44dfb4ca8a
Merged natlog2 into Discrete.log
 eberlm <eberlm@in.tum.de> parents: 
63766diff
changeset | 169 | |
| 51174 | 170 | |
| 60500 | 171 | subsection \<open>Discrete square root\<close> | 
| 51174 | 172 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 173 | qualified definition sqrt :: "nat \<Rightarrow> nat" | 
| 59349 | 174 |   where "sqrt n = Max {m. m\<^sup>2 \<le> n}"
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 175 | |
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 176 | lemma sqrt_aux: | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 177 | fixes n :: nat | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 178 |   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 | 179 | proof - | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 180 |   { fix m
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 181 | assume "m\<^sup>2 \<le> n" | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 182 | then have "m \<le> n" | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 183 | by (cases m) (simp_all add: power2_eq_square) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 184 | } note ** = this | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 185 |   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 | 186 |   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 | 187 | 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 | 188 |   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 | 189 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 190 | |
| 63766 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 191 | lemma sqrt_unique: | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 192 | assumes "m^2 \<le> n" "n < (Suc m)^2" | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 193 | shows "Discrete.sqrt n = m" | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 194 | proof - | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 195 | have "m' \<le> m" if "m'^2 \<le> n" for m' | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 196 | proof - | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 197 | note that | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 198 | also note assms(2) | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 199 | finally have "m' < Suc m" by (rule power_less_imp_less_base) simp_all | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 200 | thus "m' \<le> m" by simp | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 201 | qed | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 202 | with \<open>m^2 \<le> n\<close> sqrt_aux[of n] show ?thesis unfolding Discrete.sqrt_def | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 203 | by (intro antisym Max.boundedI Max.coboundedI) simp_all | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 204 | qed | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 205 | |
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 206 | |
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 207 | lemma sqrt_code[code]: "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 | 208 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 209 |   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 | 210 | 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 | 211 | qed | 
| 51174 | 212 | |
| 59349 | 213 | lemma sqrt_inverse_power2 [simp]: "sqrt (n\<^sup>2) = n" | 
| 51174 | 214 | proof - | 
| 215 |   have "{m. m \<le> n} \<noteq> {}" by auto
 | |
| 216 |   then have "Max {m. m \<le> n} \<le> n" by auto
 | |
| 217 | then show ?thesis | |
| 218 | by (auto simp add: sqrt_def power2_nat_le_eq_le intro: antisym) | |
| 219 | qed | |
| 220 | ||
| 59349 | 221 | lemma sqrt_zero [simp]: "sqrt 0 = 0" | 
| 58787 | 222 | using sqrt_inverse_power2 [of 0] by simp | 
| 223 | ||
| 59349 | 224 | lemma sqrt_one [simp]: "sqrt 1 = 1" | 
| 58787 | 225 | using sqrt_inverse_power2 [of 1] by simp | 
| 226 | ||
| 59349 | 227 | lemma mono_sqrt: "mono sqrt" | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 228 | proof | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 229 | fix m n :: nat | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 230 | have *: "0 * 0 \<le> m" by simp | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 231 | assume "m \<le> n" | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 232 | then show "sqrt m \<le> sqrt n" | 
| 60500 | 233 | by (auto intro!: Max_mono \<open>0 * 0 \<le> m\<close> finite_less_ub simp add: power2_eq_square sqrt_def) | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 234 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 235 | |
| 63766 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 236 | lemma mono_sqrt': "m \<le> n \<Longrightarrow> Discrete.sqrt m \<le> Discrete.sqrt n" | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 237 | using mono_sqrt unfolding mono_def by auto | 
| 
695d60817cb1
Some facts about factorial and binomial coefficients
 Manuel Eberl <eberlm@in.tum.de> parents: 
63540diff
changeset | 238 | |
| 59349 | 239 | lemma sqrt_greater_zero_iff [simp]: "sqrt n > 0 \<longleftrightarrow> n > 0" | 
| 51174 | 240 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 241 |   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 | 242 | by (rule Max_gr_iff) (fact sqrt_aux)+ | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 243 | show ?thesis | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 244 | proof | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 245 | 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 | 246 |     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 | 247 | 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 | 248 | next | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 249 | assume "0 < n" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51263diff
changeset | 250 | 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 | 251 | 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 | 252 |     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 | 253 | then show "0 < sqrt n" by (simp add: sqrt_def) | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 254 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 255 | qed | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 256 | |
| 59349 | 257 | lemma sqrt_power2_le [simp]: "(sqrt n)\<^sup>2 \<le> n" (* FIXME tune proof *) | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 258 | proof (cases "n > 0") | 
| 58787 | 259 | case False then show ?thesis by simp | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 260 | next | 
| 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 261 | 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 | 262 |   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 | 263 |   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 | 264 | using sqrt_aux [of n] by (rule mono_Max_commute) | 
| 64919 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 265 |   have "\<And>a. a * a \<le> n \<Longrightarrow> Max {m. m * m \<le> n} * a \<le> n"
 | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 266 | proof - | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 267 | fix q | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 268 | assume "q * q \<le> n" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 269 |     show "Max {m. m * m \<le> n} * q \<le> n"
 | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 270 | proof (cases "q > 0") | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 271 | case False then show ?thesis by simp | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 272 | next | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 273 | case True then have "mono (times q)" by (rule mono_times_nat) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 274 |       then have "q * Max {m. m * m \<le> n} = Max (times q ` {m. m * m \<le> n})"
 | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 275 | using sqrt_aux [of n] by (auto simp add: power2_eq_square intro: mono_Max_commute) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 276 |       then have "Max {m. m * m \<le> n} * q = Max (times q ` {m. m * m \<le> n})" by (simp add: ac_simps)
 | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67399diff
changeset | 277 |       moreover have "finite ((*) q ` {m. m * m \<le> n})"
 | 
| 64919 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 278 | by (metis (mono_tags) finite_imageI finite_less_ub le_square) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 279 | moreover have "\<exists>x. x * x \<le> n" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 280 | by (metis \<open>q * q \<le> n\<close>) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 281 | ultimately show ?thesis | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 282 | by simp (metis \<open>q * q \<le> n\<close> le_cases mult_le_mono1 mult_le_mono2 order_trans) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 283 | qed | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 284 | qed | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67399diff
changeset | 285 |   then have "Max ((*) (Max {m. m * m \<le> n}) ` {m. m * m \<le> n}) \<le> n"
 | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 286 | apply (subst Max_le_iff) | 
| 64919 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 287 | apply (metis (mono_tags) finite_imageI finite_less_ub le_square) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 288 | apply auto | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 289 | apply (metis le0 mult_0_right) | 
| 64919 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 290 | done | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 291 | with * show ?thesis by (simp add: sqrt_def power2_eq_square) | 
| 51174 | 292 | qed | 
| 293 | ||
| 59349 | 294 | lemma sqrt_le: "sqrt n \<le> n" | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 295 | using sqrt_aux [of n] by (auto simp add: sqrt_def intro: power2_nat_le_imp_le) | 
| 51174 | 296 | |
| 64919 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 297 | text \<open>Additional facts about the discrete square root, thanks to Julian Biendarra, Manuel Eberl\<close> | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 298 | |
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 299 | lemma Suc_sqrt_power2_gt: "n < (Suc (Discrete.sqrt n))^2" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 300 | using Max_ge[OF Discrete.sqrt_aux(1), of "Discrete.sqrt n + 1" n] | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 301 | by (cases "n < (Suc (Discrete.sqrt n))^2") (simp_all add: Discrete.sqrt_def) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 302 | |
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 303 | lemma le_sqrt_iff: "x \<le> Discrete.sqrt y \<longleftrightarrow> x^2 \<le> y" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 304 | proof - | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 305 | have "x \<le> Discrete.sqrt y \<longleftrightarrow> (\<exists>z. z\<^sup>2 \<le> y \<and> x \<le> z)" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 306 | using Max_ge_iff[OF Discrete.sqrt_aux, of x y] by (simp add: Discrete.sqrt_def) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 307 | also have "\<dots> \<longleftrightarrow> x^2 \<le> y" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 308 | proof safe | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 309 | fix z assume "x \<le> z" "z ^ 2 \<le> y" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 310 | thus "x^2 \<le> y" by (intro le_trans[of "x^2" "z^2" y]) (simp_all add: power2_nat_le_eq_le) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 311 | qed auto | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 312 | finally show ?thesis . | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 313 | qed | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 314 | |
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 315 | lemma le_sqrtI: "x^2 \<le> y \<Longrightarrow> x \<le> Discrete.sqrt y" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 316 | by (simp add: le_sqrt_iff) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 317 | |
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 318 | lemma sqrt_le_iff: "Discrete.sqrt y \<le> x \<longleftrightarrow> (\<forall>z. z^2 \<le> y \<longrightarrow> z \<le> x)" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 319 | using Max.bounded_iff[OF Discrete.sqrt_aux] by (simp add: Discrete.sqrt_def) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 320 | |
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 321 | lemma sqrt_leI: | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 322 | "(\<And>z. z^2 \<le> y \<Longrightarrow> z \<le> x) \<Longrightarrow> Discrete.sqrt y \<le> x" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 323 | by (simp add: sqrt_le_iff) | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 324 | |
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 325 | lemma sqrt_Suc: | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 326 | "Discrete.sqrt (Suc n) = (if \<exists>m. Suc n = m^2 then Suc (Discrete.sqrt n) else Discrete.sqrt n)" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 327 | proof cases | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 328 | assume "\<exists> m. Suc n = m^2" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 329 | then obtain m where m_def: "Suc n = m^2" by blast | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 330 | then have lhs: "Discrete.sqrt (Suc n) = m" by simp | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 331 | from m_def sqrt_power2_le[of n] | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 332 | have "(Discrete.sqrt n)^2 < m^2" by linarith | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 333 | with power2_less_imp_less have lt_m: "Discrete.sqrt n < m" by blast | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 334 | from m_def Suc_sqrt_power2_gt[of "n"] | 
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 335 | have "m^2 \<le> (Suc(Discrete.sqrt n))^2" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 336 | by linarith | 
| 64919 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 337 | with power2_nat_le_eq_le have "m \<le> Suc (Discrete.sqrt n)" by blast | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 338 | with lt_m have "m = Suc (Discrete.sqrt n)" by simp | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 339 | with lhs m_def show ?thesis by fastforce | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 340 | next | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 341 | assume asm: "\<not> (\<exists> m. Suc n = m^2)" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 342 | hence "Suc n \<noteq> (Discrete.sqrt (Suc n))^2" by simp | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 343 | with sqrt_power2_le[of "Suc n"] | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 344 | have "Discrete.sqrt (Suc n) \<le> Discrete.sqrt n" by (intro le_sqrtI) linarith | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 345 | moreover have "Discrete.sqrt (Suc n) \<ge> Discrete.sqrt n" | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 346 | by (intro monoD[OF mono_sqrt]) simp_all | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 347 | ultimately show ?thesis using asm by simp | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 348 | qed | 
| 
7e0c8924dfda
New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
 paulson <lp15@cam.ac.uk> parents: 
64449diff
changeset | 349 | |
| 51174 | 350 | end | 
| 351 | ||
| 62390 | 352 | end |