| author | chaieb | 
| Fri, 27 Mar 2009 14:44:18 +0000 | |
| changeset 30747 | b8ca7e450de3 | 
| parent 30375 | ad2a9dc516ed | 
| child 32139 | e271a64f03ff | 
| permissions | -rw-r--r-- | 
(* Author: Florian Haftmann, TU Muenchen *) header {* Pretty syntax for lattice operations *} (*<*) theory Lattice_Syntax imports Set begin notation inf (infixl "\<sqinter>" 70) and sup (infixl "\<squnion>" 65) and Inf ("\<Sqinter>_" [900] 900) and Sup ("\<Squnion>_" [900] 900) and top ("\<top>") and bot ("\<bottom>") end (*>*)