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