| author | wenzelm | 
| Sun, 15 Dec 2024 14:59:57 +0100 | |
| changeset 81595 | ed264056f5dc | 
| parent 76987 | 4c275405faae | 
| permissions | -rw-r--r-- | 
| 39157 
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
 wenzelm parents: 
37598diff
changeset | 1 | (* Title: HOL/Proofs/Extraction/Euclid.thy | 
| 25422 | 2 | Author: Markus Wenzel, TU Muenchen | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 3 | Author: Freek Wiedijk, Radboud University Nijmegen | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 4 | Author: Stefan Berghofer, TU Muenchen | 
| 25422 | 5 | *) | 
| 6 | ||
| 61986 | 7 | section \<open>Euclid's theorem\<close> | 
| 25422 | 8 | |
| 9 | theory Euclid | |
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
39157diff
changeset | 10 | imports | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66258diff
changeset | 11 | "HOL-Computational_Algebra.Primes" | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
39157diff
changeset | 12 | Util | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66258diff
changeset | 13 | "HOL-Library.Code_Target_Numeral" | 
| 67320 | 14 | "HOL-Library.Realizers" | 
| 25422 | 15 | begin | 
| 16 | ||
| 61986 | 17 | text \<open> | 
| 63361 | 18 | A constructive version of the proof of Euclid's theorem by | 
| 76987 | 19 | Markus Wenzel and Freek Wiedijk \<^cite>\<open>"Wenzel-Wiedijk-JAR2002"\<close>. | 
| 61986 | 20 | \<close> | 
| 25422 | 21 | |
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 22 | lemma factor_greater_one1: "n = m * k \<Longrightarrow> m < n \<Longrightarrow> k < n \<Longrightarrow> Suc 0 < m" | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 23 | by (induct m) auto | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 24 | |
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 25 | lemma factor_greater_one2: "n = m * k \<Longrightarrow> m < n \<Longrightarrow> k < n \<Longrightarrow> Suc 0 < k" | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 26 | by (induct k) auto | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 27 | |
| 63361 | 28 | lemma prod_mn_less_k: "0 < n \<Longrightarrow> 0 < k \<Longrightarrow> Suc 0 < m \<Longrightarrow> m * n = k \<Longrightarrow> n < k" | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 29 | by (induct m) auto | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 30 | |
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63361diff
changeset | 31 | lemma prime_eq: "prime (p::nat) \<longleftrightarrow> 1 < p \<and> (\<forall>m. m dvd p \<longrightarrow> 1 < m \<longrightarrow> m = p)" | 
| 63633 | 32 | apply (simp add: prime_nat_iff) | 
| 25422 | 33 | apply (rule iffI) | 
| 34 | apply blast | |
| 35 | apply (erule conjE) | |
| 36 | apply (rule conjI) | |
| 37 | apply assumption | |
| 38 | apply (rule allI impI)+ | |
| 39 | apply (erule allE) | |
| 40 | apply (erule impE) | |
| 41 | apply assumption | |
| 63361 | 42 | apply (case_tac "m = 0") | 
| 25422 | 43 | apply simp | 
| 63361 | 44 | apply (case_tac "m = Suc 0") | 
| 25422 | 45 | apply simp | 
| 46 | apply simp | |
| 47 | done | |
| 48 | ||
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63361diff
changeset | 49 | lemma prime_eq': "prime (p::nat) \<longleftrightarrow> 1 < p \<and> (\<forall>m k. p = m * k \<longrightarrow> 1 < m \<longrightarrow> m = p)" | 
| 37598 | 50 | by (simp add: prime_eq dvd_def HOL.all_simps [symmetric] del: HOL.all_simps) | 
| 25422 | 51 | |
| 52 | lemma not_prime_ex_mk: | |
| 53 | assumes n: "Suc 0 < n" | |
| 54 | shows "(\<exists>m k. Suc 0 < m \<and> Suc 0 < k \<and> m < n \<and> k < n \<and> n = m * k) \<or> prime n" | |
| 55 | proof - | |
| 63361 | 56 | from nat_eq_dec have "(\<exists>m<n. n = m * k) \<or> \<not> (\<exists>m<n. n = m * k)" for k | 
| 25422 | 57 | by (rule search) | 
| 63361 | 58 | then have "(\<exists>k<n. \<exists>m<n. n = m * k) \<or> \<not> (\<exists>k<n. \<exists>m<n. n = m * k)" | 
| 59 | by (rule search) | |
| 60 | then show ?thesis | |
| 25422 | 61 | proof | 
| 62 | assume "\<exists>k<n. \<exists>m<n. n = m * k" | |
| 63 | then obtain k m where k: "k<n" and m: "m<n" and nmk: "n = m * k" | |
| 64 | by iprover | |
| 65 | from nmk m k have "Suc 0 < m" by (rule factor_greater_one1) | |
| 66 | moreover from nmk m k have "Suc 0 < k" by (rule factor_greater_one2) | |
| 67 | ultimately show ?thesis using k m nmk by iprover | |
| 68 | next | |
| 69 | assume "\<not> (\<exists>k<n. \<exists>m<n. n = m * k)" | |
| 63361 | 70 | then have A: "\<forall>k<n. \<forall>m<n. n \<noteq> m * k" by iprover | 
| 25422 | 71 | have "\<forall>m k. n = m * k \<longrightarrow> Suc 0 < m \<longrightarrow> m = n" | 
| 72 | proof (intro allI impI) | |
| 73 | fix m k | |
| 74 | assume nmk: "n = m * k" | |
| 75 | assume m: "Suc 0 < m" | |
| 76 | from n m nmk have k: "0 < k" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 77 | by (cases k) auto | 
| 25422 | 78 | moreover from n have n: "0 < n" by simp | 
| 79 | moreover note m | |
| 80 | moreover from nmk have "m * k = n" by simp | |
| 81 | ultimately have kn: "k < n" by (rule prod_mn_less_k) | |
| 82 | show "m = n" | |
| 83 | proof (cases "k = Suc 0") | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 84 | case True | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 85 | with nmk show ?thesis by (simp only: mult_Suc_right) | 
| 25422 | 86 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 87 | case False | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 88 | from m have "0 < m" by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 89 | moreover note n | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 90 | moreover from False n nmk k have "Suc 0 < k" by auto | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 91 | moreover from nmk have "k * m = n" by (simp only: ac_simps) | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 92 | ultimately have mn: "m < n" by (rule prod_mn_less_k) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 93 | with kn A nmk show ?thesis by iprover | 
| 25422 | 94 | qed | 
| 95 | qed | |
| 96 | with n have "prime n" | |
| 97 | by (simp only: prime_eq' One_nat_def simp_thms) | |
| 63361 | 98 | then show ?thesis .. | 
| 25422 | 99 | qed | 
| 100 | qed | |
| 101 | ||
| 63361 | 102 | lemma dvd_factorial: "0 < m \<Longrightarrow> m \<le> n \<Longrightarrow> m dvd fact n" | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 103 | proof (induct n rule: nat_induct) | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 104 | case 0 | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 105 | then show ?case by simp | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 106 | next | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 107 | case (Suc n) | 
| 61986 | 108 | from \<open>m \<le> Suc n\<close> show ?case | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 109 | proof (rule le_SucE) | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 110 | assume "m \<le> n" | 
| 61986 | 111 | with \<open>0 < m\<close> have "m dvd fact n" by (rule Suc) | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 112 | then have "m dvd (fact n * Suc n)" by (rule dvd_mult2) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
51143diff
changeset | 113 | then show ?thesis by (simp add: mult.commute) | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 114 | next | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 115 | assume "m = Suc n" | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 116 | then have "m dvd (fact n * Suc n)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 117 | by (auto intro: dvdI simp: ac_simps) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
51143diff
changeset | 118 | then show ?thesis by (simp add: mult.commute) | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 119 | qed | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 120 | qed | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 121 | |
| 61954 | 122 | lemma dvd_prod [iff]: "n dvd (\<Prod>m::nat \<in># mset (n # ns). m)" | 
| 63830 | 123 | by (simp add: prod_mset_Un) | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 124 | |
| 63361 | 125 | definition all_prime :: "nat list \<Rightarrow> bool" | 
| 126 | where "all_prime ps \<longleftrightarrow> (\<forall>p\<in>set ps. prime p)" | |
| 37335 | 127 | |
| 128 | lemma all_prime_simps: | |
| 129 | "all_prime []" | |
| 130 | "all_prime (p # ps) \<longleftrightarrow> prime p \<and> all_prime ps" | |
| 131 | by (simp_all add: all_prime_def) | |
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 132 | |
| 63361 | 133 | lemma all_prime_append: "all_prime (ps @ qs) \<longleftrightarrow> all_prime ps \<and> all_prime qs" | 
| 37335 | 134 | by (simp add: all_prime_def ball_Un) | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 135 | |
| 37335 | 136 | lemma split_all_prime: | 
| 137 | assumes "all_prime ms" and "all_prime ns" | |
| 63361 | 138 | shows "\<exists>qs. all_prime qs \<and> | 
| 139 | (\<Prod>m::nat \<in># mset qs. m) = (\<Prod>m::nat \<in># mset ms. m) * (\<Prod>m::nat \<in># mset ns. m)" | |
| 140 | (is "\<exists>qs. ?P qs \<and> ?Q qs") | |
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 141 | proof - | 
| 37335 | 142 | from assms have "all_prime (ms @ ns)" | 
| 143 | by (simp add: all_prime_append) | |
| 63361 | 144 | moreover | 
| 145 | have "(\<Prod>m::nat \<in># mset (ms @ ns). m) = (\<Prod>m::nat \<in># mset ms. m) * (\<Prod>m::nat \<in># mset ns. m)" | |
| 63830 | 146 | using assms by (simp add: prod_mset_Un) | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 147 | ultimately have "?P (ms @ ns) \<and> ?Q (ms @ ns)" .. | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 148 | then show ?thesis .. | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 149 | qed | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 150 | |
| 37335 | 151 | lemma all_prime_nempty_g_one: | 
| 152 | assumes "all_prime ps" and "ps \<noteq> []" | |
| 61954 | 153 | shows "Suc 0 < (\<Prod>m::nat \<in># mset ps. m)" | 
| 63361 | 154 | using \<open>ps \<noteq> []\<close> \<open>all_prime ps\<close> | 
| 155 | unfolding One_nat_def [symmetric] | |
| 156 | by (induct ps rule: list_nonempty_induct) | |
| 63830 | 157 | (simp_all add: all_prime_simps prod_mset_Un prime_gt_1_nat less_1_mult del: One_nat_def) | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 158 | |
| 61954 | 159 | lemma factor_exists: "Suc 0 < n \<Longrightarrow> (\<exists>ps. all_prime ps \<and> (\<Prod>m::nat \<in># mset ps. m) = n)" | 
| 25422 | 160 | proof (induct n rule: nat_wf_ind) | 
| 161 | case (1 n) | |
| 61986 | 162 | from \<open>Suc 0 < n\<close> | 
| 25422 | 163 | have "(\<exists>m k. Suc 0 < m \<and> Suc 0 < k \<and> m < n \<and> k < n \<and> n = m * k) \<or> prime n" | 
| 164 | by (rule not_prime_ex_mk) | |
| 165 | then show ?case | |
| 63361 | 166 | proof | 
| 25422 | 167 | assume "\<exists>m k. Suc 0 < m \<and> Suc 0 < k \<and> m < n \<and> k < n \<and> n = m * k" | 
| 168 | then obtain m k where m: "Suc 0 < m" and k: "Suc 0 < k" and mn: "m < n" | |
| 63361 | 169 | and kn: "k < n" and nmk: "n = m * k" | 
| 170 | by iprover | |
| 171 | from mn and m have "\<exists>ps. all_prime ps \<and> (\<Prod>m::nat \<in># mset ps. m) = m" | |
| 172 | by (rule 1) | |
| 61954 | 173 | then obtain ps1 where "all_prime ps1" and prod_ps1_m: "(\<Prod>m::nat \<in># mset ps1. m) = m" | 
| 25422 | 174 | by iprover | 
| 63361 | 175 | from kn and k have "\<exists>ps. all_prime ps \<and> (\<Prod>m::nat \<in># mset ps. m) = k" | 
| 176 | by (rule 1) | |
| 61954 | 177 | then obtain ps2 where "all_prime ps2" and prod_ps2_k: "(\<Prod>m::nat \<in># mset ps2. m) = k" | 
| 25422 | 178 | by iprover | 
| 61986 | 179 | from \<open>all_prime ps1\<close> \<open>all_prime ps2\<close> | 
| 61954 | 180 | have "\<exists>ps. all_prime ps \<and> (\<Prod>m::nat \<in># mset ps. m) = | 
| 181 | (\<Prod>m::nat \<in># mset ps1. m) * (\<Prod>m::nat \<in># mset ps2. m)" | |
| 37335 | 182 | by (rule split_all_prime) | 
| 183 | with prod_ps1_m prod_ps2_k nmk show ?thesis by simp | |
| 25422 | 184 | next | 
| 37335 | 185 | assume "prime n" then have "all_prime [n]" by (simp add: all_prime_simps) | 
| 63830 | 186 | moreover have "(\<Prod>m::nat \<in># mset [n]. m) = n" by (simp) | 
| 61954 | 187 | ultimately have "all_prime [n] \<and> (\<Prod>m::nat \<in># mset [n]. m) = n" .. | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 188 | then show ?thesis .. | 
| 25422 | 189 | qed | 
| 190 | qed | |
| 191 | ||
| 192 | lemma prime_factor_exists: | |
| 193 | assumes N: "(1::nat) < n" | |
| 194 | shows "\<exists>p. prime p \<and> p dvd n" | |
| 195 | proof - | |
| 63361 | 196 | from N obtain ps where "all_prime ps" and prod_ps: "n = (\<Prod>m::nat \<in># mset ps. m)" | 
| 197 | using factor_exists by simp iprover | |
| 37335 | 198 | with N have "ps \<noteq> []" | 
| 63361 | 199 | by (auto simp add: all_prime_nempty_g_one) | 
| 200 | then obtain p qs where ps: "ps = p # qs" | |
| 201 | by (cases ps) simp | |
| 202 | with \<open>all_prime ps\<close> have "prime p" | |
| 203 | by (simp add: all_prime_simps) | |
| 204 | moreover from \<open>all_prime ps\<close> ps prod_ps have "p dvd n" | |
| 205 | by (simp only: dvd_prod) | |
| 25422 | 206 | ultimately show ?thesis by iprover | 
| 207 | qed | |
| 208 | ||
| 63361 | 209 | text \<open>Euclid's theorem: there are infinitely many primes.\<close> | 
| 25422 | 210 | |
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63361diff
changeset | 211 | lemma Euclid: "\<exists>p::nat. prime p \<and> n < p" | 
| 25422 | 212 | proof - | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 213 | let ?k = "fact n + (1::nat)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 214 | have "1 < ?k" by simp | 
| 63361 | 215 | then obtain p where prime: "prime p" and dvd: "p dvd ?k" | 
| 216 | using prime_factor_exists by iprover | |
| 25422 | 217 | have "n < p" | 
| 218 | proof - | |
| 219 | have "\<not> p \<le> n" | |
| 220 | proof | |
| 221 | assume pn: "p \<le> n" | |
| 61986 | 222 | from \<open>prime p\<close> have "0 < p" by (rule prime_gt_0_nat) | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 223 | then have "p dvd fact n" using pn by (rule dvd_factorial) | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 224 | with dvd have "p dvd ?k - fact n" by (rule dvd_diff_nat) | 
| 25422 | 225 | then have "p dvd 1" by simp | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 226 | with prime show False by auto | 
| 25422 | 227 | qed | 
| 228 | then show ?thesis by simp | |
| 229 | qed | |
| 230 | with prime show ?thesis by iprover | |
| 231 | qed | |
| 232 | ||
| 233 | extract Euclid | |
| 234 | ||
| 61986 | 235 | text \<open> | 
| 63361 | 236 | The program extracted from the proof of Euclid's theorem looks as follows. | 
| 237 |   @{thm [display] Euclid_def}
 | |
| 238 | The program corresponding to the proof of the factorization theorem is | |
| 239 |   @{thm [display] factor_exists_def}
 | |
| 61986 | 240 | \<close> | 
| 25422 | 241 | |
| 27982 | 242 | instantiation nat :: default | 
| 243 | begin | |
| 244 | ||
| 245 | definition "default = (0::nat)" | |
| 246 | ||
| 247 | instance .. | |
| 248 | ||
| 249 | end | |
| 25422 | 250 | |
| 27982 | 251 | instantiation list :: (type) default | 
| 252 | begin | |
| 253 | ||
| 254 | definition "default = []" | |
| 255 | ||
| 256 | instance .. | |
| 257 | ||
| 258 | end | |
| 259 | ||
| 63361 | 260 | primrec iterate :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list"
 | 
| 261 | where | |
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 262 | "iterate 0 f x = []" | 
| 63361 | 263 | | "iterate (Suc n) f x = (let y = f x in y # iterate n f y)" | 
| 37288 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 264 | |
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 265 | lemma "factor_exists 1007 = [53, 19]" by eval | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 266 | lemma "factor_exists 567 = [7, 3, 3, 3, 3]" by eval | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 267 | lemma "factor_exists 345 = [23, 5, 3]" by eval | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 268 | lemma "factor_exists 999 = [37, 3, 3, 3]" by eval | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 269 | lemma "factor_exists 876 = [73, 3, 2, 2]" by eval | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 270 | |
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 271 | lemma "iterate 4 Euclid 0 = [2, 3, 7, 71]" by eval | 
| 
2b1c6dd48995
removed dependency of Euclid on Old_Number_Theory
 haftmann parents: 
32960diff
changeset | 272 | |
| 25422 | 273 | end |