author | slotosch |
Sun, 25 May 1997 11:07:52 +0200 | |
changeset 3323 | 194ae2e0c193 |
parent 3310 | 0ceaad3c3f52 |
child 7661 | 8c3190b173aa |
permissions | -rw-r--r-- |
2640 | 1 |
(* Title: HOLCF/Porder0.thy |
298 | 2 |
ID: $Id$ |
1479 | 3 |
Author: Franz Regensburger |
298 | 4 |
Copyright 1993 Technische Universitaet Muenchen |
5 |
||
6 |
Definition of class porder (partial order) |
|
7 |
||
8 |
*) |
|
9 |
||
3310
0ceaad3c3f52
Base theory is now Arith, not Nat. (because all datatypes now require Arith).
nipkow
parents:
2850
diff
changeset
|
10 |
Porder0 = Arith + |
298 | 11 |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
12 |
(* introduce a (syntactic) class for the constant << *) |
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
13 |
axclass sq_ord<term |
298 | 14 |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
15 |
(* characteristic constant << for po *) |
2394 | 16 |
consts |
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
17 |
"<<" :: "['a,'a::sq_ord] => bool" (infixl 55) |
2394 | 18 |
|
19 |
syntax (symbols) |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
20 |
"op <<" :: "['a,'a::sq_ord] => bool" (infixl "\\<sqsubseteq>" 55) |
298 | 21 |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
22 |
axclass po < sq_ord |
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
23 |
(* class axioms: *) |
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
24 |
refl_less "x << x" |
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
25 |
antisym_less "[|x << y; y << x |] ==> x = y" |
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
26 |
trans_less "[|x << y; y << z |] ==> x << z" |
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3310
diff
changeset
|
27 |
|
298 | 28 |
end |
1274 | 29 |
|
30 |