| 
1440
 | 
     1  | 
(*  Title:      LatMorph.thy
  | 
| 
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
    Author:     Markus Wenzel, TU Muenchen
  | 
| 
 | 
     4  | 
  | 
| 
 | 
     5  | 
Some lattice morphisms.
  | 
| 
 | 
     6  | 
  | 
| 
 | 
     7  | 
TODO:
  | 
| 
 | 
     8  | 
  more morphisms (?)
  | 
| 
 | 
     9  | 
  more theorems
  | 
| 
 | 
    10  | 
*)
  | 
| 
 | 
    11  | 
  | 
| 
 | 
    12  | 
LatMorph = LatInsts +
  | 
| 
 | 
    13  | 
  | 
| 
1573
 | 
    14  | 
constdefs
  | 
| 
1440
 | 
    15  | 
  is_mono       :: "('a::le => 'b::le) => bool"
 | 
| 
1573
 | 
    16  | 
  "is_mono f == ALL x y. x [= y --> f x [= f y"
  | 
| 
 | 
    17  | 
  | 
| 
1440
 | 
    18  | 
  is_inf_morph  :: "('a::lattice => 'b::lattice) => bool"
 | 
| 
1573
 | 
    19  | 
  "is_inf_morph f == ALL x y. f (x && y) = f x && f y"
  | 
| 
1440
 | 
    20  | 
  | 
| 
1573
 | 
    21  | 
  is_sup_morph  :: "('a::lattice => 'b::lattice) => bool"
 | 
| 
 | 
    22  | 
  "is_sup_morph f == ALL x y. f (x || y) = f x || f y"
  | 
| 
 | 
    23  | 
  | 
| 
 | 
    24  | 
  is_Inf_morph  :: "('a::clattice => 'b::clattice) => bool"
 | 
| 
 | 
    25  | 
  "is_Inf_morph f == ALL A. f (Inf A) = Inf {f x |x. x:A}"
 | 
| 
 | 
    26  | 
  | 
| 
 | 
    27  | 
  is_Sup_morph  :: "('a::clattice => 'b::clattice) => bool"
 | 
| 
 | 
    28  | 
  "is_Sup_morph f == ALL A. f (Sup A) = Sup {f x |x. x:A}"
 | 
| 
1440
 | 
    29  | 
  | 
| 
 | 
    30  | 
end
  |