| 27404 |      1 | (*  Title:      HOLCF/Completion.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Brian Huffman
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | header {* Defining bifinite domains by ideal completion *}
 | 
|  |      7 | 
 | 
|  |      8 | theory Completion
 | 
|  |      9 | imports Bifinite
 | 
|  |     10 | begin
 | 
|  |     11 | 
 | 
|  |     12 | subsection {* Ideals over a preorder *}
 | 
|  |     13 | 
 | 
|  |     14 | locale preorder =
 | 
|  |     15 |   fixes r :: "'a::type \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<preceq>" 50)
 | 
|  |     16 |   assumes r_refl: "x \<preceq> x"
 | 
|  |     17 |   assumes r_trans: "\<lbrakk>x \<preceq> y; y \<preceq> z\<rbrakk> \<Longrightarrow> x \<preceq> z"
 | 
|  |     18 | begin
 | 
|  |     19 | 
 | 
|  |     20 | definition
 | 
|  |     21 |   ideal :: "'a set \<Rightarrow> bool" where
 | 
|  |     22 |   "ideal A = ((\<exists>x. x \<in> A) \<and> (\<forall>x\<in>A. \<forall>y\<in>A. \<exists>z\<in>A. x \<preceq> z \<and> y \<preceq> z) \<and>
 | 
|  |     23 |     (\<forall>x y. x \<preceq> y \<longrightarrow> y \<in> A \<longrightarrow> x \<in> A))"
 | 
|  |     24 | 
 | 
|  |     25 | lemma idealI:
 | 
|  |     26 |   assumes "\<exists>x. x \<in> A"
 | 
|  |     27 |   assumes "\<And>x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> \<exists>z\<in>A. x \<preceq> z \<and> y \<preceq> z"
 | 
|  |     28 |   assumes "\<And>x y. \<lbrakk>x \<preceq> y; y \<in> A\<rbrakk> \<Longrightarrow> x \<in> A"
 | 
|  |     29 |   shows "ideal A"
 | 
|  |     30 | unfolding ideal_def using prems by fast
 | 
|  |     31 | 
 | 
|  |     32 | lemma idealD1:
 | 
|  |     33 |   "ideal A \<Longrightarrow> \<exists>x. x \<in> A"
 | 
|  |     34 | unfolding ideal_def by fast
 | 
|  |     35 | 
 | 
|  |     36 | lemma idealD2:
 | 
|  |     37 |   "\<lbrakk>ideal A; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> \<exists>z\<in>A. x \<preceq> z \<and> y \<preceq> z"
 | 
|  |     38 | unfolding ideal_def by fast
 | 
|  |     39 | 
 | 
|  |     40 | lemma idealD3:
 | 
|  |     41 |   "\<lbrakk>ideal A; x \<preceq> y; y \<in> A\<rbrakk> \<Longrightarrow> x \<in> A"
 | 
|  |     42 | unfolding ideal_def by fast
 | 
|  |     43 | 
 | 
|  |     44 | lemma ideal_directed_finite:
 | 
|  |     45 |   assumes A: "ideal A"
 | 
|  |     46 |   shows "\<lbrakk>finite U; U \<subseteq> A\<rbrakk> \<Longrightarrow> \<exists>z\<in>A. \<forall>x\<in>U. x \<preceq> z"
 | 
|  |     47 | apply (induct U set: finite)
 | 
|  |     48 | apply (simp add: idealD1 [OF A])
 | 
|  |     49 | apply (simp, clarify, rename_tac y)
 | 
|  |     50 | apply (drule (1) idealD2 [OF A])
 | 
|  |     51 | apply (clarify, erule_tac x=z in rev_bexI)
 | 
|  |     52 | apply (fast intro: r_trans)
 | 
|  |     53 | done
 | 
|  |     54 | 
 | 
|  |     55 | lemma ideal_principal: "ideal {x. x \<preceq> z}"
 | 
|  |     56 | apply (rule idealI)
 | 
|  |     57 | apply (rule_tac x=z in exI)
 | 
|  |     58 | apply (fast intro: r_refl)
 | 
|  |     59 | apply (rule_tac x=z in bexI, fast)
 | 
|  |     60 | apply (fast intro: r_refl)
 | 
|  |     61 | apply (fast intro: r_trans)
 | 
|  |     62 | done
 | 
|  |     63 | 
 | 
|  |     64 | lemma ex_ideal: "\<exists>A. ideal A"
 | 
