| author | haftmann | 
| Mon, 25 Sep 2006 17:04:47 +0200 | |
| changeset 20709 | 645236e80885 | 
| parent 16417 | 9bc16273c2d4 | 
| child 22097 | 7ee0529c5674 | 
| permissions | -rw-r--r-- | 
| 10341 | 1 | (* ID: $Id$ *) | 
| 10295 | 2 | (* EXTRACT from HOL/ex/Primes.thy*) | 
| 3 | ||
| 11080 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 4 | (*Euclid's algorithm | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 5 | This material now appears AFTER that of Forward.thy *) | 
| 16417 | 6 | theory Primes imports Main begin | 
| 10295 | 7 | consts | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 8 | gcd :: "nat*nat \<Rightarrow> nat" | 
| 10295 | 9 | |
| 10986 | 10 | recdef gcd "measure snd" | 
| 10295 | 11 | "gcd (m, n) = (if n=0 then m else gcd(n, m mod n))" | 
| 12 | ||
| 13 | ||
| 14 | ML "Pretty.setmargin 64" | |
| 15 | ML "IsarOutput.indent := 5" (*that is, Doc/TutorialI/settings.ML*) | |
| 16 | ||
| 17 | ||
| 11234 | 18 | text {*Now in Basic.thy!
 | 
| 10295 | 19 | @{thm[display]"dvd_def"}
 | 
| 20 | \rulename{dvd_def}
 | |
| 21 | *}; | |
| 22 | ||
| 23 | ||
| 24 | (*** Euclid's Algorithm ***) | |
| 25 | ||
| 26 | lemma gcd_0 [simp]: "gcd(m,0) = m" | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 27 | apply (simp); | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 28 | done | 
| 10295 | 29 | |
| 30 | lemma gcd_non_0 [simp]: "0<n \<Longrightarrow> gcd(m,n) = gcd (n, m mod n)" | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 31 | apply (simp) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 32 | done; | 
| 10295 | 33 | |
| 34 | declare gcd.simps [simp del]; | |
| 35 | ||
| 36 | (*gcd(m,n) divides m and n. The conjunctions don't seem provable separately*) | |
| 37 | lemma gcd_dvd_both: "(gcd(m,n) dvd m) \<and> (gcd(m,n) dvd n)" | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 38 | apply (induct_tac m n rule: gcd.induct) | 
| 11080 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 39 |   --{* @{subgoals[display,indent=0,margin=65]} *}
 | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 40 | apply (case_tac "n=0") | 
| 11080 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 41 | txt{*subgoals after the case tac
 | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 42 | @{subgoals[display,indent=0,margin=65]}
 | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 43 | *}; | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 44 | apply (simp_all) | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 45 |   --{* @{subgoals[display,indent=0,margin=65]} *}
 | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 46 | by (blast dest: dvd_mod_imp_dvd) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 47 | |
| 10295 | 48 | |
| 49 | ||
| 50 | text {*
 | |
| 51 | @{thm[display] dvd_mod_imp_dvd}
 | |
| 52 | \rulename{dvd_mod_imp_dvd}
 | |
| 53 | ||
| 54 | @{thm[display] dvd_trans}
 | |
| 55 | \rulename{dvd_trans}
 | |
| 11080 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 56 | *} | 
| 10295 | 57 | |
| 58 | lemmas gcd_dvd1 [iff] = gcd_dvd_both [THEN conjunct1] | |
| 59 | lemmas gcd_dvd2 [iff] = gcd_dvd_both [THEN conjunct2]; | |
| 60 | ||
| 61 | ||
| 62 | text {*
 | |
| 63 | \begin{quote}
 | |
| 64 | @{thm[display] gcd_dvd1}
 | |
| 65 | \rulename{gcd_dvd1}
 | |
| 66 | ||
| 67 | @{thm[display] gcd_dvd2}
 | |
| 68 | \rulename{gcd_dvd2}
 | |
| 69 | \end{quote}
 | |
| 70 | *}; | |
| 71 | ||
| 72 | (*Maximality: for all m,n,k naturals, | |
| 73 | if k divides m and k divides n then k divides gcd(m,n)*) | |
| 74 | lemma gcd_greatest [rule_format]: | |
| 10790 | 75 | "k dvd m \<longrightarrow> k dvd n \<longrightarrow> k dvd gcd(m,n)" | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 76 | apply (induct_tac m n rule: gcd.induct) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 77 | apply (case_tac "n=0") | 
| 10853 | 78 | txt{*subgoals after the case tac
 | 
| 79 | @{subgoals[display,indent=0,margin=65]}
 | |
| 80 | *}; | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 81 | apply (simp_all add: dvd_mod) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 82 | done | 
| 10295 | 83 | |
| 11080 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 84 | text {*
 | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 85 | @{thm[display] dvd_mod}
 | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 86 | \rulename{dvd_mod}
 | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 87 | *} | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 88 | |
| 10853 | 89 | (*just checking the claim that case_tac "n" works too*) | 
| 90 | lemma "k dvd m \<longrightarrow> k dvd n \<longrightarrow> k dvd gcd(m,n)" | |
| 91 | apply (induct_tac m n rule: gcd.induct) | |
| 92 | apply (case_tac "n") | |
| 11080 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 93 | apply (simp_all add: dvd_mod) | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 94 | done | 
| 
22855d091249
various revisions in response to comments from Tobias
 paulson parents: 
10986diff
changeset | 95 | |
| 10853 | 96 | |
| 10295 | 97 | theorem gcd_greatest_iff [iff]: | 
| 10795 | 98 | "(k dvd gcd(m,n)) = (k dvd m \<and> k dvd n)" | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 99 | by (blast intro!: gcd_greatest intro: dvd_trans) | 
| 10295 | 100 | |
| 101 | ||
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 102 | (**** The material below was omitted from the book ****) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 103 | |
| 10295 | 104 | constdefs | 
| 10300 | 105 | is_gcd :: "[nat,nat,nat] \<Rightarrow> bool" (*gcd as a relation*) | 
| 10295 | 106 | "is_gcd p m n == p dvd m \<and> p dvd n \<and> | 
| 107 | (ALL d. d dvd m \<and> d dvd n \<longrightarrow> d dvd p)" | |
| 108 | ||
| 109 | (*Function gcd yields the Greatest Common Divisor*) | |
| 110 | lemma is_gcd: "is_gcd (gcd(m,n)) m n" | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 111 | apply (simp add: is_gcd_def gcd_greatest); | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 112 | done | 
| 10295 | 113 | |
| 114 | (*uniqueness of GCDs*) | |
| 115 | lemma is_gcd_unique: "\<lbrakk> is_gcd m a b; is_gcd n a b \<rbrakk> \<Longrightarrow> m=n" | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 116 | apply (simp add: is_gcd_def); | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 117 | apply (blast intro: dvd_anti_sym) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: 
10795diff
changeset | 118 | done | 
| 10295 | 119 | |
| 120 | ||
| 121 | text {*
 | |
| 122 | @{thm[display] dvd_anti_sym}
 | |
| 123 | \rulename{dvd_anti_sym}
 | |
| 124 | ||
| 125 | \begin{isabelle}
 | |
| 126 | proof\ (prove):\ step\ 1\isanewline | |
| 127 | \isanewline | |
| 128 | goal\ (lemma\ is_gcd_unique):\isanewline | |
| 129 | \isasymlbrakk is_gcd\ m\ a\ b;\ is_gcd\ n\ a\ b\isasymrbrakk \ \isasymLongrightarrow \ m\ =\ n\isanewline | |
| 130 | \ 1.\ \isasymlbrakk m\ dvd\ a\ \isasymand \ m\ dvd\ b\ \isasymand \ (\isasymforall d.\ d\ dvd\ a\ \isasymand \ d\ dvd\ b\ \isasymlongrightarrow \ d\ dvd\ m);\isanewline | |
| 131 | \ \ \ \ \ \ \ n\ dvd\ a\ \isasymand \ n\ dvd\ b\ \isasymand \ (\isasymforall d.\ d\ dvd\ a\ \isasymand \ d\ dvd\ b\ \isasymlongrightarrow \ d\ dvd\ n)\isasymrbrakk \isanewline | |
| 132 | \ \ \ \ \isasymLongrightarrow \ m\ =\ n | |
| 133 | \end{isabelle}
 | |
| 134 | *}; | |
| 135 | ||
| 136 | lemma gcd_assoc: "gcd(gcd(k,m),n) = gcd(k,gcd(m,n))" | |
| 137 | apply (rule is_gcd_unique) | |
| 138 | apply (rule is_gcd) | |
| 139 | apply (simp add: is_gcd_def); | |
| 140 | apply (blast intro: dvd_trans); | |
| 141 | done | |
| 142 | ||
| 143 | text{*
 | |
| 144 | \begin{isabelle}
 | |
| 145 | proof\ (prove):\ step\ 3\isanewline | |
| 146 | \isanewline | |
| 147 | goal\ (lemma\ gcd_assoc):\isanewline | |
| 148 | gcd\ (gcd\ (k,\ m),\ n)\ =\ gcd\ (k,\ gcd\ (m,\ n))\isanewline | |
| 149 | \ 1.\ gcd\ (k,\ gcd\ (m,\ n))\ dvd\ k\ \isasymand \isanewline | |
| 150 | \ \ \ \ gcd\ (k,\ gcd\ (m,\ n))\ dvd\ m\ \isasymand \ gcd\ (k,\ gcd\ (m,\ n))\ dvd\ n | |
| 151 | \end{isabelle}
 | |
| 152 | *} | |
| 153 | ||
| 154 | ||
| 155 | lemma gcd_dvd_gcd_mult: "gcd(m,n) dvd gcd(k*m, n)" | |
| 156 | apply (blast intro: dvd_trans); | |
| 157 | done | |
| 158 | ||
| 159 | (*This is half of the proof (by dvd_anti_sym) of*) | |
| 160 | lemma gcd_mult_cancel: "gcd(k,n) = 1 \<Longrightarrow> gcd(k*m, n) = gcd(m,n)" | |
| 161 | oops | |
| 162 | ||
| 163 | end |