| author | wenzelm | 
| Sat, 24 May 2008 22:04:52 +0200 | |
| changeset 26988 | 742e26213212 | 
| parent 20807 | bd3b60f9a343 | 
| child 28952 | 15a4b2cf8c34 | 
| 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  | 
lemma "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)"  | 
|
14  | 
apply simp oops  | 
|
15  | 
||
16  | 
lemma "2*u = (u::int)"  | 
|
17  | 
apply simp oops  | 
|
18  | 
||
19  | 
lemma "(i + j + 12 + (k::int)) - 15 = y"  | 
|
20  | 
apply simp oops  | 
|
21  | 
||
22  | 
lemma "(i + j + 12 + (k::int)) - 5 = y"  | 
|
23  | 
apply simp oops  | 
|
24  | 
||
25  | 
lemma "y - b < (b::int)"  | 
|
26  | 
apply simp oops  | 
|
27  | 
||
28  | 
lemma "y - (3*b + c) < (b::int) - 2*c"  | 
|
29  | 
apply simp oops  | 
|
30  | 
||
31  | 
lemma "(2*x - (u*v) + y) - v*3*u = (w::int)"  | 
|
32  | 
apply simp oops  | 
|
33  | 
||
34  | 
lemma "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::int)"  | 
|
35  | 
apply simp oops  | 
|
36  | 
||
37  | 
lemma "(2*x*u*v + (u*v)*4 + y) - v*u = (w::int)"  | 
|
38  | 
apply simp oops  | 
|
39  | 
||
40  | 
lemma "u*v - (x*u*v + (u*v)*4 + y) = (w::int)"  | 
|
41  | 
apply simp oops  | 
|
42  | 
||
43  | 
lemma "(i + j + 12 + (k::int)) = u + 15 + y"  | 
|
44  | 
apply simp oops  | 
|
45  | 
||
46  | 
lemma "(i + j*2 + 12 + (k::int)) = j + 5 + y"  | 
|
47  | 
apply simp oops  | 
|
48  | 
||
49  | 
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)"  | 
|
50  | 
apply simp oops  | 
|
51  | 
||
52  | 
lemma "a + -(b+c) + b = (d::int)"  | 
|
53  | 
apply simp oops  | 
|
54  | 
||
55  | 
lemma "a + -(b+c) - b = (d::int)"  | 
|
56  | 
apply simp oops  | 
|
57  | 
||
58  | 
(*negative numerals*)  | 
|
59  | 
lemma "(i + j + -2 + (k::int)) - (u + 5 + y) = zz"  | 
|
60  | 
apply simp oops  | 
|
61  | 
||
62  | 
lemma "(i + j + -3 + (k::int)) < u + 5 + y"  | 
|
63  | 
apply simp oops  | 
|
64  | 
||
65  | 
lemma "(i + j + 3 + (k::int)) < u + -6 + y"  | 
|
66  | 
apply simp oops  | 
|
67  | 
||
68  | 
lemma "(i + j + -12 + (k::int)) - 15 = 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  | 
||
| 14124 | 77  | 
lemma "- (2*i) + 3 + (2*i + 4) = (0::int)"  | 
78  | 
apply simp oops  | 
|
79  | 
||
80  | 
||
| 14113 | 81  | 
|
82  | 
subsection {* Arithmetic Method Tests *}
 | 
|
83  | 
||
84  | 
||
85  | 
lemma "!!a::int. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d"  | 
|
86  | 
by arith  | 
|
87  | 
||
88  | 
lemma "!!a::int. [| a < b; c < d |] ==> a-d+ 2 <= b+(-c)"  | 
|
89  | 
by arith  | 
|
90  | 
||
91  | 
lemma "!!a::int. [| a < b; c < d |] ==> a+c+ 1 < b+d"  | 
|
92  | 
by arith  | 
|
93  | 
||
94  | 
lemma "!!a::int. [| a <= b; b+b <= c |] ==> a+a <= c"  | 
|
95  | 
by arith  | 
|
96  | 
||
97  | 
lemma "!!a::int. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j"  | 
|
98  | 
by arith  | 
|
99  | 
||
100  | 
lemma "!!a::int. [| a+b < i+j; a<b; i<j |] ==> a+a - - -1 < j+j - 3"  | 
|
101  | 
by arith  | 
|
102  | 
||
103  | 
lemma "!!a::int. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k"  | 
|
104  | 
by arith  | 
|
105  | 
||
106  | 
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |]  | 
|
107  | 
==> a <= l"  | 
|
108  | 
by arith  | 
|
109  | 
||
110  | 
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |]  | 
|
111  | 
==> a+a+a+a <= l+l+l+l"  | 
|
112  | 
by arith  | 
|
113  | 
||
114  | 
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |]  | 
|
115  | 
==> a+a+a+a+a <= l+l+l+l+i"  | 
|
116  | 
by arith  | 
|
117  | 
||
118  | 
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |]  | 
|
119  | 
==> a+a+a+a+a+a <= l+l+l+l+i+l"  | 
|
120  | 
by arith  | 
|
121  | 
||
122  | 
lemma "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |]  | 
|
123  | 
==> 6*a <= 5*l+i"  | 
|
124  | 
by arith  | 
|
125  | 
||
126  | 
||
127  | 
||
| 11024 | 128  | 
subsection {* The Integers *}
 | 
