| author | streckem | 
| Fri, 08 Aug 2003 14:54:37 +0200 | |
| changeset 14141 | d3916d9183d2 | 
| parent 12338 | de0f4a63baa5 | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
| 2640 | 1 | (* Title: HOLCF/Porder0.thy | 
| 298 | 2 | ID: $Id$ | 
| 1479 | 3 | Author: Franz Regensburger | 
| 12030 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 298 | 5 | |
| 12030 | 6 | Definition of class porder (partial order). | 
| 298 | 7 | *) | 
| 8 | ||
| 7661 | 9 | Porder0 = Main + | 
| 298 | 10 | |
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 11 | (* introduce a (syntactic) class for the constant << *) | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12114diff
changeset | 12 | axclass sq_ord < type | 
| 298 | 13 | |
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 14 | (* characteristic constant << for po *) | 
| 2394 | 15 | consts | 
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 16 | "<<" :: "['a,'a::sq_ord] => bool" (infixl 55) | 
| 2394 | 17 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12030diff
changeset | 18 | syntax (xsymbols) | 
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 19 | "op <<" :: "['a,'a::sq_ord] => bool" (infixl "\\<sqsubseteq>" 55) | 
| 298 | 20 | |
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 21 | axclass po < sq_ord | 
| 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 22 | (* class axioms: *) | 
| 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 23 | refl_less "x << x" | 
| 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 24 | antisym_less "[|x << y; y << x |] ==> x = y" | 
| 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 25 | trans_less "[|x << y; y << z |] ==> x << z" | 
| 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3310diff
changeset | 26 | |
| 298 | 27 | end | 
| 1274 | 28 | |
| 29 |