| author | haftmann | 
| Wed, 02 Apr 2008 15:58:37 +0200 | |
| changeset 26515 | 4a2063a8c2d2 | 
| parent 26420 | 57a626f64875 | 
| child 26806 | 40b411ec05aa | 
| permissions | -rw-r--r-- | 
| 25904 | 1 | (* Title: HOLCF/ConvexPD.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Brian Huffman | |
| 4 | *) | |
| 5 | ||
| 6 | header {* Convex powerdomain *}
 | |
| 7 | ||
| 8 | theory ConvexPD | |
| 9 | imports UpperPD LowerPD | |
| 10 | begin | |
| 11 | ||
| 12 | subsection {* Basis preorder *}
 | |
| 13 | ||
| 14 | definition | |
| 15 | convex_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix "\<le>\<natural>" 50) where | |
| 16 | "convex_le = (\<lambda>u v. u \<le>\<sharp> v \<and> u \<le>\<flat> v)" | |
| 17 | ||
| 18 | lemma convex_le_refl [simp]: "t \<le>\<natural> t" | |
| 19 | unfolding convex_le_def by (fast intro: upper_le_refl lower_le_refl) | |
| 20 | ||
| 21 | lemma convex_le_trans: "\<lbrakk>t \<le>\<natural> u; u \<le>\<natural> v\<rbrakk> \<Longrightarrow> t \<le>\<natural> v" | |
| 22 | unfolding convex_le_def by (fast intro: upper_le_trans lower_le_trans) | |
| 23 | ||
| 24 | interpretation convex_le: preorder [convex_le] | |
| 25 | by (rule preorder.intro, rule convex_le_refl, rule convex_le_trans) | |
| 26 | ||
| 27 | lemma upper_le_minimal [simp]: "PDUnit compact_bot \<le>\<natural> t" | |
| 28 | unfolding convex_le_def Rep_PDUnit by simp | |
| 29 | ||
| 26420 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 30 | lemma PDUnit_convex_mono: "x \<sqsubseteq> y \<Longrightarrow> PDUnit x \<le>\<natural> PDUnit y" | 
| 25904 | 31 | unfolding convex_le_def by (fast intro: PDUnit_upper_mono PDUnit_lower_mono) | 
| 32 | ||
| 33 | lemma PDPlus_convex_mono: "\<lbrakk>s \<le>\<natural> t; u \<le>\<natural> v\<rbrakk> \<Longrightarrow> PDPlus s u \<le>\<natural> PDPlus t v" | |
| 34 | unfolding convex_le_def by (fast intro: PDPlus_upper_mono PDPlus_lower_mono) | |
| 35 | ||
| 36 | lemma convex_le_PDUnit_PDUnit_iff [simp]: | |
| 26420 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 37 | "(PDUnit a \<le>\<natural> PDUnit b) = a \<sqsubseteq> b" | 
| 25904 | 38 | unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit by fast | 
| 39 | ||
| 40 | lemma convex_le_PDUnit_lemma1: | |
| 26420 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 41 | "(PDUnit a \<le>\<natural> t) = (\<forall>b\<in>Rep_pd_basis t. a \<sqsubseteq> b)" | 
| 25904 | 42 | unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit | 
| 43 | using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast | |
| 44 | ||
| 45 | lemma convex_le_PDUnit_PDPlus_iff [simp]: | |
| 46 | "(PDUnit a \<le>\<natural> PDPlus t u) = (PDUnit a \<le>\<natural> t \<and> PDUnit a \<le>\<natural> u)" | |
| 47 | unfolding convex_le_PDUnit_lemma1 Rep_PDPlus by fast | |
| 48 | ||
| 49 | lemma convex_le_PDUnit_lemma2: | |
| 26420 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 50 | "(t \<le>\<natural> PDUnit b) = (\<forall>a\<in>Rep_pd_basis t. a \<sqsubseteq> b)" | 
| 25904 | 51 | unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit | 
| 52 | using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast | |
| 53 | ||
| 54 | lemma convex_le_PDPlus_PDUnit_iff [simp]: | |
| 55 | "(PDPlus t u \<le>\<natural> PDUnit a) = (t \<le>\<natural> PDUnit a \<and> u \<le>\<natural> PDUnit a)" | |
| 56 | unfolding convex_le_PDUnit_lemma2 Rep_PDPlus by fast | |
| 57 | ||
| 58 | lemma convex_le_PDPlus_lemma: | |
| 59 | assumes z: "PDPlus t u \<le>\<natural> z" | |
| 60 | shows "\<exists>v w. z = PDPlus v w \<and> t \<le>\<natural> v \<and> u \<le>\<natural> w" | |
| 61 | proof (intro exI conjI) | |
| 26420 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 62 |   let ?A = "{b\<in>Rep_pd_basis z. \<exists>a\<in>Rep_pd_basis t. a \<sqsubseteq> b}"
 | 
| 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 63 |   let ?B = "{b\<in>Rep_pd_basis z. \<exists>a\<in>Rep_pd_basis u. a \<sqsubseteq> b}"
 | 
