src/HOLCF/porder0.thy
author wenzelm
Sun, 30 Jul 2000 13:03:49 +0200
changeset 9475 b24516d96847
parent 298 3a0485439396
permissions -rw-r--r--
adapted obtain; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     1
(*  Title: 	HOLCF/porder0.thy
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     2
    ID:         $Id$
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     3
    Author: 	Franz Regensburger
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
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    23
consts	"<<"	::	"['a,'a::po] => bool"	(infixl 55)
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
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    29
refl_less	"x << x"	
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    30
				(* witness refl_less_void    *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    31
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    32
antisym_less	"[|x<<y ; y<<x |] ==> x = y"	
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    33
				(* witness antisym_less_void *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    34
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    35
trans_less	"[|x<<y ; y<<z |] ==> x<<z"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    36
				(* witness trans_less_void   *)
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
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    40
inst_void_po	"(op <<)::[void,void]=>bool = less_void"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    41
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    42
end