changeset 9685 | 6d123a7e30bd |
parent 8055 | bb15396278fb |
child 10064 | 1a77667b21ef |
--- a/src/HOL/UNITY/Union.thy Thu Aug 24 11:14:21 2000 +0200 +++ b/src/HOL/UNITY/Union.thy Thu Aug 24 12:39:24 2000 +0200 @@ -29,4 +29,10 @@ "JN x y. B" == "JN x. JN y. B" "JN x. B" == "JOIN UNIV (%x. B)" +syntax (symbols) + SKIP :: 'a program ("\\<bottom>") + "op Join" :: ['a program, 'a program] => 'a program (infixl "\\<squnion>" 65) + "@JOIN1" :: [pttrns, 'b set] => 'b set ("(3\\<Squnion> _./ _)" 10) + "@JOIN" :: [pttrn, 'a set, 'b set] => 'b set ("(3\\<Squnion> _:_./ _)" 10) + end