src/HOLCF/Porder0.ML
author slotosch
Mon, 17 Feb 1997 13:26:32 +0100
changeset 2644 2fa0f0c1c750
child 2841 c2508f4ab739
permissions -rw-r--r--
New file for theorems of Porder0 Dervie the prperties of partial orders from the axiomatic type class po
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
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    10
qed_goalw "refl_less" thy [ po_def ] "x << x"
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    11
(fn prems =>
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    12
        [
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    13
        (fast_tac (HOL_cs addSIs [ax_refl_less]) 1)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    14
        ]);
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    15
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    16
qed_goalw "antisym_less" thy [ po_def ] "[| x << y; y << x |] ==> x = y"
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    17
(fn prems =>
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    18
        [
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    19
        (fast_tac (HOL_cs addSIs prems@[ax_antisym_less]) 1)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    20
        ]);
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    21
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    22
qed_goalw "trans_less" thy [ po_def ] "[| x << y; y << z |] ==> x << z"
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    23
(fn prems =>
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    24
        [
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    25
        (fast_tac (HOL_cs addIs prems@[ax_trans_less]) 1)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    26
        ]);
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    27
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    28
(* ------------------------------------------------------------------------ *)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    29
(* minimal fixes least element                                              *)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    30
(* ------------------------------------------------------------------------ *)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    31
bind_thm("minimal2UU",allI RS (prove_goal thy "!x.uu<<x==>uu=(@u.!y.u<<y)"
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    32
(fn prems =>
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    33
        [
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    34
        (cut_facts_tac prems 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    35
        (rtac antisym_less 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    36
        (etac spec 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    37
        (res_inst_tac [("a","uu")] selectI2 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    38
	(atac 1),
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    39
	(etac spec 1)
2fa0f0c1c750 New file for theorems of Porder0
slotosch
parents:
diff changeset
    40
        ])));