src/HOLCF/Porder0.thy
author paulson
Wed, 09 Oct 1996 13:32:33 +0200
changeset 2073 fb0655539d05
parent 1479 21eb5e156d91
child 2278 d63ffafce255
permissions -rw-r--r--
New unified treatment of sequent calculi by Sara Kalvala combines the old LK and Modal with the new ILL (Int. Linear Logic)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
     1
(*  Title:      HOLCF/porder0.thy
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
     3
    Author:     Franz Regensburger
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     4
    Copyright   1993 Technische Universitaet Muenchen
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     5
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     6
Definition of class porder (partial order)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     7
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     8
The prototype theory for this class is void.thy 
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     9
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    10
*)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    11
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    12
Porder0 = Void +
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    13
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    14
(* Introduction of new class. The witness is type void. *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    15
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    16
classes po < term
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    17
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    18
(* default type is still term ! *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    19
(* void is the prototype in po *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    20
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    21
arities void :: po
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    22
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
    23
consts  "<<"    ::      "['a,'a::po] => bool"   (infixl 55)
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    24
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    25
rules
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    26
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    27
(* class axioms: justification is theory Void *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    28
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
    29
refl_less       "x << x"        
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
    30
                                (* witness refl_less_void    *)
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    31
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
    32
antisym_less    "[|x<<y ; y<<x |] ==> x = y"    
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
    33
                                (* witness antisym_less_void *)
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    34
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
    35
trans_less      "[|x<<y ; y<<z |] ==> x<<z"
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
    36
                                (* witness trans_less_void   *)
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    37
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    38
(* instance of << for the prototype void *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    39
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
    40
inst_void_po    "((op <<)::[void,void]=>bool) = less_void"
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    41
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    42
(* start 8bit 1 *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    43
(* end 8bit 1 *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    44
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    45
end 
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    46
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    47
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    48
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    49
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    50