| 14738 |      1 | (*  Title:   HOL/LOrder.thy
 | 
|  |      2 |     ID:      $Id$
 | 
|  |      3 |     Author:  Steven Obua, TU Muenchen
 | 
| 21734 |      4 | 
 | 
|  |      5 | FIXME integrate properly with lattice locales
 | 
|  |      6 | - make it a class?
 | 
|  |      7 | - get rid of the implicit there-is-a-meet/join but require eplicit ops.
 | 
|  |      8 | Also rename meet/join to inf/sup. 
 | 
| 14738 |      9 | *)
 | 
|  |     10 | 
 | 
| 21312 |     11 | header "Lattice Orders"
 | 
| 14738 |     12 | 
 | 
| 15131 |     13 | theory LOrder
 | 
| 21249 |     14 | imports Lattices
 | 
| 15131 |     15 | begin
 | 
| 14738 |     16 | 
 | 
| 21312 |     17 | text {* The theory of lattices developed here is taken from
 | 
|  |     18 | \cite{Birkhoff79}.  *}
 | 
| 14738 |     19 | 
 | 
|  |     20 | constdefs
 | 
|  |     21 |   is_meet :: "(('a::order) \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
|  |     22 |   "is_meet m == ! a b x. m a b \<le> a \<and> m a b \<le> b \<and> (x \<le> a \<and> x \<le> b \<longrightarrow> x \<le> m a b)"
 | 
|  |     23 |   is_join :: "(('a::order) \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
|  |     24 |   "is_join j == ! a b x. a \<le> j a b \<and> b \<le> j a b \<and> (a \<le> x \<and> b \<le> x \<longrightarrow> j a b \<le> x)"  
 | 
|  |     25 | 
 | 
|  |     26 | lemma is_meet_unique: 
 | 
|  |     27 |   assumes "is_meet u" "is_meet v" shows "u = v"
 | 
|  |     28 | proof -
 | 
|  |     29 |   {
 | 
|  |     30 |     fix a b :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
|  |     31 |     assume a: "is_meet a"
 | 
|  |     32 |     assume b: "is_meet b"
 | 
|  |     33 |     {
 | 
|  |     34 |       fix x y 
 | 
|  |     35 |       let ?za = "a x y"
 | 
|  |     36 |       let ?zb = "b x y"
 | 
|  |     37 |       from a have za_le: "?za <= x & ?za <= y" by (auto simp add: is_meet_def)
 | 
|  |     38 |       with b have "?za <= ?zb" by (auto simp add: is_meet_def)
 | 
|  |     39 |     }
 | 
|  |     40 |   }
 | 
|  |     41 |   note f_le = this
 | 
|  |     42 |   show "u = v" by ((rule ext)+, simp_all add: order_antisym prems f_le) 
 | 
|  |     43 | qed
 | 
|  |     44 | 
 | 
|  |     45 | lemma is_join_unique: 
 | 
|  |     46 |   assumes "is_join u" "is_join v" shows "u = v"
 | 
|  |     47 | proof -
 | 
|  |     48 |   {
 | 
|  |     49 |     fix a b :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
|  |     50 |     assume a: "is_join a"
 | 
|  |     51 |     assume b: "is_join b"
 | 
|  |     52 |     {
 | 
|  |     53 |       fix x y 
 | 
|  |     54 |       let ?za = "a x y"
 | 
|  |     55 |       let ?zb = "b x y"
 | 
|  |     56 |       from a have za_le: "x <= ?za & y <= ?za" by (auto simp add: is_join_def)
 | 
|  |     57 |       with b have "?zb <= ?za" by (auto simp add: is_join_def)
 | 
|  |     58 |     }
 | 
|  |     59 |   }
 | 
|  |     60 |   note f_le = this
 | 
|  |     61 |   show "u = v" by ((rule ext)+, simp_all add: order_antisym prems f_le) 
 | 
|  |     62 | qed
 | 
|  |     63 | 
 | 
|  |     64 | axclass join_semilorder < order
 | 
|  |     65 |   join_exists: "? j. is_join j"
 | 
|  |     66 | 
 | 
|  |     67 | axclass meet_semilorder < order
 | 
|  |     68 |   meet_exists: "? m. is_meet m"
 | 
|  |     69 | 
 | 
|  |     70 | axclass lorder < join_semilorder, meet_semilorder
 | 
|  |     71 | 
 | 
|  |     72 | constdefs
 | 
|  |     73 |   meet :: "('a::meet_semilorder) \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
|  |     74 |   "meet == THE m. is_meet m"
 | 
|  |     75 |   join :: "('a::join_semilorder) \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
|  |     76 |   "join ==  THE j. is_join j"
 | 
|  |     77 | 
 | 
|  |     78 | lemma is_meet_meet: "is_meet (meet::'a \<Rightarrow> 'a \<Rightarrow> ('a::meet_semilorder))"
 | 
|  |     79 | proof -
 | 
