author | haftmann |
Fri, 25 Sep 2009 09:50:31 +0200 | |
changeset 32705 | 04ce6bb14d85 |
parent 32139 | e271a64f03ff |
child 35787 | afdf1c4958b2 |
permissions | -rw-r--r-- |
(* Author: Florian Haftmann, TU Muenchen *) header {* Pretty syntax for lattice operations *} (*<*) theory Lattice_Syntax imports Complete_Lattice begin notation top ("\<top>") and bot ("\<bottom>") and inf (infixl "\<sqinter>" 70) and sup (infixl "\<squnion>" 65) and Inf ("\<Sqinter>_" [900] 900) and Sup ("\<Squnion>_" [900] 900) end (*>*)