129  | 
||
130  | 
text {* Addition *}
 | 
|
131  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
132  | 
lemma "(13::int) + 19 = 32"  | 
| 11024 | 133  | 
by simp  | 
134  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
135  | 
lemma "(1234::int) + 5678 = 6912"  | 
| 11024 | 136  | 
by simp  | 
137  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
138  | 
lemma "(1359::int) + -2468 = -1109"  | 
| 11024 | 139  | 
by simp  | 
140  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
141  | 
lemma "(93746::int) + -46375 = 47371"  | 
| 11024 | 142  | 
by simp  | 
143  | 
||
144  | 
||
145  | 
text {* \medskip Negation *}
 | 
|
146  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
147  | 
lemma "- (65745::int) = -65745"  | 
| 11024 | 148  | 
by simp  | 
149  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
150  | 
lemma "- (-54321::int) = 54321"  | 
| 11024 | 151  | 
by simp  | 
152  | 
||
153  | 
||
154  | 
text {* \medskip Multiplication *}
 | 
|
155  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
156  | 
lemma "(13::int) * 19 = 247"  | 
| 11024 | 157  | 
by simp  | 
158  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
159  | 
lemma "(-84::int) * 51 = -4284"  | 
| 11024 | 160  | 
by simp  | 
161  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
162  | 
lemma "(255::int) * 255 = 65025"  | 
| 11024 | 163  | 
by simp  | 
164  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
165  | 
lemma "(1359::int) * -2468 = -3354012"  | 
| 11024 | 166  | 
by simp  | 
167  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
168  | 
lemma "(89::int) * 10 \<noteq> 889"  | 
| 11024 | 169  | 
by simp  | 
170  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
171  | 
lemma "(13::int) < 18 - 4"  | 
| 11024 | 172  | 
by simp  | 
173  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
174  | 
lemma "(-345::int) < -242 + -100"  | 
| 11024 | 175  | 
by simp  | 
176  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
177  | 
lemma "(13557456::int) < 18678654"  | 
| 11024 | 178  | 
by simp  | 
179  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
180  | 
lemma "(999999::int) \<le> (1000001 + 1) - 2"  | 
| 11024 | 181  | 
by simp  | 
182  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
183  | 
lemma "(1234567::int) \<le> 1234567"  | 
| 11024 | 184  | 
by simp  | 
185  | 
||
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15013 
diff
changeset
 | 
186  | 
text{*No integer overflow!*}
 | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15013 
diff
changeset
 | 
187  | 
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
 | 
188  | 
by simp  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15013 
diff
changeset
 | 
189  | 
|
| 11024 | 190  | 
|
191  | 
text {* \medskip Quotient and Remainder *}
 | 
|
192  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
193  | 
lemma "(10::int) div 3 = 3"  | 
| 11024 | 194  | 
by simp  | 
195  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
196  | 
lemma "(10::int) mod 3 = 1"  | 
| 11024 | 197  | 
by simp  | 
198  | 
||
199  | 
text {* A negative divisor *}
 | 
|
200  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
201  | 
lemma "(10::int) div -3 = -4"  | 
| 11024 | 202  | 
by simp  | 
203  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
204  | 
lemma "(10::int) mod -3 = -2"  | 
| 11024 | 205  | 
by simp  | 
| 5545 | 206  | 
|
| 11024 | 207  | 
text {*
 | 
208  | 
  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
 | 
209  | 
convention and with ML, but not with the hardware of most computers}  | 
| 11024 | 210  | 
*}  | 
211  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
212  | 
lemma "(-10::int) div 3 = -4"  | 
| 11024 | 213  | 
by simp  | 
214  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
215  | 
lemma "(-10::int) mod 3 = 2"  | 
| 11024 | 216  | 
by simp  | 
217  | 
||
218  | 
text {* A negative dividend \emph{and} divisor *}
 | 
|
219  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
220  | 
lemma "(-10::int) div -3 = 3"  | 
| 11024 | 221  | 
by simp  | 
222  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
223  | 
lemma "(-10::int) mod -3 = -1"  | 
| 11024 | 224  | 
by simp  | 
225  | 
||
226  | 
text {* A few bigger examples *}
 | 
|
227  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
228  | 
lemma "(8452::int) mod 3 = 1"  | 
| 11024 | 229  | 
by simp  | 
230  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
231  | 
lemma "(59485::int) div 434 = 137"  | 
| 11024 | 232  | 
by simp  | 
233  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
234  | 
lemma "(1000006::int) mod 10 = 6"  | 
| 11024 | 235  | 
by simp  | 
236  | 
||
237  | 
||
238  | 
text {* \medskip Division by shifting *}
 | 