|  |     65 | by (rule exI, rule ideal_principal)
 | 
|  |     66 | 
 | 
|  |     67 | lemma directed_image_ideal:
 | 
|  |     68 |   assumes A: "ideal A"
 | 
|  |     69 |   assumes f: "\<And>x y. x \<preceq> y \<Longrightarrow> f x \<sqsubseteq> f y"
 | 
|  |     70 |   shows "directed (f ` A)"
 | 
|  |     71 | apply (rule directedI)
 | 
|  |     72 | apply (cut_tac idealD1 [OF A], fast)
 | 
|  |     73 | apply (clarify, rename_tac a b)
 | 
|  |     74 | apply (drule (1) idealD2 [OF A])
 | 
|  |     75 | apply (clarify, rename_tac c)
 | 
|  |     76 | apply (rule_tac x="f c" in rev_bexI)
 | 
|  |     77 | apply (erule imageI)
 | 
|  |     78 | apply (simp add: f)
 | 
|  |     79 | done
 | 
|  |     80 | 
 | 
|  |     81 | lemma lub_image_principal:
 | 
|  |     82 |   assumes f: "\<And>x y. x \<preceq> y \<Longrightarrow> f x \<sqsubseteq> f y"
 | 
|  |     83 |   shows "(\<Squnion>x\<in>{x. x \<preceq> y}. f x) = f y"
 | 
|  |     84 | apply (rule thelubI)
 | 
|  |     85 | apply (rule is_lub_maximal)
 | 
|  |     86 | apply (rule ub_imageI)
 | 
|  |     87 | apply (simp add: f)
 | 
|  |     88 | apply (rule imageI)
 | 
|  |     89 | apply (simp add: r_refl)
 | 
|  |     90 | done
 | 
|  |     91 | 
 | 
|  |     92 | text {* The set of ideals is a cpo *}
 | 
|  |     93 | 
 | 
|  |     94 | lemma ideal_UN:
 | 
|  |     95 |   fixes A :: "nat \<Rightarrow> 'a set"
 | 
|  |     96 |   assumes ideal_A: "\<And>i. ideal (A i)"
 | 
|  |     97 |   assumes chain_A: "\<And>i j. i \<le> j \<Longrightarrow> A i \<subseteq> A j"
 | 
|  |     98 |   shows "ideal (\<Union>i. A i)"
 | 
|  |     99 |  apply (rule idealI)
 | 
|  |    100 |    apply (cut_tac idealD1 [OF ideal_A], fast)
 | 
|  |    101 |   apply (clarify, rename_tac i j)
 | 
|  |    102 |   apply (drule subsetD [OF chain_A [OF le_maxI1]])
 | 
|  |    103 |   apply (drule subsetD [OF chain_A [OF le_maxI2]])
 | 
|  |    104 |   apply (drule (1) idealD2 [OF ideal_A])
 | 
|  |    105 |   apply blast
 | 
|  |    106 |  apply clarify
 | 
|  |    107 |  apply (drule (1) idealD3 [OF ideal_A])
 | 
|  |    108 |  apply fast
 | 
|  |    109 | done
 | 
|  |    110 | 
 | 
|  |    111 | lemma typedef_ideal_po:
 | 
|  |    112 |   fixes Abs :: "'a set \<Rightarrow> 'b::sq_ord"
 | 
|  |    113 |   assumes type: "type_definition Rep Abs {S. ideal S}"
 | 
|  |    114 |   assumes less: "\<And>x y. x \<sqsubseteq> y \<longleftrightarrow> Rep x \<subseteq> Rep y"
 | 
|  |    115 |   shows "OFCLASS('b, po_class)"
 | 
|  |    116 |  apply (intro_classes, unfold less)
 | 
|  |    117 |    apply (rule subset_refl)
 | 
|  |    118 |   apply (erule (1) subset_trans)
 | 
|  |    119 |  apply (rule type_definition.Rep_inject [OF type, THEN iffD1])
 | 
|  |    120 |  apply (erule (1) subset_antisym)
 | 
|  |    121 | done
 | 
|  |    122 | 
 | 
|  |    123 | lemma
 | 
|  |    124 |   fixes Abs :: "'a set \<Rightarrow> 'b::po"
 | 
|  |    125 |   assumes type: "type_definition Rep Abs {S. ideal S}"
 | 
