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