| 25904 | 64 | let ?v = "Abs_pd_basis ?A" | 
| 65 | let ?w = "Abs_pd_basis ?B" | |
| 66 | have Rep_v: "Rep_pd_basis ?v = ?A" | |
| 67 | apply (rule Abs_pd_basis_inverse) | |
| 68 | apply (rule Rep_pd_basis_nonempty [of t, folded ex_in_conv, THEN exE]) | |
| 69 | apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify) | |
| 70 | apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE) | |
| 71 | apply (simp add: pd_basis_def) | |
| 72 | apply fast | |
| 73 | done | |
| 74 | have Rep_w: "Rep_pd_basis ?w = ?B" | |
| 75 | apply (rule Abs_pd_basis_inverse) | |
| 76 | apply (rule Rep_pd_basis_nonempty [of u, folded ex_in_conv, THEN exE]) | |
| 77 | apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify) | |
| 78 | apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE) | |
| 79 | apply (simp add: pd_basis_def) | |
| 80 | apply fast | |
| 81 | done | |
| 82 | show "z = PDPlus ?v ?w" | |
| 83 | apply (insert z) | |
| 84 | apply (simp add: convex_le_def, erule conjE) | |
| 85 | apply (simp add: Rep_pd_basis_inject [symmetric] Rep_PDPlus) | |
| 86 | apply (simp add: Rep_v Rep_w) | |
| 87 | apply (rule equalityI) | |
| 88 | apply (rule subsetI) | |
| 89 | apply (simp only: upper_le_def) | |
| 90 | apply (drule (1) bspec, erule bexE) | |
| 91 | apply (simp add: Rep_PDPlus) | |
| 92 | apply fast | |
| 93 | apply fast | |
| 94 | done | |
| 95 | show "t \<le>\<natural> ?v" "u \<le>\<natural> ?w" | |
| 96 | apply (insert z) | |
| 97 | apply (simp_all add: convex_le_def upper_le_def lower_le_def Rep_PDPlus Rep_v Rep_w) | |
| 98 | apply fast+ | |
| 99 | done | |
| 100 | qed | |
| 101 | ||
| 102 | lemma convex_le_induct [induct set: convex_le]: | |
| 103 | assumes le: "t \<le>\<natural> u" | |
| 104 | assumes 2: "\<And>t u v. \<lbrakk>P t u; P u v\<rbrakk> \<Longrightarrow> P t v" | |
| 26420 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 105 | assumes 3: "\<And>a b. a \<sqsubseteq> b \<Longrightarrow> P (PDUnit a) (PDUnit b)" | 
| 25904 | 106 | assumes 4: "\<And>t u v w. \<lbrakk>P t v; P u w\<rbrakk> \<Longrightarrow> P (PDPlus t u) (PDPlus v w)" | 
| 107 | shows "P t u" | |
| 108 | using le apply (induct t arbitrary: u rule: pd_basis_induct) | |
| 109 | apply (erule rev_mp) | |
| 110 | apply (induct_tac u rule: pd_basis_induct1) | |
| 111 | apply (simp add: 3) | |
| 112 | apply (simp, clarify, rename_tac a b t) | |
| 113 | apply (subgoal_tac "P (PDPlus (PDUnit a) (PDUnit a)) (PDPlus (PDUnit b) t)") | |
| 114 | apply (simp add: PDPlus_absorb) | |
| 115 | apply (erule (1) 4 [OF 3]) | |
| 116 | apply (drule convex_le_PDPlus_lemma, clarify) | |
| 117 | apply (simp add: 4) | |
| 118 | done | |
| 119 | ||
| 120 | lemma approx_pd_convex_mono1: | |
| 121 | "i \<le> j \<Longrightarrow> approx_pd i t \<le>\<natural> approx_pd j t" | |
| 122 | apply (induct t rule: pd_basis_induct) | |
| 123 | apply (simp add: compact_approx_mono1) | |
| 124 | apply (simp add: PDPlus_convex_mono) | |
| 125 | done | |
| 126 | ||
| 127 | lemma approx_pd_convex_le: "approx_pd i t \<le>\<natural> t" | |
| 128 | apply (induct t rule: pd_basis_induct) | |
| 129 | apply (simp add: compact_approx_le) | |
| 130 | apply (simp add: PDPlus_convex_mono) | |
| 131 | done | |
| 132 | ||
| 133 | lemma approx_pd_convex_mono: | |
| 134 | "t \<le>\<natural> u \<Longrightarrow> approx_pd n t \<le>\<natural> approx_pd n u" | |
| 135 | apply (erule convex_le_induct) | |
| 136 | apply (erule (1) convex_le_trans) | |
| 137 | apply (simp add: compact_approx_mono) | |
| 138 | apply (simp add: PDPlus_convex_mono) | |
| 139 | done | |
| 140 | ||
| 141 | ||
| 142 | subsection {* Type definition *}
 | |
| 143 | ||
| 144 | cpodef (open) 'a convex_pd = | |
| 26407 
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
 huffman parents: 
26041diff
changeset | 145 |   "{S::'a::profinite pd_basis set. convex_le.ideal S}"
 | 
| 25904 | 146 | apply (simp add: convex_le.adm_ideal) | 
| 147 | apply (fast intro: convex_le.ideal_principal) | |
| 148 | done | |
| 149 | ||
| 150 | lemma ideal_Rep_convex_pd: "convex_le.ideal (Rep_convex_pd xs)" | |
| 151 | by (rule Rep_convex_pd [simplified]) | |
| 152 | ||
| 153 | lemma Rep_convex_pd_mono: "xs \<sqsubseteq> ys \<Longrightarrow> Rep_convex_pd xs \<subseteq> Rep_convex_pd ys" | |
| 154 | unfolding less_convex_pd_def less_set_def . | |
| 155 | ||
| 156 | ||
| 157 | subsection {* Principal ideals *}
 | |
