author | haftmann |
Tue, 23 Jun 2009 11:31:28 +0200 | |
changeset 31769 | d5f39775edd2 |
parent 30375 | ad2a9dc516ed |
child 32139 | e271a64f03ff |
permissions | -rw-r--r-- |
30330 | 1 |
(* Author: Florian Haftmann, TU Muenchen *) |
2 |
||
3 |
header {* Pretty syntax for lattice operations *} |
|
4 |
||
30331 | 5 |
(*<*) |
30330 | 6 |
theory Lattice_Syntax |
7 |
imports Set |
|
8 |
begin |
|
9 |
||
10 |
notation |
|
11 |
inf (infixl "\<sqinter>" 70) and |
|
12 |
sup (infixl "\<squnion>" 65) and |
|
13 |
Inf ("\<Sqinter>_" [900] 900) and |
|
30375 | 14 |
Sup ("\<Squnion>_" [900] 900) and |
15 |
top ("\<top>") and |
|
16 |
bot ("\<bottom>") |
|
30330 | 17 |
|
30331 | 18 |
end |
19 |
(*>*) |