src/HOLCF/ssum1.thy
author nipkow
Thu, 26 Nov 1998 12:18:08 +0100
changeset 5974 6acf3ff0f486
parent 243 c22b85994e17
permissions -rw-r--r--
Added filter_prems_tac

(*  Title: 	HOLCF/ssum1.thy
    ID:         $Id$
    Author: 	Franz Regensburger
    Copyright   1993  Technische Universitaet Muenchen

Partial ordering for the strict sum ++
*)

Ssum1 = Ssum0 +

consts

  less_ssum	:: "[('a ++ 'b),('a ++ 'b)] => bool"	

rules

  less_ssum_def "less_ssum(s1,s2) == (@z.\
\	 (! u x.s1=Isinl(u) & s2=Isinl(x) --> z = (u << x))\
\	&(! v y.s1=Isinr(v) & s2=Isinr(y) --> z = (v << y))\
\	&(! u y.s1=Isinl(u) & s2=Isinr(y) --> z = (u = UU))\
\	&(! v x.s1=Isinr(v) & s2=Isinl(x) --> z = (v = UU)))"

end