| 158 | ||
| 159 | definition | |
| 160 | convex_principal :: "'a pd_basis \<Rightarrow> 'a convex_pd" where | |
| 161 |   "convex_principal t = Abs_convex_pd {u. u \<le>\<natural> t}"
 | |
| 162 | ||
| 163 | lemma Rep_convex_principal: | |
| 164 |   "Rep_convex_pd (convex_principal t) = {u. u \<le>\<natural> t}"
 | |
| 165 | unfolding convex_principal_def | |
| 166 | apply (rule Abs_convex_pd_inverse [simplified]) | |
| 167 | apply (rule convex_le.ideal_principal) | |
| 168 | done | |
| 169 | ||
| 170 | interpretation convex_pd: | |
| 26420 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 171 | bifinite_basis [convex_le approx_pd convex_principal Rep_convex_pd] | 
| 25904 | 172 | apply unfold_locales | 
| 173 | apply (rule approx_pd_convex_le) | |
| 174 | apply (rule approx_pd_idem) | |
| 175 | apply (erule approx_pd_convex_mono) | |
| 176 | apply (rule approx_pd_convex_mono1, simp) | |
| 177 | apply (rule finite_range_approx_pd) | |
| 178 | apply (rule ex_approx_pd_eq) | |
| 26420 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 179 | apply (rule ideal_Rep_convex_pd) | 
| 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 180 | apply (rule cont_Rep_convex_pd) | 
| 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 181 | apply (rule Rep_convex_principal) | 
| 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26407diff
changeset | 182 | apply (simp only: less_convex_pd_def less_set_def) | 
| 25904 | 183 | done | 
| 184 | ||
| 185 | lemma convex_principal_less_iff [simp]: | |
| 186 | "(convex_principal t \<sqsubseteq> convex_principal u) = (t \<le>\<natural> u)" | |
| 187 | unfolding less_convex_pd_def Rep_convex_principal less_set_def | |
| 188 | by (fast intro: convex_le_refl elim: convex_le_trans) | |
| 189 | ||
| 190 | lemma convex_principal_mono: | |
| 191 | "t \<le>\<natural> u \<Longrightarrow> convex_principal t \<sqsubseteq> convex_principal u" | |
| 192 | by (rule convex_principal_less_iff [THEN iffD2]) | |
| 193 | ||
| 194 | lemma compact_convex_principal: "compact (convex_principal t)" | |
| 195 | by (rule convex_pd.compact_principal) | |
| 196 | ||
| 197 | lemma convex_pd_minimal: "convex_principal (PDUnit compact_bot) \<sqsubseteq> ys" | |
| 198 | by (induct ys rule: convex_pd.principal_induct, simp, simp) | |
| 199 | ||
| 200 | instance convex_pd :: (bifinite) pcpo | |
| 201 | by (intro_classes, fast intro: convex_pd_minimal) | |
| 202 | ||
| 203 | lemma inst_convex_pd_pcpo: "\<bottom> = convex_principal (PDUnit compact_bot)" | |
| 204 | by (rule convex_pd_minimal [THEN UU_I, symmetric]) | |
| 205 | ||
| 206 | ||
| 207 | subsection {* Approximation *}
 | |
| 208 | ||
| 26407 
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
 huffman parents: 
26041diff
changeset | 209 | instance convex_pd :: (profinite) approx .. | 
| 25904 | 210 | |
| 211 | defs (overloaded) | |
| 212 | approx_convex_pd_def: | |
| 213 | "approx \<equiv> (\<lambda>n. convex_pd.basis_fun (\<lambda>t. convex_principal (approx_pd n t)))" | |
| 214 | ||
| 215 | lemma approx_convex_principal [simp]: | |
| 216 | "approx n\<cdot>(convex_principal t) = convex_principal (approx_pd n t)" | |
| 217 | unfolding approx_convex_pd_def | |
| 218 | apply (rule convex_pd.basis_fun_principal) | |
| 219 | apply (erule convex_principal_mono [OF approx_pd_convex_mono]) | |
| 220 | done | |
| 221 | ||
| 222 | lemma chain_approx_convex_pd: | |
| 223 | "chain (approx :: nat \<Rightarrow> 'a convex_pd \<rightarrow> 'a convex_pd)" | |
| 224 | unfolding approx_convex_pd_def | |
| 225 | by (rule convex_pd.chain_basis_fun_take) | |
| 226 | ||
| 227 | lemma lub_approx_convex_pd: | |
| 228 | "(\<Squnion>i. approx i\<cdot>xs) = (xs::'a convex_pd)" | |
| 229 | unfolding approx_convex_pd_def | |
| 230 | by (rule convex_pd.lub_basis_fun_take) | |
| 231 | ||
| 232 | lemma approx_convex_pd_idem: | |
| 233 | "approx n\<cdot>(approx n\<cdot>xs) = approx n\<cdot>(xs::'a convex_pd)" | |
| 234 | apply (induct xs rule: convex_pd.principal_induct, simp) | |
| 235 | apply (simp add: approx_pd_idem) | |
| 236 | done | |
| 237 | ||
| 238 | lemma approx_eq_convex_principal: | |
| 239 | "\<exists>t\<in>Rep_convex_pd xs. approx n\<cdot>xs = convex_principal (approx_pd n t)" | |
| 240 | unfolding approx_convex_pd_def | |
| 241 | by (rule convex_pd.basis_fun_take_eq_principal) | |
| 242 | ||
| 243 | lemma finite_fixes_approx_convex_pd: | |
| 244 |   "finite {xs::'a convex_pd. approx n\<cdot>xs = xs}"
 | |
