author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 0 | 7949f97df77a |
child 249 | 492493334e0f |
permissions | -rw-r--r-- |
(* Title: HOL/ex/lex-prod.thy ID: $Id$ Author: Tobias Nipkow, TU Munich Copyright 1993 TU Munich The lexicographic product of two relations. *) LexProd = WF + Prod + consts "**" :: "[('a*'a)set, ('b*'b)set] => (('a*'b)*('a*'b))set" (infixl 70) rules lex_prod_def "ra**rb == {p. ? a a' b b'. \ \ p = <<a,b>,<a',b'>> & (<a,a'> : ra | a=a' & <b,b'> : rb)}" end