author | wenzelm |
Tue, 08 Jan 2002 20:45:45 +0100 | |
changeset 12676 | 99e7b2045d8f |
parent 12390 | 2fa13b499975 |
child 13550 | 5a176b8dda84 |
permissions | -rw-r--r-- |
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
1 |
theory Forward = Primes: |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
2 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
3 |
text{*\noindent |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
4 |
Forward proof material: of, OF, THEN, simplify, rule_format. |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
5 |
*} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
6 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
7 |
text{*\noindent |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
8 |
SKIP most developments... |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
9 |
*} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
10 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
11 |
(** Commutativity **) |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
12 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
13 |
lemma is_gcd_commute: "is_gcd k m n = is_gcd k n m" |
10958 | 14 |
apply (auto simp add: is_gcd_def); |
15 |
done |
|
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
16 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
17 |
lemma gcd_commute: "gcd(m,n) = gcd(n,m)" |
10958 | 18 |
apply (rule is_gcd_unique) |
19 |
apply (rule is_gcd) |
|
20 |
apply (subst is_gcd_commute) |
|
21 |
apply (simp add: is_gcd) |
|
22 |
done |
|
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
23 |
|
11711 | 24 |
lemma gcd_1 [simp]: "gcd(m, Suc 0) = Suc 0" |
10958 | 25 |
apply simp |
26 |
done |
|
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
27 |
|
11711 | 28 |
lemma gcd_1_left [simp]: "gcd(Suc 0, m) = Suc 0" |
29 |
apply (simp add: gcd_commute [of "Suc 0"]) |
|
10958 | 30 |
done |
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
31 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
32 |
text{*\noindent |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
33 |
as far as HERE. |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
34 |
*} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
35 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
36 |
text{*\noindent |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
37 |
SKIP THIS PROOF |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
38 |
*} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
39 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
40 |
lemma gcd_mult_distrib2: "k * gcd(m,n) = gcd(k*m, k*n)" |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
41 |
apply (induct_tac m n rule: gcd.induct) |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
42 |
apply (case_tac "n=0") |
10958 | 43 |
apply simp |
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
44 |
apply (case_tac "k=0") |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
45 |
apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2) |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
46 |
done |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
47 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
48 |
text {* |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
49 |
@{thm[display] gcd_mult_distrib2} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
50 |
\rulename{gcd_mult_distrib2} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
51 |
*}; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
52 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
53 |
text{*\noindent |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
54 |
of, simplified |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
55 |
*} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
56 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
57 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
58 |
lemmas gcd_mult_0 = gcd_mult_distrib2 [of k 1]; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
59 |
lemmas gcd_mult_1 = gcd_mult_0 [simplified]; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
60 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
61 |
text {* |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
62 |
@{thm[display] gcd_mult_distrib2 [of _ 1]} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
63 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
64 |
@{thm[display] gcd_mult_0} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
65 |
\rulename{gcd_mult_0} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
66 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
67 |
@{thm[display] gcd_mult_1} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
68 |
\rulename{gcd_mult_1} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
69 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
70 |
@{thm[display] sym} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
71 |
\rulename{sym} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
72 |
*}; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
73 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
74 |
lemmas gcd_mult = gcd_mult_1 [THEN sym]; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
75 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
76 |
lemmas gcd_mult = gcd_mult_distrib2 [of k 1, simplified, THEN sym]; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
77 |
(*better in one step!*) |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
78 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
79 |
text {* |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
80 |
more legible |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
81 |
*}; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
82 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
83 |
lemma gcd_mult [simp]: "gcd(k, k*n) = k" |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
84 |
by (rule gcd_mult_distrib2 [of k 1, simplified, THEN sym]) |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
85 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
86 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
87 |
lemmas gcd_self = gcd_mult [of k 1, simplified]; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
88 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
89 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
90 |
text {* |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
91 |
Rules handy with THEN |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
92 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
93 |
@{thm[display] iffD1} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
94 |
\rulename{iffD1} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
95 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
96 |
@{thm[display] iffD2} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
97 |
\rulename{iffD2} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
98 |
*}; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
99 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
100 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
101 |
text {* |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
102 |
again: more legible |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
103 |
*}; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
104 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
105 |
lemma gcd_self [simp]: "gcd(k,k) = k" |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
106 |
by (rule gcd_mult [of k 1, simplified]) |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
107 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
108 |
|
10958 | 109 |
text{* |
110 |
NEXT SECTION: Methods for Forward Proof |
|
111 |
||
112 |
NEW |
|
113 |
||
114 |
theorem arg_cong, useful in forward steps |
|
115 |
@{thm[display] arg_cong[no_vars]} |
|
116 |
\rulename{arg_cong} |
|
117 |
*} |
|
118 |
||
11711 | 119 |
lemma "2 \<le> u \<Longrightarrow> u*m \<noteq> Suc(u*n)" |
12390 | 120 |
apply (intro notI) |
10958 | 121 |
txt{* |
122 |
before using arg_cong |
|
123 |
@{subgoals[display,indent=0,margin=65]} |
|
124 |
*}; |
|
125 |
apply (drule_tac f="\<lambda>x. x mod u" in arg_cong) |
|
126 |
txt{* |
|
127 |
after using arg_cong |
|
128 |
@{subgoals[display,indent=0,margin=65]} |
|
129 |
*}; |
|
130 |
apply (simp add: mod_Suc) |
|
131 |
done |
|
132 |
||
133 |
text{* |
|
134 |
have just used this rule: |
|
135 |
@{thm[display] mod_Suc[no_vars]} |
|
136 |
\rulename{mod_Suc} |
|
137 |
||
138 |
@{thm[display] mult_le_mono1[no_vars]} |
|
139 |
\rulename{mult_le_mono1} |
|
140 |
*} |
|
141 |
||
142 |
||
143 |
text{* |
|
144 |
example of "insert" |
|
145 |
*} |
|
146 |
||
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
147 |
lemma relprime_dvd_mult: |
11407 | 148 |
"\<lbrakk> gcd(k,n)=1; k dvd m*n \<rbrakk> \<Longrightarrow> k dvd m" |
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
149 |
apply (insert gcd_mult_distrib2 [of m k n]) |
10958 | 150 |
apply simp |
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
151 |
apply (erule_tac t="m" in ssubst); |
10958 | 152 |
apply simp |
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
153 |
done |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
154 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
155 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
156 |
text {* |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
157 |
Another example of "insert" |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
158 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
159 |
@{thm[display] mod_div_equality} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
160 |
\rulename{mod_div_equality} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
161 |
*}; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
162 |
|
11407 | 163 |
(*MOVED to Force.thy, which now depends only on Divides.thy |
164 |
lemma div_mult_self_is_m: "0<n \<Longrightarrow> (m*n) div n = (m::nat)" |
|
165 |
*) |
|
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
166 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
167 |
lemma relprime_dvd_mult_iff: "gcd(k,n)=1 \<Longrightarrow> (k dvd m*n) = (k dvd m)"; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
168 |
by (blast intro: relprime_dvd_mult dvd_trans) |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
169 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
170 |
|
11711 | 171 |
lemma relprime_20_81: "gcd(20,81) = 1"; |
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
172 |
by (simp add: gcd.simps) |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
173 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
174 |
text {* |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
175 |
Examples of 'OF' |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
176 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
177 |
@{thm[display] relprime_dvd_mult} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
178 |
\rulename{relprime_dvd_mult} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
179 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
180 |
@{thm[display] relprime_dvd_mult [OF relprime_20_81]} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
181 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
182 |
@{thm[display] dvd_refl} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
183 |
\rulename{dvd_refl} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
184 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
185 |
@{thm[display] dvd_add} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
186 |
\rulename{dvd_add} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
187 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
188 |
@{thm[display] dvd_add [OF dvd_refl dvd_refl]} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
189 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
190 |
@{thm[display] dvd_add [OF _ dvd_refl]} |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
191 |
*}; |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
192 |
|
11711 | 193 |
lemma "\<lbrakk>(z::int) < 37; 66 < 2*z; z*z \<noteq> 1225; Q(34); Q(36)\<rbrakk> \<Longrightarrow> Q(z)"; |
194 |
apply (subgoal_tac "z = 34 \<or> z = 36") |
|
10958 | 195 |
txt{* |
196 |
the tactic leaves two subgoals: |
|
197 |
@{subgoals[display,indent=0,margin=65]} |
|
198 |
*}; |
|
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
199 |
apply blast |
11711 | 200 |
apply (subgoal_tac "z \<noteq> 35") |
10958 | 201 |
txt{* |
202 |
the tactic leaves two subgoals: |
|
203 |
@{subgoals[display,indent=0,margin=65]} |
|
204 |
*}; |
|
10846
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
205 |
apply arith |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
206 |
apply force |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
207 |
done |
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
208 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
209 |
|
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents:
diff
changeset
|
210 |
end |