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