| author | wenzelm |
| Wed, 30 Jun 1999 13:42:47 +0200 | |
| changeset 6862 | f80091bdc992 |
| parent 6791 | 2be411437c60 |
| child 6863 | 6c8bf18f9da9 |
| permissions | -rw-r--r-- |
|
6779
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
1 |
(* Title: HOL/Calculation.thy |
|
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
|
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
4 |
|
|
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
5 |
Setup transitivity rules for calculational proofs. |
|
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
6 |
*) |
|
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
7 |
|
| 6862 | 8 |
theory Calculation = Int:; |
|
6779
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
9 |
|
| 6862 | 10 |
theorems [trans] = Ord.order_trans; (* <= <= <= *) |
11 |
theorems [trans] = Ord.order_less_trans; (* < < < *) |
|
12 |
theorems [trans] = Ord.order_le_less_trans; (* <= < < *) |
|
13 |
theorems [trans] = Ord.order_less_le_trans; (* < <= < *) |
|
14 |
theorems [trans] = Ord.order_antisym; (* <= <= = *) |
|
15 |
||
16 |
theorem [trans]: "[| x <= y; y = z |] ==> x <= z"; |
|
17 |
by (rule HOL.subst[with y z]); |
|
|
6779
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
18 |
|
| 6862 | 19 |
theorem [trans]: "[| x = y; y <= z |] ==> x <= z"; |
20 |
by (rule HOL.ssubst[with x y]); |
|
21 |
||
22 |
theorem [trans]: "[| x < y; y = z |] ==> x < z"; |
|
23 |
by (rule HOL.subst[with y z]); |
|
|
6779
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
24 |
|
| 6862 | 25 |
theorem [trans]: "[| x = y; y < z |] ==> x < z"; |
26 |
by (rule HOL.ssubst[with x y]); |
|
27 |
||
28 |
theorems [trans] = HOL.subst[COMP swap_prems_rl]; (* x = x *) |
|
29 |
theorems [trans] = HOL.ssubst; (* = x x *) |
|
30 |
||
31 |
theorems [trans] = Divides.dvd_trans; (* dvd dvd dvd *) |
|
|
6779
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
32 |
|
|
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
33 |
|
| 6862 | 34 |
end; |