| author | paulson | 
| Tue, 06 Jan 2004 10:50:36 +0100 | |
| changeset 14342 | 6e564092d72d | 
| parent 14124 | 883c38e2d4c0 | 
| child 14378 | 69c4d5997669 | 
| 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  | 
||
9  | 
theory BinEx = Main:  | 
|
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  | 
||
189  | 
||
190  | 
text {* \medskip Quotient and Remainder *}
 | 
|
191  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
192  | 
lemma "(10::int) div 3 = 3"  | 
| 11024 | 193  | 
by simp  | 
194  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
195  | 
lemma "(10::int) mod 3 = 1"  | 
| 11024 | 196  | 
by simp  | 
197  | 
||
198  | 
text {* A negative divisor *}
 | 
|
199  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
200  | 
lemma "(10::int) div -3 = -4"  | 
| 11024 | 201  | 
by simp  | 
202  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
203  | 
lemma "(10::int) mod -3 = -2"  | 
| 11024 | 204  | 
by simp  | 
| 5545 | 205  | 
|
| 11024 | 206  | 
text {*
 | 
207  | 
  A negative dividend\footnote{The definition agrees with mathematical
 | 
|
208  | 
convention but not with the hardware of most computers}  | 
|
209  | 
*}  | 
|
210  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
211  | 
lemma "(-10::int) div 3 = -4"  | 
| 11024 | 212  | 
by simp  | 
213  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
214  | 
lemma "(-10::int) mod 3 = 2"  | 
| 11024 | 215  | 
by simp  | 
216  | 
||
217  | 
text {* A negative dividend \emph{and} divisor *}
 | 
|
218  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
219  | 
lemma "(-10::int) div -3 = 3"  | 
| 11024 | 220  | 
by simp  | 
221  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
222  | 
lemma "(-10::int) mod -3 = -1"  | 
| 11024 | 223  | 
by simp  | 
224  | 
||
225  | 
text {* A few bigger examples *}
 | 
|
226  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
227  | 
lemma "(8452::int) mod 3 = 1"  | 
| 11024 | 228  | 
by simp  | 
229  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
230  | 
lemma "(59485::int) div 434 = 137"  | 
| 11024 | 231  | 
by simp  | 
232  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
233  | 
lemma "(1000006::int) mod 10 = 6"  | 
| 11024 | 234  | 
by simp  | 
235  | 
||
236  | 
||
237  | 
text {* \medskip Division by shifting *}
 | 
|
238  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
239  | 
lemma "10000000 div 2 = (5000000::int)"  | 
| 11024 | 240  | 
by simp  | 
241  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
242  | 
lemma "10000001 mod 2 = (1::int)"  | 
| 11024 | 243  | 
by simp  | 
244  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
245  | 
lemma "10000055 div 32 = (312501::int)"  | 
| 11024 | 246  | 
by simp  | 
247  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
248  | 
lemma "10000055 mod 32 = (23::int)"  | 
| 11024 | 249  | 
by simp  | 
250  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
251  | 
lemma "100094 div 144 = (695::int)"  | 
| 11024 | 252  | 
by simp  | 
253  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
254  | 
lemma "100094 mod 144 = (14::int)"  | 
| 11024 | 255  | 
by simp  | 
256  | 
||
257  | 
||
| 
12613
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
258  | 
text {* \medskip Powers *}
 | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
259  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
260  | 
lemma "2 ^ 10 = (1024::int)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
261  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
262  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
263  | 
lemma "-3 ^ 7 = (-2187::int)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
264  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
265  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
266  | 
lemma "13 ^ 7 = (62748517::int)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
267  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
268  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
269  | 
lemma "3 ^ 15 = (14348907::int)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
270  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
271  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
272  | 
lemma "-5 ^ 11 = (-48828125::int)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
273  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
274  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
275  | 
|
| 11024 | 276  | 
subsection {* The Natural Numbers *}
 | 
277  | 
||
278  | 
text {* Successor *}
 | 
