| author | paulson |
| Wed, 28 Jul 2004 10:49:29 +0200 | |
| changeset 15079 | 2ef899e4526d |
| parent 15048 | 11b4dce71d73 |
| child 15131 | c69542757a4d |
| permissions | -rw-r--r-- |
| 10214 | 1 |
(* Title: HOL/NatArith.thy |
2 |
ID: $Id$ |
|
| 13297 | 3 |
Author: Tobias Nipkow and Markus Wenzel |
4 |
*) |
|
| 10214 | 5 |
|
| 13297 | 6 |
header {* More arithmetic on natural numbers *}
|
| 10214 | 7 |
|
8 |
theory NatArith = Nat |
|
9 |
files "arith_data.ML": |
|
10 |
||
11 |
setup arith_setup |
|
12 |
||
| 13297 | 13 |
|
| 11655 | 14 |
lemma pred_nat_trancl_eq_le: "((m, n) : pred_nat^*) = (m <= n)" |
| 14208 | 15 |
by (simp add: less_eq reflcl_trancl [symmetric] |
16 |
del: reflcl_trancl, arith) |
|
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11324
diff
changeset
|
17 |
|
| 10214 | 18 |
lemma nat_diff_split: |
| 10599 | 19 |
"P(a - b::nat) = ((a<b --> P 0) & (ALL d. a = b + d --> P d))" |
| 13297 | 20 |
-- {* elimination of @{text -} on @{text nat} *}
|
21 |
by (cases "a<b" rule: case_split) |
|
22 |
(auto simp add: diff_is_0_eq [THEN iffD2]) |
|
| 11324 | 23 |
|
24 |
lemma nat_diff_split_asm: |
|
25 |
"P(a - b::nat) = (~ (a < b & ~ P 0 | (EX d. a = b + d & ~ P d)))" |
|
| 13297 | 26 |
-- {* elimination of @{text -} on @{text nat} in assumptions *}
|
| 11324 | 27 |
by (simp split: nat_diff_split) |
| 10214 | 28 |
|
| 11164 | 29 |
ML {*
|
30 |
val nat_diff_split = thm "nat_diff_split"; |
|
| 11324 | 31 |
val nat_diff_split_asm = thm "nat_diff_split_asm"; |
| 13499 | 32 |
*} |
33 |
(* Careful: arith_tac produces counter examples! |
|
|
11181
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
oheimb
parents:
11164
diff
changeset
|
34 |
fun add_arith cs = cs addafter ("arith_tac", arith_tac);
|
| 14607 | 35 |
TODO: use arith_tac for force_tac in Provers/clasimp.ML *) |
| 10214 | 36 |
|
37 |
lemmas [arith_split] = nat_diff_split split_min split_max |
|
38 |
||
39 |
end |