| 245 | unfolding approx_convex_pd_def | |
| 246 | by (rule convex_pd.finite_fixes_basis_fun_take) | |
| 247 | ||
| 26407 
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
 huffman parents: 
26041diff
changeset | 248 | instance convex_pd :: (profinite) profinite | 
| 25904 | 249 | apply intro_classes | 
| 250 | apply (simp add: chain_approx_convex_pd) | |
| 251 | apply (rule lub_approx_convex_pd) | |
| 252 | apply (rule approx_convex_pd_idem) | |
| 253 | apply (rule finite_fixes_approx_convex_pd) | |
| 254 | done | |
| 255 | ||
| 26407 
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
 huffman parents: 
26041diff
changeset | 256 | instance convex_pd :: (bifinite) bifinite .. | 
| 
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
 huffman parents: 
26041diff
changeset | 257 | |
| 25904 | 258 | lemma compact_imp_convex_principal: | 
| 259 | "compact xs \<Longrightarrow> \<exists>t. xs = convex_principal t" | |
| 260 | apply (drule bifinite_compact_eq_approx) | |
| 261 | apply (erule exE) | |
| 262 | apply (erule subst) | |
| 263 | apply (cut_tac n=i and xs=xs in approx_eq_convex_principal) | |
| 264 | apply fast | |
| 265 | done | |
| 266 | ||
| 267 | lemma convex_principal_induct: | |
| 268 | "\<lbrakk>adm P; \<And>t. P (convex_principal t)\<rbrakk> \<Longrightarrow> P xs" | |
| 269 | apply (erule approx_induct, rename_tac xs) | |
| 270 | apply (cut_tac n=n and xs=xs in approx_eq_convex_principal) | |
| 271 | apply (clarify, simp) | |
| 272 | done | |
| 273 | ||
| 274 | lemma convex_principal_induct2: | |
| 275 | "\<lbrakk>\<And>ys. adm (\<lambda>xs. P xs ys); \<And>xs. adm (\<lambda>ys. P xs ys); | |
| 276 | \<And>t u. P (convex_principal t) (convex_principal u)\<rbrakk> \<Longrightarrow> P xs ys" | |
| 277 | apply (rule_tac x=ys in spec) | |
| 278 | apply (rule_tac xs=xs in convex_principal_induct, simp) | |
| 279 | apply (rule allI, rename_tac ys) | |
| 280 | apply (rule_tac xs=ys in convex_principal_induct, simp) | |
| 281 | apply simp | |
| 282 | done | |
| 283 | ||
| 284 | ||
| 285 | subsection {* Monadic unit *}
 | |
| 286 | ||
| 287 | definition | |
| 288 | convex_unit :: "'a \<rightarrow> 'a convex_pd" where | |
| 289 | "convex_unit = compact_basis.basis_fun (\<lambda>a. convex_principal (PDUnit a))" | |
| 290 | ||
| 291 | lemma convex_unit_Rep_compact_basis [simp]: | |
| 292 | "convex_unit\<cdot>(Rep_compact_basis a) = convex_principal (PDUnit a)" | |
| 293 | unfolding convex_unit_def | |
| 294 | apply (rule compact_basis.basis_fun_principal) | |
| 295 | apply (rule convex_principal_mono) | |
| 296 | apply (erule PDUnit_convex_mono) | |
| 297 | done | |
| 298 | ||
| 299 | lemma convex_unit_strict [simp]: "convex_unit\<cdot>\<bottom> = \<bottom>" | |
| 300 | unfolding inst_convex_pd_pcpo Rep_compact_bot [symmetric] by simp | |
| 301 | ||
| 302 | lemma approx_convex_unit [simp]: | |
| 303 | "approx n\<cdot>(convex_unit\<cdot>x) = convex_unit\<cdot>(approx n\<cdot>x)" | |
| 304 | apply (induct x rule: compact_basis_induct, simp) | |
| 305 | apply (simp add: approx_Rep_compact_basis) | |
| 306 | done | |
| 307 | ||
| 308 | lemma convex_unit_less_iff [simp]: | |
| 309 | "(convex_unit\<cdot>x \<sqsubseteq> convex_unit\<cdot>y) = (x \<sqsubseteq> y)" | |
| 310 | apply (rule iffI) | |
| 311 | apply (rule bifinite_less_ext) | |
| 312 | apply (drule_tac f="approx i" in monofun_cfun_arg, simp) | |
| 313 | apply (cut_tac x="approx i\<cdot>x" in compact_imp_Rep_compact_basis, simp) | |
| 314 | apply (cut_tac x="approx i\<cdot>y" in compact_imp_Rep_compact_basis, simp) | |
| 315 | apply (clarify, simp add: compact_le_def) | |
| 316 | apply (erule monofun_cfun_arg) | |
| 317 | done | |
| 318 | ||
| 319 | lemma convex_unit_eq_iff [simp]: | |
| 320 | "(convex_unit\<cdot>x = convex_unit\<cdot>y) = (x = y)" | |
| 321 | unfolding po_eq_conv by simp | |
| 322 | ||
| 323 | lemma convex_unit_strict_iff [simp]: "(convex_unit\<cdot>x = \<bottom>) = (x = \<bottom>)" | |
| 324 | unfolding convex_unit_strict [symmetric] by (rule convex_unit_eq_iff) | |
| 325 | ||
| 326 | lemma compact_convex_unit_iff [simp]: | |
| 327 | "compact (convex_unit\<cdot>x) = compact x" | |
| 328 | unfolding bifinite_compact_iff by simp | |
| 329 | ||
| 330 | ||
| 331 | subsection {* Monadic plus *}
 | |