|
279  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
280  | 
lemma "Suc 99999 = 100000"  | 
| 11024 | 281  | 
by (simp add: Suc_nat_number_of)  | 
282  | 
    -- {* not a default rewrite since sometimes we want to have @{text "Suc #nnn"} *}
 | 
|
283  | 
||
284  | 
||
285  | 
text {* \medskip Addition *}
 | 
|
286  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
287  | 
lemma "(13::nat) + 19 = 32"  | 
| 11024 | 288  | 
by simp  | 
289  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
290  | 
lemma "(1234::nat) + 5678 = 6912"  | 
| 11024 | 291  | 
by simp  | 
292  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
293  | 
lemma "(973646::nat) + 6475 = 980121"  | 
| 11024 | 294  | 
by simp  | 
295  | 
||
296  | 
||
297  | 
text {* \medskip Subtraction *}
 | 
|
298  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
299  | 
lemma "(32::nat) - 14 = 18"  | 
| 11024 | 300  | 
by simp  | 
301  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
302  | 
lemma "(14::nat) - 15 = 0"  | 
| 11024 | 303  | 
by simp  | 
| 5545 | 304  | 
|
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
305  | 
lemma "(14::nat) - 1576644 = 0"  | 
| 11024 | 306  | 
by simp  | 
307  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
308  | 
lemma "(48273776::nat) - 3873737 = 44400039"  | 
| 11024 | 309  | 
by simp  | 
310  | 
||
311  | 
||
312  | 
text {* \medskip Multiplication *}
 | 
|
313  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
314  | 
lemma "(12::nat) * 11 = 132"  | 
| 11024 | 315  | 
by simp  | 
316  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
317  | 
lemma "(647::nat) * 3643 = 2357021"  | 
| 11024 | 318  | 
by simp  | 
319  | 
||
320  | 
||
321  | 
text {* \medskip Quotient and Remainder *}
 | 
|
322  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
323  | 
lemma "(10::nat) div 3 = 3"  | 
| 11024 | 324  | 
by simp  | 
325  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
326  | 
lemma "(10::nat) mod 3 = 1"  | 
| 11024 | 327  | 
by simp  | 
328  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
329  | 
lemma "(10000::nat) div 9 = 1111"  | 
| 11024 | 330  | 
by simp  | 
331  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
332  | 
lemma "(10000::nat) mod 9 = 1"  | 
| 11024 | 333  | 
by simp  | 
334  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
335  | 
lemma "(10000::nat) div 16 = 625"  | 
| 11024 | 336  | 
by simp  | 
337  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
338  | 
lemma "(10000::nat) mod 16 = 0"  | 
| 11024 | 339  | 
by simp  | 
340  | 
||
| 5545 | 341  | 
|
| 
12613
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
342  | 
text {* \medskip Powers *}
 | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
343  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
344  | 
lemma "2 ^ 12 = (4096::nat)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
345  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
346  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
347  | 
lemma "3 ^ 10 = (59049::nat)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
348  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
349  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
350  | 
lemma "12 ^ 7 = (35831808::nat)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
351  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
352  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
353  | 
lemma "3 ^ 14 = (4782969::nat)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
354  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
355  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
356  | 
lemma "5 ^ 11 = (48828125::nat)"  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
357  | 
by simp  | 
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
358  | 
|
| 
 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 
paulson 
parents: 
11868 
diff
changeset
 | 
359  | 
|
| 11024 | 360  | 
text {* \medskip Testing the cancellation of complementary terms *}
 | 
