author | wenzelm |
Tue, 11 Jul 2006 12:16:52 +0200 | |
changeset 20070 | 3f31fb81b83a |
parent 16417 | 9bc16273c2d4 |
child 20807 | bd3b60f9a343 |
permissions | -rw-r--r-- |
5545 | 1 |
(* Title: HOL/ex/BinEx.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
*) |
|
6 |
||
11024 | 7 |
header {* Binary arithmetic examples *} |
8 |
||
16417 | 9 |
theory BinEx imports Main begin |
11024 | 10 |
|
14113 | 11 |
subsection {* Regression Testing for Cancellation Simprocs *} |
12 |
||
13 |
(*taken from HOL/Integ/int_arith1.ML *) |
|
14 |
||
15 |
||
16 |
lemma "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)" |
|
17 |
apply simp oops |
|
18 |
||
19 |
lemma "2*u = (u::int)" |
|
20 |
apply simp oops |
|
21 |
||
22 |
lemma "(i + j + 12 + (k::int)) - 15 = y" |
|
23 |
apply simp oops |
|
24 |
||
25 |
lemma "(i + j + 12 + (k::int)) - 5 = y" |
|
26 |
apply simp oops |
|
27 |
||
28 |
lemma "y - b < (b::int)" |
|
29 |
apply simp oops |
|
30 |
||
31 |
lemma "y - (3*b + c) < (b::int) - 2*c" |
|
32 |
apply simp oops |
|
33 |
||
34 |
lemma "(2*x - (u*v) + y) - v*3*u = (w::int)" |
|
35 |
apply simp oops |
|
36 |
||
37 |
lemma "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::int)" |
|
38 |
apply simp oops |
|
39 |
||
40 |
lemma "(2*x*u*v + (u*v)*4 + y) - v*u = (w::int)" |
|
41 |
apply simp oops |
|
42 |
||
43 |
lemma "u*v - (x*u*v + (u*v)*4 + y) = (w::int)" |
|
44 |
apply simp oops |
|
45 |
||
46 |
lemma "(i + j + 12 + (k::int)) = u + 15 + y" |
|
47 |
apply simp oops |
|
48 |
||
49 |
lemma "(i + j*2 + 12 + (k::int)) = j + 5 + y" |
|
50 |
apply simp oops |
|
51 |
||
52 |
lemma "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::int)" |
|
53 |
apply simp oops |
|
54 |
||
55 |
lemma "a + -(b+c) + b = (d::int)" |
|
56 |
apply simp oops |
|
57 |
||
58 |
lemma "a + -(b+c) - b = (d::int)" |
|
59 |
apply simp oops |
|
60 |
||
61 |
(*negative numerals*) |
|
62 |
lemma "(i + j + -2 + (k::int)) - (u + 5 + y) = zz" |
|
63 |
apply simp oops |
|
64 |
||
65 |
lemma "(i + j + -3 + (k::int)) < u + 5 + y" |
|
66 |
apply simp oops |
|
67 |
||
68 |
lemma "(i + j + 3 + (k::int)) < u + -6 + y" |
|
69 |
apply simp oops |
|
70 |
||
71 |
lemma "(i + j + -12 + (k::int)) - 15 = y" |
|
72 |
apply simp oops |
|
73 |
||
74 |
lemma "(i + j + 12 + (k::int)) - -15 = y" |
|
75 |
apply simp oops |
|
76 |
||
77 |
lemma "(i + j + -12 + (k::int)) - -15 = y" |
|
78 |
apply simp oops |
|
79 |
||
14124 | 80 |
lemma "- (2*i) + 3 + (2*i + 4) = (0::int)" |
81 |
apply simp oops |
|
82 |
||
83 |
||
14113 | 84 |
|
85 |
subsection {* Arithmetic Method Tests *} |
|
86 |
||
87 |
||
88 |
lemma "!!a::int. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d" |
|
89 |
by arith |
|
90 |
||
91 |
lemma "!!a::int. [| a < b; c < d |] ==> a-d+ 2 <= b+(-c)" |
|
92 |
by arith |
|
93 |
||
94 |
lemma "!!a::int. [| a < b; c < d |] ==> a+c+ 1 < b+d" |
|
95 |
by arith |
|
96 |
||
97 |
lemma "!!a::int. [| a <= b; b+b <= c |] ==> a+a <= c" |
|
98 |
by arith |
|
99 |
||
100 |
lemma "!!a::int. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j" |
|
101 |
by arith |
|
102 |
||
103 |
lemma "!!a::int. [| a+b < i+j; a<b; i<j |] ==> a+a - - -1 < j+j - 3" |
|
104 |
by arith |
|
105 |
||
106 |
lemma "!!a::int. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k" |
|
107 |
by arith |
|
108 |
||
109 |
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] |
|
110 |
==> a <= l" |
|
111 |
by arith |
|
112 |
||
113 |
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] |
|
114 |
==> a+a+a+a <= l+l+l+l" |
|
115 |
by arith |
|
116 |
||
117 |
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] |
|
118 |
==> a+a+a+a+a <= l+l+l+l+i" |
|
119 |
by arith |
|
120 |
||
121 |
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] |
|
122 |
==> a+a+a+a+a+a <= l+l+l+l+i+l" |
|
123 |
by arith |
|
124 |
||
125 |
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] |
|
126 |
==> 6*a <= 5*l+i" |
|
127 |
by arith |
|
128 |
||
129 |
||
130 |
||
11024 | 131 |
subsection {* The Integers *} |
132 |
||
133 |
text {* Addition *} |
|
134 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
135 |
lemma "(13::int) + 19 = 32" |
11024 | 136 |
by simp |
137 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
138 |
lemma "(1234::int) + 5678 = 6912" |
11024 | 139 |
by simp |
140 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
141 |
lemma "(1359::int) + -2468 = -1109" |
11024 | 142 |
by simp |
143 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
144 |
lemma "(93746::int) + -46375 = 47371" |
11024 | 145 |
by simp |
146 |
||
147 |
||
148 |
text {* \medskip Negation *} |
|
149 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
150 |
lemma "- (65745::int) = -65745" |
11024 | 151 |
by simp |
152 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
153 |
lemma "- (-54321::int) = 54321" |
11024 | 154 |
by simp |
155 |
||
156 |
||
157 |
text {* \medskip Multiplication *} |
|
158 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
159 |
lemma "(13::int) * 19 = 247" |
11024 | 160 |
by simp |
161 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
162 |
lemma "(-84::int) * 51 = -4284" |
11024 | 163 |
by simp |
164 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
165 |
lemma "(255::int) * 255 = 65025" |
11024 | 166 |
by simp |
167 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
168 |
lemma "(1359::int) * -2468 = -3354012" |
11024 | 169 |
by simp |
170 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
171 |
lemma "(89::int) * 10 \<noteq> 889" |
11024 | 172 |
by simp |
173 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
174 |
lemma "(13::int) < 18 - 4" |
11024 | 175 |
by simp |
176 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
177 |
lemma "(-345::int) < -242 + -100" |
11024 | 178 |
by simp |
179 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
180 |
lemma "(13557456::int) < 18678654" |
11024 | 181 |
by simp |
182 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
183 |
lemma "(999999::int) \<le> (1000001 + 1) - 2" |
11024 | 184 |
by simp |
185 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
186 |
lemma "(1234567::int) \<le> 1234567" |
11024 | 187 |
by simp |
188 |
||
15965
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents:
15013
diff
changeset
|
189 |
text{*No integer overflow!*} |
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents:
15013
diff
changeset
|
190 |
lemma "1234567 * (1234567::int) < 1234567*1234567*1234567" |
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents:
15013
diff
changeset
|
191 |
by simp |
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents:
15013
diff
changeset
|
192 |
|
11024 | 193 |
|
194 |
text {* \medskip Quotient and Remainder *} |
|
195 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
196 |
lemma "(10::int) div 3 = 3" |
11024 | 197 |
by simp |
198 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
199 |
lemma "(10::int) mod 3 = 1" |
11024 | 200 |
by simp |
201 |
||
202 |
text {* A negative divisor *} |
|
203 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
204 |
lemma "(10::int) div -3 = -4" |
11024 | 205 |
by simp |
206 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
207 |
lemma "(10::int) mod -3 = -2" |
11024 | 208 |
by simp |
5545 | 209 |
|
11024 | 210 |
text {* |
211 |
A negative dividend\footnote{The definition agrees with mathematical |
|
15965
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents:
15013
diff
changeset
|
212 |
convention and with ML, but not with the hardware of most computers} |
11024 | 213 |
*} |
214 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
215 |
lemma "(-10::int) div 3 = -4" |
11024 | 216 |
by simp |
217 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
218 |
lemma "(-10::int) mod 3 = 2" |
11024 | 219 |
by simp |
220 |
||
221 |
text {* A negative dividend \emph{and} divisor *} |
|
222 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
223 |
lemma "(-10::int) div -3 = 3" |
11024 | 224 |
by simp |
225 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
226 |
lemma "(-10::int) mod -3 = -1" |
11024 | 227 |
by simp |
228 |
||
229 |
text {* A few bigger examples *} |
|
230 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
231 |
lemma "(8452::int) mod 3 = 1" |
11024 | 232 |
by simp |
233 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
234 |
lemma "(59485::int) div 434 = 137" |
11024 | 235 |
by simp |
236 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
237 |
lemma "(1000006::int) mod 10 = 6" |
11024 | 238 |
by simp |
239 |
||
240 |
||
241 |
text {* \medskip Division by shifting *} |
|
242 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
243 |
lemma "10000000 div 2 = (5000000::int)" |
11024 | 244 |
by simp |
245 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
246 |
lemma "10000001 mod 2 = (1::int)" |
11024 | 247 |
by simp |
248 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
249 |
lemma "10000055 div 32 = (312501::int)" |
11024 | 250 |
by simp |
251 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
252 |
lemma "10000055 mod 32 = (23::int)" |
11024 | 253 |
by simp |
254 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
255 |
lemma "100094 div 144 = (695::int)" |
11024 | 256 |
by simp |
257 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
258 |
lemma "100094 mod 144 = (14::int)" |
11024 | 259 |
by simp |
260 |
||
261 |
||
12613
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
262 |
text {* \medskip Powers *} |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
263 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
264 |
lemma "2 ^ 10 = (1024::int)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
265 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
266 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
267 |
lemma "-3 ^ 7 = (-2187::int)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
268 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
269 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
270 |
lemma "13 ^ 7 = (62748517::int)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
271 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
272 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
273 |
lemma "3 ^ 15 = (14348907::int)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
274 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
275 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
276 |
lemma "-5 ^ 11 = (-48828125::int)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
277 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
278 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
279 |
|
11024 | 280 |
subsection {* The Natural Numbers *} |
281 |
||
282 |
text {* Successor *} |
|
283 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
284 |
lemma "Suc 99999 = 100000" |
11024 | 285 |
by (simp add: Suc_nat_number_of) |
286 |
-- {* not a default rewrite since sometimes we want to have @{text "Suc #nnn"} *} |
|
287 |
||
288 |
||
289 |
text {* \medskip Addition *} |
|
290 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
291 |
lemma "(13::nat) + 19 = 32" |
11024 | 292 |
by simp |
293 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
294 |
lemma "(1234::nat) + 5678 = 6912" |
11024 | 295 |
by simp |
296 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
297 |
lemma "(973646::nat) + 6475 = 980121" |
11024 | 298 |
by simp |
299 |
||
300 |
||
301 |
text {* \medskip Subtraction *} |
|
302 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
303 |
lemma "(32::nat) - 14 = 18" |
11024 | 304 |
by simp |
305 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
306 |
lemma "(14::nat) - 15 = 0" |
11024 | 307 |
by simp |
5545 | 308 |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
309 |
lemma "(14::nat) - 1576644 = 0" |
11024 | 310 |
by simp |
311 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
312 |
lemma "(48273776::nat) - 3873737 = 44400039" |
11024 | 313 |
by simp |
314 |
||
315 |
||
316 |
text {* \medskip Multiplication *} |
|
317 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
318 |
lemma "(12::nat) * 11 = 132" |
11024 | 319 |
by simp |
320 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
321 |
lemma "(647::nat) * 3643 = 2357021" |
11024 | 322 |
by simp |
323 |
||
324 |
||
325 |
text {* \medskip Quotient and Remainder *} |
|
326 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
327 |
lemma "(10::nat) div 3 = 3" |
11024 | 328 |
by simp |
329 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
330 |
lemma "(10::nat) mod 3 = 1" |
11024 | 331 |
by simp |
332 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
333 |
lemma "(10000::nat) div 9 = 1111" |
11024 | 334 |
by simp |
335 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
336 |
lemma "(10000::nat) mod 9 = 1" |
11024 | 337 |
by simp |
338 |
||
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
339 |
lemma "(10000::nat) div 16 = 625" |
11024 | 340 |
by simp |
341 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
342 |
lemma "(10000::nat) mod 16 = 0" |
11024 | 343 |
by simp |
344 |
||
5545 | 345 |
|
12613
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
346 |
text {* \medskip Powers *} |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
347 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
348 |
lemma "2 ^ 12 = (4096::nat)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
349 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
350 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
351 |
lemma "3 ^ 10 = (59049::nat)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
352 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
353 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
354 |
lemma "12 ^ 7 = (35831808::nat)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
355 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
356 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
357 |
lemma "3 ^ 14 = (4782969::nat)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
358 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
359 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
360 |
lemma "5 ^ 11 = (48828125::nat)" |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
361 |
by simp |
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
362 |
|
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents:
11868
diff
changeset
|
363 |
|
11024 | 364 |
text {* \medskip Testing the cancellation of complementary terms *} |
365 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
366 |
lemma "y + (x + -x) = (0::int) + y" |
11024 | 367 |
by simp |
368 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
369 |
lemma "y + (-x + (- y + x)) = (0::int)" |
11024 | 370 |
by simp |
371 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
372 |
lemma "-x + (y + (- y + x)) = (0::int)" |
11024 | 373 |
by simp |
374 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
375 |
lemma "x + (x + (- x + (- x + (- y + - z)))) = (0::int) - y - z" |
11024 | 376 |
by simp |
377 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
378 |
lemma "x + x - x - x - y - z = (0::int) - y - z" |
11024 | 379 |
by simp |
380 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
381 |
lemma "x + y + z - (x + z) = y - (0::int)" |
11024 | 382 |
by simp |
383 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
384 |
lemma "x + (y + (y + (y + (-x + -x)))) = (0::int) + y - x + y + y" |
11024 | 385 |
by simp |
386 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
387 |
lemma "x + (y + (y + (y + (-y + -x)))) = y + (0::int) + y" |
11024 | 388 |
by simp |
389 |
||
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11704
diff
changeset
|
390 |
lemma "x + y - x + z - x - y - z + x < (1::int)" |
11024 | 391 |
by simp |
392 |
||
393 |
||
15013 | 394 |
text{*The proofs about arithmetic yielding normal forms have been deleted: |
395 |
they are irrelevant with the new treatment of numerals.*} |
|
13192 | 396 |
|
5545 | 397 |
end |