author | paulson |
Thu, 25 Sep 1997 13:23:41 +0200 | |
changeset 3715 | 6e074b41c735 |
parent 3460 | 5d71eed16fbe |
child 3842 | b55686a7b22c |
permissions | -rw-r--r-- |
2644 | 1 |
(* Title: HOLCF/Porder0.ML |
2 |
ID: $Id$ |
|
3 |
Author: Oscar Slotosch |
|
4 |
Copyright 1997 Technische Universitaet Muenchen |
|
5 |
||
6 |
derive the characteristic axioms for the characteristic constants *) |
|
7 |
||
8 |
open Porder0; |
|
9 |
||
2841
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
2644
diff
changeset
|
10 |
AddIffs [refl_less]; |
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
2644
diff
changeset
|
11 |
|
2644 | 12 |
(* ------------------------------------------------------------------------ *) |
13 |
(* minimal fixes least element *) |
|
14 |
(* ------------------------------------------------------------------------ *) |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3026
diff
changeset
|
15 |
bind_thm("minimal2UU",allI RS (prove_goal thy "!x::'a::po.uu<<x==>uu=(@u.!y.u<<y)" |
2644 | 16 |
(fn prems => |
17 |
[ |
|
18 |
(cut_facts_tac prems 1), |
|
19 |
(rtac antisym_less 1), |
|
20 |
(etac spec 1), |
|
21 |
(res_inst_tac [("a","uu")] selectI2 1), |
|
22 |
(atac 1), |
|
23 |
(etac spec 1) |
|
24 |
]))); |
|
3026
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
25 |
|
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
26 |
(* ------------------------------------------------------------------------ *) |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
27 |
(* 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
|
28 |
(* ------------------------------------------------------------------------ *) |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
29 |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3026
diff
changeset
|
30 |
qed_goal "antisym_less_inverse" thy "(x::'a::po)=y ==> x << y & y << x" |
3026
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
31 |
(fn prems => |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
32 |
[ |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
33 |
(cut_facts_tac prems 1), |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
34 |
(rtac conjI 1), |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
35 |
((rtac subst 1) THEN (rtac refl_less 2) THEN (atac 1)), |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
36 |
((rtac subst 1) THEN (rtac refl_less 2) THEN (etac sym 1)) |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
37 |
]); |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
38 |
|
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
39 |
|
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
40 |
qed_goal "box_less" thy |
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3026
diff
changeset
|
41 |
"[| (a::'a::po) << b; c << a; b << d|] ==> c << d" |
3026
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
42 |
(fn prems => |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
43 |
[ |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
44 |
(cut_facts_tac prems 1), |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
45 |
(etac trans_less 1), |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
46 |
(etac trans_less 1), |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
47 |
(atac 1) |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
48 |
]); |
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
slotosch
parents:
2841
diff
changeset
|
49 |
|
3460 | 50 |
goal Porder0.thy "((x::'a::po)=y) = (x << y & y << x)"; |
51 |
by(fast_tac (HOL_cs addSEs [antisym_less_inverse] addSIs [antisym_less]) 1); |
|
52 |
qed "po_eq_conv"; |