| 332 | ||
| 333 | definition | |
| 334 | convex_plus :: "'a convex_pd \<rightarrow> 'a convex_pd \<rightarrow> 'a convex_pd" where | |
| 335 | "convex_plus = convex_pd.basis_fun (\<lambda>t. convex_pd.basis_fun (\<lambda>u. | |
| 336 | convex_principal (PDPlus t u)))" | |
| 337 | ||
| 338 | abbreviation | |
| 339 | convex_add :: "'a convex_pd \<Rightarrow> 'a convex_pd \<Rightarrow> 'a convex_pd" | |
| 340 | (infixl "+\<natural>" 65) where | |
| 341 | "xs +\<natural> ys == convex_plus\<cdot>xs\<cdot>ys" | |
| 342 | ||
| 343 | lemma convex_plus_principal [simp]: | |
| 344 | "convex_plus\<cdot>(convex_principal t)\<cdot>(convex_principal u) = | |
| 345 | convex_principal (PDPlus t u)" | |
| 346 | unfolding convex_plus_def | |
| 347 | by (simp add: convex_pd.basis_fun_principal | |
| 348 | convex_pd.basis_fun_mono PDPlus_convex_mono) | |
| 349 | ||
| 350 | lemma approx_convex_plus [simp]: | |
| 351 | "approx n\<cdot>(convex_plus\<cdot>xs\<cdot>ys) = convex_plus\<cdot>(approx n\<cdot>xs)\<cdot>(approx n\<cdot>ys)" | |
| 352 | by (induct xs ys rule: convex_principal_induct2, simp, simp, simp) | |
| 353 | ||
| 354 | lemma convex_plus_commute: "convex_plus\<cdot>xs\<cdot>ys = convex_plus\<cdot>ys\<cdot>xs" | |
| 355 | apply (induct xs ys rule: convex_principal_induct2, simp, simp) | |
| 356 | apply (simp add: PDPlus_commute) | |
| 357 | done | |
| 358 | ||
| 359 | lemma convex_plus_assoc: | |
| 360 | "convex_plus\<cdot>(convex_plus\<cdot>xs\<cdot>ys)\<cdot>zs = convex_plus\<cdot>xs\<cdot>(convex_plus\<cdot>ys\<cdot>zs)" | |
| 361 | apply (induct xs ys arbitrary: zs rule: convex_principal_induct2, simp, simp) | |
| 362 | apply (rule_tac xs=zs in convex_principal_induct, simp) | |
| 363 | apply (simp add: PDPlus_assoc) | |
| 364 | done | |
| 365 | ||
| 366 | lemma convex_plus_absorb: "convex_plus\<cdot>xs\<cdot>xs = xs" | |
| 367 | apply (induct xs rule: convex_principal_induct, simp) | |
| 368 | apply (simp add: PDPlus_absorb) | |
| 369 | done | |
| 370 | ||
| 371 | lemma convex_unit_less_plus_iff [simp]: | |
| 372 | "(convex_unit\<cdot>x \<sqsubseteq> convex_plus\<cdot>ys\<cdot>zs) = | |
| 373 | (convex_unit\<cdot>x \<sqsubseteq> ys \<and> convex_unit\<cdot>x \<sqsubseteq> zs)" | |
| 374 | apply (rule iffI) | |
| 375 | apply (subgoal_tac | |
| 376 | "adm (\<lambda>f. f\<cdot>(convex_unit\<cdot>x) \<sqsubseteq> f\<cdot>ys \<and> f\<cdot>(convex_unit\<cdot>x) \<sqsubseteq> f\<cdot>zs)") | |
| 25925 | 377 | apply (drule admD, rule chain_approx) | 
| 25904 | 378 | apply (drule_tac f="approx i" in monofun_cfun_arg) | 
| 379 | apply (cut_tac x="approx i\<cdot>x" in compact_imp_Rep_compact_basis, simp) | |
| 380 | apply (cut_tac xs="approx i\<cdot>ys" in compact_imp_convex_principal, simp) | |
| 381 | apply (cut_tac xs="approx i\<cdot>zs" in compact_imp_convex_principal, simp) | |
| 382 | apply (clarify, simp) | |
| 383 | apply simp | |
| 384 | apply simp | |
| 385 | apply (erule conjE) | |
| 386 | apply (subst convex_plus_absorb [of "convex_unit\<cdot>x", symmetric]) | |
| 387 | apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) | |
| 388 | done | |
| 389 | ||
| 390 | lemma convex_plus_less_unit_iff [simp]: | |
| 391 | "(convex_plus\<cdot>xs\<cdot>ys \<sqsubseteq> convex_unit\<cdot>z) = | |
| 392 | (xs \<sqsubseteq> convex_unit\<cdot>z \<and> ys \<sqsubseteq> convex_unit\<cdot>z)" | |
| 393 | apply (rule iffI) | |
| 394 | apply (subgoal_tac | |
| 395 | "adm (\<lambda>f. f\<cdot>xs \<sqsubseteq> f\<cdot>(convex_unit\<cdot>z) \<and> f\<cdot>ys \<sqsubseteq> f\<cdot>(convex_unit\<cdot>z))") | |
| 25925 | 396 | apply (drule admD, rule chain_approx) | 
| 25904 | 397 | apply (drule_tac f="approx i" in monofun_cfun_arg) | 
| 398 | apply (cut_tac xs="approx i\<cdot>xs" in compact_imp_convex_principal, simp) | |
| 399 | apply (cut_tac xs="approx i\<cdot>ys" in compact_imp_convex_principal, simp) | |
| 400 | apply (cut_tac x="approx i\<cdot>z" in compact_imp_Rep_compact_basis, simp) | |
| 401 | apply (clarify, simp) | |
| 402 | apply simp | |
| 403 | apply simp | |
| 404 | apply (erule conjE) | |
| 405 | apply (subst convex_plus_absorb [of "convex_unit\<cdot>z", symmetric]) | |
| 406 | apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) | |
| 407 | done | |
| 408 | ||
| 409 | ||
| 410 | subsection {* Induction rules *}
 | |
