author | wenzelm |
Fri, 06 Oct 2000 17:35:58 +0200 | |
changeset 10168 | 50be659d4222 |
parent 9945 | a0efbd7c88dc |
permissions | -rw-r--r-- |
7032 | 1 |
(* Title: HOL/NatBin.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1999 University of Cambridge |
|
5 |
||
6 |
Binary arithmetic for the natural numbers |
|
7 |
*) |
|
8 |
||
9 |
(** nat (coercion from int to nat) **) |
|
10 |
||
11 |
Goal "nat (number_of w) = number_of w"; |
|
12 |
by (simp_tac (simpset() addsimps [nat_number_of_def]) 1); |
|
13 |
qed "nat_number_of"; |
|
14 |
Addsimps [nat_number_of]; |
|
15 |
||
16 |
(*These rewrites should one day be re-oriented...*) |
|
17 |
||
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
18 |
Goal "#0 = (0::nat)"; |
9425 | 19 |
by (simp_tac (HOL_basic_ss addsimps [nat_0, nat_number_of_def]) 1); |
7032 | 20 |
qed "numeral_0_eq_0"; |
21 |
||
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
22 |
Goal "#1 = (1::nat)"; |
9425 | 23 |
by (simp_tac (HOL_basic_ss addsimps [nat_1, nat_number_of_def]) 1); |
7032 | 24 |
qed "numeral_1_eq_1"; |
25 |
||
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
26 |
Goal "#2 = (2::nat)"; |
9425 | 27 |
by (simp_tac (HOL_basic_ss addsimps [nat_2, nat_number_of_def]) 1); |
7032 | 28 |
qed "numeral_2_eq_2"; |
29 |
||
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
30 |
bind_thm ("zero_eq_numeral_0", numeral_0_eq_0 RS sym); |
7032 | 31 |
|
32 |
(** int (coercion from nat to int) **) |
|
33 |
||
34 |
(*"neg" is used in rewrite rules for binary comparisons*) |
|
35 |
Goal "int (number_of v :: nat) = \ |
|
36 |
\ (if neg (number_of v) then #0 \ |
|
37 |
\ else (number_of v :: int))"; |
|
38 |
by (simp_tac |
|
39 |
(simpset_of Int.thy addsimps [neg_nat, nat_number_of_def, |
|
40 |
not_neg_nat, int_0]) 1); |
|
41 |
qed "int_nat_number_of"; |
|
42 |
Addsimps [int_nat_number_of]; |
|
43 |
||
44 |
||
45 |
(** Successor **) |
|
46 |
||
47 |
Goal "(#0::int) <= z ==> Suc (nat z) = nat (#1 + z)"; |
|
7086 | 48 |
by (rtac sym 1); |
7032 | 49 |
by (asm_simp_tac (simpset() addsimps [nat_eq_iff]) 1); |
50 |
qed "Suc_nat_eq_nat_zadd1"; |
|
51 |
||
52 |
Goal "Suc (number_of v) = \ |
|
53 |
\ (if neg (number_of v) then #1 else number_of (bin_succ v))"; |
|
54 |
by (simp_tac |
|
55 |
(simpset_of Int.thy addsimps [neg_nat, nat_1, not_neg_eq_ge_0, |
|
56 |
nat_number_of_def, int_Suc, |
|
57 |
Suc_nat_eq_nat_zadd1, number_of_succ]) 1); |
|
58 |
qed "Suc_nat_number_of"; |
|
8737
f9733879ff25
instantiates new simprocs for numerals of type "nat"
paulson
parents:
8698
diff
changeset
|
59 |
Addsimps [Suc_nat_number_of]; |
f9733879ff25
instantiates new simprocs for numerals of type "nat"
paulson
parents:
8698
diff
changeset
|
60 |
|
f9733879ff25
instantiates new simprocs for numerals of type "nat"
paulson
parents:
8698
diff
changeset
|
61 |
Goal "Suc (number_of v + n) = \ |
f9733879ff25
instantiates new simprocs for numerals of type "nat"
paulson
parents:
8698
diff
changeset
|
62 |
\ (if neg (number_of v) then #1+n else number_of (bin_succ v) + n)"; |
f9733879ff25
instantiates new simprocs for numerals of type "nat"
paulson
parents:
8698
diff
changeset
|
63 |
by (Simp_tac 1); |
f9733879ff25
instantiates new simprocs for numerals of type "nat"
paulson
parents:
8698
diff
changeset
|
64 |
qed "Suc_nat_number_of_add"; |
7032 | 65 |
|
66 |
Goal "Suc #0 = #1"; |
|
8737
f9733879ff25
instantiates new simprocs for numerals of type "nat"
paulson
parents:
8698
diff
changeset
|
67 |
by (Simp_tac 1); |
7032 | 68 |
qed "Suc_numeral_0_eq_1"; |
69 |
||
70 |
Goal "Suc #1 = #2"; |
|
8737
f9733879ff25
instantiates new simprocs for numerals of type "nat"
paulson
parents:
8698
diff
changeset
|
71 |
by (Simp_tac 1); |
7032 | 72 |
qed "Suc_numeral_1_eq_2"; |
73 |
||
74 |
(** Addition **) |
|
75 |
||
8028 | 76 |
Goal "[| (#0::int) <= z; #0 <= z' |] ==> nat (z+z') = nat z + nat z'"; |
7032 | 77 |
by (rtac (inj_int RS injD) 1); |
78 |
by (asm_simp_tac (simpset() addsimps [zadd_int RS sym]) 1); |
|
8028 | 79 |
qed "nat_add_distrib"; |
7032 | 80 |
|
81 |
(*"neg" is used in rewrite rules for binary comparisons*) |
|
82 |
Goal "(number_of v :: nat) + number_of v' = \ |
|
83 |
\ (if neg (number_of v) then number_of v' \ |
|
84 |
\ else if neg (number_of v') then number_of v \ |
|
85 |
\ else number_of (bin_add v v'))"; |
|
86 |
by (simp_tac |
|
87 |
(simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, |
|
8028 | 88 |
nat_add_distrib RS sym, number_of_add]) 1); |
7032 | 89 |
qed "add_nat_number_of"; |
90 |
||
91 |
Addsimps [add_nat_number_of]; |
|
92 |
||
93 |
||
94 |
(** Subtraction **) |
|
95 |
||
8028 | 96 |
Goal "[| (#0::int) <= z'; z' <= z |] ==> nat (z-z') = nat z - nat z'"; |
7032 | 97 |
by (rtac (inj_int RS injD) 1); |
98 |
by (asm_simp_tac (simpset() addsimps [zdiff_int RS sym, nat_le_eq_zle]) 1); |
|
8028 | 99 |
qed "nat_diff_distrib"; |
7032 | 100 |
|
101 |
||
102 |
Goal "nat z - nat z' = \ |
|
103 |
\ (if neg z' then nat z \ |
|
104 |
\ else let d = z-z' in \ |
|
105 |
\ if neg d then 0 else nat d)"; |
|
8028 | 106 |
by (simp_tac (simpset() addsimps [Let_def, nat_diff_distrib RS sym, |
7032 | 107 |
neg_eq_less_0, not_neg_eq_ge_0]) 1); |
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
108 |
by (simp_tac (simpset() addsimps [diff_is_0_eq, nat_le_eq_zle]) 1); |
7032 | 109 |
qed "diff_nat_eq_if"; |
110 |
||
111 |
Goalw [nat_number_of_def] |
|
112 |
"(number_of v :: nat) - number_of v' = \ |
|
113 |
\ (if neg (number_of v') then number_of v \ |
|
114 |
\ else let d = number_of (bin_add v (bin_minus v')) in \ |
|
115 |
\ if neg d then #0 else nat d)"; |
|
116 |
by (simp_tac |
|
117 |
(simpset_of Int.thy delcongs [if_weak_cong] |
|
118 |
addsimps [not_neg_eq_ge_0, nat_0, |
|
119 |
diff_nat_eq_if, diff_number_of_eq]) 1); |
|
120 |
qed "diff_nat_number_of"; |
|
121 |
||
122 |
Addsimps [diff_nat_number_of]; |
|
123 |
||
124 |
||
125 |
(** Multiplication **) |
|
126 |
||
8028 | 127 |
Goal "(#0::int) <= z ==> nat (z*z') = nat z * nat z'"; |
7032 | 128 |
by (case_tac "#0 <= z'" 1); |
9064
eacebbcafe57
tidied a proof using new lemmas for signs of products
paulson
parents:
8935
diff
changeset
|
129 |
by (asm_full_simp_tac (simpset() addsimps [zmult_le_0_iff]) 2); |
7032 | 130 |
by (rtac (inj_int RS injD) 1); |
9064
eacebbcafe57
tidied a proof using new lemmas for signs of products
paulson
parents:
8935
diff
changeset
|
131 |
by (asm_simp_tac (simpset() addsimps [zmult_int RS sym, |
eacebbcafe57
tidied a proof using new lemmas for signs of products
paulson
parents:
8935
diff
changeset
|
132 |
int_0_le_mult_iff]) 1); |
8028 | 133 |
qed "nat_mult_distrib"; |
7032 | 134 |
|
9945 | 135 |
Goal "z <= (#0::int) ==> nat(z*z') = nat(-z) * nat(-z')"; |
136 |
by (rtac trans 1); |
|
137 |
by (rtac nat_mult_distrib 2); |
|
138 |
by Auto_tac; |
|
139 |
qed "nat_mult_distrib_neg"; |
|
140 |
||
7032 | 141 |
Goal "(number_of v :: nat) * number_of v' = \ |
142 |
\ (if neg (number_of v) then #0 else number_of (bin_mult v v'))"; |
|
143 |
by (simp_tac |
|
144 |
(simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, |
|
8028 | 145 |
nat_mult_distrib RS sym, number_of_mult, |
7032 | 146 |
nat_0]) 1); |
147 |
qed "mult_nat_number_of"; |
|
148 |
||
149 |
Addsimps [mult_nat_number_of]; |
|
150 |
||
151 |
||
152 |
(** Quotient **) |
|
153 |
||
8028 | 154 |
Goal "(#0::int) <= z ==> nat (z div z') = nat z div nat z'"; |
7032 | 155 |
by (case_tac "#0 <= z'" 1); |
156 |
by (auto_tac (claset(), |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
157 |
simpset() addsimps [div_nonneg_neg_le0, DIVISION_BY_ZERO_DIV])); |
7032 | 158 |
by (zdiv_undefined_case_tac "z' = #0" 1); |
159 |
by (simp_tac (simpset() addsimps [numeral_0_eq_0, DIVISION_BY_ZERO_DIV]) 1); |
|
160 |
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset())); |
|
161 |
by (rename_tac "m m'" 1); |
|
162 |
by (subgoal_tac "#0 <= int m div int m'" 1); |
|
9945 | 163 |
by (asm_full_simp_tac |
164 |
(simpset() addsimps [numeral_0_eq_0, pos_imp_zdiv_nonneg_iff]) 2); |
|
7032 | 165 |
by (rtac (inj_int RS injD) 1); |
166 |
by (Asm_simp_tac 1); |
|
167 |
by (res_inst_tac [("r", "int (m mod m')")] quorem_div 1); |
|
168 |
by (Force_tac 2); |
|
9945 | 169 |
by (asm_full_simp_tac |
170 |
(simpset() addsimps [nat_less_iff RS sym, quorem_def, |
|
171 |
numeral_0_eq_0, zadd_int, zmult_int]) 1); |
|
7032 | 172 |
by (rtac (mod_div_equality RS sym RS trans) 1); |
173 |
by (asm_simp_tac (simpset() addsimps add_ac@mult_ac) 1); |
|
8028 | 174 |
qed "nat_div_distrib"; |
7032 | 175 |
|
176 |
Goal "(number_of v :: nat) div number_of v' = \ |
|
177 |
\ (if neg (number_of v) then #0 \ |
|
178 |
\ else nat (number_of v div number_of v'))"; |
|
179 |
by (simp_tac |
|
180 |
(simpset_of Int.thy addsimps [not_neg_eq_ge_0, nat_number_of_def, neg_nat, |
|
8028 | 181 |
nat_div_distrib RS sym, nat_0]) 1); |
7032 | 182 |
qed "div_nat_number_of"; |
183 |
||
184 |
Addsimps [div_nat_number_of]; |
|
185 |
||
186 |
||
187 |
(** Remainder **) |
|
188 |
||
189 |
(*Fails if z'<0: the LHS collapses to (nat z) but the RHS doesn't*) |
|
8028 | 190 |
Goal "[| (#0::int) <= z; #0 <= z' |] ==> nat (z mod z') = nat z mod nat z'"; |
7032 | 191 |
by (zdiv_undefined_case_tac "z' = #0" 1); |
192 |
by (simp_tac (simpset() addsimps [numeral_0_eq_0, DIVISION_BY_ZERO_MOD]) 1); |
|
193 |
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset())); |
|
194 |
by (rename_tac "m m'" 1); |
|
195 |
by (subgoal_tac "#0 <= int m mod int m'" 1); |
|
9945 | 196 |
by (asm_full_simp_tac |
197 |
(simpset() addsimps [nat_less_iff, numeral_0_eq_0, pos_mod_sign]) 2); |
|
7032 | 198 |
by (rtac (inj_int RS injD) 1); |
199 |
by (Asm_simp_tac 1); |
|
200 |
by (res_inst_tac [("q", "int (m div m')")] quorem_mod 1); |
|
201 |
by (Force_tac 2); |
|
9945 | 202 |
by (asm_full_simp_tac |
203 |
(simpset() addsimps [nat_less_iff RS sym, quorem_def, |
|
204 |
numeral_0_eq_0, zadd_int, zmult_int]) 1); |
|
7032 | 205 |
by (rtac (mod_div_equality RS sym RS trans) 1); |
206 |
by (asm_simp_tac (simpset() addsimps add_ac@mult_ac) 1); |
|
8028 | 207 |
qed "nat_mod_distrib"; |
7032 | 208 |
|
209 |
Goal "(number_of v :: nat) mod number_of v' = \ |
|
210 |
\ (if neg (number_of v) then #0 \ |
|
211 |
\ else if neg (number_of v') then number_of v \ |
|
212 |
\ else nat (number_of v mod number_of v'))"; |
|
213 |
by (simp_tac |
|
214 |
(simpset_of Int.thy addsimps [not_neg_eq_ge_0, nat_number_of_def, |
|
215 |
neg_nat, nat_0, DIVISION_BY_ZERO_MOD, |
|
8028 | 216 |
nat_mod_distrib RS sym]) 1); |
7032 | 217 |
qed "mod_nat_number_of"; |
218 |
||
219 |
Addsimps [mod_nat_number_of]; |
|
220 |
||
221 |
||
222 |
(*** Comparisons ***) |
|
223 |
||
224 |
(** Equals (=) **) |
|
225 |
||
226 |
Goal "[| (#0::int) <= z; #0 <= z' |] ==> (nat z = nat z') = (z=z')"; |
|
227 |
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset())); |
|
228 |
qed "eq_nat_nat_iff"; |
|
229 |
||
230 |
(*"neg" is used in rewrite rules for binary comparisons*) |
|
231 |
Goal "((number_of v :: nat) = number_of v') = \ |
|
8795 | 232 |
\ (if neg (number_of v) then (iszero (number_of v') | neg (number_of v')) \ |
233 |
\ else if neg (number_of v') then iszero (number_of v) \ |
|
234 |
\ else iszero (number_of (bin_add v (bin_minus v'))))"; |
|
7032 | 235 |
by (simp_tac |
236 |
(simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, |
|
237 |
eq_nat_nat_iff, eq_number_of_eq, nat_0]) 1); |
|
8795 | 238 |
by (simp_tac (simpset_of Int.thy addsimps [nat_eq_iff, nat_eq_iff2, |
239 |
iszero_def]) 1); |
|
9425 | 240 |
by (simp_tac (simpset () addsimps [not_neg_eq_ge_0 RS sym]) 1); |
7032 | 241 |
qed "eq_nat_number_of"; |
242 |
||
243 |
Addsimps [eq_nat_number_of]; |
|
244 |
||
245 |
(** Less-than (<) **) |
|
246 |
||
247 |
(*"neg" is used in rewrite rules for binary comparisons*) |
|
248 |
Goal "((number_of v :: nat) < number_of v') = \ |
|
249 |
\ (if neg (number_of v) then neg (number_of (bin_minus v')) \ |
|
250 |
\ else neg (number_of (bin_add v (bin_minus v'))))"; |
|
251 |
by (simp_tac |
|
252 |
(simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, |
|
253 |
nat_less_eq_zless, less_number_of_eq_neg, |
|
254 |
nat_0]) 1); |
|
255 |
by (simp_tac (simpset_of Int.thy addsimps [neg_eq_less_int0, zminus_zless, |
|
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
256 |
number_of_minus, zless_nat_eq_int_zless]) 1); |
7032 | 257 |
qed "less_nat_number_of"; |
258 |
||
259 |
Addsimps [less_nat_number_of]; |
|
260 |
||
261 |
||
262 |
(** Less-than-or-equals (<=) **) |
|
263 |
||
264 |
Goal "(number_of x <= (number_of y::nat)) = \ |
|
265 |
\ (~ number_of y < (number_of x::nat))"; |
|
266 |
by (rtac (linorder_not_less RS sym) 1); |
|
267 |
qed "le_nat_number_of_eq_not_less"; |
|
268 |
||
269 |
Addsimps [le_nat_number_of_eq_not_less]; |
|
270 |
||
271 |
(*** New versions of existing theorems involving 0, 1, 2 ***) |
|
272 |
||
273 |
(*Maps n to #n for n = 0, 1, 2*) |
|
274 |
val numeral_sym_ss = |
|
275 |
HOL_ss addsimps [numeral_0_eq_0 RS sym, |
|
276 |
numeral_1_eq_1 RS sym, |
|
277 |
numeral_2_eq_2 RS sym, |
|
278 |
Suc_numeral_1_eq_2, Suc_numeral_0_eq_1]; |
|
279 |
||
9425 | 280 |
fun rename_numerals th = simplify numeral_sym_ss (Thm.transfer (the_context ()) th); |
7032 | 281 |
|
282 |
(*Maps #n to n for n = 0, 1, 2*) |
|
283 |
val numeral_ss = |
|
284 |
simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1, numeral_2_eq_2]; |
|
285 |
||
286 |
(** Nat **) |
|
287 |
||
288 |
Goal "#0 < n ==> n = Suc(n - #1)"; |
|
289 |
by (asm_full_simp_tac numeral_ss 1); |
|
290 |
qed "Suc_pred'"; |
|
291 |
||
292 |
(*Expresses a natural number constant as the Suc of another one. |
|
293 |
NOT suitable for rewriting because n recurs in the condition.*) |
|
294 |
bind_thm ("expand_Suc", inst "n" "number_of ?v" Suc_pred'); |
|
295 |
||
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
296 |
(** NatDef & Nat **) |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
297 |
|
9425 | 298 |
Addsimps (map rename_numerals [min_0L, min_0R, max_0L, max_0R]); |
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
299 |
|
9425 | 300 |
AddIffs (map rename_numerals |
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
301 |
[Suc_not_Zero, Zero_not_Suc, zero_less_Suc, not_less0, less_one, |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
302 |
le0, le_0_eq, |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
303 |
neq0_conv, zero_neq_conv, not_gr0]); |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
304 |
|
7032 | 305 |
(** Arith **) |
306 |
||
8776 | 307 |
(*Identity laws for + - * *) |
308 |
val basic_renamed_arith_simps = |
|
9425 | 309 |
map rename_numerals |
8776 | 310 |
[diff_0, diff_0_eq_0, add_0, add_0_right, |
311 |
mult_0, mult_0_right, mult_1, mult_1_right]; |
|
312 |
||
313 |
(*Non-trivial simplifications*) |
|
314 |
val other_renamed_arith_simps = |
|
9425 | 315 |
map rename_numerals |
9633
a71a83253997
better rules for cancellation of common factors across comparisons
paulson
parents:
9425
diff
changeset
|
316 |
[diff_is_0_eq, zero_is_diff_eq, zero_less_diff, |
8776 | 317 |
mult_is_0, zero_is_mult, zero_less_mult_iff, mult_eq_1_iff]; |
318 |
||
319 |
Addsimps (basic_renamed_arith_simps @ other_renamed_arith_simps); |
|
7032 | 320 |
|
9425 | 321 |
AddIffs (map rename_numerals [add_is_0, zero_is_add, add_gr_0]); |
7032 | 322 |
|
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
323 |
Goal "Suc n = n + #1"; |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
324 |
by (asm_simp_tac numeral_ss 1); |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
325 |
qed "Suc_eq_add_numeral_1"; |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
326 |
|
7032 | 327 |
(* These two can be useful when m = number_of... *) |
328 |
||
329 |
Goal "(m::nat) + n = (if m=#0 then n else Suc ((m - #1) + n))"; |
|
8442
96023903c2df
case_tac now subsumes both boolean and datatype cases;
wenzelm
parents:
8423
diff
changeset
|
330 |
by (case_tac "m" 1); |
7032 | 331 |
by (ALLGOALS (asm_simp_tac numeral_ss)); |
332 |
qed "add_eq_if"; |
|
333 |
||
334 |
Goal "(m::nat) * n = (if m=#0 then #0 else n + ((m - #1) * n))"; |
|
8442
96023903c2df
case_tac now subsumes both boolean and datatype cases;
wenzelm
parents:
8423
diff
changeset
|
335 |
by (case_tac "m" 1); |
7032 | 336 |
by (ALLGOALS (asm_simp_tac numeral_ss)); |
337 |
qed "mult_eq_if"; |
|
338 |
||
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
339 |
Goal "(p ^ m :: nat) = (if m=#0 then #1 else p * (p ^ (m - #1)))"; |
8442
96023903c2df
case_tac now subsumes both boolean and datatype cases;
wenzelm
parents:
8423
diff
changeset
|
340 |
by (case_tac "m" 1); |
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
341 |
by (ALLGOALS (asm_simp_tac numeral_ss)); |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
342 |
qed "power_eq_if"; |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
343 |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
344 |
Goal "[| #0<n; #0<m |] ==> m - n < (m::nat)"; |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
345 |
by (asm_full_simp_tac (numeral_ss addsimps [diff_less]) 1); |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
346 |
qed "diff_less'"; |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
347 |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
348 |
Addsimps [inst "n" "number_of ?v" diff_less']; |
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
349 |
|
7032 | 350 |
(*various theorems that aren't in the default simpset*) |
9425 | 351 |
bind_thm ("add_is_one'", rename_numerals add_is_1); |
352 |
bind_thm ("one_is_add'", rename_numerals one_is_add); |
|
353 |
bind_thm ("zero_induct'", rename_numerals zero_induct); |
|
354 |
bind_thm ("diff_self_eq_0'", rename_numerals diff_self_eq_0); |
|
355 |
bind_thm ("mult_eq_self_implies_10'", rename_numerals mult_eq_self_implies_10); |
|
356 |
bind_thm ("le_pred_eq'", rename_numerals le_pred_eq); |
|
357 |
bind_thm ("less_pred_eq'", rename_numerals less_pred_eq); |
|
7032 | 358 |
|
359 |
(** Divides **) |
|
360 |
||
9425 | 361 |
Addsimps (map rename_numerals [mod_1, mod_0, div_1, div_0]); |
362 |
AddIffs (map rename_numerals [dvd_1_left, dvd_0_right]); |
|
7032 | 363 |
|
364 |
(*useful?*) |
|
9425 | 365 |
bind_thm ("mod_self'", rename_numerals mod_self); |
366 |
bind_thm ("div_self'", rename_numerals div_self); |
|
367 |
bind_thm ("div_less'", rename_numerals div_less); |
|
368 |
bind_thm ("mod_mult_self_is_zero'", rename_numerals mod_mult_self_is_0); |
|
7032 | 369 |
|
370 |
(** Power **) |
|
371 |
||
372 |
Goal "(p::nat) ^ #0 = #1"; |
|
373 |
by (simp_tac numeral_ss 1); |
|
374 |
qed "power_zero"; |
|
8792 | 375 |
|
376 |
Goal "(p::nat) ^ #1 = p"; |
|
377 |
by (simp_tac numeral_ss 1); |
|
378 |
qed "power_one"; |
|
379 |
Addsimps [power_zero, power_one]; |
|
380 |
||
381 |
Goal "(p::nat) ^ #2 = p*p"; |
|
382 |
by (simp_tac numeral_ss 1); |
|
383 |
qed "power_two"; |
|
7032 | 384 |
|
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
385 |
Goal "#0 < (i::nat) ==> #0 < i^n"; |
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
386 |
by (asm_simp_tac numeral_ss 1); |
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
387 |
qed "zero_less_power'"; |
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
388 |
Addsimps [zero_less_power']; |
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
389 |
|
9425 | 390 |
bind_thm ("binomial_zero", rename_numerals binomial_0); |
391 |
bind_thm ("binomial_Suc'", rename_numerals binomial_Suc); |
|
392 |
bind_thm ("binomial_n_n'", rename_numerals binomial_n_n); |
|
7032 | 393 |
|
394 |
(*binomial_0_Suc doesn't work well on numerals*) |
|
9425 | 395 |
Addsimps (map rename_numerals [binomial_n_0, binomial_zero, binomial_1]); |
7032 | 396 |
|
9425 | 397 |
Addsimps [rename_numerals card_Pow]; |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
398 |
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
399 |
(*** Comparisons involving (0::nat) ***) |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
400 |
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
401 |
Goal "(number_of v = (0::nat)) = \ |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
402 |
\ (if neg (number_of v) then True else iszero (number_of v))"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
403 |
by (simp_tac (simpset() addsimps [numeral_0_eq_0 RS sym]) 1); |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
404 |
qed "eq_number_of_0"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
405 |
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
406 |
Goal "((0::nat) = number_of v) = \ |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
407 |
\ (if neg (number_of v) then True else iszero (number_of v))"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
408 |
by (rtac ([eq_sym_conv, eq_number_of_0] MRS trans) 1); |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
409 |
qed "eq_0_number_of"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
410 |
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
411 |
Goal "((0::nat) < number_of v) = neg (number_of (bin_minus v))"; |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
412 |
by (simp_tac (simpset() addsimps [numeral_0_eq_0 RS sym]) 1); |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
413 |
qed "less_0_number_of"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
414 |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
415 |
(*Simplification already handles n<0, n<=0 and 0<=n.*) |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
416 |
Addsimps [eq_number_of_0, eq_0_number_of, less_0_number_of]; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
417 |
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
418 |
Goal "neg (number_of v) ==> number_of v = (0::nat)"; |
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
419 |
by (asm_simp_tac (simpset() addsimps [numeral_0_eq_0 RS sym]) 1); |
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
420 |
qed "neg_imp_number_of_eq_0"; |
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
421 |
|
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
422 |
|
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
423 |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
424 |
(*** Comparisons involving Suc ***) |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
425 |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
426 |
Goal "(number_of v = Suc n) = \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
427 |
\ (let pv = number_of (bin_pred v) in \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
428 |
\ if neg pv then False else nat pv = n)"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
429 |
by (simp_tac |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
430 |
(simpset_of Int.thy addsimps |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
431 |
[Let_def, neg_eq_less_0, linorder_not_less, number_of_pred, |
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
432 |
nat_number_of_def, zadd_0] @ zadd_ac) 1); |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
433 |
by (res_inst_tac [("x", "number_of v")] spec 1); |
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
434 |
by (auto_tac (claset(), simpset() addsimps [nat_eq_iff])); |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
435 |
qed "eq_number_of_Suc"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
436 |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
437 |
Goal "(Suc n = number_of v) = \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
438 |
\ (let pv = number_of (bin_pred v) in \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
439 |
\ if neg pv then False else nat pv = n)"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
440 |
by (rtac ([eq_sym_conv, eq_number_of_Suc] MRS trans) 1); |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
441 |
qed "Suc_eq_number_of"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
442 |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
443 |
Goal "(number_of v < Suc n) = \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
444 |
\ (let pv = number_of (bin_pred v) in \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
445 |
\ if neg pv then True else nat pv < n)"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
446 |
by (simp_tac |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
447 |
(simpset_of Int.thy addsimps |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
448 |
[Let_def, neg_eq_less_0, linorder_not_less, number_of_pred, |
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
449 |
nat_number_of_def, zadd_0] @ zadd_ac) 1); |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
450 |
by (res_inst_tac [("x", "number_of v")] spec 1); |
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
451 |
by (auto_tac (claset(), simpset() addsimps [nat_less_iff])); |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
452 |
qed "less_number_of_Suc"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
453 |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
454 |
Goal "(Suc n < number_of v) = \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
455 |
\ (let pv = number_of (bin_pred v) in \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
456 |
\ if neg pv then False else n < nat pv)"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
457 |
by (simp_tac |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
458 |
(simpset_of Int.thy addsimps |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
459 |
[Let_def, neg_eq_less_0, linorder_not_less, number_of_pred, |
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
460 |
nat_number_of_def, zadd_0] @ zadd_ac) 1); |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
461 |
by (res_inst_tac [("x", "number_of v")] spec 1); |
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
462 |
by (auto_tac (claset(), simpset() addsimps [zless_nat_eq_int_zless])); |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
463 |
qed "less_Suc_number_of"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
464 |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
465 |
Goal "(number_of v <= Suc n) = \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
466 |
\ (let pv = number_of (bin_pred v) in \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
467 |
\ if neg pv then True else nat pv <= n)"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
468 |
by (simp_tac |
9425 | 469 |
(simpset () addsimps |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
470 |
[Let_def, less_Suc_number_of, linorder_not_less RS sym]) 1); |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
471 |
qed "le_number_of_Suc"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
472 |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
473 |
Goal "(Suc n <= number_of v) = \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
474 |
\ (let pv = number_of (bin_pred v) in \ |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
475 |
\ if neg pv then False else n <= nat pv)"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
476 |
by (simp_tac |
9425 | 477 |
(simpset () addsimps |
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
478 |
[Let_def, less_number_of_Suc, linorder_not_less RS sym]) 1); |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
479 |
qed "le_Suc_number_of"; |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
480 |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
481 |
Addsimps [eq_number_of_Suc, Suc_eq_number_of, |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
482 |
less_number_of_Suc, less_Suc_number_of, |
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
483 |
le_number_of_Suc, le_Suc_number_of]; |
8116 | 484 |
|
8257 | 485 |
(* Push int(.) inwards: *) |
8116 | 486 |
Addsimps [int_Suc,zadd_int RS sym]; |
8798 | 487 |
|
488 |
Goal "(m+m = n+n) = (m = (n::int))"; |
|
489 |
by Auto_tac; |
|
490 |
val lemma1 = result(); |
|
491 |
||
492 |
Goal "m+m ~= int 1 + n + n"; |
|
493 |
by Auto_tac; |
|
494 |
by (dres_inst_tac [("f", "%x. x mod #2")] arg_cong 1); |
|
495 |
by (full_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); |
|
496 |
val lemma2 = result(); |
|
497 |
||
498 |
Goal "((number_of (v BIT x) ::int) = number_of (w BIT y)) = \ |
|
499 |
\ (x=y & (((number_of v) ::int) = number_of w))"; |
|
500 |
by (simp_tac (simpset_of Int.thy addsimps |
|
501 |
[number_of_BIT, lemma1, lemma2, eq_commute]) 1); |
|
502 |
qed "eq_number_of_BIT_BIT"; |
|
503 |
||
504 |
Goal "((number_of (v BIT x) ::int) = number_of Pls) = \ |
|
505 |
\ (x=False & (((number_of v) ::int) = number_of Pls))"; |
|
506 |
by (simp_tac (simpset_of Int.thy addsimps |
|
507 |
[number_of_BIT, number_of_Pls, eq_commute]) 1); |
|
508 |
by (res_inst_tac [("x", "number_of v")] spec 1); |
|
509 |
by Safe_tac; |
|
510 |
by (ALLGOALS Full_simp_tac); |
|
511 |
by (dres_inst_tac [("f", "%x. x mod #2")] arg_cong 1); |
|
512 |
by (full_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); |
|
513 |
qed "eq_number_of_BIT_Pls"; |
|
514 |
||
515 |
Goal "((number_of (v BIT x) ::int) = number_of Min) = \ |
|
516 |
\ (x=True & (((number_of v) ::int) = number_of Min))"; |
|
517 |
by (simp_tac (simpset_of Int.thy addsimps |
|
518 |
[number_of_BIT, number_of_Min, eq_commute]) 1); |
|
519 |
by (res_inst_tac [("x", "number_of v")] spec 1); |
|
520 |
by Auto_tac; |
|
521 |
by (dres_inst_tac [("f", "%x. x mod #2")] arg_cong 1); |
|
522 |
by Auto_tac; |
|
523 |
qed "eq_number_of_BIT_Min"; |
|
524 |
||
525 |
Goal "(number_of Pls ::int) ~= number_of Min"; |
|
526 |
by Auto_tac; |
|
527 |
qed "eq_number_of_Pls_Min"; |
|
9945 | 528 |
|
529 |
||
530 |
(*** Further lemmas about "nat" ***) |
|
531 |
||
532 |
Goal "nat (abs (w * z)) = nat (abs w) * nat (abs z)"; |
|
533 |
by (case_tac "z=#0 | w=#0" 1); |
|
534 |
by Auto_tac; |
|
535 |
by (simp_tac (simpset() addsimps [zabs_def, nat_mult_distrib RS sym, |
|
536 |
nat_mult_distrib_neg RS sym, zmult_less_0_iff]) 1); |
|
537 |
by (arith_tac 1); |
|
538 |
qed "nat_abs_mult_distrib"; |