| author | wenzelm | 
| Thu, 15 Apr 2010 16:55:49 +0200 | |
| changeset 36150 | 123f721c9a37 | 
| parent 35787 | afdf1c4958b2 | 
| child 41080 | 294956ff285b | 
| permissions | -rw-r--r-- | 
| 30330 | 1  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
2  | 
||
3  | 
header {* Pretty syntax for lattice operations *}
 | 
|
4  | 
||
5  | 
theory Lattice_Syntax  | 
|
| 32139 | 6  | 
imports Complete_Lattice  | 
| 30330 | 7  | 
begin  | 
8  | 
||
9  | 
notation  | 
|
| 32139 | 10  | 
  top ("\<top>") and
 | 
11  | 
  bot ("\<bottom>") and
 | 
|
| 30330 | 12  | 
inf (infixl "\<sqinter>" 70) and  | 
13  | 
sup (infixl "\<squnion>" 65) and  | 
|
14  | 
  Inf  ("\<Sqinter>_" [900] 900) and
 | 
|
| 32139 | 15  | 
  Sup  ("\<Squnion>_" [900] 900)
 | 
| 30330 | 16  | 
|
| 30331 | 17  | 
end  |