| 411 | ||
| 412 | lemma convex_pd_induct1: | |
| 413 | assumes P: "adm P" | |
| 414 | assumes unit: "\<And>x. P (convex_unit\<cdot>x)" | |
| 415 | assumes insert: | |
| 416 | "\<And>x ys. \<lbrakk>P (convex_unit\<cdot>x); P ys\<rbrakk> \<Longrightarrow> P (convex_plus\<cdot>(convex_unit\<cdot>x)\<cdot>ys)" | |
| 417 | shows "P (xs::'a convex_pd)" | |
| 418 | apply (induct xs rule: convex_principal_induct, rule P) | |
| 419 | apply (induct_tac t rule: pd_basis_induct1) | |
| 420 | apply (simp only: convex_unit_Rep_compact_basis [symmetric]) | |
| 421 | apply (rule unit) | |
| 422 | apply (simp only: convex_unit_Rep_compact_basis [symmetric] | |
| 423 | convex_plus_principal [symmetric]) | |
| 424 | apply (erule insert [OF unit]) | |
| 425 | done | |
| 426 | ||
| 427 | lemma convex_pd_induct: | |
| 428 | assumes P: "adm P" | |
| 429 | assumes unit: "\<And>x. P (convex_unit\<cdot>x)" | |
| 430 | assumes plus: "\<And>xs ys. \<lbrakk>P xs; P ys\<rbrakk> \<Longrightarrow> P (convex_plus\<cdot>xs\<cdot>ys)" | |
| 431 | shows "P (xs::'a convex_pd)" | |
| 432 | apply (induct xs rule: convex_principal_induct, rule P) | |
| 433 | apply (induct_tac t rule: pd_basis_induct) | |
| 434 | apply (simp only: convex_unit_Rep_compact_basis [symmetric] unit) | |
| 435 | apply (simp only: convex_plus_principal [symmetric] plus) | |
| 436 | done | |
| 437 | ||
| 438 | ||
| 439 | subsection {* Monadic bind *}
 | |
| 440 | ||
| 441 | definition | |
| 442 | convex_bind_basis :: | |
| 443 |   "'a pd_basis \<Rightarrow> ('a \<rightarrow> 'b convex_pd) \<rightarrow> 'b convex_pd" where
 | |
| 444 | "convex_bind_basis = fold_pd | |
| 445 | (\<lambda>a. \<Lambda> f. f\<cdot>(Rep_compact_basis a)) | |
| 446 | (\<lambda>x y. \<Lambda> f. convex_plus\<cdot>(x\<cdot>f)\<cdot>(y\<cdot>f))" | |
| 447 | ||
| 26041 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 haftmann parents: 
25925diff
changeset | 448 | lemma ACI_convex_bind: "ab_semigroup_idem_mult (\<lambda>x y. \<Lambda> f. convex_plus\<cdot>(x\<cdot>f)\<cdot>(y\<cdot>f))" | 
| 25904 | 449 | apply unfold_locales | 
| 26041 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 haftmann parents: 
25925diff
changeset | 450 | apply (simp add: convex_plus_assoc) | 
| 25904 | 451 | apply (simp add: convex_plus_commute) | 
| 452 | apply (simp add: convex_plus_absorb eta_cfun) | |
| 453 | done | |
| 454 | ||
| 455 | lemma convex_bind_basis_simps [simp]: | |
| 456 | "convex_bind_basis (PDUnit a) = | |
| 457 | (\<Lambda> f. f\<cdot>(Rep_compact_basis a))" | |
| 458 | "convex_bind_basis (PDPlus t u) = | |
| 459 | (\<Lambda> f. convex_plus\<cdot>(convex_bind_basis t\<cdot>f)\<cdot>(convex_bind_basis u\<cdot>f))" | |
| 460 | unfolding convex_bind_basis_def | |
| 461 | apply - | |
| 26041 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 haftmann parents: 
25925diff
changeset | 462 | apply (rule ab_semigroup_idem_mult.fold_pd_PDUnit [OF ACI_convex_bind]) | 
| 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 haftmann parents: 
25925diff
changeset | 463 | apply (rule ab_semigroup_idem_mult.fold_pd_PDPlus [OF ACI_convex_bind]) | 
| 25904 | 464 | done | 
| 465 | ||
| 466 | lemma monofun_LAM: | |
| 467 | "\<lbrakk>cont f; cont g; \<And>x. f x \<sqsubseteq> g x\<rbrakk> \<Longrightarrow> (\<Lambda> x. f x) \<sqsubseteq> (\<Lambda> x. g x)" | |
| 468 | by (simp add: expand_cfun_less) | |
| 469 | ||
| 470 | lemma convex_bind_basis_mono: | |
| 471 | "t \<le>\<natural> u \<Longrightarrow> convex_bind_basis t \<sqsubseteq> convex_bind_basis u" | |
| 472 | apply (erule convex_le_induct) | |
| 473 | apply (erule (1) trans_less) | |
| 474 | apply (simp add: monofun_LAM compact_le_def monofun_cfun) | |
| 475 | apply (simp add: monofun_LAM compact_le_def monofun_cfun) | |
| 476 | done | |
| 477 | ||
| 478 | definition | |
| 479 |   convex_bind :: "'a convex_pd \<rightarrow> ('a \<rightarrow> 'b convex_pd) \<rightarrow> 'b convex_pd" where
 | |
