author | wenzelm |
Thu, 20 Sep 2007 20:56:33 +0200 | |
changeset 24665 | e5bea50b9b89 |
parent 24328 | 83afe527504d |
child 30686 | 47a32dd1b86e |
permissions | -rw-r--r-- |
23193 | 1 |
(* Title: HOL/ex/Arith_Examples.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tjark Weber |
|
4 |
*) |
|
5 |
||
23218 | 6 |
header {* Arithmetic *} |
23193 | 7 |
|
8 |
theory Arith_Examples imports Main begin |
|
9 |
||
10 |
text {* |
|
23218 | 11 |
The @{text arith} method is used frequently throughout the Isabelle |
23193 | 12 |
distribution. This file merely contains some additional tests and special |
13 |
corner cases. Some rather technical remarks: |
|
14 |
||
23218 | 15 |
@{ML fast_arith_tac} is a very basic version of the tactic. It performs no |
23193 | 16 |
meta-to-object-logic conversion, and only some splitting of operators. |
23218 | 17 |
@{ML simple_arith_tac} performs meta-to-object-logic conversion, full |
18 |
splitting of operators, and NNF normalization of the goal. The @{text arith} |
|
19 |
method combines them both, and tries other methods (e.g.~@{text presburger}) |
|
23193 | 20 |
as well. This is the one that you should use in your proofs! |
21 |
||
24093 | 22 |
An @{text arith}-based simproc is available as well (see @{ML |
23 |
LinArith.lin_arith_simproc}), which---for performance |
|
24 |
reasons---however does even less splitting than @{ML fast_arith_tac} |
|
25 |
at the moment (namely inequalities only). (On the other hand, it |
|
26 |
does take apart conjunctions, which @{ML fast_arith_tac} currently |
|
27 |
does not do.) |
|
23193 | 28 |
*} |
29 |
||
23196 | 30 |
(* |
23193 | 31 |
ML {* set trace_arith; *} |
23196 | 32 |
*) |
23193 | 33 |
|
23218 | 34 |
subsection {* Splitting of Operators: @{term max}, @{term min}, @{term abs}, |
23193 | 35 |
@{term HOL.minus}, @{term nat}, @{term Divides.mod}, |
36 |
@{term Divides.div} *} |
|
37 |
||
38 |
lemma "(i::nat) <= max i j" |
|
24075 | 39 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 40 |
|
41 |
lemma "(i::int) <= max i j" |
|
24075 | 42 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 43 |
|
44 |
lemma "min i j <= (i::nat)" |
|
24075 | 45 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 46 |
|
47 |
lemma "min i j <= (i::int)" |
|
24075 | 48 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 49 |
|
50 |
lemma "min (i::nat) j <= max i j" |
|
24075 | 51 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 52 |
|
53 |
lemma "min (i::int) j <= max i j" |
|
24075 | 54 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 55 |
|
23208 | 56 |
lemma "min (i::nat) j + max i j = i + j" |
24075 | 57 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23208 | 58 |
|
59 |
lemma "min (i::int) j + max i j = i + j" |
|
24075 | 60 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23208 | 61 |
|
23193 | 62 |
lemma "(i::nat) < j ==> min i j < max i j" |
24075 | 63 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 64 |
|
65 |
lemma "(i::int) < j ==> min i j < max i j" |
|
24075 | 66 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 67 |
|
68 |
lemma "(0::int) <= abs i" |
|
24075 | 69 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 70 |
|
71 |
lemma "(i::int) <= abs i" |
|
24075 | 72 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 73 |
|
74 |
lemma "abs (abs (i::int)) = abs i" |
|
24075 | 75 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 76 |
|
77 |
text {* Also testing subgoals with bound variables. *} |
|
78 |
||
79 |
lemma "!!x. (x::nat) <= y ==> x - y = 0" |
|
24075 | 80 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 81 |
|
82 |
lemma "!!x. (x::nat) - y = 0 ==> x <= y" |
|
24075 | 83 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 84 |
|
85 |
lemma "!!x. ((x::nat) <= y) = (x - y = 0)" |
|
24075 | 86 |
by (tactic {* simple_arith_tac @{context} 1 *}) |
23193 | 87 |
|
88 |
lemma "[| (x::nat) < y; d < 1 |] ==> x - y = d" |
|
24075 | 89 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 90 |
|
91 |
lemma "[| (x::nat) < y; d < 1 |] ==> x - y - x = d - x" |
|
24075 | 92 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 93 |
|
94 |
lemma "(x::int) < y ==> x - y < 0" |
|
24075 | 95 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 96 |
|
97 |
lemma "nat (i + j) <= nat i + nat j" |
|
24075 | 98 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 99 |
|
100 |
lemma "i < j ==> nat (i - j) = 0" |
|
24075 | 101 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 102 |
|
103 |
lemma "(i::nat) mod 0 = i" |
|
23198 | 104 |
(* FIXME: need to replace 0 by its numeral representation *) |
105 |
apply (subst nat_numeral_0_eq_0 [symmetric]) |
|
24075 | 106 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23198 | 107 |
|
108 |
lemma "(i::nat) mod 1 = 0" |
|
109 |
(* FIXME: need to replace 1 by its numeral representation *) |
|
110 |
apply (subst nat_numeral_1_eq_1 [symmetric]) |
|
24075 | 111 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 112 |
|
23198 | 113 |
lemma "(i::nat) mod 42 <= 41" |
24075 | 114 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23198 | 115 |
|
116 |
lemma "(i::int) mod 0 = i" |
|
117 |
(* FIXME: need to replace 0 by its numeral representation *) |
|
118 |
apply (subst numeral_0_eq_0 [symmetric]) |
|
24075 | 119 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23198 | 120 |
|
121 |
lemma "(i::int) mod 1 = 0" |
|
122 |
(* FIXME: need to replace 1 by its numeral representation *) |
|
123 |
apply (subst numeral_1_eq_1 [symmetric]) |
|
124 |
(* FIXME: arith does not know about iszero *) |
|
24093 | 125 |
apply (tactic {* lin_arith_pre_tac @{context} 1 *}) |
23193 | 126 |
oops |
127 |
||
23198 | 128 |
lemma "(i::int) mod 42 <= 41" |
129 |
(* FIXME: arith does not know about iszero *) |
|
24093 | 130 |
apply (tactic {* lin_arith_pre_tac @{context} 1 *}) |
23193 | 131 |
oops |
132 |
||
24328
83afe527504d
fixed a bug in demult: -a in (-a * b) is no longer treated as atomic
webertj
parents:
24093
diff
changeset
|
133 |
lemma "-(i::int) * 1 = 0 ==> i = 0" |
83afe527504d
fixed a bug in demult: -a in (-a * b) is no longer treated as atomic
webertj
parents:
24093
diff
changeset
|
134 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
83afe527504d
fixed a bug in demult: -a in (-a * b) is no longer treated as atomic
webertj
parents:
24093
diff
changeset
|
135 |
|
83afe527504d
fixed a bug in demult: -a in (-a * b) is no longer treated as atomic
webertj
parents:
24093
diff
changeset
|
136 |
lemma "[| (0::int) < abs i; abs i * 1 < abs i * j |] ==> 1 < abs i * j" |
83afe527504d
fixed a bug in demult: -a in (-a * b) is no longer treated as atomic
webertj
parents:
24093
diff
changeset
|
137 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
83afe527504d
fixed a bug in demult: -a in (-a * b) is no longer treated as atomic
webertj
parents:
24093
diff
changeset
|
138 |
|
23218 | 139 |
|
140 |
subsection {* Meta-Logic *} |
|
23193 | 141 |
|
142 |
lemma "x < Suc y == x <= y" |
|
24075 | 143 |
by (tactic {* simple_arith_tac @{context} 1 *}) |
23193 | 144 |
|
145 |
lemma "((x::nat) == z ==> x ~= y) ==> x ~= y | z ~= y" |
|
24075 | 146 |
by (tactic {* simple_arith_tac @{context} 1 *}) |
23193 | 147 |
|
23218 | 148 |
|
149 |
subsection {* Various Other Examples *} |
|
23193 | 150 |
|
23198 | 151 |
lemma "(x < Suc y) = (x <= y)" |
24075 | 152 |
by (tactic {* simple_arith_tac @{context} 1 *}) |
23198 | 153 |
|
23193 | 154 |
lemma "[| (x::nat) < y; y < z |] ==> x < z" |
24075 | 155 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 156 |
|
157 |
lemma "(x::nat) < y & y < z ==> x < z" |
|
24075 | 158 |
by (tactic {* simple_arith_tac @{context} 1 *}) |
23193 | 159 |
|
23208 | 160 |
text {* This example involves no arithmetic at all, but is solved by |
161 |
preprocessing (i.e. NNF normalization) alone. *} |
|
162 |
||
163 |
lemma "(P::bool) = Q ==> Q = P" |
|
24075 | 164 |
by (tactic {* simple_arith_tac @{context} 1 *}) |
23208 | 165 |
|
166 |
lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> min (x::nat) y = 0" |
|
24075 | 167 |
by (tactic {* simple_arith_tac @{context} 1 *}) |
23208 | 168 |
|
169 |
lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> max (x::nat) y = x + y" |
|
24075 | 170 |
by (tactic {* simple_arith_tac @{context} 1 *}) |
23208 | 171 |
|
23193 | 172 |
lemma "[| (x::nat) ~= y; a + 2 = b; a < y; y < b; a < x; x < b |] ==> False" |
24075 | 173 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 174 |
|
175 |
lemma "[| (x::nat) > y; y > z; z > x |] ==> False" |
|
24075 | 176 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 177 |
|
178 |
lemma "(x::nat) - 5 > y ==> y < x" |
|
24075 | 179 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 180 |
|
181 |
lemma "(x::nat) ~= 0 ==> 0 < x" |
|
24075 | 182 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 183 |
|
184 |
lemma "[| (x::nat) ~= y; x <= y |] ==> x < y" |
|
24075 | 185 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 186 |
|
23196 | 187 |
lemma "[| (x::nat) < y; P (x - y) |] ==> P 0" |
24075 | 188 |
by (tactic {* simple_arith_tac @{context} 1 *}) |
23193 | 189 |
|
190 |
lemma "(x - y) - (x::nat) = (x - x) - y" |
|
24075 | 191 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 192 |
|
193 |
lemma "[| (a::nat) < b; c < d |] ==> (a - b) = (c - d)" |
|
24075 | 194 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 195 |
|
196 |
lemma "((a::nat) - (b - (c - (d - e)))) = (a - (b - (c - (d - e))))" |
|
24075 | 197 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 198 |
|
23198 | 199 |
lemma "(n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) | |
200 |
(n = n' & n' < m) | (n = m & m < n') | |
|
201 |
(n' < m & m < n) | (n' < m & m = n) | |
|
202 |
(n' < n & n < m) | (n' = n & n < m) | (n' = m & m < n) | |
|
203 |
(m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) | |
|
204 |
(m = n & n < n') | (m = n' & n' < n) | |
|
205 |
(n' = m & m = (n::nat))" |
|
206 |
(* FIXME: this should work in principle, but is extremely slow because *) |
|
207 |
(* preprocessing negates the goal and tries to compute its negation *) |
|
208 |
(* normal form, which creates lots of separate cases for this *) |
|
209 |
(* disjunction of conjunctions *) |
|
210 |
(* by (tactic {* simple_arith_tac 1 *}) *) |
|
211 |
oops |
|
212 |
||
213 |
lemma "2 * (x::nat) ~= 1" |
|
23208 | 214 |
(* FIXME: this is beyond the scope of the decision procedure at the moment, *) |
215 |
(* because its negation is satisfiable in the rationals? *) |
|
23198 | 216 |
(* by (tactic {* fast_arith_tac 1 *}) *) |
217 |
oops |
|
218 |
||
219 |
text {* Constants. *} |
|
220 |
||
221 |
lemma "(0::nat) < 1" |
|
24075 | 222 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23198 | 223 |
|
224 |
lemma "(0::int) < 1" |
|
24075 | 225 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23198 | 226 |
|
227 |
lemma "(47::nat) + 11 < 08 * 15" |
|
24075 | 228 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23198 | 229 |
|
230 |
lemma "(47::int) + 11 < 08 * 15" |
|
24075 | 231 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23198 | 232 |
|
23193 | 233 |
text {* Splitting of inequalities of different type. *} |
234 |
||
235 |
lemma "[| (a::nat) ~= b; (i::int) ~= j; a < 2; b < 2 |] ==> |
|
236 |
a + b <= nat (max (abs i) (abs j))" |
|
24075 | 237 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 238 |
|
23198 | 239 |
text {* Again, but different order. *} |
240 |
||
23193 | 241 |
lemma "[| (i::int) ~= j; (a::nat) ~= b; a < 2; b < 2 |] ==> |
242 |
a + b <= nat (max (abs i) (abs j))" |
|
24075 | 243 |
by (tactic {* fast_arith_tac @{context} 1 *}) |
23193 | 244 |
|
23196 | 245 |
(* |
23193 | 246 |
ML {* reset trace_arith; *} |
23196 | 247 |
*) |
23193 | 248 |
|
249 |
end |