|
239  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
240  | 
lemma "10000000 div 2 = (5000000::int)"  | 
| 11024 | 241  | 
by simp  | 
242  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
243  | 
lemma "10000001 mod 2 = (1::int)"  | 
| 11024 | 244  | 
by simp  | 
245  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
246  | 
lemma "10000055 div 32 = (312501::int)"  | 
| 11024 | 247  | 
by simp  | 
248  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
249  | 
lemma "10000055 mod 32 = (23::int)"  | 
| 11024 | 250  | 
by simp  | 
251  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
252  | 
lemma "100094 div 144 = (695::int)"  | 
| 11024 | 253  | 
by simp  | 
254  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
255  | 
lemma "100094 mod 144 = (14::int)"  | 
| 11024 | 256  | 
by simp  | 
257  | 
||
258  | 
||
| 
12613
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
259  | 
text {* \medskip Powers *}
 | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
260  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
261  | 
lemma "2 ^ 10 = (1024::int)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
262  | 
by simp  | 
| 
 
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 "-3 ^ 7 = (-2187::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 "13 ^ 7 = (62748517::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 "3 ^ 15 = (14348907::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 "-5 ^ 11 = (-48828125::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  | 
|
| 11024 | 277  | 
subsection {* The Natural Numbers *}
 | 
278  | 
||
279  | 
text {* Successor *}
 | 
|
280  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
281  | 
lemma "Suc 99999 = 100000"  | 
| 11024 | 282  | 
by (simp add: Suc_nat_number_of)  | 
| 20807 | 283  | 
    -- {* not a default rewrite since sometimes we want to have @{text "Suc nnn"} *}
 | 
| 11024 | 284  | 
|
285  | 
||
286  | 
text {* \medskip Addition *}
 | 
|
287  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
288  | 
lemma "(13::nat) + 19 = 32"  | 
| 11024 | 289  | 
by simp  | 
290  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
291  | 
lemma "(1234::nat) + 5678 = 6912"  | 
| 11024 | 292  | 
by simp  | 
293  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
294  | 
lemma "(973646::nat) + 6475 = 980121"  | 
| 11024 | 295  | 
by simp  | 
296  | 
||
297  | 
||
298  | 
text {* \medskip Subtraction *}
 | 
|
299  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
300  | 
lemma "(32::nat) - 14 = 18"  | 
| 11024 | 301  | 
by simp  | 
302  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
303  | 
lemma "(14::nat) - 15 = 0"  | 
| 11024 | 304  | 
by simp  | 
| 5545 | 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) - 1576644 = 0"  | 
| 11024 | 307  | 
by simp  | 
308  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
309  | 
lemma "(48273776::nat) - 3873737 = 44400039"  | 
| 11024 | 310  | 
by simp  | 
311  | 
||
312  | 
||
313  | 
text {* \medskip Multiplication *}
 | 
|
314  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
315  | 
lemma "(12::nat) * 11 = 132"  | 
| 11024 | 316  | 
by simp  | 
317  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
318  | 
lemma "(647::nat) * 3643 = 2357021"  | 
| 11024 | 319  | 
by simp  | 
320  | 
||
321  | 
||
322  | 
text {* \medskip Quotient and Remainder *}
 | 
|
323  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
324  | 
lemma "(10::nat) div 3 = 3"  | 
| 11024 | 325  | 
by simp  | 
326  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
327  | 
lemma "(10::nat) mod 3 = 1"  | 
| 11024 | 328  | 
by simp  | 
329  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
330  | 
lemma "(10000::nat) div 9 = 1111"  | 
| 11024 | 331  | 
by simp  | 
332  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
333  | 
lemma "(10000::nat) mod 9 = 1"  | 
| 11024 | 334  | 
by simp  | 
335  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
336  | 
lemma "(10000::nat) div 16 = 625"  | 
| 11024 | 337  | 
by simp  | 
338  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
339  | 
lemma "(10000::nat) mod 16 = 0"  | 
| 11024 | 340  | 
by simp  | 
341  | 
||
| 5545 | 342  | 
|
| 
12613
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
343  | 
text {* \medskip Powers *}
 | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
344  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
345  | 
lemma "2 ^ 12 = (4096::nat)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
346  | 
by simp  | 
| 
 
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 "3 ^ 10 = (59049::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 "12 ^ 7 = (35831808::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 "3 ^ 14 = (4782969::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 "5 ^ 11 = (48828125::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  | 
|
| 11024 | 361  | 
text {* \medskip Testing the cancellation of complementary terms *}
 | 
362  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
363  | 
lemma "y + (x + -x) = (0::int) + y"  | 
| 11024 | 364  | 
by simp  | 
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 + (- y + x)) = (0::int)"  | 
| 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 "-x + (y + (- 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 + (x + (- x + (- x + (- y + - z)))) = (0::int) - y - z"  | 
| 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 + y + z - (x + z) = y - (0::int)"  | 
| 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 + (y + (y + (-x + -x)))) = (0::int) + y - x + y + y"  | 
| 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 + (-y + -x)))) = y + (0::int) + 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 - x + z - x - y - z + x < (1::int)"  | 
| 11024 | 388  | 
by simp  | 
389  | 
||
| 5545 | 390  | 
end  |