| 480 | "convex_bind = convex_pd.basis_fun convex_bind_basis" | |
| 481 | ||
| 482 | lemma convex_bind_principal [simp]: | |
| 483 | "convex_bind\<cdot>(convex_principal t) = convex_bind_basis t" | |
| 484 | unfolding convex_bind_def | |
| 485 | apply (rule convex_pd.basis_fun_principal) | |
| 486 | apply (erule convex_bind_basis_mono) | |
| 487 | done | |
| 488 | ||
| 489 | lemma convex_bind_unit [simp]: | |
| 490 | "convex_bind\<cdot>(convex_unit\<cdot>x)\<cdot>f = f\<cdot>x" | |
| 491 | by (induct x rule: compact_basis_induct, simp, simp) | |
| 492 | ||
| 493 | lemma convex_bind_plus [simp]: | |
| 494 | "convex_bind\<cdot>(convex_plus\<cdot>xs\<cdot>ys)\<cdot>f = | |
| 495 | convex_plus\<cdot>(convex_bind\<cdot>xs\<cdot>f)\<cdot>(convex_bind\<cdot>ys\<cdot>f)" | |
| 496 | by (induct xs ys rule: convex_principal_induct2, simp, simp, simp) | |
| 497 | ||
| 498 | lemma convex_bind_strict [simp]: "convex_bind\<cdot>\<bottom>\<cdot>f = f\<cdot>\<bottom>" | |
| 499 | unfolding convex_unit_strict [symmetric] by (rule convex_bind_unit) | |
| 500 | ||
| 501 | ||
| 502 | subsection {* Map and join *}
 | |
| 503 | ||
| 504 | definition | |
| 505 |   convex_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a convex_pd \<rightarrow> 'b convex_pd" where
 | |
| 506 | "convex_map = (\<Lambda> f xs. convex_bind\<cdot>xs\<cdot>(\<Lambda> x. convex_unit\<cdot>(f\<cdot>x)))" | |
| 507 | ||
| 508 | definition | |
| 509 | convex_join :: "'a convex_pd convex_pd \<rightarrow> 'a convex_pd" where | |
| 510 | "convex_join = (\<Lambda> xss. convex_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))" | |
| 511 | ||
| 512 | lemma convex_map_unit [simp]: | |
| 513 | "convex_map\<cdot>f\<cdot>(convex_unit\<cdot>x) = convex_unit\<cdot>(f\<cdot>x)" | |
| 514 | unfolding convex_map_def by simp | |
| 515 | ||
| 516 | lemma convex_map_plus [simp]: | |
| 517 | "convex_map\<cdot>f\<cdot>(convex_plus\<cdot>xs\<cdot>ys) = | |
| 518 | convex_plus\<cdot>(convex_map\<cdot>f\<cdot>xs)\<cdot>(convex_map\<cdot>f\<cdot>ys)" | |
| 519 | unfolding convex_map_def by simp | |
| 520 | ||
| 521 | lemma convex_join_unit [simp]: | |
| 522 | "convex_join\<cdot>(convex_unit\<cdot>xs) = xs" | |
| 523 | unfolding convex_join_def by simp | |
| 524 | ||
| 525 | lemma convex_join_plus [simp]: | |
| 526 | "convex_join\<cdot>(convex_plus\<cdot>xss\<cdot>yss) = | |
| 527 | convex_plus\<cdot>(convex_join\<cdot>xss)\<cdot>(convex_join\<cdot>yss)" | |
| 528 | unfolding convex_join_def by simp | |
| 529 | ||
| 530 | lemma convex_map_ident: "convex_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs" | |
| 531 | by (induct xs rule: convex_pd_induct, simp_all) | |
| 532 | ||
| 533 | lemma convex_map_map: | |
| 534 | "convex_map\<cdot>f\<cdot>(convex_map\<cdot>g\<cdot>xs) = convex_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs" | |
| 535 | by (induct xs rule: convex_pd_induct, simp_all) | |
| 536 | ||
| 537 | lemma convex_join_map_unit: | |
| 538 | "convex_join\<cdot>(convex_map\<cdot>convex_unit\<cdot>xs) = xs" | |
| 539 | by (induct xs rule: convex_pd_induct, simp_all) | |
| 540 | ||
| 541 | lemma convex_join_map_join: | |
| 542 | "convex_join\<cdot>(convex_map\<cdot>convex_join\<cdot>xsss) = convex_join\<cdot>(convex_join\<cdot>xsss)" | |
| 543 | by (induct xsss rule: convex_pd_induct, simp_all) | |
| 544 | ||
| 545 | lemma convex_join_map_map: | |
| 546 | "convex_join\<cdot>(convex_map\<cdot>(convex_map\<cdot>f)\<cdot>xss) = | |
| 547 | convex_map\<cdot>f\<cdot>(convex_join\<cdot>xss)" | |
| 548 | by (induct xss rule: convex_pd_induct, simp_all) | |
| 549 | ||
| 550 | lemma convex_map_approx: "convex_map\<cdot>(approx n)\<cdot>xs = approx n\<cdot>xs" | |
| 551 | by (induct xs rule: convex_pd_induct, simp_all) | |
| 552 | ||
| 553 | ||
| 554 | subsection {* Conversions to other powerdomains *}
 | |