|  |     80 |   from meet_exists obtain k::"'a \<Rightarrow> 'a \<Rightarrow> 'a" where "is_meet k" ..
 | 
|  |     81 |   with is_meet_unique[of _ k] show ?thesis
 | 
|  |     82 |     by (simp add: meet_def theI[of is_meet])    
 | 
|  |     83 | qed
 | 
|  |     84 | 
 | 
|  |     85 | lemma meet_unique: "(is_meet m) = (m = meet)" 
 | 
|  |     86 | by (insert is_meet_meet, auto simp add: is_meet_unique)
 | 
|  |     87 | 
 | 
|  |     88 | lemma is_join_join: "is_join (join::'a \<Rightarrow> 'a \<Rightarrow> ('a::join_semilorder))"
 | 
|  |     89 | proof -
 | 
|  |     90 |   from join_exists obtain k::"'a \<Rightarrow> 'a \<Rightarrow> 'a" where "is_join k" ..
 | 
|  |     91 |   with is_join_unique[of _ k] show ?thesis
 | 
|  |     92 |     by (simp add: join_def theI[of is_join])    
 | 
|  |     93 | qed
 | 
|  |     94 | 
 | 
|  |     95 | lemma join_unique: "(is_join j) = (j = join)"
 | 
|  |     96 | by (insert is_join_join, auto simp add: is_join_unique)
 | 
|  |     97 | 
 | 
| 21733 |     98 | interpretation meet_semilat:
 | 
|  |     99 |   lower_semilattice ["op \<le> \<Colon> 'a\<Colon>meet_semilorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" meet]
 | 
| 21380 |    100 | proof unfold_locales
 | 
| 21733 |    101 |   fix x y z :: "'a\<Colon>meet_semilorder"
 | 
| 21380 |    102 |   from is_meet_meet have "is_meet meet" by blast
 | 
|  |    103 |   note meet = this is_meet_def
 | 
|  |    104 |   from meet show "meet x y \<le> x" by blast
 | 
|  |    105 |   from meet show "meet x y \<le> y" by blast
 | 
|  |    106 |   from meet show "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> meet y z" by blast
 | 
| 21733 |    107 | qed
 | 
|  |    108 | 
 | 
|  |    109 | interpretation join_semilat:
 | 
|  |    110 |   upper_semilattice ["op \<le> \<Colon> 'a\<Colon>join_semilorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" join]
 | 
|  |    111 | proof unfold_locales
 | 
|  |    112 |   fix x y z :: "'a\<Colon>join_semilorder"
 | 
| 21380 |    113 |   from is_join_join have "is_join join" by blast
 | 
|  |    114 |   note join = this is_join_def
 | 
|  |    115 |   from join show "x \<le> join x y" by blast
 | 
|  |    116 |   from join show "y \<le> join x y" by blast
 | 
| 21733 |    117 |   from join show "x \<le> z \<Longrightarrow> y \<le> z \<Longrightarrow> join x y \<le> z" by blast
 | 
| 21380 |    118 | qed
 | 
|  |    119 | 
 | 
| 21733 |    120 | declare
 | 
|  |    121 |  join_semilat.antisym_intro[rule del] meet_semilat.antisym_intro[rule del]
 | 
| 21734 |    122 |  join_semilat.le_supE[rule del] meet_semilat.le_infE[rule del]
 | 
| 14738 |    123 | 
 | 
| 21733 |    124 | interpretation meet_join_lat:
 | 
|  |    125 |   lattice ["op \<le> \<Colon> 'a\<Colon>lorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" meet join]
 | 
|  |    126 | by unfold_locales
 | 
| 14738 |    127 | 
 | 
| 21733 |    128 | lemmas meet_left_le = meet_semilat.inf_le1
 | 
|  |    129 | lemmas meet_right_le = meet_semilat.inf_le2
 | 
| 21734 |    130 | lemmas le_meetI[rule del] = meet_semilat.le_infI
 | 
| 21733 |    131 | (* intro! breaks a proof in Hyperreal/SEQ and NumberTheory/IntPrimes *)
 | 
|  |    132 | lemmas join_left_le = join_semilat.sup_ge1
 | 
|  |    133 | lemmas join_right_le = join_semilat.sup_ge2
 | 
| 21734 |    134 | lemmas join_leI[rule del] = join_semilat.le_supI
 | 
| 21312 |    135 | 
 | 
| 21733 |    136 | lemmas meet_join_le = meet_left_le meet_right_le join_left_le join_right_le
 | 
| 21312 |    137 | 
 | 
| 21734 |    138 | lemmas le_meet = meet_semilat.le_inf_iff
 | 
| 21312 |    139 | 
 | 
| 21734 |    140 | lemmas le_join = join_semilat.ge_sup_conv
 | 
| 14738 |    141 | 
 | 
|  |    142 | lemma is_meet_min: "is_meet (min::'a \<Rightarrow> 'a \<Rightarrow> ('a::linorder))"
 | 
