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