| author | wenzelm |
| Tue, 12 Jan 1999 15:17:37 +0100 | |
| changeset 6093 | 87bf8c03b169 |
| parent 6070 | 032babd0120b |
| child 6153 | bff90585cce5 |
| permissions | -rw-r--r-- |
| 1793 | 1 |
(* Title: ZF/Arith.ML |
| 0 | 2 |
ID: $Id$ |
| 1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
| 0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
| 1609 | 6 |
Arithmetic operators and their definitions |
| 0 | 7 |
|
8 |
Proofs about elementary arithmetic: addition, multiplication, etc. |
|
9 |
*) |
|
10 |
||
11 |
(*"Difference" is subtraction of natural numbers. |
|
12 |
There are no negative numbers; we have |
|
13 |
m #- n = 0 iff m<=n and m #- n = succ(k) iff m>n. |
|
14 |
Also, rec(m, 0, %z w.z) is pred(m). |
|
15 |
*) |
|
16 |
||
| 2469 | 17 |
Addsimps [rec_type, nat_0_le, nat_le_refl]; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
18 |
val nat_typechecks = [rec_type, nat_0I, nat_1I, nat_succI, Ord_nat]; |
| 0 | 19 |
|
| 5137 | 20 |
Goal "[| 0<k; k: nat |] ==> EX j: nat. k = succ(j)"; |
| 1708 | 21 |
by (etac rev_mp 1); |
| 6070 | 22 |
by (induct_tac "k" 1); |
| 2469 | 23 |
by (Simp_tac 1); |
| 3016 | 24 |
by (Blast_tac 1); |
| 1708 | 25 |
val lemma = result(); |
26 |
||
27 |
(* [| 0 < k; k: nat; !!j. [| j: nat; k = succ(j) |] ==> Q |] ==> Q *) |
|
28 |
bind_thm ("zero_lt_natE", lemma RS bexE);
|
|
29 |
||
30 |
||
| 0 | 31 |
(** Addition **) |
32 |
||
| 6070 | 33 |
Goal "[| m:nat; n:nat |] ==> m #+ n : nat"; |
34 |
by (induct_tac "m" 1); |
|
35 |
by Auto_tac; |
|
36 |
qed "add_type"; |
|
37 |
Addsimps [add_type]; |
|
| 2469 | 38 |
|
| 0 | 39 |
(** Multiplication **) |
40 |
||
| 6070 | 41 |
Goal "[| m:nat; n:nat |] ==> m #* n : nat"; |
42 |
by (induct_tac "m" 1); |
|
43 |
by Auto_tac; |
|
44 |
qed "mult_type"; |
|
45 |
Addsimps [mult_type]; |
|
| 0 | 46 |
|
| 2469 | 47 |
|
| 0 | 48 |
(** Difference **) |
49 |
||
| 6070 | 50 |
Goal "[| m:nat; n:nat |] ==> m #- n : nat"; |
51 |
by (induct_tac "n" 1); |
|
52 |
by Auto_tac; |
|
53 |
by (fast_tac (claset() addIs [nat_case_type]) 1); |
|
54 |
qed "diff_type"; |
|
55 |
Addsimps [diff_type]; |
|
| 0 | 56 |
|
| 6070 | 57 |
Goal "n:nat ==> 0 #- n = 0"; |
58 |
by (induct_tac "n" 1); |
|
59 |
by Auto_tac; |
|
60 |
qed "diff_0_eq_0"; |
|
| 0 | 61 |
|
| 6070 | 62 |
(*Must simplify BEFORE the induction: else we get a critical pair*) |
63 |
Goal "[| m:nat; n:nat |] ==> succ(m) #- succ(n) = m #- n"; |
|
64 |
by (Asm_simp_tac 1); |
|
65 |
by (induct_tac "n" 1); |
|
66 |
by Auto_tac; |
|
67 |
qed "diff_succ_succ"; |
|
| 0 | 68 |
|
| 6070 | 69 |
Addsimps [diff_0_eq_0, diff_succ_succ]; |
| 2469 | 70 |
|
| 6070 | 71 |
(*This defining property is no longer wanted*) |
72 |
Delsimps [diff_SUCC]; |
|
| 0 | 73 |
|
74 |
val prems = goal Arith.thy |
|
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
75 |
"[| m:nat; n:nat |] ==> m #- n le m"; |
|
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
76 |
by (rtac (prems MRS diff_induct) 1); |
|
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
77 |
by (etac leE 3); |
| 0 | 78 |
by (ALLGOALS |
| 5529 | 79 |
(asm_simp_tac (simpset() addsimps prems @ [le_iff, nat_into_Ord]))); |
| 760 | 80 |
qed "diff_le_self"; |
| 0 | 81 |
|
82 |
(*** Simplification over add, mult, diff ***) |
|
83 |
||
84 |
val arith_typechecks = [add_type, mult_type, diff_type]; |
|
85 |
||
86 |
||
87 |
(*** Addition ***) |
|
88 |
||
89 |
(*Associative law for addition*) |
|
| 6070 | 90 |
Goal "m:nat ==> (m #+ n) #+ k = m #+ (n #+ k)"; |
91 |
by (induct_tac "m" 1); |
|
92 |
by Auto_tac; |
|
93 |
qed "add_assoc"; |
|
| 0 | 94 |
|
95 |
(*The following two lemmas are used for add_commute and sometimes |
|
96 |
elsewhere, since they are safe for rewriting.*) |
|
| 6070 | 97 |
Goal "m:nat ==> m #+ 0 = m"; |
98 |
by (induct_tac "m" 1); |
|
99 |
by Auto_tac; |
|
100 |
qed "add_0_right"; |
|
| 0 | 101 |
|
| 6070 | 102 |
Goal "m:nat ==> m #+ succ(n) = succ(m #+ n)"; |
103 |
by (induct_tac "m" 1); |
|
104 |
by Auto_tac; |
|
105 |
qed "add_succ_right"; |
|
| 2469 | 106 |
|
107 |
Addsimps [add_0_right, add_succ_right]; |
|
| 0 | 108 |
|
109 |
(*Commutative law for addition*) |
|
| 6070 | 110 |
Goal "[| m:nat; n:nat |] ==> m #+ n = n #+ m"; |
111 |
by (induct_tac "n" 1); |
|
112 |
by Auto_tac; |
|
113 |
qed "add_commute"; |
|
| 435 | 114 |
|
| 437 | 115 |
(*for a/c rewriting*) |
| 6070 | 116 |
Goal "[| m:nat; n:nat |] ==> m#+(n#+k)=n#+(m#+k)"; |
117 |
by (asm_simp_tac (simpset() addsimps [add_assoc RS sym, add_commute]) 1); |
|
118 |
qed "add_left_commute"; |
|
| 435 | 119 |
|
120 |
(*Addition is an AC-operator*) |
|
121 |
val add_ac = [add_assoc, add_commute, add_left_commute]; |
|
| 0 | 122 |
|
123 |
(*Cancellation law on the left*) |
|
| 6070 | 124 |
Goal "[| k #+ m = k #+ n; k:nat |] ==> m=n"; |
125 |
by (etac rev_mp 1); |
|
126 |
by (induct_tac "k" 1); |
|
127 |
by Auto_tac; |
|
| 760 | 128 |
qed "add_left_cancel"; |
| 0 | 129 |
|
130 |
(*** Multiplication ***) |
|
131 |
||
132 |
(*right annihilation in product*) |
|
| 6070 | 133 |
Goal "m:nat ==> m #* 0 = 0"; |
134 |
by (induct_tac "m" 1); |
|
135 |
by Auto_tac; |
|
136 |
qed "mult_0_right"; |
|
| 0 | 137 |
|
138 |
(*right successor law for multiplication*) |
|
| 6070 | 139 |
Goal "[| m:nat; n:nat |] ==> m #* succ(n) = m #+ (m #* n)"; |
140 |
by (induct_tac "m" 1); |
|
141 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps add_ac))); |
|
142 |
qed "mult_succ_right"; |
|
| 2469 | 143 |
|
144 |
Addsimps [mult_0_right, mult_succ_right]; |
|
| 0 | 145 |
|
| 5137 | 146 |
Goal "n:nat ==> 1 #* n = n"; |
| 2469 | 147 |
by (Asm_simp_tac 1); |
| 1793 | 148 |
qed "mult_1"; |
149 |
||
| 5137 | 150 |
Goal "n:nat ==> n #* 1 = n"; |
| 2469 | 151 |
by (Asm_simp_tac 1); |
| 1793 | 152 |
qed "mult_1_right"; |
153 |
||
| 6070 | 154 |
Addsimps [mult_1, mult_1_right]; |
155 |
||
| 0 | 156 |
(*Commutative law for multiplication*) |
| 6070 | 157 |
Goal "[| m:nat; n:nat |] ==> m #* n = n #* m"; |
158 |
by (induct_tac "m" 1); |
|
159 |
by Auto_tac; |
|
160 |
qed "mult_commute"; |
|
| 0 | 161 |
|
162 |
(*addition distributes over multiplication*) |
|
| 6070 | 163 |
Goal "[| m:nat; k:nat |] ==> (m #+ n) #* k = (m #* k) #+ (n #* k)"; |
164 |
by (induct_tac "m" 1); |
|
165 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_assoc RS sym]))); |
|
166 |
qed "add_mult_distrib"; |
|
| 0 | 167 |
|
168 |
(*Distributive law on the left; requires an extra typing premise*) |
|
| 6070 | 169 |
Goal "[| m:nat; n:nat; k:nat |] ==> k #* (m #+ n) = (k #* m) #+ (k #* n)"; |
170 |
by (induct_tac "m" 1); |
|
171 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps add_ac))); |
|
172 |
qed "add_mult_distrib_left"; |
|
| 0 | 173 |
|
174 |
(*Associative law for multiplication*) |
|
| 6070 | 175 |
Goal "[| m:nat; n:nat; k:nat |] ==> (m #* n) #* k = m #* (n #* k)"; |
176 |
by (induct_tac "m" 1); |
|
177 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_mult_distrib]))); |
|
178 |
qed "mult_assoc"; |
|
| 0 | 179 |
|
| 437 | 180 |
(*for a/c rewriting*) |
| 6070 | 181 |
Goal "[| m:nat; n:nat; k:nat |] ==> m #* (n #* k) = n #* (m #* k)"; |
182 |
by (rtac (mult_commute RS trans) 1); |
|
183 |
by (rtac (mult_assoc RS trans) 3); |
|
184 |
by (rtac (mult_commute RS subst_context) 6); |
|
185 |
by (REPEAT (ares_tac [mult_type] 1)); |
|
186 |
qed "mult_left_commute"; |
|
| 437 | 187 |
|
188 |
val mult_ac = [mult_assoc,mult_commute,mult_left_commute]; |
|
189 |
||
| 0 | 190 |
|
191 |
(*** Difference ***) |
|
192 |
||
| 6070 | 193 |
Goal "m:nat ==> m #- m = 0"; |
194 |
by (induct_tac "m" 1); |
|
195 |
by Auto_tac; |
|
196 |
qed "diff_self_eq_0"; |
|
| 0 | 197 |
|
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
198 |
(*Addition is the inverse of subtraction*) |
| 5137 | 199 |
Goal "[| n le m; m:nat |] ==> n #+ (m#-n) = m"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
200 |
by (forward_tac [lt_nat_in_nat] 1); |
| 127 | 201 |
by (etac nat_succI 1); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
202 |
by (etac rev_mp 1); |
| 0 | 203 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
|
| 2469 | 204 |
by (ALLGOALS Asm_simp_tac); |
| 760 | 205 |
qed "add_diff_inverse"; |
| 0 | 206 |
|
| 5504 | 207 |
Goal "[| n le m; m:nat |] ==> (m#-n) #+ n = m"; |
208 |
by (forward_tac [lt_nat_in_nat] 1); |
|
209 |
by (etac nat_succI 1); |
|
210 |
by (asm_simp_tac (simpset() addsimps [add_commute, add_diff_inverse]) 1); |
|
211 |
qed "add_diff_inverse2"; |
|
212 |
||
| 1609 | 213 |
(*Proof is IDENTICAL to that above*) |
| 5137 | 214 |
Goal "[| n le m; m:nat |] ==> succ(m) #- n = succ(m#-n)"; |
| 1609 | 215 |
by (forward_tac [lt_nat_in_nat] 1); |
216 |
by (etac nat_succI 1); |
|
217 |
by (etac rev_mp 1); |
|
218 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
|
|
| 2469 | 219 |
by (ALLGOALS Asm_simp_tac); |
| 1609 | 220 |
qed "diff_succ"; |
221 |
||
| 5341 | 222 |
Goal "[| m: nat; n: nat |] ==> 0 < n #- m <-> m<n"; |
223 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
|
|
224 |
by (ALLGOALS Asm_simp_tac); |
|
225 |
qed "zero_less_diff"; |
|
226 |
Addsimps [zero_less_diff]; |
|
227 |
||
228 |
||
| 1708 | 229 |
(** Subtraction is the inverse of addition. **) |
230 |
||
| 6070 | 231 |
Goal "[| m:nat; n:nat |] ==> (n#+m) #- n = m"; |
232 |
by (induct_tac "n" 1); |
|
233 |
by Auto_tac; |
|
| 760 | 234 |
qed "diff_add_inverse"; |
| 0 | 235 |
|
| 5137 | 236 |
Goal "[| m:nat; n:nat |] ==> (m#+n) #- n = m"; |
| 437 | 237 |
by (res_inst_tac [("m1","m")] (add_commute RS ssubst) 1);
|
238 |
by (REPEAT (ares_tac [diff_add_inverse] 1)); |
|
| 760 | 239 |
qed "diff_add_inverse2"; |
| 437 | 240 |
|
| 5137 | 241 |
Goal "[| k:nat; m: nat; n: nat |] ==> (k#+m) #- (k#+n) = m #- n"; |
| 6070 | 242 |
by (induct_tac "k" 1); |
| 2469 | 243 |
by (ALLGOALS Asm_simp_tac); |
| 1708 | 244 |
qed "diff_cancel"; |
245 |
||
| 5137 | 246 |
Goal "[| k:nat; m: nat; n: nat |] ==> (m#+k) #- (n#+k) = m #- n"; |
| 1708 | 247 |
val add_commute_k = read_instantiate [("n","k")] add_commute;
|
| 4091 | 248 |
by (asm_simp_tac (simpset() addsimps [add_commute_k, diff_cancel]) 1); |
| 1708 | 249 |
qed "diff_cancel2"; |
250 |
||
| 6070 | 251 |
Goal "[| m:nat; n:nat |] ==> n #- (n#+m) = 0"; |
252 |
by (induct_tac "n" 1); |
|
253 |
by Auto_tac; |
|
| 760 | 254 |
qed "diff_add_0"; |
| 0 | 255 |
|
| 1708 | 256 |
(** Difference distributes over multiplication **) |
257 |
||
| 5137 | 258 |
Goal "[| m:nat; n: nat; k:nat |] ==> (m #- n) #* k = (m #* k) #- (n #* k)"; |
| 1708 | 259 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
|
| 4091 | 260 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_cancel]))); |
| 1708 | 261 |
qed "diff_mult_distrib" ; |
262 |
||
| 5137 | 263 |
Goal "[| m:nat; n: nat; k:nat |] ==> k #* (m #- n) = (k #* m) #- (k #* n)"; |
| 1708 | 264 |
val mult_commute_k = read_instantiate [("m","k")] mult_commute;
|
| 4091 | 265 |
by (asm_simp_tac (simpset() addsimps |
| 1793 | 266 |
[mult_commute_k, diff_mult_distrib]) 1); |
| 1708 | 267 |
qed "diff_mult_distrib2" ; |
268 |
||
| 0 | 269 |
(*** Remainder ***) |
270 |
||
| 5137 | 271 |
Goal "[| 0<n; n le m; m:nat |] ==> m #- n < m"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
272 |
by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); |
| 0 | 273 |
by (etac rev_mp 1); |
274 |
by (etac rev_mp 1); |
|
275 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
|
|
| 6070 | 276 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_le_self]))); |
| 760 | 277 |
qed "div_termination"; |
| 0 | 278 |
|
| 1461 | 279 |
val div_rls = (*for mod and div*) |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
280 |
nat_typechecks @ |
|
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
281 |
[Ord_transrec_type, apply_type, div_termination RS ltD, if_type, |
| 435 | 282 |
nat_into_Ord, not_lt_iff_le RS iffD1]; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
283 |
|
| 6070 | 284 |
val div_ss = simpset() addsimps [nat_into_Ord, div_termination RS ltD, |
285 |
not_lt_iff_le RS iffD2]; |
|
| 0 | 286 |
|
287 |
(*Type checking depends upon termination!*) |
|
| 5137 | 288 |
Goalw [mod_def] "[| 0<n; m:nat; n:nat |] ==> m mod n : nat"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
289 |
by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1)); |
| 760 | 290 |
qed "mod_type"; |
| 0 | 291 |
|
| 5137 | 292 |
Goal "[| 0<n; m<n |] ==> m mod n = m"; |
| 0 | 293 |
by (rtac (mod_def RS def_transrec RS trans) 1); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
294 |
by (asm_simp_tac div_ss 1); |
| 760 | 295 |
qed "mod_less"; |
| 0 | 296 |
|
| 5137 | 297 |
Goal "[| 0<n; n le m; m:nat |] ==> m mod n = (m#-n) mod n"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
298 |
by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); |
| 0 | 299 |
by (rtac (mod_def RS def_transrec RS trans) 1); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
300 |
by (asm_simp_tac div_ss 1); |
| 760 | 301 |
qed "mod_geq"; |
| 0 | 302 |
|
| 2469 | 303 |
Addsimps [mod_type, mod_less, mod_geq]; |
304 |
||
| 0 | 305 |
(*** Quotient ***) |
306 |
||
307 |
(*Type checking depends upon termination!*) |
|
| 5067 | 308 |
Goalw [div_def] |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
309 |
"[| 0<n; m:nat; n:nat |] ==> m div n : nat"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
310 |
by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1)); |
| 760 | 311 |
qed "div_type"; |
| 0 | 312 |
|
| 5137 | 313 |
Goal "[| 0<n; m<n |] ==> m div n = 0"; |
| 0 | 314 |
by (rtac (div_def RS def_transrec RS trans) 1); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
315 |
by (asm_simp_tac div_ss 1); |
| 760 | 316 |
qed "div_less"; |
| 0 | 317 |
|
| 5137 | 318 |
Goal "[| 0<n; n le m; m:nat |] ==> m div n = succ((m#-n) div n)"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
319 |
by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); |
| 0 | 320 |
by (rtac (div_def RS def_transrec RS trans) 1); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
321 |
by (asm_simp_tac div_ss 1); |
| 760 | 322 |
qed "div_geq"; |
| 0 | 323 |
|
| 2469 | 324 |
Addsimps [div_type, div_less, div_geq]; |
325 |
||
| 1609 | 326 |
(*A key result*) |
| 5137 | 327 |
Goal "[| 0<n; m:nat; n:nat |] ==> (m div n)#*n #+ m mod n = m"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
328 |
by (etac complete_induct 1); |
| 437 | 329 |
by (excluded_middle_tac "x<n" 1); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
330 |
(*case x<n*) |
| 2469 | 331 |
by (Asm_simp_tac 2); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
332 |
(*case n le x*) |
|
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
333 |
by (asm_full_simp_tac |
| 4091 | 334 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, add_assoc, |
| 1461 | 335 |
div_termination RS ltD, add_diff_inverse]) 1); |
| 760 | 336 |
qed "mod_div_equality"; |
| 0 | 337 |
|
| 6068 | 338 |
|
339 |
(*** Further facts about mod (mainly for mutilated chess board) ***) |
|
| 1609 | 340 |
|
| 6068 | 341 |
Goal "[| 0<n; m:nat; n:nat |] \ |
342 |
\ ==> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))"; |
|
| 1609 | 343 |
by (etac complete_induct 1); |
344 |
by (excluded_middle_tac "succ(x)<n" 1); |
|
| 1623 | 345 |
(* case succ(x) < n *) |
| 4091 | 346 |
by (asm_simp_tac (simpset() addsimps [mod_less, nat_le_refl RS lt_trans, |
| 1623 | 347 |
succ_neq_self]) 2); |
| 4091 | 348 |
by (asm_simp_tac (simpset() addsimps [ltD RS mem_imp_not_eq]) 2); |
| 1623 | 349 |
(* case n le succ(x) *) |
| 1609 | 350 |
by (asm_full_simp_tac |
| 4091 | 351 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, mod_geq]) 1); |
| 1623 | 352 |
by (etac leE 1); |
| 4091 | 353 |
by (asm_simp_tac (simpset() addsimps [div_termination RS ltD, diff_succ, |
| 1623 | 354 |
mod_geq]) 1); |
| 4091 | 355 |
by (asm_simp_tac (simpset() addsimps [mod_less, diff_self_eq_0]) 1); |
| 1609 | 356 |
qed "mod_succ"; |
357 |
||
| 5137 | 358 |
Goal "[| 0<n; m:nat; n:nat |] ==> m mod n < n"; |
| 1609 | 359 |
by (etac complete_induct 1); |
360 |
by (excluded_middle_tac "x<n" 1); |
|
361 |
(*case x<n*) |
|
| 4091 | 362 |
by (asm_simp_tac (simpset() addsimps [mod_less]) 2); |
| 1609 | 363 |
(*case n le x*) |
364 |
by (asm_full_simp_tac |
|
| 4091 | 365 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, |
| 1609 | 366 |
mod_geq, div_termination RS ltD]) 1); |
367 |
qed "mod_less_divisor"; |
|
368 |
||
369 |
||
| 6068 | 370 |
Goal "[| k: nat; b<2 |] ==> k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)"; |
| 1609 | 371 |
by (subgoal_tac "k mod 2: 2" 1); |
| 4091 | 372 |
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2); |
| 1623 | 373 |
by (dtac ltD 1); |
| 5137 | 374 |
by Auto_tac; |
| 1609 | 375 |
qed "mod2_cases"; |
376 |
||
| 5137 | 377 |
Goal "m:nat ==> succ(succ(m)) mod 2 = m mod 2"; |
| 1609 | 378 |
by (subgoal_tac "m mod 2: 2" 1); |
| 4091 | 379 |
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2); |
380 |
by (asm_simp_tac (simpset() addsimps [mod_succ] setloop Step_tac) 1); |
|
| 1609 | 381 |
qed "mod2_succ_succ"; |
382 |
||
| 5137 | 383 |
Goal "m:nat ==> (m#+m) mod 2 = 0"; |
| 6070 | 384 |
by (induct_tac "m" 1); |
| 4091 | 385 |
by (simp_tac (simpset() addsimps [mod_less]) 1); |
386 |
by (asm_simp_tac (simpset() addsimps [mod2_succ_succ, add_succ_right]) 1); |
|
| 1609 | 387 |
qed "mod2_add_self"; |
388 |
||
| 0 | 389 |
|
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
390 |
(**** Additional theorems about "le" ****) |
| 0 | 391 |
|
| 5137 | 392 |
Goal "[| m:nat; n:nat |] ==> m le m #+ n"; |
| 6070 | 393 |
by (induct_tac "m" 1); |
| 2469 | 394 |
by (ALLGOALS Asm_simp_tac); |
| 760 | 395 |
qed "add_le_self"; |
|
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
396 |
|
| 5137 | 397 |
Goal "[| m:nat; n:nat |] ==> m le n #+ m"; |
| 2033 | 398 |
by (stac add_commute 1); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
399 |
by (REPEAT (ares_tac [add_le_self] 1)); |
| 760 | 400 |
qed "add_le_self2"; |
|
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
401 |
|
| 1708 | 402 |
(*** Monotonicity of Addition ***) |
|
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
403 |
|
| 1708 | 404 |
(*strict, in 1st argument; proof is by rule induction on 'less than'*) |
| 5137 | 405 |
Goal "[| i<j; j:nat; k:nat |] ==> i#+k < j#+k"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
406 |
by (forward_tac [lt_nat_in_nat] 1); |
| 127 | 407 |
by (assume_tac 1); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
408 |
by (etac succ_lt_induct 1); |
| 4091 | 409 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [leI]))); |
| 760 | 410 |
qed "add_lt_mono1"; |
|
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
411 |
|
|
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
412 |
(*strict, in both arguments*) |
| 5137 | 413 |
Goal "[| i<j; k<l; j:nat; l:nat |] ==> i#+k < j#+l"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
414 |
by (rtac (add_lt_mono1 RS lt_trans) 1); |
|
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
415 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1)); |
| 2033 | 416 |
by (EVERY [stac add_commute 1, |
417 |
stac add_commute 3, |
|
| 1461 | 418 |
rtac add_lt_mono1 5]); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
419 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1)); |
| 760 | 420 |
qed "add_lt_mono"; |
|
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
421 |
|
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
422 |
(*A [clumsy] way of lifting < monotonicity to le monotonicity *) |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5147
diff
changeset
|
423 |
val lt_mono::ford::prems = Goal |
| 1461 | 424 |
"[| !!i j. [| i<j; j:k |] ==> f(i) < f(j); \ |
425 |
\ !!i. i:k ==> Ord(f(i)); \ |
|
426 |
\ i le j; j:k \ |
|
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
427 |
\ |] ==> f(i) le f(j)"; |
|
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
428 |
by (cut_facts_tac prems 1); |
| 3016 | 429 |
by (blast_tac (le_cs addSIs [lt_mono,ford] addSEs [leE]) 1); |
| 760 | 430 |
qed "Ord_lt_mono_imp_le_mono"; |
|
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
431 |
|
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
432 |
(*le monotonicity, 1st argument*) |
| 5137 | 433 |
Goal "[| i le j; j:nat; k:nat |] ==> i#+k le j#+k"; |
| 3840 | 434 |
by (res_inst_tac [("f", "%j. j#+k")] Ord_lt_mono_imp_le_mono 1);
|
| 435 | 435 |
by (REPEAT (ares_tac [add_lt_mono1, add_type RS nat_into_Ord] 1)); |
| 760 | 436 |
qed "add_le_mono1"; |
|
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
437 |
|
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
438 |
(* le monotonicity, BOTH arguments*) |
| 5137 | 439 |
Goal "[| i le j; k le l; j:nat; l:nat |] ==> i#+k le j#+l"; |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
440 |
by (rtac (add_le_mono1 RS le_trans) 1); |
|
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
441 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); |
| 2033 | 442 |
by (EVERY [stac add_commute 1, |
443 |
stac add_commute 3, |
|
| 1461 | 444 |
rtac add_le_mono1 5]); |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
445 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); |
| 760 | 446 |
qed "add_le_mono"; |
| 1609 | 447 |
|
| 1708 | 448 |
(*** Monotonicity of Multiplication ***) |
449 |
||
| 5137 | 450 |
Goal "[| i le j; j:nat; k:nat |] ==> i#*k le j#*k"; |
| 1708 | 451 |
by (forward_tac [lt_nat_in_nat] 1); |
| 6070 | 452 |
by (induct_tac "k" 2); |
| 4091 | 453 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_le_mono]))); |
| 1708 | 454 |
qed "mult_le_mono1"; |
455 |
||
456 |
(* le monotonicity, BOTH arguments*) |
|
| 5137 | 457 |
Goal "[| i le j; k le l; j:nat; l:nat |] ==> i#*k le j#*l"; |
| 1708 | 458 |
by (rtac (mult_le_mono1 RS le_trans) 1); |
459 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); |
|
| 2033 | 460 |
by (EVERY [stac mult_commute 1, |
461 |
stac mult_commute 3, |
|
| 1708 | 462 |
rtac mult_le_mono1 5]); |
463 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); |
|
464 |
qed "mult_le_mono"; |
|
465 |
||
466 |
(*strict, in 1st argument; proof is by induction on k>0*) |
|
| 5137 | 467 |
Goal "[| i<j; 0<k; j:nat; k:nat |] ==> k#*i < k#*j"; |
| 1793 | 468 |
by (etac zero_lt_natE 1); |
| 1708 | 469 |
by (forward_tac [lt_nat_in_nat] 2); |
| 2469 | 470 |
by (ALLGOALS Asm_simp_tac); |
| 6070 | 471 |
by (induct_tac "x" 1); |
| 4091 | 472 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_lt_mono]))); |
| 1708 | 473 |
qed "mult_lt_mono2"; |
474 |
||
| 5137 | 475 |
Goal "[| i<j; 0<c; i:nat; j:nat; c:nat |] ==> i#*c < j#*c"; |
| 4839 | 476 |
by (asm_simp_tac (simpset() addsimps [mult_lt_mono2, mult_commute]) 1); |
477 |
qed "mult_lt_mono1"; |
|
478 |
||
| 5137 | 479 |
Goal "[| m: nat; n: nat |] ==> 0 < m#*n <-> 0<m & 0<n"; |
| 4091 | 480 |
by (best_tac (claset() addEs [natE] addss (simpset())) 1); |
| 1708 | 481 |
qed "zero_lt_mult_iff"; |
482 |
||
| 5137 | 483 |
Goal "[| m: nat; n: nat |] ==> m#*n = 1 <-> m=1 & n=1"; |
| 4091 | 484 |
by (best_tac (claset() addEs [natE] addss (simpset())) 1); |
| 1793 | 485 |
qed "mult_eq_1_iff"; |
486 |
||
| 1708 | 487 |
(*Cancellation law for division*) |
| 5137 | 488 |
Goal "[| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> (k#*m) div (k#*n) = m div n"; |
| 1708 | 489 |
by (eres_inst_tac [("i","m")] complete_induct 1);
|
490 |
by (excluded_middle_tac "x<n" 1); |
|
| 4091 | 491 |
by (asm_simp_tac (simpset() addsimps [div_less, zero_lt_mult_iff, |
| 1793 | 492 |
mult_lt_mono2]) 2); |
| 1708 | 493 |
by (asm_full_simp_tac |
| 4091 | 494 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, |
| 1708 | 495 |
zero_lt_mult_iff, le_refl RS mult_le_mono, div_geq, |
496 |
diff_mult_distrib2 RS sym, |
|
| 1793 | 497 |
div_termination RS ltD]) 1); |
| 1708 | 498 |
qed "div_cancel"; |
499 |
||
| 5137 | 500 |
Goal "[| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> \ |
| 1708 | 501 |
\ (k#*m) mod (k#*n) = k #* (m mod n)"; |
502 |
by (eres_inst_tac [("i","m")] complete_induct 1);
|
|
503 |
by (excluded_middle_tac "x<n" 1); |
|
| 4091 | 504 |
by (asm_simp_tac (simpset() addsimps [mod_less, zero_lt_mult_iff, |
| 1793 | 505 |
mult_lt_mono2]) 2); |
| 1708 | 506 |
by (asm_full_simp_tac |
| 4091 | 507 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, |
| 1708 | 508 |
zero_lt_mult_iff, le_refl RS mult_le_mono, mod_geq, |
509 |
diff_mult_distrib2 RS sym, |
|
| 1793 | 510 |
div_termination RS ltD]) 1); |
| 1708 | 511 |
qed "mult_mod_distrib"; |
512 |
||
| 6070 | 513 |
(*Lemma for gcd*) |
| 5137 | 514 |
Goal "[| m = m#*n; m: nat; n: nat |] ==> n=1 | m=0"; |
| 1793 | 515 |
by (rtac disjCI 1); |
516 |
by (dtac sym 1); |
|
517 |
by (rtac Ord_linear_lt 1 THEN REPEAT_SOME (ares_tac [nat_into_Ord,nat_1I])); |
|
| 6070 | 518 |
by (dtac (nat_into_Ord RS Ord_0_lt RSN (2,mult_lt_mono2)) 2); |
519 |
by Auto_tac; |
|
| 1793 | 520 |
qed "mult_eq_self_implies_10"; |
| 1708 | 521 |
|
| 2469 | 522 |
(*Thanks to Sten Agerholm*) |
| 5504 | 523 |
Goal "[|m#+n le m#+k; m:nat; n:nat; k:nat|] ==> n le k"; |
| 2493 | 524 |
by (etac rev_mp 1); |
| 6070 | 525 |
by (induct_tac "m" 1); |
| 2469 | 526 |
by (Asm_simp_tac 1); |
|
3736
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents:
3207
diff
changeset
|
527 |
by Safe_tac; |
| 4091 | 528 |
by (asm_full_simp_tac (simpset() addsimps [not_le_iff_lt,nat_into_Ord]) 1); |
| 2469 | 529 |
qed "add_le_elim1"; |
530 |
||
| 5504 | 531 |
Goal "[| m<n; n: nat |] ==> EX k: nat. n = succ(m#+k)"; |
532 |
by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1); |
|
533 |
be rev_mp 1; |
|
| 6070 | 534 |
by (induct_tac "n" 1); |
| 5504 | 535 |
by (ALLGOALS (simp_tac (simpset() addsimps [le_iff]))); |
536 |
by (blast_tac (claset() addSEs [leE] |
|
537 |
addSIs [add_0_right RS sym, add_succ_right RS sym]) 1); |
|
538 |
qed_spec_mp "less_imp_Suc_add"; |