| author | desharna | 
| Mon, 10 Jun 2024 13:44:46 +0200 | |
| changeset 80321 | 31b9dfbe534c | 
| parent 69947 | 77a92e8d5167 | 
| permissions | -rw-r--r-- | 
| 69909 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 1 | (* Title: Dual_Ordered_Lattice.thy | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 2 | Authors: Makarius; Peter Gammie; Brian Huffman; Florian Haftmann, TU Muenchen | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 3 | *) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 4 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 5 | section \<open>Type of dual ordered lattices\<close> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 6 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 7 | theory Dual_Ordered_Lattice | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 8 | imports Main | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 9 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 10 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 11 | text \<open> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 12 | The \<^emph>\<open>dual\<close> of an ordered structure is an isomorphic copy of the | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 13 | underlying type, with the \<open>\<le>\<close> relation defined as the inverse | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 14 | of the original one. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 15 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 16 | The class of lattices is closed under formation of dual structures. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 17 | This means that for any theorem of lattice theory, the dualized | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 18 | statement holds as well; this important fact simplifies many proofs | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 19 | of lattice theory. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 20 | \<close> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 21 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 22 | typedef 'a dual = "UNIV :: 'a set" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 23 | morphisms undual dual .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 24 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 25 | setup_lifting type_definition_dual | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 26 | |
| 69947 | 27 | code_datatype dual | 
| 28 | ||
| 69909 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 29 | lemma dual_eqI: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 30 | "x = y" if "undual x = undual y" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 31 | using that by transfer assumption | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 32 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 33 | lemma dual_eq_iff: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 34 | "x = y \<longleftrightarrow> undual x = undual y" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 35 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 36 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 37 | lemma eq_dual_iff [iff]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 38 | "dual x = dual y \<longleftrightarrow> x = y" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 39 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 40 | |
| 69947 | 41 | lemma undual_dual [simp, code]: | 
| 69909 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 42 | "undual (dual x) = x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 43 | by transfer rule | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 44 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 45 | lemma dual_undual [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 46 | "dual (undual x) = x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 47 | by transfer rule | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 48 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 49 | lemma undual_comp_dual [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 50 | "undual \<circ> dual = id" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 51 | by (simp add: fun_eq_iff) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 52 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 53 | lemma dual_comp_undual [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 54 | "dual \<circ> undual = id" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 55 | by (simp add: fun_eq_iff) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 56 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 57 | lemma inj_dual: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 58 | "inj dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 59 | by (rule injI) simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 60 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 61 | lemma inj_undual: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 62 | "inj undual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 63 | by (rule injI) (rule dual_eqI) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 64 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 65 | lemma surj_dual: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 66 | "surj dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 67 | by (rule surjI [of _ undual]) simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 68 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 69 | lemma surj_undual: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 70 | "surj undual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 71 | by (rule surjI [of _ dual]) simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 72 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 73 | lemma bij_dual: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 74 | "bij dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 75 | using inj_dual surj_dual by (rule bijI) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 76 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 77 | lemma bij_undual: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 78 | "bij undual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 79 | using inj_undual surj_undual by (rule bijI) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 80 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 81 | instance dual :: (finite) finite | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 82 | proof | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 83 | from finite have "finite (range dual :: 'a dual set)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 84 | by (rule finite_imageI) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 85 | then show "finite (UNIV :: 'a dual set)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 86 | by (simp add: surj_dual) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 87 | qed | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 88 | |
| 69947 | 89 | instantiation dual :: (equal) equal | 
| 90 | begin | |
| 91 | ||
| 92 | lift_definition equal_dual :: "'a dual \<Rightarrow> 'a dual \<Rightarrow> bool" | |
| 93 | is HOL.equal . | |
| 94 | ||
| 95 | instance | |
| 96 | by (standard; transfer) (simp add: equal) | |
| 97 | ||
| 98 | end | |
| 99 | ||
| 69909 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 100 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 101 | subsection \<open>Pointwise ordering\<close> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 102 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 103 | instantiation dual :: (ord) ord | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 104 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 105 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 106 | lift_definition less_eq_dual :: "'a dual \<Rightarrow> 'a dual \<Rightarrow> bool" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 107 | is "(\<ge>)" . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 108 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 109 | lift_definition less_dual :: "'a dual \<Rightarrow> 'a dual \<Rightarrow> bool" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 110 | is "(>)" . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 111 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 112 | instance .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 113 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 114 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 115 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 116 | lemma dual_less_eqI: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 117 | "x \<le> y" if "undual y \<le> undual x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 118 | using that by transfer assumption | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 119 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 120 | lemma dual_less_eq_iff: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 121 | "x \<le> y \<longleftrightarrow> undual y \<le> undual x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 122 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 123 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 124 | lemma less_eq_dual_iff [iff]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 125 | "dual x \<le> dual y \<longleftrightarrow> y \<le> x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 126 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 127 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 128 | lemma dual_lessI: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 129 | "x < y" if "undual y < undual x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 130 | using that by transfer assumption | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 131 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 132 | lemma dual_less_iff: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 133 | "x < y \<longleftrightarrow> undual y < undual x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 134 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 135 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 136 | lemma less_dual_iff [iff]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 137 | "dual x < dual y \<longleftrightarrow> y < x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 138 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 139 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 140 | instance dual :: (preorder) preorder | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 141 | by (standard; transfer) (auto simp add: less_le_not_le intro: order_trans) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 142 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 143 | instance dual :: (order) order | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 144 | by (standard; transfer) simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 145 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 146 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 147 | subsection \<open>Binary infimum and supremum\<close> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 148 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 149 | instantiation dual :: (sup) inf | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 150 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 151 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 152 | lift_definition inf_dual :: "'a dual \<Rightarrow> 'a dual \<Rightarrow> 'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 153 | is sup . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 154 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 155 | instance .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 156 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 157 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 158 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 159 | lemma undual_inf_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 160 | "undual (inf x y) = sup (undual x) (undual y)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 161 | by (fact inf_dual.rep_eq) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 162 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 163 | lemma dual_sup_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 164 | "dual (sup x y) = inf (dual x) (dual y)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 165 | by transfer rule | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 166 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 167 | instantiation dual :: (inf) sup | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 168 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 169 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 170 | lift_definition sup_dual :: "'a dual \<Rightarrow> 'a dual \<Rightarrow> 'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 171 | is inf . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 172 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 173 | instance .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 174 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 175 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 176 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 177 | lemma undual_sup_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 178 | "undual (sup x y) = inf (undual x) (undual y)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 179 | by (fact sup_dual.rep_eq) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 180 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 181 | lemma dual_inf_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 182 | "dual (inf x y) = sup (dual x) (dual y)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 183 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 184 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 185 | instance dual :: (semilattice_sup) semilattice_inf | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 186 | by (standard; transfer) simp_all | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 187 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 188 | instance dual :: (semilattice_inf) semilattice_sup | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 189 | by (standard; transfer) simp_all | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 190 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 191 | instance dual :: (lattice) lattice .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 192 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 193 | instance dual :: (distrib_lattice) distrib_lattice | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 194 | by (standard; transfer) (fact inf_sup_distrib1) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 195 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 196 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 197 | subsection \<open>Top and bottom elements\<close> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 198 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 199 | instantiation dual :: (top) bot | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 200 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 201 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 202 | lift_definition bot_dual :: "'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 203 | is top . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 204 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 205 | instance .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 206 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 207 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 208 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 209 | lemma undual_bot_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 210 | "undual bot = top" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 211 | by (fact bot_dual.rep_eq) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 212 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 213 | lemma dual_top_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 214 | "dual top = bot" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 215 | by transfer rule | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 216 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 217 | instantiation dual :: (bot) top | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 218 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 219 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 220 | lift_definition top_dual :: "'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 221 | is bot . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 222 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 223 | instance .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 224 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 225 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 226 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 227 | lemma undual_top_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 228 | "undual top = bot" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 229 | by (fact top_dual.rep_eq) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 230 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 231 | lemma dual_bot_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 232 | "dual bot = top" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 233 | by transfer rule | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 234 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 235 | instance dual :: (order_top) order_bot | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 236 | by (standard; transfer) simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 237 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 238 | instance dual :: (order_bot) order_top | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 239 | by (standard; transfer) simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 240 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 241 | instance dual :: (bounded_lattice_top) bounded_lattice_bot .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 242 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 243 | instance dual :: (bounded_lattice_bot) bounded_lattice_top .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 244 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 245 | instance dual :: (bounded_lattice) bounded_lattice .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 246 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 247 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 248 | subsection \<open>Complement\<close> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 249 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 250 | instantiation dual :: (uminus) uminus | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 251 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 252 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 253 | lift_definition uminus_dual :: "'a dual \<Rightarrow> 'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 254 | is uminus . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 255 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 256 | instance .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 257 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 258 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 259 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 260 | lemma undual_uminus_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 261 | "undual (- x) = - undual x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 262 | by (fact uminus_dual.rep_eq) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 263 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 264 | lemma dual_uminus_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 265 | "dual (- x) = - dual x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 266 | by transfer rule | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 267 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 268 | instantiation dual :: (boolean_algebra) boolean_algebra | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 269 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 270 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 271 | lift_definition minus_dual :: "'a dual \<Rightarrow> 'a dual \<Rightarrow> 'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 272 | is "\<lambda>x y. - (y - x)" . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 273 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 274 | instance | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 275 | by (standard; transfer) (simp_all add: diff_eq ac_simps) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 276 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 277 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 278 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 279 | lemma undual_minus_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 280 | "undual (x - y) = - (undual y - undual x)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 281 | by (fact minus_dual.rep_eq) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 282 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 283 | lemma dual_minus_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 284 | "dual (x - y) = - (dual y - dual x)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 285 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 286 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 287 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 288 | subsection \<open>Complete lattice operations\<close> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 289 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 290 | text \<open> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 291 | The class of complete lattices is closed under formation of dual | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 292 | structures. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 293 | \<close> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 294 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 295 | instantiation dual :: (Sup) Inf | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 296 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 297 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 298 | lift_definition Inf_dual :: "'a dual set \<Rightarrow> 'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 299 | is Sup . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 300 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 301 | instance .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 302 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 303 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 304 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 305 | lemma undual_Inf_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 306 | "undual (Inf A) = Sup (undual ` A)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 307 | by (fact Inf_dual.rep_eq) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 308 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 309 | lemma dual_Sup_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 310 | "dual (Sup A) = Inf (dual ` A)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 311 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 312 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 313 | instantiation dual :: (Inf) Sup | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 314 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 315 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 316 | lift_definition Sup_dual :: "'a dual set \<Rightarrow> 'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 317 | is Inf . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 318 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 319 | instance .. | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 320 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 321 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 322 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 323 | lemma undual_Sup_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 324 | "undual (Sup A) = Inf (undual ` A)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 325 | by (fact Sup_dual.rep_eq) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 326 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 327 | lemma dual_Inf_eq [simp]: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 328 | "dual (Inf A) = Sup (dual ` A)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 329 | by transfer simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 330 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 331 | instance dual :: (complete_lattice) complete_lattice | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 332 | by (standard; transfer) (auto intro: Inf_lower Sup_upper Inf_greatest Sup_least) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 333 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 334 | context | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 335 | fixes f :: "'a::complete_lattice \<Rightarrow> 'a" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 336 | and g :: "'a dual \<Rightarrow> 'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 337 | assumes "mono f" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 338 | defines "g \<equiv> dual \<circ> f \<circ> undual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 339 | begin | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 340 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 341 | private lemma mono_dual: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 342 | "mono g" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 343 | proof | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 344 | fix x y :: "'a dual" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 345 | assume "x \<le> y" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 346 | then have "undual y \<le> undual x" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 347 | by (simp add: dual_less_eq_iff) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 348 | with \<open>mono f\<close> have "f (undual y) \<le> f (undual x)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 349 | by (rule monoD) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 350 | then have "(dual \<circ> f \<circ> undual) x \<le> (dual \<circ> f \<circ> undual) y" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 351 | by simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 352 | then show "g x \<le> g y" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 353 | by (simp add: g_def) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 354 | qed | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 355 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 356 | lemma lfp_dual_gfp: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 357 | "lfp f = undual (gfp g)" (is "?lhs = ?rhs") | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 358 | proof (rule antisym) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 359 | have "dual (undual (g (gfp g))) \<le> dual (f (undual (gfp g)))" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 360 | by (simp add: g_def) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 361 | with mono_dual have "f (undual (gfp g)) \<le> undual (gfp g)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 362 | by (simp add: gfp_unfold [where f = g, symmetric] dual_less_eq_iff) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 363 | then show "?lhs \<le> ?rhs" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 364 | by (rule lfp_lowerbound) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 365 | from \<open>mono f\<close> have "dual (lfp f) \<le> dual (undual (gfp g))" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 366 | by (simp add: lfp_fixpoint gfp_upperbound g_def) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 367 | then show "?rhs \<le> ?lhs" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 368 | by (simp only: less_eq_dual_iff) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 369 | qed | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 370 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 371 | lemma gfp_dual_lfp: | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 372 | "gfp f = undual (lfp g)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 373 | proof - | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 374 | have "mono (\<lambda>x. undual (undual x))" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 375 | by (rule monoI) (simp add: dual_less_eq_iff) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 376 | moreover have "mono (\<lambda>a. dual (dual (f a)))" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 377 | using \<open>mono f\<close> by (auto intro: monoI dest: monoD) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 378 | moreover have "gfp f = gfp (\<lambda>x. undual (undual (dual (dual (f x)))))" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 379 | by simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 380 | ultimately have "undual (undual (gfp (\<lambda>x. dual | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 381 | (dual (f (undual (undual x))))))) = | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 382 | gfp (\<lambda>x. undual (undual (dual (dual (f x)))))" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 383 | by (subst gfp_rolling [where g = "\<lambda>x. undual (undual x)"]) simp_all | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 384 | then have "gfp f = | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 385 | undual | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 386 | (undual | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 387 | (gfp (\<lambda>x. dual (dual (f (undual (undual x)))))))" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 388 | by simp | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 389 | also have "\<dots> = undual (undual (gfp (dual \<circ> g \<circ> undual)))" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 390 | by (simp add: comp_def g_def) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 391 | also have "\<dots> = undual (lfp g)" | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 392 | using mono_dual by (simp only: Dual_Ordered_Lattice.lfp_dual_gfp) | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 393 | finally show ?thesis . | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 394 | qed | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 395 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 396 | end | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 397 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 398 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 399 | text \<open>Finally\<close> | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 400 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 401 | lifting_update dual.lifting | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 402 | lifting_forget dual.lifting | 
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 403 | |
| 
5382f5691a11
proper theory for type of dual ordered lattice in distribution
 haftmann parents: diff
changeset | 404 | end |