361  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
362  | 
lemma "y + (x + -x) = (0::int) + y"  | 
| 11024 | 363  | 
by simp  | 
364  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
365  | 
lemma "y + (-x + (- y + x)) = (0::int)"  | 
| 11024 | 366  | 
by simp  | 
367  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
368  | 
lemma "-x + (y + (- y + x)) = (0::int)"  | 
| 11024 | 369  | 
by simp  | 
370  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
371  | 
lemma "x + (x + (- x + (- x + (- y + - z)))) = (0::int) - y - z"  | 
| 11024 | 372  | 
by simp  | 
373  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
374  | 
lemma "x + x - x - x - y - z = (0::int) - y - z"  | 
| 11024 | 375  | 
by simp  | 
376  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
377  | 
lemma "x + y + z - (x + z) = y - (0::int)"  | 
| 11024 | 378  | 
by simp  | 
379  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
380  | 
lemma "x + (y + (y + (y + (-x + -x)))) = (0::int) + y - x + y + y"  | 
| 11024 | 381  | 
by simp  | 
382  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
383  | 
lemma "x + (y + (y + (y + (-y + -x)))) = y + (0::int) + y"  | 
| 11024 | 384  | 
by simp  | 
385  | 
||
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
386  | 
lemma "x + y - x + z - x - y - z + x < (1::int)"  | 
| 11024 | 387  | 
by simp  | 
388  | 
||
389  | 
||
390  | 
subsection {* Normal form of bit strings *}
 | 
