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