author | nipkow |
Fri, 10 Sep 2004 00:19:15 +0200 | |
changeset 15195 | 197e00ce3f20 |
parent 14981 | e73f8140af78 |
child 15562 | 8455c9671494 |
permissions | -rw-r--r-- |
2644 | 1 |
(* Title: HOLCF/Porder0.ML |
2 |
ID: $Id$ |
|
3 |
Author: Oscar Slotosch |
|
4 |
||
12030 | 5 |
derive the characteristic axioms for the characteristic constants |
9245 | 6 |
*) |
2644 | 7 |
|
2841
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
2644
diff
changeset
|
8 |
AddIffs [refl_less]; |
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
2644
diff
changeset
|
9 |
|
2644 | 10 |
(* ------------------------------------------------------------------------ *) |
11 |
(* minimal fixes least element *) |
|
12 |
(* ------------------------------------------------------------------------ *) |
|
9245 | 13 |
Goal "!x::'a::po. uu<<x ==> uu=(@u.!y. u<<y)"; |
9969 | 14 |
by (blast_tac (claset() addIs [someI2,antisym_less]) 1); |
9245 | 15 |
bind_thm ("minimal2UU", allI RS result()); |
3026
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
16 |
|
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
17 |
(* ------------------------------------------------------------------------ *) |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
18 |
(* the reverse law of anti--symmetrie of << *) |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
19 |
(* ------------------------------------------------------------------------ *) |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
20 |
|
9169 | 21 |
Goal "(x::'a::po)=y ==> x << y & y << x"; |
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
22 |
by (Blast_tac 1); |
9169 | 23 |
qed "antisym_less_inverse"; |
3026
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
24 |
|
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
25 |
|
9169 | 26 |
Goal "[| (a::'a::po) << b; c << a; b << d|] ==> c << d"; |
27 |
by (etac trans_less 1); |
|
28 |
by (etac trans_less 1); |
|
29 |
by (atac 1); |
|
30 |
qed "box_less"; |
|
3026
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
31 |
|
5068 | 32 |
Goal "((x::'a::po)=y) = (x << y & y << x)"; |
4423 | 33 |
by (fast_tac (HOL_cs addSEs [antisym_less_inverse] addSIs [antisym_less]) 1); |
3460 | 34 |
qed "po_eq_conv"; |