| 555 | ||
| 556 | text {* Convex to upper *}
 | |
| 557 | ||
| 558 | lemma convex_le_imp_upper_le: "t \<le>\<natural> u \<Longrightarrow> t \<le>\<sharp> u" | |
| 559 | unfolding convex_le_def by simp | |
| 560 | ||
| 561 | definition | |
| 562 | convex_to_upper :: "'a convex_pd \<rightarrow> 'a upper_pd" where | |
| 563 | "convex_to_upper = convex_pd.basis_fun upper_principal" | |
| 564 | ||
| 565 | lemma convex_to_upper_principal [simp]: | |
| 566 | "convex_to_upper\<cdot>(convex_principal t) = upper_principal t" | |
| 567 | unfolding convex_to_upper_def | |
| 568 | apply (rule convex_pd.basis_fun_principal) | |
| 569 | apply (rule upper_principal_mono) | |
| 570 | apply (erule convex_le_imp_upper_le) | |
| 571 | done | |
| 572 | ||
| 573 | lemma convex_to_upper_unit [simp]: | |
| 574 | "convex_to_upper\<cdot>(convex_unit\<cdot>x) = upper_unit\<cdot>x" | |
| 575 | by (induct x rule: compact_basis_induct, simp, simp) | |
| 576 | ||
| 577 | lemma convex_to_upper_plus [simp]: | |
| 578 | "convex_to_upper\<cdot>(convex_plus\<cdot>xs\<cdot>ys) = | |
| 579 | upper_plus\<cdot>(convex_to_upper\<cdot>xs)\<cdot>(convex_to_upper\<cdot>ys)" | |
| 580 | by (induct xs ys rule: convex_principal_induct2, simp, simp, simp) | |
| 581 | ||
| 582 | lemma approx_convex_to_upper: | |
| 583 | "approx i\<cdot>(convex_to_upper\<cdot>xs) = convex_to_upper\<cdot>(approx i\<cdot>xs)" | |
| 584 | by (induct xs rule: convex_pd_induct, simp, simp, simp) | |
| 585 | ||
| 586 | text {* Convex to lower *}
 | |
| 587 | ||
| 588 | lemma convex_le_imp_lower_le: "t \<le>\<natural> u \<Longrightarrow> t \<le>\<flat> u" | |
| 589 | unfolding convex_le_def by simp | |
| 590 | ||
| 591 | definition | |
| 592 | convex_to_lower :: "'a convex_pd \<rightarrow> 'a lower_pd" where | |
| 593 | "convex_to_lower = convex_pd.basis_fun lower_principal" | |
| 594 | ||
| 595 | lemma convex_to_lower_principal [simp]: | |
| 596 | "convex_to_lower\<cdot>(convex_principal t) = lower_principal t" | |
| 597 | unfolding convex_to_lower_def | |
| 598 | apply (rule convex_pd.basis_fun_principal) | |
| 599 | apply (rule lower_principal_mono) | |
| 600 | apply (erule convex_le_imp_lower_le) | |
| 601 | done | |
| 602 | ||
| 603 | lemma convex_to_lower_unit [simp]: | |
| 604 | "convex_to_lower\<cdot>(convex_unit\<cdot>x) = lower_unit\<cdot>x" | |
| 605 | by (induct x rule: compact_basis_induct, simp, simp) | |
| 606 | ||
| 607 | lemma convex_to_lower_plus [simp]: | |
| 608 | "convex_to_lower\<cdot>(convex_plus\<cdot>xs\<cdot>ys) = | |
| 609 | lower_plus\<cdot>(convex_to_lower\<cdot>xs)\<cdot>(convex_to_lower\<cdot>ys)" | |
| 610 | by (induct xs ys rule: convex_principal_induct2, simp, simp, simp) | |
| 611 | ||
| 612 | lemma approx_convex_to_lower: | |
| 613 | "approx i\<cdot>(convex_to_lower\<cdot>xs) = convex_to_lower\<cdot>(approx i\<cdot>xs)" | |
| 614 | by (induct xs rule: convex_pd_induct, simp, simp, simp) | |
| 615 | ||
| 616 | text {* Ordering property *}
 | |
| 617 | ||
| 618 | lemma convex_pd_less_iff: | |
| 619 | "(xs \<sqsubseteq> ys) = | |
| 620 | (convex_to_upper\<cdot>xs \<sqsubseteq> convex_to_upper\<cdot>ys \<and> | |
| 621 | convex_to_lower\<cdot>xs \<sqsubseteq> convex_to_lower\<cdot>ys)" | |
| 622 | apply (safe elim!: monofun_cfun_arg) | |
| 623 | apply (rule bifinite_less_ext) | |
| 624 | apply (drule_tac f="approx i" in monofun_cfun_arg) | |
| 625 | apply (drule_tac f="approx i" in monofun_cfun_arg) | |
| 626 | apply (cut_tac xs="approx i\<cdot>xs" in compact_imp_convex_principal, simp) | |
| 627 | apply (cut_tac xs="approx i\<cdot>ys" in compact_imp_convex_principal, simp) | |
| 628 | apply clarify | |
| 629 | apply (simp add: approx_convex_to_upper approx_convex_to_lower convex_le_def) | |
| 630 | done | |
| 631 | ||
| 632 | end |