|
391  | 
||
392  | 
text {*
 | 
|
393  | 
Definition of normal form for proving that binary arithmetic on  | 
|
394  | 
normalized operands yields normalized results. Normal means no  | 
|
395  | 
leading 0s on positive numbers and no leading 1s on negatives.  | 
|
396  | 
*}  | 
|
| 5545 | 397  | 
|
| 11024 | 398  | 
consts normal :: "bin set"  | 
| 11637 | 399  | 
inductive normal  | 
400  | 
intros  | 
|
| 13491 | 401  | 
Pls [simp]: "bin.Pls: normal"  | 
402  | 
Min [simp]: "bin.Min: normal"  | 
|
403  | 
BIT_F [simp]: "w: normal ==> w \<noteq> bin.Pls ==> w BIT False : normal"  | 
|
404  | 
BIT_T [simp]: "w: normal ==> w \<noteq> bin.Min ==> w BIT True : normal"  | 
|
| 11024 | 405  | 
|
406  | 
text {*
 | 
|
407  | 
\medskip Binary arithmetic on normalized operands yields normalized  | 
|
408  | 
results.  | 
|
409  | 
*}  | 
|
410  | 
||
411  | 
lemma normal_BIT_I [simp]: "w BIT b \<in> normal ==> w BIT b BIT c \<in> normal"  | 
|
412  | 
apply (case_tac c)  | 
|
413  | 
apply auto  | 
|
414  | 
done  | 
|
415  | 
||
416  | 
lemma normal_BIT_D: "w BIT b \<in> normal ==> w \<in> normal"  | 
|
417  | 
apply (erule normal.cases)  | 
|
418  | 
apply auto  | 
|
419  | 
done  | 
|
420  | 
||
421  | 
lemma NCons_normal [simp]: "w \<in> normal ==> NCons w b \<in> normal"  | 
|
422  | 
apply (induct w)  | 
|
423  | 
apply (auto simp add: NCons_Pls NCons_Min)  | 
|
424  | 
done  | 
|
425  | 
||
| 13491 | 426  | 
lemma NCons_True: "NCons w True \<noteq> bin.Pls"  | 
| 11024 | 427  | 
apply (induct w)  | 
428  | 
apply auto  | 
|
429  | 
done  | 
|
430  | 
||
| 13491 | 431  | 
lemma NCons_False: "NCons w False \<noteq> bin.Min"  | 
| 11024 | 432  | 
apply (induct w)  | 
433  | 
apply auto  | 
|
434  | 
done  | 
|
| 5545 | 435  | 
|
| 11024 | 436  | 
lemma bin_succ_normal [simp]: "w \<in> normal ==> bin_succ w \<in> normal"  | 
437  | 
apply (erule normal.induct)  | 
|
438  | 
apply (case_tac [4] w)  | 
|
439  | 
apply (auto simp add: NCons_True bin_succ_BIT)  | 
|
440  | 
done  | 
|
441  | 
||
442  | 
lemma bin_pred_normal [simp]: "w \<in> normal ==> bin_pred w \<in> normal"  | 
|
443  | 
apply (erule normal.induct)  | 
|
444  | 
apply (case_tac [3] w)  | 
|
445  | 
apply (auto simp add: NCons_False bin_pred_BIT)  | 
|
446  | 
done  | 
|
447  | 
||
448  | 
lemma bin_add_normal [rule_format]:  | 
|
449  | 
"w \<in> normal --> (\<forall>z. z \<in> normal --> bin_add w z \<in> normal)"  | 
|
450  | 
apply (induct w)  | 
|
451  | 
apply simp  | 
|
452  | 
apply simp  | 
|
453  | 
apply (rule impI)  | 
|
454  | 
apply (rule allI)  | 
|
455  | 
apply (induct_tac z)  | 
|
456  | 
apply (simp_all add: bin_add_BIT)  | 
|
457  | 
apply (safe dest!: normal_BIT_D)  | 
|
458  | 
apply simp_all  | 
|
459  | 
done  | 
|
460  | 
||
| 13491 | 461  | 
lemma normal_Pls_eq_0: "w \<in> normal ==> (w = bin.Pls) = (number_of w = (0::int))"  | 
| 11024 | 462  | 
apply (erule normal.induct)  | 
463  | 
apply auto  | 
|
464  | 
done  | 
|
| 13192 | 465  | 
|
| 11024 | 466  | 
lemma bin_minus_normal: "w \<in> normal ==> bin_minus w \<in> normal"  | 
467  | 
apply (erule normal.induct)  | 
|
468  | 
apply (simp_all add: bin_minus_BIT)  | 
|
469  | 
apply (rule normal.intros)  | 
|
| 13187 | 470  | 
apply assumption  | 
| 11024 | 471  | 
apply (simp add: normal_Pls_eq_0)  | 
| 14113 | 472  | 
apply (simp only: number_of_minus zminus_0 iszero_def  | 
| 13192 | 473  | 
zminus_equation [of _ "0"])  | 
474  | 
apply (simp add: eq_commute)  | 
|
475  | 
done  | 
|
| 13187 | 476  | 
|
| 13192 | 477  | 
(*The previous command should have finished the proof but the lin-arith  | 
| 13187 | 478  | 
procedure at the end of simplificatioon fails.  | 
479  | 
Problem: lin-arith correctly derives the contradictory thm  | 
|
480  | 
"number_of w + 1 + - 0 \<le> 0 + number_of w" [..]  | 
|
481  | 
which its local simplification setup should turn into False.  | 
|
482  | 
But on the way we get  | 
|
483  | 
||
484  | 
Procedure "int_add_eval_numerals" produced rewrite rule:  | 
|
| 13491 | 485  | 
number_of ?v3 + 1 \<equiv> number_of (bin_add ?v3 (bin.Pls BIT True))  | 
| 13187 | 486  | 
|
487  | 
and eventually we arrive not at false but at  | 
|
488  | 
||
| 13491 | 489  | 
"\<not> neg (number_of (bin_add w (bin_minus (bin_add w (bin.Pls BIT True)))))"  | 
| 13187 | 490  | 
|
| 13192 | 491  | 
The simplification with eq_commute should now be obsolete.  | 
492  | 
*)  | 
|
| 11024 | 493  | 
|
494  | 
lemma bin_mult_normal [rule_format]:  | 
|
495  | 
"w \<in> normal ==> z \<in> normal --> bin_mult w z \<in> normal"  | 
|
496  | 
apply (erule normal.induct)  | 
|
497  | 
apply (simp_all add: bin_minus_normal bin_mult_BIT)  | 
|
498  | 
apply (safe dest!: normal_BIT_D)  | 
|
499  | 
apply (simp add: bin_add_normal)  | 
|
500  | 
done  | 
|
| 13192 | 501  | 
|
| 5545 | 502  | 
end  |