|  |    126 |   assumes less: "\<And>x y. x \<sqsubseteq> y \<longleftrightarrow> Rep x \<subseteq> Rep y"
 | 
|  |    127 |   assumes S: "chain S"
 | 
|  |    128 |   shows typedef_ideal_lub: "range S <<| Abs (\<Union>i. Rep (S i))"
 | 
|  |    129 |     and typedef_ideal_rep_contlub: "Rep (\<Squnion>i. S i) = (\<Union>i. Rep (S i))"
 | 
|  |    130 | proof -
 | 
|  |    131 |   have 1: "ideal (\<Union>i. Rep (S i))"
 | 
|  |    132 |     apply (rule ideal_UN)
 | 
|  |    133 |      apply (rule type_definition.Rep [OF type, unfolded mem_Collect_eq])
 | 
|  |    134 |     apply (subst less [symmetric])
 | 
|  |    135 |     apply (erule chain_mono [OF S])
 | 
|  |    136 |     done
 | 
|  |    137 |   hence 2: "Rep (Abs (\<Union>i. Rep (S i))) = (\<Union>i. Rep (S i))"
 | 
|  |    138 |     by (simp add: type_definition.Abs_inverse [OF type])
 | 
|  |    139 |   show 3: "range S <<| Abs (\<Union>i. Rep (S i))"
 | 
|  |    140 |     apply (rule is_lubI)
 | 
|  |    141 |      apply (rule is_ubI)
 | 
|  |    142 |      apply (simp add: less 2, fast)
 | 
|  |    143 |     apply (simp add: less 2 is_ub_def, fast)
 | 
|  |    144 |     done
 | 
|  |    145 |   hence 4: "(\<Squnion>i. S i) = Abs (\<Union>i. Rep (S i))"
 | 
|  |    146 |     by (rule thelubI)
 | 
|  |    147 |   show 5: "Rep (\<Squnion>i. S i) = (\<Union>i. Rep (S i))"
 | 
|  |    148 |     by (simp add: 4 2)
 | 
|  |    149 | qed
 | 
|  |    150 | 
 | 
|  |    151 | lemma typedef_ideal_cpo:
 | 
|  |    152 |   fixes Abs :: "'a set \<Rightarrow> 'b::po"
 | 
|  |    153 |   assumes type: "type_definition Rep Abs {S. ideal S}"
 | 
|  |    154 |   assumes less: "\<And>x y. x \<sqsubseteq> y \<longleftrightarrow> Rep x \<subseteq> Rep y"
 | 
|  |    155 |   shows "OFCLASS('b, cpo_class)"
 | 
|  |    156 | by (default, rule exI, erule typedef_ideal_lub [OF type less])
 | 
|  |    157 | 
 | 
|  |    158 | end
 | 
|  |    159 | 
 | 
|  |    160 | interpretation sq_le: preorder ["sq_le :: 'a::po \<Rightarrow> 'a \<Rightarrow> bool"]
 | 
|  |    161 | apply unfold_locales
 | 
|  |    162 | apply (rule refl_less)
 | 
|  |    163 | apply (erule (1) trans_less)
 | 
|  |    164 | done
 | 
|  |    165 | 
 | 
|  |    166 | subsection {* Defining functions in terms of basis elements *}
 | 
|  |    167 | 
 | 
|  |    168 | lemma finite_directed_contains_lub:
 | 
|  |    169 |   "\<lbrakk>finite S; directed S\<rbrakk> \<Longrightarrow> \<exists>u\<in>S. S <<| u"
 | 
|  |    170 | apply (drule (1) directed_finiteD, rule subset_refl)
 | 
|  |    171 | apply (erule bexE)
 | 
|  |    172 | apply (rule rev_bexI, assumption)
 | 
|  |    173 | apply (erule (1) is_lub_maximal)
 | 
|  |    174 | done
 | 
|  |    175 | 
 | 
|  |    176 | lemma lub_finite_directed_in_self:
 | 
|  |    177 |   "\<lbrakk>finite S; directed S\<rbrakk> \<Longrightarrow> lub S \<in> S"
 | 
|  |    178 | apply (drule (1) finite_directed_contains_lub, clarify)
 | 
|  |    179 | apply (drule thelubI, simp)
 | 
|  |    180 | done
 | 
|  |    181 | 
 | 
|  |    182 | lemma finite_directed_has_lub: "\<lbrakk>finite S; directed S\<rbrakk> \<Longrightarrow> \<exists>u. S <<| u"
 | 
