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