author | nipkow |
Thu, 10 May 2001 17:28:40 +0200 | |
changeset 11295 | 66925f23ac7f |
parent 11181 | d04f57b91166 |
child 11324 | 82406bd816a5 |
permissions | -rw-r--r-- |
10214 | 1 |
(* Title: HOL/NatArith.thy |
2 |
ID: $Id$ |
|
3 |
||
4 |
Setup arithmetic proof procedures. |
|
5 |
*) |
|
6 |
||
7 |
theory NatArith = Nat |
|
8 |
files "arith_data.ML": |
|
9 |
||
10 |
setup arith_setup |
|
11 |
||
12 |
(*elimination of `-' on nat*) |
|
13 |
lemma nat_diff_split: |
|
10599 | 14 |
"P(a - b::nat) = ((a<b --> P 0) & (ALL d. a = b + d --> P d))" |
10214 | 15 |
by (cases "a < b" rule: case_split) (auto simp add: diff_is_0_eq [THEN iffD2]) |
16 |
||
11164 | 17 |
ML {* |
18 |
val nat_diff_split = thm "nat_diff_split"; |
|
19 |
||
20 |
(* TODO: use this for force_tac in Provers/clasip.ML *) |
|
11181
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
oheimb
parents:
11164
diff
changeset
|
21 |
fun add_arith cs = cs addafter ("arith_tac", arith_tac); |
11164 | 22 |
*} |
10214 | 23 |
|
24 |
lemmas [arith_split] = nat_diff_split split_min split_max |
|
25 |
||
11164 | 26 |
|
10214 | 27 |
end |