src/HOLCF/Universal.thy
changeset 40774 0437dbc127b3
parent 40773 6c12f5e24e34
child 40775 ed7a4eadb2f6
equal deleted inserted replaced
40773:6c12f5e24e34 40774:0437dbc127b3
     1 (*  Title:      HOLCF/Universal.thy
       
     2     Author:     Brian Huffman
       
     3 *)
       
     4 
       
     5 header {* A universal bifinite domain *}
       
     6 
       
     7 theory Universal
       
     8 imports Completion Deflation Nat_Bijection
       
     9 begin
       
    10 
       
    11 subsection {* Basis for universal domain *}
       
    12 
       
    13 subsubsection {* Basis datatype *}
       
    14 
       
    15 types ubasis = nat
       
    16 
       
    17 definition
       
    18   node :: "nat \<Rightarrow> ubasis \<Rightarrow> ubasis set \<Rightarrow> ubasis"
       
    19 where
       
    20   "node i a S = Suc (prod_encode (i, prod_encode (a, set_encode S)))"
       
    21 
       
    22 lemma node_not_0 [simp]: "node i a S \<noteq> 0"
       
    23 unfolding node_def by simp
       
    24 
       
    25 lemma node_gt_0 [simp]: "0 < node i a S"
       
    26 unfolding node_def by simp
       
    27 
       
    28 lemma node_inject [simp]:
       
    29   "\<lbrakk>finite S; finite T\<rbrakk>
       
    30     \<Longrightarrow> node i a S = node j b T \<longleftrightarrow> i = j \<and> a = b \<and> S = T"
       
    31 unfolding node_def by (simp add: prod_encode_eq set_encode_eq)
       
    32 
       
    33 lemma node_gt0: "i < node i a S"
       
    34 unfolding node_def less_Suc_eq_le
       
    35 by (rule le_prod_encode_1)
       
    36 
       
    37 lemma node_gt1: "a < node i a S"
       
    38 unfolding node_def less_Suc_eq_le
       
    39 by (rule order_trans [OF le_prod_encode_1 le_prod_encode_2])
       
    40 
       
    41 lemma nat_less_power2: "n < 2^n"
       
    42 by (induct n) simp_all
       
    43 
       
    44 lemma node_gt2: "\<lbrakk>finite S; b \<in> S\<rbrakk> \<Longrightarrow> b < node i a S"
       
    45 unfolding node_def less_Suc_eq_le set_encode_def
       
    46 apply (rule order_trans [OF _ le_prod_encode_2])
       
    47 apply (rule order_trans [OF _ le_prod_encode_2])
       
    48 apply (rule order_trans [where y="setsum (op ^ 2) {b}"])
       
    49 apply (simp add: nat_less_power2 [THEN order_less_imp_le])
       
    50 apply (erule setsum_mono2, simp, simp)
       
    51 done
       
    52 
       
    53 lemma eq_prod_encode_pairI:
       
    54   "\<lbrakk>fst (prod_decode x) = a; snd (prod_decode x) = b\<rbrakk> \<Longrightarrow> x = prod_encode (a, b)"
       
    55 by (erule subst, erule subst, simp)
       
    56 
       
    57 lemma node_cases:
       
    58   assumes 1: "x = 0 \<Longrightarrow> P"
       
    59   assumes 2: "\<And>i a S. \<lbrakk>finite S; x = node i a S\<rbrakk> \<Longrightarrow> P"
       
    60   shows "P"
       
    61  apply (cases x)
       
    62   apply (erule 1)
       
    63  apply (rule 2)
       
    64   apply (rule finite_set_decode)
       
    65  apply (simp add: node_def)
       
    66  apply (rule eq_prod_encode_pairI [OF refl])
       
    67  apply (rule eq_prod_encode_pairI [OF refl refl])
       
    68 done
       
    69 
       
    70 lemma node_induct:
       
    71   assumes 1: "P 0"
       
    72   assumes 2: "\<And>i a S. \<lbrakk>P a; finite S; \<forall>b\<in>S. P b\<rbrakk> \<Longrightarrow> P (node i a S)"
       
    73   shows "P x"
       
    74  apply (induct x rule: nat_less_induct)
       
    75  apply (case_tac n rule: node_cases)
       
    76   apply (simp add: 1)
       
    77  apply (simp add: 2 node_gt1 node_gt2)
       
    78 done
       
    79 
       
    80 subsubsection {* Basis ordering *}
       
    81 
       
    82 inductive
       
    83   ubasis_le :: "nat \<Rightarrow> nat \<Rightarrow> bool"
       
    84 where
       
    85   ubasis_le_refl: "ubasis_le a a"
       
    86 | ubasis_le_trans:
       
    87     "\<lbrakk>ubasis_le a b; ubasis_le b c\<rbrakk> \<Longrightarrow> ubasis_le a c"
       
    88 | ubasis_le_lower:
       
    89     "finite S \<Longrightarrow> ubasis_le a (node i a S)"
       
    90 | ubasis_le_upper:
       
    91     "\<lbrakk>finite S; b \<in> S; ubasis_le a b\<rbrakk> \<Longrightarrow> ubasis_le (node i a S) b"
       
    92 
       
    93 lemma ubasis_le_minimal: "ubasis_le 0 x"
       
    94 apply (induct x rule: node_induct)
       
    95 apply (rule ubasis_le_refl)
       
    96 apply (erule ubasis_le_trans)
       
    97 apply (erule ubasis_le_lower)
       
    98 done
       
    99 
       
   100 interpretation udom: preorder ubasis_le
       
   101 apply default
       
   102 apply (rule ubasis_le_refl)
       
   103 apply (erule (1) ubasis_le_trans)
       
   104 done
       
   105 
       
   106 subsubsection {* Generic take function *}
       
   107 
       
   108 function
       
   109   ubasis_until :: "(ubasis \<Rightarrow> bool) \<Rightarrow> ubasis \<Rightarrow> ubasis"
       
   110 where
       
   111   "ubasis_until P 0 = 0"
       
   112 | "finite S \<Longrightarrow> ubasis_until P (node i a S) =
       
   113     (if P (node i a S) then node i a S else ubasis_until P a)"
       
   114     apply clarify
       
   115     apply (rule_tac x=b in node_cases)
       
   116      apply simp
       
   117     apply simp
       
   118     apply fast
       
   119    apply simp
       
   120   apply simp
       
   121  apply simp
       
   122 done
       
   123 
       
   124 termination ubasis_until
       
   125 apply (relation "measure snd")
       
   126 apply (rule wf_measure)
       
   127 apply (simp add: node_gt1)
       
   128 done
       
   129 
       
   130 lemma ubasis_until: "P 0 \<Longrightarrow> P (ubasis_until P x)"
       
   131 by (induct x rule: node_induct) simp_all
       
   132 
       
   133 lemma ubasis_until': "0 < ubasis_until P x \<Longrightarrow> P (ubasis_until P x)"
       
   134 by (induct x rule: node_induct) auto
       
   135 
       
   136 lemma ubasis_until_same: "P x \<Longrightarrow> ubasis_until P x = x"
       
   137 by (induct x rule: node_induct) simp_all
       
   138 
       
   139 lemma ubasis_until_idem:
       
   140   "P 0 \<Longrightarrow> ubasis_until P (ubasis_until P x) = ubasis_until P x"
       
   141 by (rule ubasis_until_same [OF ubasis_until])
       
   142 
       
   143 lemma ubasis_until_0:
       
   144   "\<forall>x. x \<noteq> 0 \<longrightarrow> \<not> P x \<Longrightarrow> ubasis_until P x = 0"
       
   145 by (induct x rule: node_induct) simp_all
       
   146 
       
   147 lemma ubasis_until_less: "ubasis_le (ubasis_until P x) x"
       
   148 apply (induct x rule: node_induct)
       
   149 apply (simp add: ubasis_le_refl)
       
   150 apply (simp add: ubasis_le_refl)
       
   151 apply (rule impI)
       
   152 apply (erule ubasis_le_trans)
       
   153 apply (erule ubasis_le_lower)
       
   154 done
       
   155 
       
   156 lemma ubasis_until_chain:
       
   157   assumes PQ: "\<And>x. P x \<Longrightarrow> Q x"
       
   158   shows "ubasis_le (ubasis_until P x) (ubasis_until Q x)"
       
   159 apply (induct x rule: node_induct)
       
   160 apply (simp add: ubasis_le_refl)
       
   161 apply (simp add: ubasis_le_refl)
       
   162 apply (simp add: PQ)
       
   163 apply clarify
       
   164 apply (rule ubasis_le_trans)
       
   165 apply (rule ubasis_until_less)
       
   166 apply (erule ubasis_le_lower)
       
   167 done
       
   168 
       
   169 lemma ubasis_until_mono:
       
   170   assumes "\<And>i a S b. \<lbrakk>finite S; P (node i a S); b \<in> S; ubasis_le a b\<rbrakk> \<Longrightarrow> P b"
       
   171   shows "ubasis_le a b \<Longrightarrow> ubasis_le (ubasis_until P a) (ubasis_until P b)"
       
   172 proof (induct set: ubasis_le)
       
   173   case (ubasis_le_refl a) show ?case by (rule ubasis_le.ubasis_le_refl)
       
   174 next
       
   175   case (ubasis_le_trans a b c) thus ?case by - (rule ubasis_le.ubasis_le_trans)
       
   176 next
       
   177   case (ubasis_le_lower S a i) thus ?case
       
   178     apply (clarsimp simp add: ubasis_le_refl)
       
   179     apply (rule ubasis_le_trans [OF ubasis_until_less])
       
   180     apply (erule ubasis_le.ubasis_le_lower)
       
   181     done
       
   182 next
       
   183   case (ubasis_le_upper S b a i) thus ?case
       
   184     apply clarsimp
       
   185     apply (subst ubasis_until_same)
       
   186      apply (erule (3) prems)
       
   187     apply (erule (2) ubasis_le.ubasis_le_upper)
       
   188     done
       
   189 qed
       
   190 
       
   191 lemma finite_range_ubasis_until:
       
   192   "finite {x. P x} \<Longrightarrow> finite (range (ubasis_until P))"
       
   193 apply (rule finite_subset [where B="insert 0 {x. P x}"])
       
   194 apply (clarsimp simp add: ubasis_until')
       
   195 apply simp
       
   196 done
       
   197 
       
   198 
       
   199 subsection {* Defining the universal domain by ideal completion *}
       
   200 
       
   201 typedef (open) udom = "{S. udom.ideal S}"
       
   202 by (fast intro: udom.ideal_principal)
       
   203 
       
   204 instantiation udom :: below
       
   205 begin
       
   206 
       
   207 definition
       
   208   "x \<sqsubseteq> y \<longleftrightarrow> Rep_udom x \<subseteq> Rep_udom y"
       
   209 
       
   210 instance ..
       
   211 end
       
   212 
       
   213 instance udom :: po
       
   214 using type_definition_udom below_udom_def
       
   215 by (rule udom.typedef_ideal_po)
       
   216 
       
   217 instance udom :: cpo
       
   218 using type_definition_udom below_udom_def
       
   219 by (rule udom.typedef_ideal_cpo)
       
   220 
       
   221 definition
       
   222   udom_principal :: "nat \<Rightarrow> udom" where
       
   223   "udom_principal t = Abs_udom {u. ubasis_le u t}"
       
   224 
       
   225 lemma ubasis_countable: "\<exists>f::ubasis \<Rightarrow> nat. inj f"
       
   226 by (rule exI, rule inj_on_id)
       
   227 
       
   228 interpretation udom:
       
   229   ideal_completion ubasis_le udom_principal Rep_udom
       
   230 using type_definition_udom below_udom_def
       
   231 using udom_principal_def ubasis_countable
       
   232 by (rule udom.typedef_ideal_completion)
       
   233 
       
   234 text {* Universal domain is pointed *}
       
   235 
       
   236 lemma udom_minimal: "udom_principal 0 \<sqsubseteq> x"
       
   237 apply (induct x rule: udom.principal_induct)
       
   238 apply (simp, simp add: ubasis_le_minimal)
       
   239 done
       
   240 
       
   241 instance udom :: pcpo
       
   242 by intro_classes (fast intro: udom_minimal)
       
   243 
       
   244 lemma inst_udom_pcpo: "\<bottom> = udom_principal 0"
       
   245 by (rule udom_minimal [THEN UU_I, symmetric])
       
   246 
       
   247 
       
   248 subsection {* Compact bases of domains *}
       
   249 
       
   250 typedef (open) 'a compact_basis = "{x::'a::pcpo. compact x}"
       
   251 by auto
       
   252 
       
   253 lemma compact_Rep_compact_basis: "compact (Rep_compact_basis a)"
       
   254 by (rule Rep_compact_basis [unfolded mem_Collect_eq])
       
   255 
       
   256 instantiation compact_basis :: (pcpo) below
       
   257 begin
       
   258 
       
   259 definition
       
   260   compact_le_def:
       
   261     "(op \<sqsubseteq>) \<equiv> (\<lambda>x y. Rep_compact_basis x \<sqsubseteq> Rep_compact_basis y)"
       
   262 
       
   263 instance ..
       
   264 end
       
   265 
       
   266 instance compact_basis :: (pcpo) po
       
   267 using type_definition_compact_basis compact_le_def
       
   268 by (rule typedef_po)
       
   269 
       
   270 definition
       
   271   approximants :: "'a \<Rightarrow> 'a compact_basis set" where
       
   272   "approximants = (\<lambda>x. {a. Rep_compact_basis a \<sqsubseteq> x})"
       
   273 
       
   274 definition
       
   275   compact_bot :: "'a::pcpo compact_basis" where
       
   276   "compact_bot = Abs_compact_basis \<bottom>"
       
   277 
       
   278 lemma Rep_compact_bot [simp]: "Rep_compact_basis compact_bot = \<bottom>"
       
   279 unfolding compact_bot_def by (simp add: Abs_compact_basis_inverse)
       
   280 
       
   281 lemma compact_bot_minimal [simp]: "compact_bot \<sqsubseteq> a"
       
   282 unfolding compact_le_def Rep_compact_bot by simp
       
   283 
       
   284 
       
   285 subsection {* Universality of \emph{udom} *}
       
   286 
       
   287 text {* We use a locale to parameterize the construction over a chain
       
   288 of approx functions on the type to be embedded. *}
       
   289 
       
   290 locale approx_chain =
       
   291   fixes approx :: "nat \<Rightarrow> 'a::pcpo \<rightarrow> 'a"
       
   292   assumes chain_approx [simp]: "chain (\<lambda>i. approx i)"
       
   293   assumes lub_approx [simp]: "(\<Squnion>i. approx i) = ID"
       
   294   assumes finite_deflation_approx: "\<And>i. finite_deflation (approx i)"
       
   295 begin
       
   296 
       
   297 subsubsection {* Choosing a maximal element from a finite set *}
       
   298 
       
   299 lemma finite_has_maximal:
       
   300   fixes A :: "'a compact_basis set"
       
   301   shows "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y"
       
   302 proof (induct rule: finite_ne_induct)
       
   303   case (singleton x)
       
   304     show ?case by simp
       
   305 next
       
   306   case (insert a A)
       
   307   from `\<exists>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y`
       
   308   obtain x where x: "x \<in> A"
       
   309            and x_eq: "\<And>y. \<lbrakk>y \<in> A; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> x = y" by fast
       
   310   show ?case
       
   311   proof (intro bexI ballI impI)
       
   312     fix y
       
   313     assume "y \<in> insert a A" and "(if x \<sqsubseteq> a then a else x) \<sqsubseteq> y"
       
   314     thus "(if x \<sqsubseteq> a then a else x) = y"
       
   315       apply auto
       
   316       apply (frule (1) below_trans)
       
   317       apply (frule (1) x_eq)
       
   318       apply (rule below_antisym, assumption)
       
   319       apply simp
       
   320       apply (erule (1) x_eq)
       
   321       done
       
   322   next
       
   323     show "(if x \<sqsubseteq> a then a else x) \<in> insert a A"
       
   324       by (simp add: x)
       
   325   qed
       
   326 qed
       
   327 
       
   328 definition
       
   329   choose :: "'a compact_basis set \<Rightarrow> 'a compact_basis"
       
   330 where
       
   331   "choose A = (SOME x. x \<in> {x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y})"
       
   332 
       
   333 lemma choose_lemma:
       
   334   "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> choose A \<in> {x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y}"
       
   335 unfolding choose_def
       
   336 apply (rule someI_ex)
       
   337 apply (frule (1) finite_has_maximal, fast)
       
   338 done
       
   339 
       
   340 lemma maximal_choose:
       
   341   "\<lbrakk>finite A; y \<in> A; choose A \<sqsubseteq> y\<rbrakk> \<Longrightarrow> choose A = y"
       
   342 apply (cases "A = {}", simp)
       
   343 apply (frule (1) choose_lemma, simp)
       
   344 done
       
   345 
       
   346 lemma choose_in: "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> choose A \<in> A"
       
   347 by (frule (1) choose_lemma, simp)
       
   348 
       
   349 function
       
   350   choose_pos :: "'a compact_basis set \<Rightarrow> 'a compact_basis \<Rightarrow> nat"
       
   351 where
       
   352   "choose_pos A x =
       
   353     (if finite A \<and> x \<in> A \<and> x \<noteq> choose A
       
   354       then Suc (choose_pos (A - {choose A}) x) else 0)"
       
   355 by auto
       
   356 
       
   357 termination choose_pos
       
   358 apply (relation "measure (card \<circ> fst)", simp)
       
   359 apply clarsimp
       
   360 apply (rule card_Diff1_less)
       
   361 apply assumption
       
   362 apply (erule choose_in)
       
   363 apply clarsimp
       
   364 done
       
   365 
       
   366 declare choose_pos.simps [simp del]
       
   367 
       
   368 lemma choose_pos_choose: "finite A \<Longrightarrow> choose_pos A (choose A) = 0"
       
   369 by (simp add: choose_pos.simps)
       
   370 
       
   371 lemma inj_on_choose_pos [OF refl]:
       
   372   "\<lbrakk>card A = n; finite A\<rbrakk> \<Longrightarrow> inj_on (choose_pos A) A"
       
   373  apply (induct n arbitrary: A)
       
   374   apply simp
       
   375  apply (case_tac "A = {}", simp)
       
   376  apply (frule (1) choose_in)
       
   377  apply (rule inj_onI)
       
   378  apply (drule_tac x="A - {choose A}" in meta_spec, simp)
       
   379  apply (simp add: choose_pos.simps)
       
   380  apply (simp split: split_if_asm)
       
   381  apply (erule (1) inj_onD, simp, simp)
       
   382 done
       
   383 
       
   384 lemma choose_pos_bounded [OF refl]:
       
   385   "\<lbrakk>card A = n; finite A; x \<in> A\<rbrakk> \<Longrightarrow> choose_pos A x < n"
       
   386 apply (induct n arbitrary: A)
       
   387 apply simp
       
   388  apply (case_tac "A = {}", simp)
       
   389  apply (frule (1) choose_in)
       
   390 apply (subst choose_pos.simps)
       
   391 apply simp
       
   392 done
       
   393 
       
   394 lemma choose_pos_lessD:
       
   395   "\<lbrakk>choose_pos A x < choose_pos A y; finite A; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> \<not> x \<sqsubseteq> y"
       
   396  apply (induct A x arbitrary: y rule: choose_pos.induct)
       
   397  apply simp
       
   398  apply (case_tac "x = choose A")
       
   399   apply simp
       
   400   apply (rule notI)
       
   401   apply (frule (2) maximal_choose)
       
   402   apply simp
       
   403  apply (case_tac "y = choose A")
       
   404   apply (simp add: choose_pos_choose)
       
   405  apply (drule_tac x=y in meta_spec)
       
   406  apply simp
       
   407  apply (erule meta_mp)
       
   408  apply (simp add: choose_pos.simps)
       
   409 done
       
   410 
       
   411 subsubsection {* Properties of approx function *}
       
   412 
       
   413 lemma deflation_approx: "deflation (approx i)"
       
   414 using finite_deflation_approx by (rule finite_deflation_imp_deflation)
       
   415 
       
   416 lemma approx_idem: "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
       
   417 using deflation_approx by (rule deflation.idem)
       
   418 
       
   419 lemma approx_below: "approx i\<cdot>x \<sqsubseteq> x"
       
   420 using deflation_approx by (rule deflation.below)
       
   421 
       
   422 lemma finite_range_approx: "finite (range (\<lambda>x. approx i\<cdot>x))"
       
   423 apply (rule finite_deflation.finite_range)
       
   424 apply (rule finite_deflation_approx)
       
   425 done
       
   426 
       
   427 lemma compact_approx: "compact (approx n\<cdot>x)"
       
   428 apply (rule finite_deflation.compact)
       
   429 apply (rule finite_deflation_approx)
       
   430 done
       
   431 
       
   432 lemma compact_eq_approx: "compact x \<Longrightarrow> \<exists>i. approx i\<cdot>x = x"
       
   433 by (rule admD2, simp_all)
       
   434 
       
   435 subsubsection {* Compact basis take function *}
       
   436 
       
   437 primrec
       
   438   cb_take :: "nat \<Rightarrow> 'a compact_basis \<Rightarrow> 'a compact_basis" where
       
   439   "cb_take 0 = (\<lambda>x. compact_bot)"
       
   440 | "cb_take (Suc n) = (\<lambda>a. Abs_compact_basis (approx n\<cdot>(Rep_compact_basis a)))"
       
   441 
       
   442 declare cb_take.simps [simp del]
       
   443 
       
   444 lemma cb_take_zero [simp]: "cb_take 0 a = compact_bot"
       
   445 by (simp only: cb_take.simps)
       
   446 
       
   447 lemma Rep_cb_take:
       
   448   "Rep_compact_basis (cb_take (Suc n) a) = approx n\<cdot>(Rep_compact_basis a)"
       
   449 by (simp add: Abs_compact_basis_inverse cb_take.simps(2) compact_approx)
       
   450 
       
   451 lemmas approx_Rep_compact_basis = Rep_cb_take [symmetric]
       
   452 
       
   453 lemma cb_take_covers: "\<exists>n. cb_take n x = x"
       
   454 apply (subgoal_tac "\<exists>n. cb_take (Suc n) x = x", fast)
       
   455 apply (simp add: Rep_compact_basis_inject [symmetric])
       
   456 apply (simp add: Rep_cb_take)
       
   457 apply (rule compact_eq_approx)
       
   458 apply (rule compact_Rep_compact_basis)
       
   459 done
       
   460 
       
   461 lemma cb_take_less: "cb_take n x \<sqsubseteq> x"
       
   462 unfolding compact_le_def
       
   463 by (cases n, simp, simp add: Rep_cb_take approx_below)
       
   464 
       
   465 lemma cb_take_idem: "cb_take n (cb_take n x) = cb_take n x"
       
   466 unfolding Rep_compact_basis_inject [symmetric]
       
   467 by (cases n, simp, simp add: Rep_cb_take approx_idem)
       
   468 
       
   469 lemma cb_take_mono: "x \<sqsubseteq> y \<Longrightarrow> cb_take n x \<sqsubseteq> cb_take n y"
       
   470 unfolding compact_le_def
       
   471 by (cases n, simp, simp add: Rep_cb_take monofun_cfun_arg)
       
   472 
       
   473 lemma cb_take_chain_le: "m \<le> n \<Longrightarrow> cb_take m x \<sqsubseteq> cb_take n x"
       
   474 unfolding compact_le_def
       
   475 apply (cases m, simp, cases n, simp)
       
   476 apply (simp add: Rep_cb_take, rule chain_mono, simp, simp)
       
   477 done
       
   478 
       
   479 lemma finite_range_cb_take: "finite (range (cb_take n))"
       
   480 apply (cases n)
       
   481 apply (subgoal_tac "range (cb_take 0) = {compact_bot}", simp, force)
       
   482 apply (rule finite_imageD [where f="Rep_compact_basis"])
       
   483 apply (rule finite_subset [where B="range (\<lambda>x. approx (n - 1)\<cdot>x)"])
       
   484 apply (clarsimp simp add: Rep_cb_take)
       
   485 apply (rule finite_range_approx)
       
   486 apply (rule inj_onI, simp add: Rep_compact_basis_inject)
       
   487 done
       
   488 
       
   489 subsubsection {* Rank of basis elements *}
       
   490 
       
   491 definition
       
   492   rank :: "'a compact_basis \<Rightarrow> nat"
       
   493 where
       
   494   "rank x = (LEAST n. cb_take n x = x)"
       
   495 
       
   496 lemma compact_approx_rank: "cb_take (rank x) x = x"
       
   497 unfolding rank_def
       
   498 apply (rule LeastI_ex)
       
   499 apply (rule cb_take_covers)
       
   500 done
       
   501 
       
   502 lemma rank_leD: "rank x \<le> n \<Longrightarrow> cb_take n x = x"
       
   503 apply (rule below_antisym [OF cb_take_less])
       
   504 apply (subst compact_approx_rank [symmetric])
       
   505 apply (erule cb_take_chain_le)
       
   506 done
       
   507 
       
   508 lemma rank_leI: "cb_take n x = x \<Longrightarrow> rank x \<le> n"
       
   509 unfolding rank_def by (rule Least_le)
       
   510 
       
   511 lemma rank_le_iff: "rank x \<le> n \<longleftrightarrow> cb_take n x = x"
       
   512 by (rule iffI [OF rank_leD rank_leI])
       
   513 
       
   514 lemma rank_compact_bot [simp]: "rank compact_bot = 0"
       
   515 using rank_leI [of 0 compact_bot] by simp
       
   516 
       
   517 lemma rank_eq_0_iff [simp]: "rank x = 0 \<longleftrightarrow> x = compact_bot"
       
   518 using rank_le_iff [of x 0] by auto
       
   519 
       
   520 definition
       
   521   rank_le :: "'a compact_basis \<Rightarrow> 'a compact_basis set"
       
   522 where
       
   523   "rank_le x = {y. rank y \<le> rank x}"
       
   524 
       
   525 definition
       
   526   rank_lt :: "'a compact_basis \<Rightarrow> 'a compact_basis set"
       
   527 where
       
   528   "rank_lt x = {y. rank y < rank x}"
       
   529 
       
   530 definition
       
   531   rank_eq :: "'a compact_basis \<Rightarrow> 'a compact_basis set"
       
   532 where
       
   533   "rank_eq x = {y. rank y = rank x}"
       
   534 
       
   535 lemma rank_eq_cong: "rank x = rank y \<Longrightarrow> rank_eq x = rank_eq y"
       
   536 unfolding rank_eq_def by simp
       
   537 
       
   538 lemma rank_lt_cong: "rank x = rank y \<Longrightarrow> rank_lt x = rank_lt y"
       
   539 unfolding rank_lt_def by simp
       
   540 
       
   541 lemma rank_eq_subset: "rank_eq x \<subseteq> rank_le x"
       
   542 unfolding rank_eq_def rank_le_def by auto
       
   543 
       
   544 lemma rank_lt_subset: "rank_lt x \<subseteq> rank_le x"
       
   545 unfolding rank_lt_def rank_le_def by auto
       
   546 
       
   547 lemma finite_rank_le: "finite (rank_le x)"
       
   548 unfolding rank_le_def
       
   549 apply (rule finite_subset [where B="range (cb_take (rank x))"])
       
   550 apply clarify
       
   551 apply (rule range_eqI)
       
   552 apply (erule rank_leD [symmetric])
       
   553 apply (rule finite_range_cb_take)
       
   554 done
       
   555 
       
   556 lemma finite_rank_eq: "finite (rank_eq x)"
       
   557 by (rule finite_subset [OF rank_eq_subset finite_rank_le])
       
   558 
       
   559 lemma finite_rank_lt: "finite (rank_lt x)"
       
   560 by (rule finite_subset [OF rank_lt_subset finite_rank_le])
       
   561 
       
   562 lemma rank_lt_Int_rank_eq: "rank_lt x \<inter> rank_eq x = {}"
       
   563 unfolding rank_lt_def rank_eq_def rank_le_def by auto
       
   564 
       
   565 lemma rank_lt_Un_rank_eq: "rank_lt x \<union> rank_eq x = rank_le x"
       
   566 unfolding rank_lt_def rank_eq_def rank_le_def by auto
       
   567 
       
   568 subsubsection {* Sequencing basis elements *}
       
   569 
       
   570 definition
       
   571   place :: "'a compact_basis \<Rightarrow> nat"
       
   572 where
       
   573   "place x = card (rank_lt x) + choose_pos (rank_eq x) x"
       
   574 
       
   575 lemma place_bounded: "place x < card (rank_le x)"
       
   576 unfolding place_def
       
   577  apply (rule ord_less_eq_trans)
       
   578   apply (rule add_strict_left_mono)
       
   579   apply (rule choose_pos_bounded)
       
   580    apply (rule finite_rank_eq)
       
   581   apply (simp add: rank_eq_def)
       
   582  apply (subst card_Un_disjoint [symmetric])
       
   583     apply (rule finite_rank_lt)
       
   584    apply (rule finite_rank_eq)
       
   585   apply (rule rank_lt_Int_rank_eq)
       
   586  apply (simp add: rank_lt_Un_rank_eq)
       
   587 done
       
   588 
       
   589 lemma place_ge: "card (rank_lt x) \<le> place x"
       
   590 unfolding place_def by simp
       
   591 
       
   592 lemma place_rank_mono:
       
   593   fixes x y :: "'a compact_basis"
       
   594   shows "rank x < rank y \<Longrightarrow> place x < place y"
       
   595 apply (rule less_le_trans [OF place_bounded])
       
   596 apply (rule order_trans [OF _ place_ge])
       
   597 apply (rule card_mono)
       
   598 apply (rule finite_rank_lt)
       
   599 apply (simp add: rank_le_def rank_lt_def subset_eq)
       
   600 done
       
   601 
       
   602 lemma place_eqD: "place x = place y \<Longrightarrow> x = y"
       
   603  apply (rule linorder_cases [where x="rank x" and y="rank y"])
       
   604    apply (drule place_rank_mono, simp)
       
   605   apply (simp add: place_def)
       
   606   apply (rule inj_on_choose_pos [where A="rank_eq x", THEN inj_onD])
       
   607      apply (rule finite_rank_eq)
       
   608     apply (simp cong: rank_lt_cong rank_eq_cong)
       
   609    apply (simp add: rank_eq_def)
       
   610   apply (simp add: rank_eq_def)
       
   611  apply (drule place_rank_mono, simp)
       
   612 done
       
   613 
       
   614 lemma inj_place: "inj place"
       
   615 by (rule inj_onI, erule place_eqD)
       
   616 
       
   617 subsubsection {* Embedding and projection on basis elements *}
       
   618 
       
   619 definition
       
   620   sub :: "'a compact_basis \<Rightarrow> 'a compact_basis"
       
   621 where
       
   622   "sub x = (case rank x of 0 \<Rightarrow> compact_bot | Suc k \<Rightarrow> cb_take k x)"
       
   623 
       
   624 lemma rank_sub_less: "x \<noteq> compact_bot \<Longrightarrow> rank (sub x) < rank x"
       
   625 unfolding sub_def
       
   626 apply (cases "rank x", simp)
       
   627 apply (simp add: less_Suc_eq_le)
       
   628 apply (rule rank_leI)
       
   629 apply (rule cb_take_idem)
       
   630 done
       
   631 
       
   632 lemma place_sub_less: "x \<noteq> compact_bot \<Longrightarrow> place (sub x) < place x"
       
   633 apply (rule place_rank_mono)
       
   634 apply (erule rank_sub_less)
       
   635 done
       
   636 
       
   637 lemma sub_below: "sub x \<sqsubseteq> x"
       
   638 unfolding sub_def by (cases "rank x", simp_all add: cb_take_less)
       
   639 
       
   640 lemma rank_less_imp_below_sub: "\<lbrakk>x \<sqsubseteq> y; rank x < rank y\<rbrakk> \<Longrightarrow> x \<sqsubseteq> sub y"
       
   641 unfolding sub_def
       
   642 apply (cases "rank y", simp)
       
   643 apply (simp add: less_Suc_eq_le)
       
   644 apply (subgoal_tac "cb_take nat x \<sqsubseteq> cb_take nat y")
       
   645 apply (simp add: rank_leD)
       
   646 apply (erule cb_take_mono)
       
   647 done
       
   648 
       
   649 function
       
   650   basis_emb :: "'a compact_basis \<Rightarrow> ubasis"
       
   651 where
       
   652   "basis_emb x = (if x = compact_bot then 0 else
       
   653     node (place x) (basis_emb (sub x))
       
   654       (basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y}))"
       
   655 by auto
       
   656 
       
   657 termination basis_emb
       
   658 apply (relation "measure place", simp)
       
   659 apply (simp add: place_sub_less)
       
   660 apply simp
       
   661 done
       
   662 
       
   663 declare basis_emb.simps [simp del]
       
   664 
       
   665 lemma basis_emb_compact_bot [simp]: "basis_emb compact_bot = 0"
       
   666 by (simp add: basis_emb.simps)
       
   667 
       
   668 lemma fin1: "finite {y. place y < place x \<and> x \<sqsubseteq> y}"
       
   669 apply (subst Collect_conj_eq)
       
   670 apply (rule finite_Int)
       
   671 apply (rule disjI1)
       
   672 apply (subgoal_tac "finite (place -` {n. n < place x})", simp)
       
   673 apply (rule finite_vimageI [OF _ inj_place])
       
   674 apply (simp add: lessThan_def [symmetric])
       
   675 done
       
   676 
       
   677 lemma fin2: "finite (basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y})"
       
   678 by (rule finite_imageI [OF fin1])
       
   679 
       
   680 lemma rank_place_mono:
       
   681   "\<lbrakk>place x < place y; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> rank x < rank y"
       
   682 apply (rule linorder_cases, assumption)
       
   683 apply (simp add: place_def cong: rank_lt_cong rank_eq_cong)
       
   684 apply (drule choose_pos_lessD)
       
   685 apply (rule finite_rank_eq)
       
   686 apply (simp add: rank_eq_def)
       
   687 apply (simp add: rank_eq_def)
       
   688 apply simp
       
   689 apply (drule place_rank_mono, simp)
       
   690 done
       
   691 
       
   692 lemma basis_emb_mono:
       
   693   "x \<sqsubseteq> y \<Longrightarrow> ubasis_le (basis_emb x) (basis_emb y)"
       
   694 proof (induct "max (place x) (place y)" arbitrary: x y rule: less_induct)
       
   695   case less
       
   696   show ?case proof (rule linorder_cases)
       
   697     assume "place x < place y"
       
   698     then have "rank x < rank y"
       
   699       using `x \<sqsubseteq> y` by (rule rank_place_mono)
       
   700     with `place x < place y` show ?case
       
   701       apply (case_tac "y = compact_bot", simp)
       
   702       apply (simp add: basis_emb.simps [of y])
       
   703       apply (rule ubasis_le_trans [OF _ ubasis_le_lower [OF fin2]])
       
   704       apply (rule less)
       
   705        apply (simp add: less_max_iff_disj)
       
   706        apply (erule place_sub_less)
       
   707       apply (erule rank_less_imp_below_sub [OF `x \<sqsubseteq> y`])
       
   708       done
       
   709   next
       
   710     assume "place x = place y"
       
   711     hence "x = y" by (rule place_eqD)
       
   712     thus ?case by (simp add: ubasis_le_refl)
       
   713   next
       
   714     assume "place x > place y"
       
   715     with `x \<sqsubseteq> y` show ?case
       
   716       apply (case_tac "x = compact_bot", simp add: ubasis_le_minimal)
       
   717       apply (simp add: basis_emb.simps [of x])
       
   718       apply (rule ubasis_le_upper [OF fin2], simp)
       
   719       apply (rule less)
       
   720        apply (simp add: less_max_iff_disj)
       
   721        apply (erule place_sub_less)
       
   722       apply (erule rev_below_trans)
       
   723       apply (rule sub_below)
       
   724       done
       
   725   qed
       
   726 qed
       
   727 
       
   728 lemma inj_basis_emb: "inj basis_emb"
       
   729  apply (rule inj_onI)
       
   730  apply (case_tac "x = compact_bot")
       
   731   apply (case_tac [!] "y = compact_bot")
       
   732     apply simp
       
   733    apply (simp add: basis_emb.simps)
       
   734   apply (simp add: basis_emb.simps)
       
   735  apply (simp add: basis_emb.simps)
       
   736  apply (simp add: fin2 inj_eq [OF inj_place])
       
   737 done
       
   738 
       
   739 definition
       
   740   basis_prj :: "ubasis \<Rightarrow> 'a compact_basis"
       
   741 where
       
   742   "basis_prj x = inv basis_emb
       
   743     (ubasis_until (\<lambda>x. x \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> ubasis)) x)"
       
   744 
       
   745 lemma basis_prj_basis_emb: "\<And>x. basis_prj (basis_emb x) = x"
       
   746 unfolding basis_prj_def
       
   747  apply (subst ubasis_until_same)
       
   748   apply (rule rangeI)
       
   749  apply (rule inv_f_f)
       
   750  apply (rule inj_basis_emb)
       
   751 done
       
   752 
       
   753 lemma basis_prj_node:
       
   754   "\<lbrakk>finite S; node i a S \<notin> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)\<rbrakk>
       
   755     \<Longrightarrow> basis_prj (node i a S) = (basis_prj a :: 'a compact_basis)"
       
   756 unfolding basis_prj_def by simp
       
   757 
       
   758 lemma basis_prj_0: "basis_prj 0 = compact_bot"
       
   759 apply (subst basis_emb_compact_bot [symmetric])
       
   760 apply (rule basis_prj_basis_emb)
       
   761 done
       
   762 
       
   763 lemma node_eq_basis_emb_iff:
       
   764   "finite S \<Longrightarrow> node i a S = basis_emb x \<longleftrightarrow>
       
   765     x \<noteq> compact_bot \<and> i = place x \<and> a = basis_emb (sub x) \<and>
       
   766         S = basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y}"
       
   767 apply (cases "x = compact_bot", simp)
       
   768 apply (simp add: basis_emb.simps [of x])
       
   769 apply (simp add: fin2)
       
   770 done
       
   771 
       
   772 lemma basis_prj_mono: "ubasis_le a b \<Longrightarrow> basis_prj a \<sqsubseteq> basis_prj b"
       
   773 proof (induct a b rule: ubasis_le.induct)
       
   774   case (ubasis_le_refl a) show ?case by (rule below_refl)
       
   775 next
       
   776   case (ubasis_le_trans a b c) thus ?case by - (rule below_trans)
       
   777 next
       
   778   case (ubasis_le_lower S a i) thus ?case
       
   779     apply (cases "node i a S \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)")
       
   780      apply (erule rangeE, rename_tac x)
       
   781      apply (simp add: basis_prj_basis_emb)
       
   782      apply (simp add: node_eq_basis_emb_iff)
       
   783      apply (simp add: basis_prj_basis_emb)
       
   784      apply (rule sub_below)
       
   785     apply (simp add: basis_prj_node)
       
   786     done
       
   787 next
       
   788   case (ubasis_le_upper S b a i) thus ?case
       
   789     apply (cases "node i a S \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)")
       
   790      apply (erule rangeE, rename_tac x)
       
   791      apply (simp add: basis_prj_basis_emb)
       
   792      apply (clarsimp simp add: node_eq_basis_emb_iff)
       
   793      apply (simp add: basis_prj_basis_emb)
       
   794     apply (simp add: basis_prj_node)
       
   795     done
       
   796 qed
       
   797 
       
   798 lemma basis_emb_prj_less: "ubasis_le (basis_emb (basis_prj x)) x"
       
   799 unfolding basis_prj_def
       
   800  apply (subst f_inv_into_f [where f=basis_emb])
       
   801   apply (rule ubasis_until)
       
   802   apply (rule range_eqI [where x=compact_bot])
       
   803   apply simp
       
   804  apply (rule ubasis_until_less)
       
   805 done
       
   806 
       
   807 end
       
   808 
       
   809 sublocale approx_chain \<subseteq> compact_basis!:
       
   810   ideal_completion below Rep_compact_basis
       
   811     "approximants :: 'a \<Rightarrow> 'a compact_basis set"
       
   812 proof
       
   813   fix w :: "'a"
       
   814   show "below.ideal (approximants w)"
       
   815   proof (rule below.idealI)
       
   816     show "\<exists>x. x \<in> approximants w"
       
   817       unfolding approximants_def
       
   818       apply (rule_tac x="Abs_compact_basis (approx 0\<cdot>w)" in exI)
       
   819       apply (simp add: Abs_compact_basis_inverse approx_below compact_approx)
       
   820       done
       
   821   next
       
   822     fix x y :: "'a compact_basis"
       
   823     assume "x \<in> approximants w" "y \<in> approximants w"
       
   824     thus "\<exists>z \<in> approximants w. x \<sqsubseteq> z \<and> y \<sqsubseteq> z"
       
   825       unfolding approximants_def
       
   826       apply simp
       
   827       apply (cut_tac a=x in compact_Rep_compact_basis)
       
   828       apply (cut_tac a=y in compact_Rep_compact_basis)
       
   829       apply (drule compact_eq_approx)
       
   830       apply (drule compact_eq_approx)
       
   831       apply (clarify, rename_tac i j)
       
   832       apply (rule_tac x="Abs_compact_basis (approx (max i j)\<cdot>w)" in exI)
       
   833       apply (simp add: compact_le_def)
       
   834       apply (simp add: Abs_compact_basis_inverse approx_below compact_approx)
       
   835       apply (erule subst, erule subst)
       
   836       apply (simp add: monofun_cfun chain_mono [OF chain_approx])
       
   837       done
       
   838   next
       
   839     fix x y :: "'a compact_basis"
       
   840     assume "x \<sqsubseteq> y" "y \<in> approximants w" thus "x \<in> approximants w"
       
   841       unfolding approximants_def
       
   842       apply simp
       
   843       apply (simp add: compact_le_def)
       
   844       apply (erule (1) below_trans)
       
   845       done
       
   846   qed
       
   847 next
       
   848   fix Y :: "nat \<Rightarrow> 'a"
       
   849   assume Y: "chain Y"
       
   850   show "approximants (\<Squnion>i. Y i) = (\<Union>i. approximants (Y i))"
       
   851     unfolding approximants_def
       
   852     apply safe
       
   853     apply (simp add: compactD2 [OF compact_Rep_compact_basis Y])
       
   854     apply (erule below_lub [OF Y])
       
   855     done
       
   856 next
       
   857   fix a :: "'a compact_basis"
       
   858   show "approximants (Rep_compact_basis a) = {b. b \<sqsubseteq> a}"
       
   859     unfolding approximants_def compact_le_def ..
       
   860 next
       
   861   fix x y :: "'a"
       
   862   assume "approximants x \<subseteq> approximants y" thus "x \<sqsubseteq> y"
       
   863     apply (subgoal_tac "(\<Squnion>i. approx i\<cdot>x) \<sqsubseteq> y")
       
   864     apply (simp add: lub_distribs)
       
   865     apply (rule admD, simp, simp)
       
   866     apply (drule_tac c="Abs_compact_basis (approx i\<cdot>x)" in subsetD)
       
   867     apply (simp add: approximants_def Abs_compact_basis_inverse
       
   868                      approx_below compact_approx)
       
   869     apply (simp add: approximants_def Abs_compact_basis_inverse compact_approx)
       
   870     done
       
   871 next
       
   872   show "\<exists>f::'a compact_basis \<Rightarrow> nat. inj f"
       
   873     by (rule exI, rule inj_place)
       
   874 qed
       
   875 
       
   876 subsubsection {* EP-pair from any bifinite domain into \emph{udom} *}
       
   877 
       
   878 context approx_chain begin
       
   879 
       
   880 definition
       
   881   udom_emb :: "'a \<rightarrow> udom"
       
   882 where
       
   883   "udom_emb = compact_basis.basis_fun (\<lambda>x. udom_principal (basis_emb x))"
       
   884 
       
   885 definition
       
   886   udom_prj :: "udom \<rightarrow> 'a"
       
   887 where
       
   888   "udom_prj = udom.basis_fun (\<lambda>x. Rep_compact_basis (basis_prj x))"
       
   889 
       
   890 lemma udom_emb_principal:
       
   891   "udom_emb\<cdot>(Rep_compact_basis x) = udom_principal (basis_emb x)"
       
   892 unfolding udom_emb_def
       
   893 apply (rule compact_basis.basis_fun_principal)
       
   894 apply (rule udom.principal_mono)
       
   895 apply (erule basis_emb_mono)
       
   896 done
       
   897 
       
   898 lemma udom_prj_principal:
       
   899   "udom_prj\<cdot>(udom_principal x) = Rep_compact_basis (basis_prj x)"
       
   900 unfolding udom_prj_def
       
   901 apply (rule udom.basis_fun_principal)
       
   902 apply (rule compact_basis.principal_mono)
       
   903 apply (erule basis_prj_mono)
       
   904 done
       
   905 
       
   906 lemma ep_pair_udom: "ep_pair udom_emb udom_prj"
       
   907  apply default
       
   908   apply (rule compact_basis.principal_induct, simp)
       
   909   apply (simp add: udom_emb_principal udom_prj_principal)
       
   910   apply (simp add: basis_prj_basis_emb)
       
   911  apply (rule udom.principal_induct, simp)
       
   912  apply (simp add: udom_emb_principal udom_prj_principal)
       
   913  apply (rule basis_emb_prj_less)
       
   914 done
       
   915 
       
   916 end
       
   917 
       
   918 abbreviation "udom_emb \<equiv> approx_chain.udom_emb"
       
   919 abbreviation "udom_prj \<equiv> approx_chain.udom_prj"
       
   920 
       
   921 lemmas ep_pair_udom = approx_chain.ep_pair_udom
       
   922 
       
   923 subsection {* Chain of approx functions for type \emph{udom} *}
       
   924 
       
   925 definition
       
   926   udom_approx :: "nat \<Rightarrow> udom \<rightarrow> udom"
       
   927 where
       
   928   "udom_approx i =
       
   929     udom.basis_fun (\<lambda>x. udom_principal (ubasis_until (\<lambda>y. y \<le> i) x))"
       
   930 
       
   931 lemma udom_approx_mono:
       
   932   "ubasis_le a b \<Longrightarrow>
       
   933     udom_principal (ubasis_until (\<lambda>y. y \<le> i) a) \<sqsubseteq>
       
   934     udom_principal (ubasis_until (\<lambda>y. y \<le> i) b)"
       
   935 apply (rule udom.principal_mono)
       
   936 apply (rule ubasis_until_mono)
       
   937 apply (frule (2) order_less_le_trans [OF node_gt2])
       
   938 apply (erule order_less_imp_le)
       
   939 apply assumption
       
   940 done
       
   941 
       
   942 lemma adm_mem_finite: "\<lbrakk>cont f; finite S\<rbrakk> \<Longrightarrow> adm (\<lambda>x. f x \<in> S)"
       
   943 by (erule adm_subst, induct set: finite, simp_all)
       
   944 
       
   945 lemma udom_approx_principal:
       
   946   "udom_approx i\<cdot>(udom_principal x) =
       
   947     udom_principal (ubasis_until (\<lambda>y. y \<le> i) x)"
       
   948 unfolding udom_approx_def
       
   949 apply (rule udom.basis_fun_principal)
       
   950 apply (erule udom_approx_mono)
       
   951 done
       
   952 
       
   953 lemma finite_deflation_udom_approx: "finite_deflation (udom_approx i)"
       
   954 proof
       
   955   fix x show "udom_approx i\<cdot>(udom_approx i\<cdot>x) = udom_approx i\<cdot>x"
       
   956     by (induct x rule: udom.principal_induct, simp)
       
   957        (simp add: udom_approx_principal ubasis_until_idem)
       
   958 next
       
   959   fix x show "udom_approx i\<cdot>x \<sqsubseteq> x"
       
   960     by (induct x rule: udom.principal_induct, simp)
       
   961        (simp add: udom_approx_principal ubasis_until_less)
       
   962 next
       
   963   have *: "finite (range (\<lambda>x. udom_principal (ubasis_until (\<lambda>y. y \<le> i) x)))"
       
   964     apply (subst range_composition [where f=udom_principal])
       
   965     apply (simp add: finite_range_ubasis_until)
       
   966     done
       
   967   show "finite {x. udom_approx i\<cdot>x = x}"
       
   968     apply (rule finite_range_imp_finite_fixes)
       
   969     apply (rule rev_finite_subset [OF *])
       
   970     apply (clarsimp, rename_tac x)
       
   971     apply (induct_tac x rule: udom.principal_induct)
       
   972     apply (simp add: adm_mem_finite *)
       
   973     apply (simp add: udom_approx_principal)
       
   974     done
       
   975 qed
       
   976 
       
   977 interpretation udom_approx: finite_deflation "udom_approx i"
       
   978 by (rule finite_deflation_udom_approx)
       
   979 
       
   980 lemma chain_udom_approx [simp]: "chain (\<lambda>i. udom_approx i)"
       
   981 unfolding udom_approx_def
       
   982 apply (rule chainI)
       
   983 apply (rule udom.basis_fun_mono)
       
   984 apply (erule udom_approx_mono)
       
   985 apply (erule udom_approx_mono)
       
   986 apply (rule udom.principal_mono)
       
   987 apply (rule ubasis_until_chain, simp)
       
   988 done
       
   989 
       
   990 lemma lub_udom_approx [simp]: "(\<Squnion>i. udom_approx i) = ID"
       
   991 apply (rule cfun_eqI, simp add: contlub_cfun_fun)
       
   992 apply (rule below_antisym)
       
   993 apply (rule lub_below)
       
   994 apply (simp)
       
   995 apply (rule udom_approx.below)
       
   996 apply (rule_tac x=x in udom.principal_induct)
       
   997 apply (simp add: lub_distribs)
       
   998 apply (rule_tac i=a in below_lub)
       
   999 apply simp
       
  1000 apply (simp add: udom_approx_principal)
       
  1001 apply (simp add: ubasis_until_same ubasis_le_refl)
       
  1002 done
       
  1003  
       
  1004 lemma udom_approx: "approx_chain udom_approx"
       
  1005 proof
       
  1006   show "chain (\<lambda>i. udom_approx i)"
       
  1007     by (rule chain_udom_approx)
       
  1008   show "(\<Squnion>i. udom_approx i) = ID"
       
  1009     by (rule lub_udom_approx)
       
  1010 qed
       
  1011 
       
  1012 hide_const (open) node
       
  1013 
       
  1014 end