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