author | wenzelm |
Tue, 27 Apr 1999 10:44:42 +0200 | |
changeset 6515 | 18e113be12ee |
parent 6402 | 2b23e14dd386 |
child 6855 | 36de02d1a257 |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/Ord.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
2624 | 6 |
Type classes for order signatures and orders. |
923 | 7 |
*) |
8 |
||
9 |
Ord = HOL + |
|
10 |
||
5889 | 11 |
(*FIXME move to HOL.thy*) |
6515 | 12 |
setup hol_setup |
5889 | 13 |
setup Classical.setup |
5953 | 14 |
setup Blast.setup |
15 |
setup Clasimp.setup |
|
6515 | 16 |
setup simpdata_setup |
5953 | 17 |
|
5889 | 18 |
|
923 | 19 |
axclass |
20 |
ord < term |
|
21 |
||
3947 | 22 |
global |
23 |
||
3820 | 24 |
syntax |
25 |
"op <" :: ['a::ord, 'a] => bool ("op <") |
|
26 |
"op <=" :: ['a::ord, 'a] => bool ("op <=") |
|
27 |
||
923 | 28 |
consts |
2259 | 29 |
"op <" :: ['a::ord, 'a] => bool ("(_/ < _)" [50, 51] 50) |
30 |
"op <=" :: ['a::ord, 'a] => bool ("(_/ <= _)" [50, 51] 50) |
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
965
diff
changeset
|
31 |
mono :: ['a::ord => 'b::ord] => bool (*monotonicity*) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
965
diff
changeset
|
32 |
min, max :: ['a::ord, 'a] => 'a |
923 | 33 |
|
2608 | 34 |
Least :: ('a::ord=>bool) => 'a (binder "LEAST " 10) |
35 |
||
3820 | 36 |
syntax (symbols) |
37 |
"op <=" :: ['a::ord, 'a] => bool ("op \\<le>") |
|
38 |
"op <=" :: ['a::ord, 'a] => bool ("(_/ \\<le> _)" [50, 51] 50) |
|
2259 | 39 |
|
40 |
||
3947 | 41 |
local |
42 |
||
923 | 43 |
defs |
44 |
mono_def "mono(f) == (!A B. A <= B --> f(A) <= f(B))" |
|
965 | 45 |
min_def "min a b == (if a <= b then a else b)" |
46 |
max_def "max a b == (if a <= b then b else a)" |
|
3143
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
2624
diff
changeset
|
47 |
Least_def "Least P == @x. P(x) & (ALL y. P(y) --> x <= y)" |
2608 | 48 |
|
49 |
||
50 |
axclass order < ord |
|
51 |
order_refl "x <= x" |
|
52 |
order_trans "[| x <= y; y <= z |] ==> x <= z" |
|
53 |
order_antisym "[| x <= y; y <= x |] ==> x = y" |
|
54 |
order_less_le "x < y = (x <= y & x ~= y)" |
|
923 | 55 |
|
4640 | 56 |
axclass linorder < order |
57 |
linorder_linear "x <= y | y <= x" |
|
58 |
||
6402 | 59 |
(*bounded quantifiers*) |
60 |
syntax |
|
61 |
"@lessAll" :: [idt, 'a, bool] => bool ("(3! _<_./ _)" [0, 0, 10] 10) |
|
62 |
"@lessEx" :: [idt, 'a, bool] => bool ("(3? _<_./ _)" [0, 0, 10] 10) |
|
63 |
"@leAll" :: [idt, 'a, bool] => bool ("(3! _<=_./ _)" [0, 0, 10] 10) |
|
64 |
"@leEx" :: [idt, 'a, bool] => bool ("(3? _<=_./ _)" [0, 0, 10] 10) |
|
65 |
translations |
|
66 |
"! x<y. P" => "! x. x < y --> P" |
|
67 |
"! x<=y. P" => "! x. x <= y --> P" |
|
68 |
"? x<y. P" => "? x. x < y & P" |
|
69 |
"? x<=y. P" => "! x. x <= y & P" |
|
70 |
||
923 | 71 |
end |