|  |    183 | by (drule (1) finite_directed_contains_lub, fast)
 | 
|  |    184 | 
 | 
|  |    185 | lemma is_ub_thelub0: "\<lbrakk>\<exists>u. S <<| u; x \<in> S\<rbrakk> \<Longrightarrow> x \<sqsubseteq> lub S"
 | 
|  |    186 | apply (erule exE, drule lubI)
 | 
|  |    187 | apply (drule is_lubD1)
 | 
|  |    188 | apply (erule (1) is_ubD)
 | 
|  |    189 | done
 | 
|  |    190 | 
 | 
|  |    191 | lemma is_lub_thelub0: "\<lbrakk>\<exists>u. S <<| u; S <| x\<rbrakk> \<Longrightarrow> lub S \<sqsubseteq> x"
 | 
|  |    192 | by (erule exE, drule lubI, erule is_lub_lub)
 | 
|  |    193 | 
 | 
|  |    194 | locale basis_take = preorder +
 | 
|  |    195 |   fixes take :: "nat \<Rightarrow> 'a::type \<Rightarrow> 'a"
 | 
|  |    196 |   assumes take_less: "take n a \<preceq> a"
 | 
|  |    197 |   assumes take_take: "take n (take n a) = take n a"
 | 
|  |    198 |   assumes take_mono: "a \<preceq> b \<Longrightarrow> take n a \<preceq> take n b"
 | 
|  |    199 |   assumes take_chain: "take n a \<preceq> take (Suc n) a"
 | 
|  |    200 |   assumes finite_range_take: "finite (range (take n))"
 | 
|  |    201 |   assumes take_covers: "\<exists>n. take n a = a"
 | 
|  |    202 | begin
 | 
|  |    203 | 
 | 
|  |    204 | lemma take_chain_less: "m < n \<Longrightarrow> take m a \<preceq> take n a"
 | 
|  |    205 | by (erule less_Suc_induct, rule take_chain, erule (1) r_trans)
 | 
|  |    206 | 
 | 
|  |    207 | lemma take_chain_le: "m \<le> n \<Longrightarrow> take m a \<preceq> take n a"
 | 
|  |    208 | by (cases "m = n", simp add: r_refl, simp add: take_chain_less)
 | 
|  |    209 | 
 | 
|  |    210 | end
 | 
|  |    211 | 
 | 
|  |    212 | locale ideal_completion = basis_take +
 | 
|  |    213 |   fixes principal :: "'a::type \<Rightarrow> 'b::cpo"
 | 
|  |    214 |   fixes rep :: "'b::cpo \<Rightarrow> 'a::type set"
 | 
|  |    215 |   assumes ideal_rep: "\<And>x. preorder.ideal r (rep x)"
 | 
|  |    216 |   assumes rep_contlub: "\<And>Y. chain Y \<Longrightarrow> rep (\<Squnion>i. Y i) = (\<Union>i. rep (Y i))"
 | 
|  |    217 |   assumes rep_principal: "\<And>a. rep (principal a) = {b. b \<preceq> a}"
 | 
|  |    218 |   assumes subset_repD: "\<And>x y. rep x \<subseteq> rep y \<Longrightarrow> x \<sqsubseteq> y"
 | 
|  |    219 | begin
 | 
|  |    220 | 
 | 
|  |    221 | lemma finite_take_rep: "finite (take n ` rep x)"
 | 
|  |    222 | by (rule finite_subset [OF image_mono [OF subset_UNIV] finite_range_take])
 | 
|  |    223 | 
 | 
|  |    224 | lemma basis_fun_lemma0:
 | 
|  |    225 |   fixes f :: "'a::type \<Rightarrow> 'c::cpo"
 | 
|  |    226 |   assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
 | 
|  |    227 |   shows "\<exists>u. f ` take i ` rep x <<| u"
 | 
|  |    228 | apply (rule finite_directed_has_lub)
 | 
|  |    229 | apply (rule finite_imageI)
 | 
|  |    230 | apply (rule finite_take_rep)
 | 
|  |    231 | apply (subst image_image)
 | 
|  |    232 | apply (rule directed_image_ideal)
 | 
|  |    233 | apply (rule ideal_rep)
 | 
|  |    234 | apply (rule f_mono)
 | 
|  |    235 | apply (erule take_mono)
 | 
|  |    236 | done
 | 
|  |    237 | 
 | 
