src/HOLCF/Porder0.ML
author nipkow
Thu, 26 Nov 1998 12:18:08 +0100
changeset 5974 6acf3ff0f486
parent 5068 fb28eaa07e01
child 9169 85a47aa21f74
permissions -rw-r--r--
Added filter_prems_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2644
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
     1
(*  Title:      HOLCF/Porder0.ML
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
     2
    ID:         $Id$
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
     3
    Author:     Oscar Slotosch
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
     4
    Copyright   1997 Technische Universitaet Muenchen
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
     5
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
     6
    derive the characteristic axioms for the characteristic constants *)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
     7
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
     8
open Porder0;
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
     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
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    12
(* ------------------------------------------------------------------------ *)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    13
(* minimal fixes least element                                              *)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    14
(* ------------------------------------------------------------------------ *)
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3460
diff changeset
    15
bind_thm("minimal2UU",allI RS (prove_goal thy "!x::'a::po. uu<<x==>uu=(@u.!y. u<<y)"
2644
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    16
(fn prems =>
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    17
        [
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    18
        (cut_facts_tac prems 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    19
        (rtac antisym_less 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    20
        (etac spec 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    21
        (res_inst_tac [("a","uu")] selectI2 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    22
	(atac 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    23
	(etac spec 1)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    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
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
    50
Goal "((x::'a::po)=y) = (x << y & y << x)";
4423
a129b817b58a expandshort;
wenzelm
parents: 3842
diff changeset
    51
by (fast_tac (HOL_cs addSEs [antisym_less_inverse] addSIs [antisym_less]) 1);
3460
5d71eed16fbe Tuned Franz's proofs.
nipkow
parents: 3323
diff changeset
    52
qed "po_eq_conv";