|  |    143 | by (auto simp add: is_meet_def min_def)
 | 
|  |    144 | 
 | 
|  |    145 | lemma is_join_max: "is_join (max::'a \<Rightarrow> 'a \<Rightarrow> ('a::linorder))"
 | 
|  |    146 | by (auto simp add: is_join_def max_def)
 | 
|  |    147 | 
 | 
|  |    148 | instance linorder \<subseteq> meet_semilorder
 | 
|  |    149 | proof
 | 
|  |    150 |   from is_meet_min show "? (m::'a\<Rightarrow>'a\<Rightarrow>('a::linorder)). is_meet m" by auto
 | 
|  |    151 | qed
 | 
|  |    152 | 
 | 
|  |    153 | instance linorder \<subseteq> join_semilorder
 | 
|  |    154 | proof
 | 
|  |    155 |   from is_join_max show "? (j::'a\<Rightarrow>'a\<Rightarrow>('a::linorder)). is_join j" by auto 
 | 
|  |    156 | qed
 | 
|  |    157 |     
 | 
|  |    158 | instance linorder \<subseteq> lorder ..
 | 
|  |    159 | 
 | 
|  |    160 | lemma meet_min: "meet = (min :: 'a\<Rightarrow>'a\<Rightarrow>('a::linorder))" 
 | 
|  |    161 | by (simp add: is_meet_meet is_meet_min is_meet_unique)
 | 
|  |    162 | 
 | 
|  |    163 | lemma join_max: "join = (max :: 'a\<Rightarrow>'a\<Rightarrow>('a::linorder))"
 | 
|  |    164 | by (simp add: is_join_join is_join_max is_join_unique)
 | 
|  |    165 | 
 | 
| 21733 |    166 | lemmas meet_idempotent = meet_semilat.inf_idem
 | 
|  |    167 | lemmas join_idempotent = join_semilat.sup_idem
 | 
|  |    168 | lemmas meet_comm = meet_semilat.inf_commute
 | 
|  |    169 | lemmas join_comm = join_semilat.sup_commute
 | 
| 21734 |    170 | lemmas meet_leI1[rule del] = meet_semilat.le_infI1
 | 
|  |    171 | lemmas meet_leI2[rule del] = meet_semilat.le_infI2
 | 
|  |    172 | lemmas le_joinI1[rule del] = join_semilat.le_supI1
 | 
|  |    173 | lemmas le_joinI2[rule del] = join_semilat.le_supI2
 | 
| 21733 |    174 | lemmas meet_assoc = meet_semilat.inf_assoc
 | 
|  |    175 | lemmas join_assoc = join_semilat.sup_assoc
 | 
|  |    176 | lemmas meet_left_comm = meet_semilat.inf_left_commute
 | 
|  |    177 | lemmas meet_left_idempotent = meet_semilat.inf_left_idem
 | 
|  |    178 | lemmas join_left_comm = join_semilat.sup_left_commute
 | 
|  |    179 | lemmas join_left_idempotent= join_semilat.sup_left_idem
 | 
| 14738 |    180 |     
 | 
|  |    181 | lemmas meet_aci = meet_assoc meet_comm meet_left_comm meet_left_idempotent
 | 
|  |    182 | lemmas join_aci = join_assoc join_comm join_left_comm join_left_idempotent
 | 
|  |    183 | 
 | 
| 21734 |    184 | lemmas le_def_meet = meet_semilat.le_iff_inf
 | 
|  |    185 | lemmas le_def_join = join_semilat.le_iff_sup
 | 
| 21312 |    186 | 
 | 
| 21733 |    187 | lemmas join_absorp2 = join_semilat.sup_absorb2
 | 
|  |    188 | lemmas join_absorp1 = join_semilat.sup_absorb1
 | 
| 21312 |    189 | 
 | 
| 21733 |    190 | lemmas meet_absorp1 = meet_semilat.inf_absorb1
 | 
|  |    191 | lemmas meet_absorp2 = meet_semilat.inf_absorb2
 | 
| 14738 |    192 | 
 | 
| 21734 |    193 | interpretation meet_join_lat:
 | 
|  |    194 |   lattice ["op \<le> \<Colon> 'a\<Colon>{meet_semilorder,join_semilorder} \<Rightarrow> 'a \<Rightarrow> bool" "op <" meet join]
 | 
|  |    195 | by unfold_locales
 | 
| 14738 |    196 | 
 | 
| 21734 |    197 | lemmas meet_join_absorp = meet_join_lat.inf_sup_absorb
 | 
|  |    198 | lemmas join_meet_absorp = meet_join_lat.sup_inf_absorb
 | 
| 14738 |    199 | 
 | 
| 21734 |    200 | lemmas distrib_join_le = meet_join_lat.distrib_sup_le
 | 
|  |    201 | lemmas distrib_meet_le = meet_join_lat.distrib_inf_le
 | 
| 14738 |    202 | 
 | 
| 15131 |    203 | end
 |