| author | wenzelm | 
| Wed, 26 May 1999 16:51:05 +0200 | |
| changeset 6732 | cf9f66ca9ee3 | 
| parent 5068 | fb28eaa07e01 | 
| child 9169 | 85a47aa21f74 | 
| 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: 
2644diff
changeset | 10 | AddIffs [refl_less]; | 
| 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 nipkow parents: 
2644diff
changeset | 11 | |
| 2644 | 12 | (* ------------------------------------------------------------------------ *) | 
| 13 | (* minimal fixes least element *) | |
| 14 | (* ------------------------------------------------------------------------ *) | |
| 3842 | 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: 
2841diff
changeset | 25 | |
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 26 | (* ------------------------------------------------------------------------ *) | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 27 | (* the reverse law of anti--symmetrie of << *) | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 28 | (* ------------------------------------------------------------------------ *) | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 29 | |
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3026diff
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: 
2841diff
changeset | 31 | (fn prems => | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 32 | [ | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 33 | (cut_facts_tac prems 1), | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 34 | (rtac conjI 1), | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
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: 
2841diff
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: 
2841diff
changeset | 37 | ]); | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 38 | |
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 39 | |
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 40 | qed_goal "box_less" thy | 
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3026diff
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: 
2841diff
changeset | 42 | (fn prems => | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 43 | [ | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 44 | (cut_facts_tac prems 1), | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 45 | (etac trans_less 1), | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 46 | (etac trans_less 1), | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 47 | (atac 1) | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 48 | ]); | 
| 
7a5611f66b72
moved antisym_less_inverse,box_less from Porder.ML to Porder0.ML
 slotosch parents: 
2841diff
changeset | 49 | |
| 5068 | 50 | Goal "((x::'a::po)=y) = (x << y & y << x)"; | 
| 4423 | 51 | by (fast_tac (HOL_cs addSEs [antisym_less_inverse] addSIs [antisym_less]) 1); | 
| 3460 | 52 | qed "po_eq_conv"; |