# HG changeset patch # User paulson # Date 979121351 -3600 # Node ID 623141a08705c4a516a96a4a72a850a885c20f45 # Parent 3696bc935bbd00b2c37a56ab2071b92216752184 reformatting, and splitting the end of "Primes" to create "Forward" diff -r 3696bc935bbd -r 623141a08705 doc-src/TutorialI/Rules/Forward.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/TutorialI/Rules/Forward.thy Wed Jan 10 11:09:11 2001 +0100 @@ -0,0 +1,202 @@ +theory Forward = Primes: + +text{*\noindent +Forward proof material: of, OF, THEN, simplify, rule_format. +*} + +text{*\noindent +SKIP most developments... +*} + +(** Commutativity **) + +lemma is_gcd_commute: "is_gcd k m n = is_gcd k n m" + apply (auto simp add: is_gcd_def); + done + +lemma gcd_commute: "gcd(m,n) = gcd(n,m)" + apply (rule is_gcd_unique) + apply (rule is_gcd) + apply (subst is_gcd_commute) + apply (simp add: is_gcd) + done + +lemma gcd_1 [simp]: "gcd(m,1) = 1" + apply (simp) + done + +lemma gcd_1_left [simp]: "gcd(1,m) = 1" + apply (simp add: gcd_commute [of 1]) + done + +text{*\noindent +as far as HERE. +*} + + +text {* +@{thm[display] gcd_1} +\rulename{gcd_1} + +@{thm[display] gcd_1_left} +\rulename{gcd_1_left} +*}; + +text{*\noindent +SKIP THIS PROOF +*} + +lemma gcd_mult_distrib2: "k * gcd(m,n) = gcd(k*m, k*n)" +apply (induct_tac m n rule: gcd.induct) +apply (case_tac "n=0") +apply (simp) +apply (case_tac "k=0") +apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2) +done + +text {* +@{thm[display] gcd_mult_distrib2} +\rulename{gcd_mult_distrib2} +*}; + +text{*\noindent +of, simplified +*} + + +lemmas gcd_mult_0 = gcd_mult_distrib2 [of k 1]; +lemmas gcd_mult_1 = gcd_mult_0 [simplified]; + +text {* +@{thm[display] gcd_mult_distrib2 [of _ 1]} + +@{thm[display] gcd_mult_0} +\rulename{gcd_mult_0} + +@{thm[display] gcd_mult_1} +\rulename{gcd_mult_1} + +@{thm[display] sym} +\rulename{sym} +*}; + +lemmas gcd_mult = gcd_mult_1 [THEN sym]; + +lemmas gcd_mult = gcd_mult_distrib2 [of k 1, simplified, THEN sym]; + (*better in one step!*) + +text {* +more legible +*}; + +lemma gcd_mult [simp]: "gcd(k, k*n) = k" +by (rule gcd_mult_distrib2 [of k 1, simplified, THEN sym]) + + +lemmas gcd_self = gcd_mult [of k 1, simplified]; + + +text {* +Rules handy with THEN + +@{thm[display] iffD1} +\rulename{iffD1} + +@{thm[display] iffD2} +\rulename{iffD2} +*}; + + +text {* +again: more legible +*}; + +lemma gcd_self [simp]: "gcd(k,k) = k" +by (rule gcd_mult [of k 1, simplified]) + + +lemma relprime_dvd_mult: + "\ gcd(k,n)=1; k dvd m*n \ \ k dvd m"; +apply (insert gcd_mult_distrib2 [of m k n]) +apply (simp) +apply (erule_tac t="m" in ssubst); +apply (simp) +done + + +text {* +Another example of "insert" + +@{thm[display] mod_div_equality} +\rulename{mod_div_equality} +*}; + +lemma div_mult_self_is_m: + "0 (m*n) div n = (m::nat)" +apply (insert mod_div_equality [of "m*n" n]) +apply (simp) +done + +lemma relprime_dvd_mult_iff: "gcd(k,n)=1 \ (k dvd m*n) = (k dvd m)"; +by (blast intro: relprime_dvd_mult dvd_trans) + + +lemma relprime_20_81: "gcd(#20,#81) = 1"; +by (simp add: gcd.simps) + + + +text {* +Examples of 'OF' + +@{thm[display] relprime_dvd_mult} +\rulename{relprime_dvd_mult} + +@{thm[display] relprime_dvd_mult [OF relprime_20_81]} + +@{thm[display] dvd_refl} +\rulename{dvd_refl} + +@{thm[display] dvd_add} +\rulename{dvd_add} + +@{thm[display] dvd_add [OF dvd_refl dvd_refl]} + +@{thm[display] dvd_add [OF _ dvd_refl]} +*}; + +lemma "\(z::int) < #37; #66 < #2*z; z*z \ #1225; Q(#34); Q(#36)\ \ Q(z)"; +apply (subgoal_tac "z = #34 \ z = #36") +apply blast +apply (subgoal_tac "z \ #35") +apply arith +apply force +done + +text +{* +proof\ (prove):\ step\ 1\isanewline +\isanewline +goal\ (lemma):\isanewline +\isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36\isasymrbrakk \ \isasymLongrightarrow \ Q\ z\isanewline +\ 1.\ \isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36;\isanewline +\ \ \ \ \ \ \ z\ =\ \#34\ \isasymor \ z\ =\ \#36\isasymrbrakk \isanewline +\ \ \ \ \isasymLongrightarrow \ Q\ z\isanewline +\ 2.\ \isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36\isasymrbrakk \isanewline +\ \ \ \ \isasymLongrightarrow \ z\ =\ \#34\ \isasymor \ z\ =\ \#36 + + + +proof\ (prove):\ step\ 3\isanewline +\isanewline +goal\ (lemma):\isanewline +\isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36\isasymrbrakk \ \isasymLongrightarrow \ Q\ z\isanewline +\ 1.\ \isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36;\isanewline +\ \ \ \ \ \ \ z\ \isasymnoteq \ \#35\isasymrbrakk \isanewline +\ \ \ \ \isasymLongrightarrow \ z\ =\ \#34\ \isasymor \ z\ =\ \#36\isanewline +\ 2.\ \isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36\isasymrbrakk \isanewline +\ \ \ \ \isasymLongrightarrow \ z\ \isasymnoteq \ \#35 +*} + + +end diff -r 3696bc935bbd -r 623141a08705 doc-src/TutorialI/Rules/Primes.thy --- a/doc-src/TutorialI/Rules/Primes.thy Wed Jan 10 11:08:29 2001 +0100 +++ b/doc-src/TutorialI/Rules/Primes.thy Wed Jan 10 11:09:11 2001 +0100 @@ -1,9 +1,10 @@ (* ID: $Id$ *) (* EXTRACT from HOL/ex/Primes.thy*) +(*Euclid's algorithm *) theory Primes = Main: consts - gcd :: "nat*nat \ nat" (*Euclid's algorithm *) + gcd :: "nat*nat \ nat" recdef gcd "measure ((\(m,n).n) ::nat*nat \ nat)" "gcd (m, n) = (if n=0 then m else gcd(n, m mod n))" @@ -24,22 +25,22 @@ (*** Euclid's Algorithm ***) lemma gcd_0 [simp]: "gcd(m,0) = m" - apply (simp); - done +apply (simp); +done lemma gcd_non_0 [simp]: "0 gcd(m,n) = gcd (n, m mod n)" - apply (simp) - done; +apply (simp) +done; declare gcd.simps [simp del]; (*gcd(m,n) divides m and n. The conjunctions don't seem provable separately*) lemma gcd_dvd_both: "(gcd(m,n) dvd m) \ (gcd(m,n) dvd n)" - apply (induct_tac m n rule: gcd.induct) - apply (case_tac "n=0") - apply (simp_all) - apply (blast dest: dvd_mod_imp_dvd) - done +apply (induct_tac m n rule: gcd.induct) +apply (case_tac "n=0") +apply (simp_all) +by (blast dest: dvd_mod_imp_dvd) + text {* @@ -78,17 +79,18 @@ if k divides m and k divides n then k divides gcd(m,n)*) lemma gcd_greatest [rule_format]: "k dvd m \ k dvd n \ k dvd gcd(m,n)" - apply (induct_tac m n rule: gcd.induct) - apply (case_tac "n=0") - apply (simp_all add: dvd_mod); - done; +apply (induct_tac m n rule: gcd.induct) +apply (case_tac "n=0") +apply (simp_all add: dvd_mod) +done theorem gcd_greatest_iff [iff]: "(k dvd gcd(m,n)) = (k dvd m \ k dvd n)" - apply (blast intro!: gcd_greatest intro: dvd_trans); - done; +by (blast intro!: gcd_greatest intro: dvd_trans) +(**** The material below was omitted from the book ****) + constdefs is_gcd :: "[nat,nat,nat] \ bool" (*gcd as a relation*) "is_gcd p m n == p dvd m \ p dvd n \ @@ -96,14 +98,14 @@ (*Function gcd yields the Greatest Common Divisor*) lemma is_gcd: "is_gcd (gcd(m,n)) m n" - apply (simp add: is_gcd_def gcd_greatest); - done +apply (simp add: is_gcd_def gcd_greatest); +done (*uniqueness of GCDs*) lemma is_gcd_unique: "\ is_gcd m a b; is_gcd n a b \ \ m=n" - apply (simp add: is_gcd_def); - apply (blast intro: dvd_anti_sym) - done +apply (simp add: is_gcd_def); +apply (blast intro: dvd_anti_sym) +done text {* @@ -148,205 +150,4 @@ lemma gcd_mult_cancel: "gcd(k,n) = 1 \ gcd(k*m, n) = gcd(m,n)" oops - - -text{*\noindent -Forward proof material: of, OF, THEN, simplify. -*} - -text{*\noindent -SKIP most developments... -*} - -(** Commutativity **) - -lemma is_gcd_commute: "is_gcd k m n = is_gcd k n m" - apply (auto simp add: is_gcd_def); - done - -lemma gcd_commute: "gcd(m,n) = gcd(n,m)" - apply (rule is_gcd_unique) - apply (rule is_gcd) - apply (subst is_gcd_commute) - apply (simp add: is_gcd) - done - -lemma gcd_1 [simp]: "gcd(m,1) = 1" - apply (simp) - done - -lemma gcd_1_left [simp]: "gcd(1,m) = 1" - apply (simp add: gcd_commute [of 1]) - done - -text{*\noindent -as far as HERE. -*} - - -text {* -@{thm[display] gcd_1} -\rulename{gcd_1} - -@{thm[display] gcd_1_left} -\rulename{gcd_1_left} -*}; - -text{*\noindent -SKIP THIS PROOF -*} - -lemma gcd_mult_distrib2: "k * gcd(m,n) = gcd(k*m, k*n)" - apply (induct_tac m n rule: gcd.induct) - apply (case_tac "n=0") - apply (simp) - apply (case_tac "k=0") - apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2) - done - -text {* -@{thm[display] gcd_mult_distrib2} -\rulename{gcd_mult_distrib2} -*}; - -text{*\noindent -of, simplified -*} - - -lemmas gcd_mult_0 = gcd_mult_distrib2 [of k 1]; -lemmas gcd_mult_1 = gcd_mult_0 [simplified]; - -text {* -@{thm[display] gcd_mult_distrib2 [of _ 1]} - -@{thm[display] gcd_mult_0} -\rulename{gcd_mult_0} - -@{thm[display] gcd_mult_1} -\rulename{gcd_mult_1} - -@{thm[display] sym} -\rulename{sym} -*}; - -lemmas gcd_mult = gcd_mult_1 [THEN sym]; - -lemmas gcd_mult = gcd_mult_distrib2 [of k 1, simplified, THEN sym]; - (*better in one step!*) - -text {* -more legible -*}; - -lemma gcd_mult [simp]: "gcd(k, k*n) = k" - apply (rule gcd_mult_distrib2 [of k 1, simplified, THEN sym]) - done - -lemmas gcd_self = gcd_mult [of k 1, simplified]; - - -text {* -Rules handy with THEN - -@{thm[display] iffD1} -\rulename{iffD1} - -@{thm[display] iffD2} -\rulename{iffD2} -*}; - - -text {* -again: more legible -*}; - -lemma gcd_self [simp]: "gcd(k,k) = k" - apply (rule gcd_mult [of k 1, simplified]) - done - -lemma relprime_dvd_mult: - "\ gcd(k,n)=1; k dvd m*n \ \ k dvd m"; - apply (insert gcd_mult_distrib2 [of m k n]) - apply (simp) - apply (erule_tac t="m" in ssubst); - apply (simp) - done - - -text {* -Another example of "insert" - -@{thm[display] mod_div_equality} -\rulename{mod_div_equality} -*}; - -lemma div_mult_self_is_m: - "0 (m*n) div n = (m::nat)" - apply (insert mod_div_equality [of "m*n" n]) - apply (simp) - done - -lemma relprime_dvd_mult_iff: "gcd(k,n)=1 \ (k dvd m*n) = (k dvd m)"; - apply (blast intro: relprime_dvd_mult dvd_trans) - done - -lemma relprime_20_81: "gcd(#20,#81) = 1"; - apply (simp add: gcd.simps) - done - - -text {* -Examples of 'OF' - -@{thm[display] relprime_dvd_mult} -\rulename{relprime_dvd_mult} - -@{thm[display] relprime_dvd_mult [OF relprime_20_81]} - -@{thm[display] dvd_refl} -\rulename{dvd_refl} - -@{thm[display] dvd_add} -\rulename{dvd_add} - -@{thm[display] dvd_add [OF dvd_refl dvd_refl]} - -@{thm[display] dvd_add [OF _ dvd_refl]} -*}; - -lemma "\(z::int) < #37; #66 < #2*z; z*z \ #1225; Q(#34); Q(#36)\ \ Q(z)"; -apply (subgoal_tac "z = #34 \ z = #36") -apply blast -apply (subgoal_tac "z \ #35") -apply arith -apply force -done - -text -{* -proof\ (prove):\ step\ 1\isanewline -\isanewline -goal\ (lemma):\isanewline -\isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36\isasymrbrakk \ \isasymLongrightarrow \ Q\ z\isanewline -\ 1.\ \isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36;\isanewline -\ \ \ \ \ \ \ z\ =\ \#34\ \isasymor \ z\ =\ \#36\isasymrbrakk \isanewline -\ \ \ \ \isasymLongrightarrow \ Q\ z\isanewline -\ 2.\ \isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36\isasymrbrakk \isanewline -\ \ \ \ \isasymLongrightarrow \ z\ =\ \#34\ \isasymor \ z\ =\ \#36 - - - -proof\ (prove):\ step\ 3\isanewline -\isanewline -goal\ (lemma):\isanewline -\isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36\isasymrbrakk \ \isasymLongrightarrow \ Q\ z\isanewline -\ 1.\ \isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36;\isanewline -\ \ \ \ \ \ \ z\ \isasymnoteq \ \#35\isasymrbrakk \isanewline -\ \ \ \ \isasymLongrightarrow \ z\ =\ \#34\ \isasymor \ z\ =\ \#36\isanewline -\ 2.\ \isasymlbrakk z\ <\ \#37;\ \#66\ <\ \#2\ *\ z;\ z\ *\ z\ \isasymnoteq \ \#1225;\ Q\ \#34;\ Q\ \#36\isasymrbrakk \isanewline -\ \ \ \ \isasymLongrightarrow \ z\ \isasymnoteq \ \#35 -*} - - end