|  |    238 | lemma basis_fun_lemma1:
 | 
|  |    239 |   fixes f :: "'a::type \<Rightarrow> 'c::cpo"
 | 
|  |    240 |   assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
 | 
|  |    241 |   shows "chain (\<lambda>i. lub (f ` take i ` rep x))"
 | 
|  |    242 |  apply (rule chainI)
 | 
|  |    243 |  apply (rule is_lub_thelub0)
 | 
|  |    244 |   apply (rule basis_fun_lemma0, erule f_mono)
 | 
|  |    245 |  apply (rule is_ubI, clarsimp, rename_tac a)
 | 
|  |    246 |  apply (rule sq_le.trans_less [OF f_mono [OF take_chain]])
 | 
|  |    247 |  apply (rule is_ub_thelub0)
 | 
|  |    248 |   apply (rule basis_fun_lemma0, erule f_mono)
 | 
|  |    249 |  apply simp
 | 
|  |    250 | done
 | 
|  |    251 | 
 | 
|  |    252 | lemma basis_fun_lemma2:
 | 
|  |    253 |   fixes f :: "'a::type \<Rightarrow> 'c::cpo"
 | 
|  |    254 |   assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
 | 
|  |    255 |   shows "f ` rep x <<| (\<Squnion>i. lub (f ` take i ` rep x))"
 | 
|  |    256 |  apply (rule is_lubI)
 | 
|  |    257 |  apply (rule ub_imageI, rename_tac a)
 | 
|  |    258 |   apply (cut_tac a=a in take_covers, erule exE, rename_tac i)
 | 
|  |    259 |   apply (erule subst)
 | 
|  |    260 |   apply (rule rev_trans_less)
 | 
|  |    261 |    apply (rule_tac x=i in is_ub_thelub)
 | 
|  |    262 |    apply (rule basis_fun_lemma1, erule f_mono)
 | 
|  |    263 |   apply (rule is_ub_thelub0)
 | 
|  |    264 |    apply (rule basis_fun_lemma0, erule f_mono)
 | 
|  |    265 |   apply simp
 | 
|  |    266 |  apply (rule is_lub_thelub [OF _ ub_rangeI])
 | 
|  |    267 |   apply (rule basis_fun_lemma1, erule f_mono)
 | 
|  |    268 |  apply (rule is_lub_thelub0)
 | 
|  |    269 |   apply (rule basis_fun_lemma0, erule f_mono)
 | 
|  |    270 |  apply (rule is_ubI, clarsimp, rename_tac a)
 | 
|  |    271 |  apply (rule sq_le.trans_less [OF f_mono [OF take_less]])
 | 
|  |    272 |  apply (erule (1) ub_imageD)
 | 
|  |    273 | done
 | 
|  |    274 | 
 | 
|  |    275 | lemma basis_fun_lemma:
 | 
|  |    276 |   fixes f :: "'a::type \<Rightarrow> 'c::cpo"
 | 
|  |    277 |   assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
 | 
|  |    278 |   shows "\<exists>u. f ` rep x <<| u"
 | 
|  |    279 | by (rule exI, rule basis_fun_lemma2, erule f_mono)
 | 
|  |    280 | 
 | 
|  |    281 | lemma rep_mono: "x \<sqsubseteq> y \<Longrightarrow> rep x \<subseteq> rep y"
 | 
|  |    282 | apply (frule bin_chain)
 | 
|  |    283 | apply (drule rep_contlub)
 | 
|  |    284 | apply (simp only: thelubI [OF lub_bin_chain])
 | 
|  |    285 | apply (rule subsetI, rule UN_I [where a=0], simp_all)
 | 
|  |    286 | done
 | 
|  |    287 | 
 | 
|  |    288 | lemma less_def: "x \<sqsubseteq> y \<longleftrightarrow> rep x \<subseteq> rep y"
 | 
|  |    289 | by (rule iffI [OF rep_mono subset_repD])
 | 
|  |    290 | 
 | 
|  |    291 | lemma rep_eq: "rep x = {a. principal a \<sqsubseteq> x}"
 | 
|  |    292 | unfolding less_def rep_principal
 | 
|  |    293 | apply safe
 | 
|  |    294 | apply (erule (1) idealD3 [OF ideal_rep])
 | 
|  |    295 | apply (erule subsetD, simp add: r_refl)
 | 
|  |    296 | done
 | 
|  |    297 | 
 | 
