| author | wenzelm |
| Thu, 27 Jul 2000 18:23:12 +0200 | |
| changeset 9450 | c97dba47e504 |
| parent 9425 | fd6866d90ec1 |
| child 9633 | a71a83253997 |
| 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 |
|
135 |
Goal "(number_of v :: nat) * number_of v' = \ |
|
136 |
\ (if neg (number_of v) then #0 else number_of (bin_mult v v'))"; |
|
137 |
by (simp_tac |
|
138 |
(simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, |
|
| 8028 | 139 |
nat_mult_distrib RS sym, number_of_mult, |
| 7032 | 140 |
nat_0]) 1); |
141 |
qed "mult_nat_number_of"; |
|
142 |
||
143 |
Addsimps [mult_nat_number_of]; |
|
144 |
||
145 |
||
146 |
(** Quotient **) |
|
147 |
||
| 8028 | 148 |
Goal "(#0::int) <= z ==> nat (z div z') = nat z div nat z'"; |
| 7032 | 149 |
by (case_tac "#0 <= z'" 1); |
150 |
by (auto_tac (claset(), |
|
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
151 |
simpset() addsimps [div_nonneg_neg_le0, DIVISION_BY_ZERO_DIV])); |
| 7032 | 152 |
by (zdiv_undefined_case_tac "z' = #0" 1); |
153 |
by (simp_tac (simpset() addsimps [numeral_0_eq_0, DIVISION_BY_ZERO_DIV]) 1); |
|
154 |
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset())); |
|
155 |
by (rename_tac "m m'" 1); |
|
156 |
by (subgoal_tac "#0 <= int m div int m'" 1); |
|
157 |
by (asm_simp_tac (simpset() addsimps [nat_less_iff RS sym, numeral_0_eq_0, |
|
158 |
pos_imp_zdiv_nonneg_iff]) 2); |
|
159 |
by (rtac (inj_int RS injD) 1); |
|
160 |
by (Asm_simp_tac 1); |
|
161 |
by (res_inst_tac [("r", "int (m mod m')")] quorem_div 1);
|
|
162 |
by (Force_tac 2); |
|
163 |
by (asm_simp_tac (simpset() addsimps [nat_less_iff RS sym, quorem_def, |
|
| 8698 | 164 |
numeral_0_eq_0, zadd_int, zmult_int]) 1); |
| 7032 | 165 |
by (rtac (mod_div_equality RS sym RS trans) 1); |
166 |
by (asm_simp_tac (simpset() addsimps add_ac@mult_ac) 1); |
|
| 8028 | 167 |
qed "nat_div_distrib"; |
| 7032 | 168 |
|
169 |
Goal "(number_of v :: nat) div number_of v' = \ |
|
170 |
\ (if neg (number_of v) then #0 \ |
|
171 |
\ else nat (number_of v div number_of v'))"; |
|
172 |
by (simp_tac |
|
173 |
(simpset_of Int.thy addsimps [not_neg_eq_ge_0, nat_number_of_def, neg_nat, |
|
| 8028 | 174 |
nat_div_distrib RS sym, nat_0]) 1); |
| 7032 | 175 |
qed "div_nat_number_of"; |
176 |
||
177 |
Addsimps [div_nat_number_of]; |
|
178 |
||
179 |
||
180 |
(** Remainder **) |
|
181 |
||
182 |
(*Fails if z'<0: the LHS collapses to (nat z) but the RHS doesn't*) |
|
| 8028 | 183 |
Goal "[| (#0::int) <= z; #0 <= z' |] ==> nat (z mod z') = nat z mod nat z'"; |
| 7032 | 184 |
by (zdiv_undefined_case_tac "z' = #0" 1); |
185 |
by (simp_tac (simpset() addsimps [numeral_0_eq_0, DIVISION_BY_ZERO_MOD]) 1); |
|
186 |
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset())); |
|
187 |
by (rename_tac "m m'" 1); |
|
188 |
by (subgoal_tac "#0 <= int m mod int m'" 1); |
|
189 |
by (asm_simp_tac (simpset() addsimps [nat_less_iff RS sym, numeral_0_eq_0, |
|
190 |
pos_mod_sign]) 2); |
|
191 |
by (rtac (inj_int RS injD) 1); |
|
192 |
by (Asm_simp_tac 1); |
|
193 |
by (res_inst_tac [("q", "int (m div m')")] quorem_mod 1);
|
|
194 |
by (Force_tac 2); |
|
195 |
by (asm_simp_tac (simpset() addsimps [nat_less_iff RS sym, quorem_def, |
|
| 8698 | 196 |
numeral_0_eq_0, zadd_int, zmult_int]) 1); |
| 7032 | 197 |
by (rtac (mod_div_equality RS sym RS trans) 1); |
198 |
by (asm_simp_tac (simpset() addsimps add_ac@mult_ac) 1); |
|
| 8028 | 199 |
qed "nat_mod_distrib"; |
| 7032 | 200 |
|
201 |
Goal "(number_of v :: nat) mod number_of v' = \ |
|
202 |
\ (if neg (number_of v) then #0 \ |
|
203 |
\ else if neg (number_of v') then number_of v \ |
|
204 |
\ else nat (number_of v mod number_of v'))"; |
|
205 |
by (simp_tac |
|
206 |
(simpset_of Int.thy addsimps [not_neg_eq_ge_0, nat_number_of_def, |
|
207 |
neg_nat, nat_0, DIVISION_BY_ZERO_MOD, |
|
| 8028 | 208 |
nat_mod_distrib RS sym]) 1); |
| 7032 | 209 |
qed "mod_nat_number_of"; |
210 |
||
211 |
Addsimps [mod_nat_number_of]; |
|
212 |
||
213 |
||
214 |
(*** Comparisons ***) |
|
215 |
||
216 |
(** Equals (=) **) |
|
217 |
||
218 |
Goal "[| (#0::int) <= z; #0 <= z' |] ==> (nat z = nat z') = (z=z')"; |
|
219 |
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset())); |
|
220 |
qed "eq_nat_nat_iff"; |
|
221 |
||
222 |
(*"neg" is used in rewrite rules for binary comparisons*) |
|
223 |
Goal "((number_of v :: nat) = number_of v') = \ |
|
| 8795 | 224 |
\ (if neg (number_of v) then (iszero (number_of v') | neg (number_of v')) \ |
225 |
\ else if neg (number_of v') then iszero (number_of v) \ |
|
226 |
\ else iszero (number_of (bin_add v (bin_minus v'))))"; |
|
| 7032 | 227 |
by (simp_tac |
228 |
(simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, |
|
229 |
eq_nat_nat_iff, eq_number_of_eq, nat_0]) 1); |
|
| 8795 | 230 |
by (simp_tac (simpset_of Int.thy addsimps [nat_eq_iff, nat_eq_iff2, |
231 |
iszero_def]) 1); |
|
| 9425 | 232 |
by (simp_tac (simpset () addsimps [not_neg_eq_ge_0 RS sym]) 1); |
| 7032 | 233 |
qed "eq_nat_number_of"; |
234 |
||
235 |
Addsimps [eq_nat_number_of]; |
|
236 |
||
237 |
(** Less-than (<) **) |
|
238 |
||
239 |
(*"neg" is used in rewrite rules for binary comparisons*) |
|
240 |
Goal "((number_of v :: nat) < number_of v') = \ |
|
241 |
\ (if neg (number_of v) then neg (number_of (bin_minus v')) \ |
|
242 |
\ else neg (number_of (bin_add v (bin_minus v'))))"; |
|
243 |
by (simp_tac |
|
244 |
(simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, |
|
245 |
nat_less_eq_zless, less_number_of_eq_neg, |
|
246 |
nat_0]) 1); |
|
247 |
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
|
248 |
number_of_minus, zless_nat_eq_int_zless]) 1); |
| 7032 | 249 |
qed "less_nat_number_of"; |
250 |
||
251 |
Addsimps [less_nat_number_of]; |
|
252 |
||
253 |
||
254 |
(** Less-than-or-equals (<=) **) |
|
255 |
||
256 |
Goal "(number_of x <= (number_of y::nat)) = \ |
|
257 |
\ (~ number_of y < (number_of x::nat))"; |
|
258 |
by (rtac (linorder_not_less RS sym) 1); |
|
259 |
qed "le_nat_number_of_eq_not_less"; |
|
260 |
||
261 |
Addsimps [le_nat_number_of_eq_not_less]; |
|
262 |
||
263 |
(*** New versions of existing theorems involving 0, 1, 2 ***) |
|
264 |
||
265 |
(*Maps n to #n for n = 0, 1, 2*) |
|
266 |
val numeral_sym_ss = |
|
267 |
HOL_ss addsimps [numeral_0_eq_0 RS sym, |
|
268 |
numeral_1_eq_1 RS sym, |
|
269 |
numeral_2_eq_2 RS sym, |
|
270 |
Suc_numeral_1_eq_2, Suc_numeral_0_eq_1]; |
|
271 |
||
| 9425 | 272 |
fun rename_numerals th = simplify numeral_sym_ss (Thm.transfer (the_context ()) th); |
| 7032 | 273 |
|
274 |
(*Maps #n to n for n = 0, 1, 2*) |
|
275 |
val numeral_ss = |
|
276 |
simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1, numeral_2_eq_2]; |
|
277 |
||
278 |
(** Nat **) |
|
279 |
||
280 |
Goal "#0 < n ==> n = Suc(n - #1)"; |
|
281 |
by (asm_full_simp_tac numeral_ss 1); |
|
282 |
qed "Suc_pred'"; |
|
283 |
||
284 |
(*Expresses a natural number constant as the Suc of another one. |
|
285 |
NOT suitable for rewriting because n recurs in the condition.*) |
|
286 |
bind_thm ("expand_Suc", inst "n" "number_of ?v" Suc_pred');
|
|
287 |
||
|
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
288 |
(** NatDef & Nat **) |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
289 |
|
| 9425 | 290 |
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
|
291 |
|
| 9425 | 292 |
AddIffs (map rename_numerals |
|
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
293 |
[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
|
294 |
le0, le_0_eq, |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
295 |
neq0_conv, zero_neq_conv, not_gr0]); |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
296 |
|
| 7032 | 297 |
(** Arith **) |
298 |
||
| 8776 | 299 |
(*Identity laws for + - * *) |
300 |
val basic_renamed_arith_simps = |
|
| 9425 | 301 |
map rename_numerals |
| 8776 | 302 |
[diff_0, diff_0_eq_0, add_0, add_0_right, |
303 |
mult_0, mult_0_right, mult_1, mult_1_right]; |
|
304 |
||
305 |
(*Non-trivial simplifications*) |
|
306 |
val other_renamed_arith_simps = |
|
| 9425 | 307 |
map rename_numerals |
| 8776 | 308 |
[add_pred, diff_is_0_eq, zero_is_diff_eq, zero_less_diff, |
309 |
mult_is_0, zero_is_mult, zero_less_mult_iff, mult_eq_1_iff]; |
|
310 |
||
311 |
Addsimps (basic_renamed_arith_simps @ other_renamed_arith_simps); |
|
| 7032 | 312 |
|
| 9425 | 313 |
AddIffs (map rename_numerals [add_is_0, zero_is_add, add_gr_0]); |
| 7032 | 314 |
|
|
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
315 |
Goal "Suc n = n + #1"; |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
316 |
by (asm_simp_tac numeral_ss 1); |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
317 |
qed "Suc_eq_add_numeral_1"; |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
318 |
|
| 7032 | 319 |
(* These two can be useful when m = number_of... *) |
320 |
||
321 |
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
|
322 |
by (case_tac "m" 1); |
| 7032 | 323 |
by (ALLGOALS (asm_simp_tac numeral_ss)); |
324 |
qed "add_eq_if"; |
|
325 |
||
326 |
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
|
327 |
by (case_tac "m" 1); |
| 7032 | 328 |
by (ALLGOALS (asm_simp_tac numeral_ss)); |
329 |
qed "mult_eq_if"; |
|
330 |
||
|
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
331 |
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
|
332 |
by (case_tac "m" 1); |
|
7056
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
333 |
by (ALLGOALS (asm_simp_tac numeral_ss)); |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
334 |
qed "power_eq_if"; |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
335 |
|
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
336 |
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
|
337 |
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
|
338 |
qed "diff_less'"; |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
339 |
|
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
340 |
Addsimps [inst "n" "number_of ?v" diff_less']; |
|
522a7013d7df
more existing theorems renamed to use #0; also new results
paulson
parents:
7032
diff
changeset
|
341 |
|
| 7032 | 342 |
(*various theorems that aren't in the default simpset*) |
| 9425 | 343 |
bind_thm ("add_is_one'", rename_numerals add_is_1);
|
344 |
bind_thm ("one_is_add'", rename_numerals one_is_add);
|
|
345 |
bind_thm ("zero_induct'", rename_numerals zero_induct);
|
|
346 |
bind_thm ("diff_self_eq_0'", rename_numerals diff_self_eq_0);
|
|
347 |
bind_thm ("mult_eq_self_implies_10'", rename_numerals mult_eq_self_implies_10);
|
|
348 |
bind_thm ("le_pred_eq'", rename_numerals le_pred_eq);
|
|
349 |
bind_thm ("less_pred_eq'", rename_numerals less_pred_eq);
|
|
| 7032 | 350 |
|
351 |
(** Divides **) |
|
352 |
||
| 9425 | 353 |
Addsimps (map rename_numerals [mod_1, mod_0, div_1, div_0]); |
354 |
AddIffs (map rename_numerals [dvd_1_left, dvd_0_right]); |
|
| 7032 | 355 |
|
356 |
(*useful?*) |
|
| 9425 | 357 |
bind_thm ("mod_self'", rename_numerals mod_self);
|
358 |
bind_thm ("div_self'", rename_numerals div_self);
|
|
359 |
bind_thm ("div_less'", rename_numerals div_less);
|
|
360 |
bind_thm ("mod_mult_self_is_zero'", rename_numerals mod_mult_self_is_0);
|
|
| 7032 | 361 |
|
362 |
(** Power **) |
|
363 |
||
364 |
Goal "(p::nat) ^ #0 = #1"; |
|
365 |
by (simp_tac numeral_ss 1); |
|
366 |
qed "power_zero"; |
|
| 8792 | 367 |
|
368 |
Goal "(p::nat) ^ #1 = p"; |
|
369 |
by (simp_tac numeral_ss 1); |
|
370 |
qed "power_one"; |
|
371 |
Addsimps [power_zero, power_one]; |
|
372 |
||
373 |
Goal "(p::nat) ^ #2 = p*p"; |
|
374 |
by (simp_tac numeral_ss 1); |
|
375 |
qed "power_two"; |
|
| 7032 | 376 |
|
|
8864
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
377 |
Goal "#0 < (i::nat) ==> #0 < i^n"; |
|
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
378 |
by (asm_simp_tac numeral_ss 1); |
|
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
379 |
qed "zero_less_power'"; |
|
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
380 |
Addsimps [zero_less_power']; |
|
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
381 |
|
| 9425 | 382 |
bind_thm ("binomial_zero", rename_numerals binomial_0);
|
383 |
bind_thm ("binomial_Suc'", rename_numerals binomial_Suc);
|
|
384 |
bind_thm ("binomial_n_n'", rename_numerals binomial_n_n);
|
|
| 7032 | 385 |
|
386 |
(*binomial_0_Suc doesn't work well on numerals*) |
|
| 9425 | 387 |
Addsimps (map rename_numerals [binomial_n_0, binomial_zero, binomial_1]); |
| 7032 | 388 |
|
| 9425 | 389 |
Addsimps [rename_numerals card_Pow]; |
|
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
390 |
|
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
391 |
(*** Comparisons involving (0::nat) ***) |
|
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
392 |
|
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
393 |
Goal "(number_of v = (0::nat)) = \ |
|
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
394 |
\ (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
|
395 |
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
|
396 |
qed "eq_number_of_0"; |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
397 |
|
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
398 |
Goal "((0::nat) = number_of v) = \ |
|
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
399 |
\ (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
|
400 |
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
|
401 |
qed "eq_0_number_of"; |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
402 |
|
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
403 |
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
|
404 |
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
|
405 |
qed "less_0_number_of"; |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
406 |
|
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
407 |
(*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
|
408 |
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
|
409 |
|
|
8935
548901d05a0e
added type constraint ::nat because 0 is now overloaded
paulson
parents:
8864
diff
changeset
|
410 |
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
|
411 |
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
|
412 |
qed "neg_imp_number_of_eq_0"; |
|
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
413 |
|
|
a12ccd629e2c
tidying, especially to remove zcompare_rls from proofs
paulson
parents:
8798
diff
changeset
|
414 |
|
|
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
415 |
|
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
416 |
(*** Comparisons involving Suc ***) |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
417 |
|
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
418 |
Goal "(number_of v = Suc n) = \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
419 |
\ (let pv = number_of (bin_pred v) in \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
420 |
\ 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
|
421 |
by (simp_tac |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
422 |
(simpset_of Int.thy addsimps |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
423 |
[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
|
424 |
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
|
425 |
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
|
426 |
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
|
427 |
qed "eq_number_of_Suc"; |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
428 |
|
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
429 |
Goal "(Suc n = number_of v) = \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
430 |
\ (let pv = number_of (bin_pred v) in \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
431 |
\ 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
|
432 |
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
|
433 |
qed "Suc_eq_number_of"; |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
434 |
|
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
435 |
Goal "(number_of v < Suc n) = \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
436 |
\ (let pv = number_of (bin_pred v) in \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
437 |
\ 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
|
438 |
by (simp_tac |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
439 |
(simpset_of Int.thy addsimps |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
440 |
[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
|
441 |
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
|
442 |
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
|
443 |
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
|
444 |
qed "less_number_of_Suc"; |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
445 |
|
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
446 |
Goal "(Suc n < number_of v) = \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
447 |
\ (let pv = number_of (bin_pred v) in \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
448 |
\ 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
|
449 |
by (simp_tac |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
450 |
(simpset_of Int.thy addsimps |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
451 |
[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
|
452 |
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
|
453 |
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
|
454 |
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
|
455 |
qed "less_Suc_number_of"; |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
456 |
|
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
457 |
Goal "(number_of v <= Suc n) = \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
458 |
\ (let pv = number_of (bin_pred v) in \ |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
459 |
\ 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
|
460 |
by (simp_tac |
| 9425 | 461 |
(simpset () addsimps |
|
7519
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
462 |
[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
|
463 |
qed "le_number_of_Suc"; |
|
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 "(Suc n <= number_of v) = \ |
|
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 False else n <= nat pv)"; |
|
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_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
|
471 |
qed "le_Suc_number_of"; |
|
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 |
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
|
474 |
less_number_of_Suc, less_Suc_number_of, |
|
8e4a21d1ba4f
simplification of relations involving 0, Suc and natural-number numerals
paulson
parents:
7127
diff
changeset
|
475 |
le_number_of_Suc, le_Suc_number_of]; |
| 8116 | 476 |
|
| 8257 | 477 |
(* Push int(.) inwards: *) |
| 8116 | 478 |
Addsimps [int_Suc,zadd_int RS sym]; |
| 8798 | 479 |
|
480 |
Goal "(m+m = n+n) = (m = (n::int))"; |
|
481 |
by Auto_tac; |
|
482 |
val lemma1 = result(); |
|
483 |
||
484 |
Goal "m+m ~= int 1 + n + n"; |
|
485 |
by Auto_tac; |
|
486 |
by (dres_inst_tac [("f", "%x. x mod #2")] arg_cong 1);
|
|
487 |
by (full_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); |
|
488 |
val lemma2 = result(); |
|
489 |
||
490 |
Goal "((number_of (v BIT x) ::int) = number_of (w BIT y)) = \ |
|
491 |
\ (x=y & (((number_of v) ::int) = number_of w))"; |
|
492 |
by (simp_tac (simpset_of Int.thy addsimps |
|
493 |
[number_of_BIT, lemma1, lemma2, eq_commute]) 1); |
|
494 |
qed "eq_number_of_BIT_BIT"; |
|
495 |
||
496 |
Goal "((number_of (v BIT x) ::int) = number_of Pls) = \ |
|
497 |
\ (x=False & (((number_of v) ::int) = number_of Pls))"; |
|
498 |
by (simp_tac (simpset_of Int.thy addsimps |
|
499 |
[number_of_BIT, number_of_Pls, eq_commute]) 1); |
|
500 |
by (res_inst_tac [("x", "number_of v")] spec 1);
|
|
501 |
by Safe_tac; |
|
502 |
by (ALLGOALS Full_simp_tac); |
|
503 |
by (dres_inst_tac [("f", "%x. x mod #2")] arg_cong 1);
|
|
504 |
by (full_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); |
|
505 |
qed "eq_number_of_BIT_Pls"; |
|
506 |
||
507 |
Goal "((number_of (v BIT x) ::int) = number_of Min) = \ |
|
508 |
\ (x=True & (((number_of v) ::int) = number_of Min))"; |
|
509 |
by (simp_tac (simpset_of Int.thy addsimps |
|
510 |
[number_of_BIT, number_of_Min, eq_commute]) 1); |
|
511 |
by (res_inst_tac [("x", "number_of v")] spec 1);
|
|
512 |
by Auto_tac; |
|
513 |
by (dres_inst_tac [("f", "%x. x mod #2")] arg_cong 1);
|
|
514 |
by Auto_tac; |
|
515 |
qed "eq_number_of_BIT_Min"; |
|
516 |
||
517 |
Goal "(number_of Pls ::int) ~= number_of Min"; |
|
518 |
by Auto_tac; |
|
519 |
qed "eq_number_of_Pls_Min"; |