author | wenzelm |
Thu, 01 Jul 1999 17:41:16 +0200 | |
changeset 6872 | b250da153b1e |
parent 6863 | 6c8bf18f9da9 |
child 6873 | b123f5522ea1 |
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 |
|
6863 | 10 |
theorems [trans] = Ord.order_antisym; (* <= <= = *) |
6862 | 11 |
theorems [trans] = Ord.order_trans; (* <= <= <= *) |
12 |
theorems [trans] = Ord.order_less_trans; (* < < < *) |
|
13 |
theorems [trans] = Ord.order_le_less_trans; (* <= < < *) |
|
14 |
theorems [trans] = Ord.order_less_le_trans; (* < <= < *) |
|
15 |
||
6863 | 16 |
theorem [trans]: "[| x <= y; y = z |] ==> x <= z"; (* <= = <= *) |
6862 | 17 |
by (rule HOL.subst[with y z]); |
6779
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
18 |
|
6863 | 19 |
theorem [trans]: "[| x = y; y <= z |] ==> x <= z"; (* = <= <= *) |
6862 | 20 |
by (rule HOL.ssubst[with x y]); |
21 |
||
6863 | 22 |
theorem [trans]: "[| x < y; y = z |] ==> x < z"; (* < = < *) |
6862 | 23 |
by (rule HOL.subst[with y z]); |
6779
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
diff
changeset
|
24 |
|
6863 | 25 |
theorem [trans]: "[| x = y; y < z |] ==> x < z"; (* = < < *) |
6862 | 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; |