|  |    298 | lemma mem_rep_iff_principal_less: "a \<in> rep x \<longleftrightarrow> principal a \<sqsubseteq> x"
 | 
|  |    299 | by (simp add: rep_eq)
 | 
|  |    300 | 
 | 
|  |    301 | lemma principal_less_iff_mem_rep: "principal a \<sqsubseteq> x \<longleftrightarrow> a \<in> rep x"
 | 
|  |    302 | by (simp add: rep_eq)
 | 
|  |    303 | 
 | 
|  |    304 | lemma principal_less_iff [simp]: "principal a \<sqsubseteq> principal b \<longleftrightarrow> a \<preceq> b"
 | 
|  |    305 | by (simp add: principal_less_iff_mem_rep rep_principal)
 | 
|  |    306 | 
 | 
|  |    307 | lemma principal_eq_iff: "principal a = principal b \<longleftrightarrow> a \<preceq> b \<and> b \<preceq> a"
 | 
|  |    308 | unfolding po_eq_conv [where 'a='b] principal_less_iff ..
 | 
|  |    309 | 
 | 
|  |    310 | lemma repD: "a \<in> rep x \<Longrightarrow> principal a \<sqsubseteq> x"
 | 
|  |    311 | by (simp add: rep_eq)
 | 
|  |    312 | 
 | 
|  |    313 | lemma principal_mono: "a \<preceq> b \<Longrightarrow> principal a \<sqsubseteq> principal b"
 | 
|  |    314 | by (simp only: principal_less_iff)
 | 
|  |    315 | 
 | 
|  |    316 | lemma lessI: "(\<And>a. principal a \<sqsubseteq> x \<Longrightarrow> principal a \<sqsubseteq> u) \<Longrightarrow> x \<sqsubseteq> u"
 | 
|  |    317 | unfolding principal_less_iff_mem_rep
 | 
|  |    318 | by (simp add: less_def subset_eq)
 | 
|  |    319 | 
 | 
|  |    320 | lemma lub_principal_rep: "principal ` rep x <<| x"
 | 
|  |    321 | apply (rule is_lubI)
 | 
|  |    322 | apply (rule ub_imageI)
 | 
|  |    323 | apply (erule repD)
 | 
|  |    324 | apply (subst less_def)
 | 
|  |    325 | apply (rule subsetI)
 | 
|  |    326 | apply (drule (1) ub_imageD)
 | 
|  |    327 | apply (simp add: rep_eq)
 | 
|  |    328 | done
 | 
|  |    329 | 
 | 
|  |    330 | definition
 | 
|  |    331 |   basis_fun :: "('a::type \<Rightarrow> 'c::cpo) \<Rightarrow> 'b \<rightarrow> 'c" where
 | 
|  |    332 |   "basis_fun = (\<lambda>f. (\<Lambda> x. lub (f ` rep x)))"
 | 
|  |    333 | 
 | 
|  |    334 | lemma basis_fun_beta:
 | 
|  |    335 |   fixes f :: "'a::type \<Rightarrow> 'c::cpo"
 | 
|  |    336 |   assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
 | 
|  |    337 |   shows "basis_fun f\<cdot>x = lub (f ` rep x)"
 | 
|  |    338 | unfolding basis_fun_def
 | 
|  |    339 | proof (rule beta_cfun)
 | 
|  |    340 |   have lub: "\<And>x. \<exists>u. f ` rep x <<| u"
 | 
|  |    341 |     using f_mono by (rule basis_fun_lemma)
 | 
|  |    342 |   show cont: "cont (\<lambda>x. lub (f ` rep x))"
 | 
|  |    343 |     apply (rule contI2)
 | 
|  |    344 |      apply (rule monofunI)
 | 
|  |    345 |      apply (rule is_lub_thelub0 [OF lub ub_imageI])
 | 
|  |    346 |      apply (rule is_ub_thelub0 [OF lub imageI])
 | 
|  |    347 |      apply (erule (1) subsetD [OF rep_mono])
 | 
|  |    348 |     apply (rule is_lub_thelub0 [OF lub ub_imageI])
 | 
|  |    349 |     apply (simp add: rep_contlub, clarify)
 | 
|  |    350 |     apply (erule rev_trans_less [OF is_ub_thelub])
 | 
|  |    351 |     apply (erule is_ub_thelub0 [OF lub imageI])
 | 
|  |    352 |     done
 | 
|  |    353 | qed
 | 
|  |    354 | 
 | 
|  |    355 | lemma basis_fun_principal:
 | 
|  |    356 |   fixes f :: "'a::type \<Rightarrow> 'c::cpo"
 | 
|  |    357 |   assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
 | 
|  |    358 |   shows "basis_fun f\<cdot>(principal a) = f a"
 | 
|  |    359 | apply (subst basis_fun_beta, erule f_mono)
 | 
|  |    360 | apply (subst rep_principal)
 | 
|  |    361 | apply (rule lub_image_principal, erule f_mono)
 | 
|  |    362 | done
 | 
|  |    363 | 
 | 
|  |    364 | lemma basis_fun_mono:
 | 
|  |    365 |   assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
 | 
|  |    366 |   assumes g_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> g a \<sqsubseteq> g b"
 | 
|  |    367 |   assumes less: "\<And>a. f a \<sqsubseteq> g a"
 | 
|  |    368 |   shows "basis_fun f \<sqsubseteq> basis_fun g"
 | 
|  |    369 |  apply (rule less_cfun_ext)
 | 
|  |    370 |  apply (simp only: basis_fun_beta f_mono g_mono)
 | 
|  |    371 |  apply (rule is_lub_thelub0)
 | 
|  |    372 |   apply (rule basis_fun_lemma, erule f_mono)
 | 
|  |    373 |  apply (rule ub_imageI, rename_tac a)
 | 
|  |    374 |  apply (rule sq_le.trans_less [OF less])
 | 
|  |    375 |  apply (rule is_ub_thelub0)
 | 
|  |    376 |   apply (rule basis_fun_lemma, erule g_mono)
 | 
|  |    377 |  apply (erule imageI)
 | 
|  |    378 | done
 | 
|  |    379 | 
 | 
|  |    380 | lemma compact_principal [simp]: "compact (principal a)"
 | 
|  |    381 | by (rule compactI2, simp add: principal_less_iff_mem_rep rep_contlub)
 | 
|  |    382 | 
 | 
|  |    383 | definition
 | 
|  |    384 |   completion_approx :: "nat \<Rightarrow> 'b \<rightarrow> 'b" where
 | 
|  |    385 |   "completion_approx = (\<lambda>i. basis_fun (\<lambda>a. principal (take i a)))"
 | 
|  |    386 | 
 | 
|  |    387 | lemma completion_approx_beta:
 | 
|  |    388 |   "completion_approx i\<cdot>x = (\<Squnion>a\<in>rep x. principal (take i a))"
 | 
|  |    389 | unfolding completion_approx_def
 | 
|  |    390 | by (simp add: basis_fun_beta principal_mono take_mono)
 | 
|  |    391 | 
 | 
|  |    392 | lemma completion_approx_principal:
 | 
|  |    393 |   "completion_approx i\<cdot>(principal a) = principal (take i a)"
 | 
|  |    394 | unfolding completion_approx_def
 | 
|  |    395 | by (simp add: basis_fun_principal principal_mono take_mono)
 | 
|  |    396 | 
 | 
|  |    397 | lemma chain_completion_approx: "chain completion_approx"
 | 
|  |    398 | unfolding completion_approx_def
 | 
|  |    399 | apply (rule chainI)
 | 
|  |    400 | apply (rule basis_fun_mono)
 | 
|  |    401 | apply (erule principal_mono [OF take_mono])
 | 
|  |    402 | apply (erule principal_mono [OF take_mono])
 | 
|  |    403 | apply (rule principal_mono [OF take_chain])
 | 
|  |    404 | done
 | 
|  |    405 | 
 | 
|  |    406 | lemma lub_completion_approx: "(\<Squnion>i. completion_approx i\<cdot>x) = x"
 | 
|  |    407 | unfolding completion_approx_beta
 | 
|  |    408 |  apply (subst image_image [where f=principal, symmetric])
 | 
|  |    409 |  apply (rule unique_lub [OF _ lub_principal_rep])
 | 
|  |    410 |  apply (rule basis_fun_lemma2, erule principal_mono)
 | 
|  |    411 | done
 | 
|  |    412 | 
 | 
|  |    413 | lemma completion_approx_eq_principal:
 | 
|  |    414 |   "\<exists>a\<in>rep x. completion_approx i\<cdot>x = principal (take i a)"
 | 
|  |    415 | unfolding completion_approx_beta
 | 
|  |    416 |  apply (subst image_image [where f=principal, symmetric])
 | 
|  |    417 |  apply (subgoal_tac "finite (principal ` take i ` rep x)")
 | 
|  |    418 |   apply (subgoal_tac "directed (principal ` take i ` rep x)")
 | 
|  |    419 |    apply (drule (1) lub_finite_directed_in_self, fast)
 | 
|  |    420 |   apply (subst image_image)
 | 
|  |    421 |   apply (rule directed_image_ideal)
 | 
|  |    422 |    apply (rule ideal_rep)
 | 
|  |    423 |   apply (erule principal_mono [OF take_mono])
 | 
|  |    424 |  apply (rule finite_imageI)
 | 
|  |    425 |  apply (rule finite_take_rep)
 | 
|  |    426 | done
 | 
|  |    427 | 
 | 
|  |    428 | lemma completion_approx_idem:
 | 
|  |    429 |   "completion_approx i\<cdot>(completion_approx i\<cdot>x) = completion_approx i\<cdot>x"
 | 
|  |    430 | using completion_approx_eq_principal [where i=i and x=x]
 | 
|  |    431 | by (auto simp add: completion_approx_principal take_take)
 | 
|  |    432 | 
 | 
|  |    433 | lemma finite_fixes_completion_approx:
 | 
|  |    434 |   "finite {x. completion_approx i\<cdot>x = x}" (is "finite ?S")
 | 
|  |    435 | apply (subgoal_tac "?S \<subseteq> principal ` range (take i)")
 | 
|  |    436 | apply (erule finite_subset)
 | 
|  |    437 | apply (rule finite_imageI)
 | 
|  |    438 | apply (rule finite_range_take)
 | 
|  |    439 | apply (clarify, erule subst)
 | 
|  |    440 | apply (cut_tac x=x and i=i in completion_approx_eq_principal)
 | 
|  |    441 | apply fast
 | 
|  |    442 | done
 | 
|  |    443 | 
 | 
|  |    444 | lemma principal_induct:
 | 
|  |    445 |   assumes adm: "adm P"
 | 
|  |    446 |   assumes P: "\<And>a. P (principal a)"
 | 
|  |    447 |   shows "P x"
 | 
|  |    448 |  apply (subgoal_tac "P (\<Squnion>i. completion_approx i\<cdot>x)")
 | 
|  |    449 |  apply (simp add: lub_completion_approx)
 | 
|  |    450 |  apply (rule admD [OF adm])
 | 
|  |    451 |   apply (simp add: chain_completion_approx)
 | 
|  |    452 |  apply (cut_tac x=x and i=i in completion_approx_eq_principal)
 | 
|  |    453 |  apply (clarify, simp add: P)
 | 
|  |    454 | done
 | 
|  |    455 | 
 | 
|  |    456 | lemma principal_induct2:
 | 
|  |    457 |   "\<lbrakk>\<And>y. adm (\<lambda>x. P x y); \<And>x. adm (\<lambda>y. P x y);
 | 
|  |    458 |     \<And>a b. P (principal a) (principal b)\<rbrakk> \<Longrightarrow> P x y"
 | 
|  |    459 | apply (rule_tac x=y in spec)
 | 
|  |    460 | apply (rule_tac x=x in principal_induct, simp)
 | 
|  |    461 | apply (rule allI, rename_tac y)
 | 
|  |    462 | apply (rule_tac x=y in principal_induct, simp)
 | 
|  |    463 | apply simp
 | 
|  |    464 | done
 | 
|  |    465 | 
 | 
|  |    466 | lemma compact_imp_principal: "compact x \<Longrightarrow> \<exists>a. x = principal a"
 | 
|  |    467 | apply (drule adm_compact_neq [OF _ cont_id])
 | 
|  |    468 | apply (drule admD2 [where Y="\<lambda>n. completion_approx n\<cdot>x"])
 | 
|  |    469 | apply (simp add: chain_completion_approx)
 | 
|  |    470 | apply (simp add: lub_completion_approx)
 | 
|  |    471 | apply (erule exE, erule ssubst)
 | 
|  |    472 | apply (cut_tac i=i and x=x in completion_approx_eq_principal)
 | 
|  |    473 | apply (clarify, erule exI)
 | 
|  |    474 | done
 | 
|  |    475 | 
 | 
|  |    476 | end
 | 
|  |    477 | 
 | 
|  |    478 | end
 |