src/HOLCF/Universal.thy
changeset 27411 60fad3219d32
child 28889 1a1447cb6b71
equal deleted inserted replaced
27410:22f75653163f 27411:60fad3219d32
       
     1 (*  Title:      HOLCF/Universal.thy
       
     2     ID:         $Id$
       
     3     Author:     Brian Huffman
       
     4 *)
       
     5 
       
     6 theory Universal
       
     7 imports CompactBasis NatIso
       
     8 begin
       
     9 
       
    10 subsection {* Basis datatype *}
       
    11 
       
    12 types ubasis = nat
       
    13 
       
    14 definition
       
    15   node :: "nat \<Rightarrow> ubasis \<Rightarrow> ubasis set \<Rightarrow> ubasis"
       
    16 where
       
    17   "node i x A = Suc (prod2nat (i, prod2nat (x, set2nat A)))"
       
    18 
       
    19 lemma node_not_0 [simp]: "node i x A \<noteq> 0"
       
    20 unfolding node_def by simp
       
    21 
       
    22 lemma node_gt_0 [simp]: "0 < node i x A"
       
    23 unfolding node_def by simp
       
    24 
       
    25 lemma node_inject [simp]:
       
    26   "\<lbrakk>finite A; finite B\<rbrakk>
       
    27     \<Longrightarrow> node i x A = node j y B \<longleftrightarrow> i = j \<and> x = y \<and> A = B"
       
    28 unfolding node_def by simp
       
    29 
       
    30 lemma node_gt0: "i < node i x A"
       
    31 unfolding node_def less_Suc_eq_le
       
    32 by (rule le_prod2nat_1)
       
    33 
       
    34 lemma node_gt1: "x < node i x A"
       
    35 unfolding node_def less_Suc_eq_le
       
    36 by (rule order_trans [OF le_prod2nat_1 le_prod2nat_2])
       
    37 
       
    38 lemma nat_less_power2: "n < 2^n"
       
    39 by (induct n) simp_all
       
    40 
       
    41 lemma node_gt2: "\<lbrakk>finite A; y \<in> A\<rbrakk> \<Longrightarrow> y < node i x A"
       
    42 unfolding node_def less_Suc_eq_le set2nat_def
       
    43 apply (rule order_trans [OF _ le_prod2nat_2])
       
    44 apply (rule order_trans [OF _ le_prod2nat_2])
       
    45 apply (rule order_trans [where y="setsum (op ^ 2) {y}"])
       
    46 apply (simp add: nat_less_power2 [THEN order_less_imp_le])
       
    47 apply (erule setsum_mono2, simp, simp)
       
    48 done
       
    49 
       
    50 lemma eq_prod2nat_pairI:
       
    51   "\<lbrakk>fst (nat2prod x) = a; snd (nat2prod x) = b\<rbrakk> \<Longrightarrow> x = prod2nat (a, b)"
       
    52 by (erule subst, erule subst, simp)
       
    53 
       
    54 lemma node_cases:
       
    55   assumes 1: "x = 0 \<Longrightarrow> P"
       
    56   assumes 2: "\<And>i y A. \<lbrakk>finite A; x = node i y A\<rbrakk> \<Longrightarrow> P"
       
    57   shows "P"
       
    58  apply (cases x)
       
    59   apply (erule 1)
       
    60  apply (rule 2)
       
    61   apply (rule finite_nat2set)
       
    62  apply (simp add: node_def)
       
    63  apply (rule eq_prod2nat_pairI [OF refl])
       
    64  apply (rule eq_prod2nat_pairI [OF refl refl])
       
    65 done
       
    66 
       
    67 lemma node_induct:
       
    68   assumes 1: "P 0"
       
    69   assumes 2: "\<And>i x A. \<lbrakk>P x; finite A; \<forall>y\<in>A. P y\<rbrakk> \<Longrightarrow> P (node i x A)"
       
    70   shows "P x"
       
    71  apply (induct x rule: nat_less_induct)
       
    72  apply (case_tac n rule: node_cases)
       
    73   apply (simp add: 1)
       
    74  apply (simp add: 2 node_gt1 node_gt2)
       
    75 done
       
    76 
       
    77 subsection {* Basis ordering *}
       
    78 
       
    79 inductive
       
    80   ubasis_le :: "nat \<Rightarrow> nat \<Rightarrow> bool"
       
    81 where
       
    82   ubasis_le_refl: "ubasis_le x x"
       
    83 | ubasis_le_trans:
       
    84     "\<lbrakk>ubasis_le x y; ubasis_le y z\<rbrakk> \<Longrightarrow> ubasis_le x z"
       
    85 | ubasis_le_lower:
       
    86     "finite A \<Longrightarrow> ubasis_le x (node i x A)"
       
    87 | ubasis_le_upper:
       
    88     "\<lbrakk>finite A; y \<in> A; ubasis_le x y\<rbrakk> \<Longrightarrow> ubasis_le (node i x A) y"
       
    89 
       
    90 lemma ubasis_le_minimal: "ubasis_le 0 x"
       
    91 apply (induct x rule: node_induct)
       
    92 apply (rule ubasis_le_refl)
       
    93 apply (erule ubasis_le_trans)
       
    94 apply (erule ubasis_le_lower)
       
    95 done
       
    96 
       
    97 subsubsection {* Generic take function *}
       
    98 
       
    99 function
       
   100   ubasis_until :: "(ubasis \<Rightarrow> bool) \<Rightarrow> ubasis \<Rightarrow> ubasis"
       
   101 where
       
   102   "ubasis_until P 0 = 0"
       
   103 | "finite A \<Longrightarrow> ubasis_until P (node i x A) =
       
   104     (if P (node i x A) then node i x A else ubasis_until P x)"
       
   105     apply clarify
       
   106     apply (rule_tac x=b in node_cases)
       
   107      apply simp
       
   108     apply simp
       
   109     apply fast
       
   110    apply simp
       
   111   apply simp
       
   112  apply simp
       
   113 done
       
   114 
       
   115 termination ubasis_until
       
   116 apply (relation "measure snd")
       
   117 apply (rule wf_measure)
       
   118 apply (simp add: node_gt1)
       
   119 done
       
   120 
       
   121 lemma ubasis_until: "P 0 \<Longrightarrow> P (ubasis_until P x)"
       
   122 by (induct x rule: node_induct) simp_all
       
   123 
       
   124 lemma ubasis_until': "0 < ubasis_until P x \<Longrightarrow> P (ubasis_until P x)"
       
   125 by (induct x rule: node_induct) auto
       
   126 
       
   127 lemma ubasis_until_same: "P x \<Longrightarrow> ubasis_until P x = x"
       
   128 by (induct x rule: node_induct) simp_all
       
   129 
       
   130 lemma ubasis_until_idem:
       
   131   "P 0 \<Longrightarrow> ubasis_until P (ubasis_until P x) = ubasis_until P x"
       
   132 by (rule ubasis_until_same [OF ubasis_until])
       
   133 
       
   134 lemma ubasis_until_0:
       
   135   "\<forall>x. x \<noteq> 0 \<longrightarrow> \<not> P x \<Longrightarrow> ubasis_until P x = 0"
       
   136 by (induct x rule: node_induct) simp_all
       
   137 
       
   138 lemma ubasis_until_less: "ubasis_le (ubasis_until P x) x"
       
   139 apply (induct x rule: node_induct)
       
   140 apply (simp add: ubasis_le_refl)
       
   141 apply (simp add: ubasis_le_refl)
       
   142 apply (rule impI)
       
   143 apply (erule ubasis_le_trans)
       
   144 apply (erule ubasis_le_lower)
       
   145 done
       
   146 
       
   147 lemma ubasis_until_chain:
       
   148   assumes PQ: "\<And>x. P x \<Longrightarrow> Q x"
       
   149   shows "ubasis_le (ubasis_until P x) (ubasis_until Q x)"
       
   150 apply (induct x rule: node_induct)
       
   151 apply (simp add: ubasis_le_refl)
       
   152 apply (simp add: ubasis_le_refl)
       
   153 apply (simp add: PQ)
       
   154 apply clarify
       
   155 apply (rule ubasis_le_trans)
       
   156 apply (rule ubasis_until_less)
       
   157 apply (erule ubasis_le_lower)
       
   158 done
       
   159 
       
   160 lemma ubasis_until_mono:
       
   161   assumes "\<And>i x A y. \<lbrakk>finite A; P (node i x A); y \<in> A; ubasis_le x y\<rbrakk> \<Longrightarrow> P y"
       
   162   shows "ubasis_le x y \<Longrightarrow> ubasis_le (ubasis_until P x) (ubasis_until P y)"
       
   163  apply (induct set: ubasis_le)
       
   164     apply (rule ubasis_le_refl)
       
   165    apply (erule (1) ubasis_le_trans)
       
   166   apply (simp add: ubasis_le_refl)
       
   167   apply (rule impI)
       
   168   apply (rule ubasis_le_trans)
       
   169    apply (rule ubasis_until_less)
       
   170   apply (erule ubasis_le_lower)
       
   171  apply simp
       
   172  apply (rule impI)
       
   173  apply (subst ubasis_until_same)
       
   174   apply (erule (3) prems)
       
   175  apply (erule (2) ubasis_le_upper)
       
   176 done
       
   177 
       
   178 lemma finite_range_ubasis_until:
       
   179   "finite {x. P x} \<Longrightarrow> finite (range (ubasis_until P))"
       
   180 apply (rule finite_subset [where B="insert 0 {x. P x}"])
       
   181 apply (clarsimp simp add: ubasis_until')
       
   182 apply simp
       
   183 done
       
   184 
       
   185 subsubsection {* Take function for @{typ ubasis} *}
       
   186 
       
   187 definition
       
   188   ubasis_take :: "nat \<Rightarrow> ubasis \<Rightarrow> ubasis"
       
   189 where
       
   190   "ubasis_take n = ubasis_until (\<lambda>x. x \<le> n)"
       
   191 
       
   192 lemma ubasis_take_le: "ubasis_take n x \<le> n"
       
   193 unfolding ubasis_take_def by (rule ubasis_until, rule le0)
       
   194 
       
   195 lemma ubasis_take_same: "x \<le> n \<Longrightarrow> ubasis_take n x = x"
       
   196 unfolding ubasis_take_def by (rule ubasis_until_same)
       
   197 
       
   198 lemma ubasis_take_idem: "ubasis_take n (ubasis_take n x) = ubasis_take n x"
       
   199 by (rule ubasis_take_same [OF ubasis_take_le])
       
   200 
       
   201 lemma ubasis_take_0 [simp]: "ubasis_take 0 x = 0"
       
   202 unfolding ubasis_take_def by (simp add: ubasis_until_0)
       
   203 
       
   204 lemma ubasis_take_less: "ubasis_le (ubasis_take n x) x"
       
   205 unfolding ubasis_take_def by (rule ubasis_until_less)
       
   206 
       
   207 lemma ubasis_take_chain: "ubasis_le (ubasis_take n x) (ubasis_take (Suc n) x)"
       
   208 unfolding ubasis_take_def by (rule ubasis_until_chain) simp
       
   209 
       
   210 lemma ubasis_take_mono:
       
   211   assumes "ubasis_le x y"
       
   212   shows "ubasis_le (ubasis_take n x) (ubasis_take n y)"
       
   213 unfolding ubasis_take_def
       
   214  apply (rule ubasis_until_mono [OF _ prems])
       
   215  apply (frule (2) order_less_le_trans [OF node_gt2])
       
   216  apply (erule order_less_imp_le)
       
   217 done
       
   218 
       
   219 lemma finite_range_ubasis_take: "finite (range (ubasis_take n))"
       
   220 apply (rule finite_subset [where B="{..n}"])
       
   221 apply (simp add: subset_eq ubasis_take_le)
       
   222 apply simp
       
   223 done
       
   224 
       
   225 lemma ubasis_take_covers: "\<exists>n. ubasis_take n x = x"
       
   226 apply (rule exI [where x=x])
       
   227 apply (simp add: ubasis_take_same)
       
   228 done
       
   229 
       
   230 interpretation udom: preorder [ubasis_le]
       
   231 apply default
       
   232 apply (rule ubasis_le_refl)
       
   233 apply (erule (1) ubasis_le_trans)
       
   234 done
       
   235 
       
   236 interpretation udom: basis_take [ubasis_le ubasis_take]
       
   237 apply default
       
   238 apply (rule ubasis_take_less)
       
   239 apply (rule ubasis_take_idem)
       
   240 apply (erule ubasis_take_mono)
       
   241 apply (rule ubasis_take_chain)
       
   242 apply (rule finite_range_ubasis_take)
       
   243 apply (rule ubasis_take_covers)
       
   244 done
       
   245 
       
   246 subsection {* Defining the universal domain by ideal completion *}
       
   247 
       
   248 typedef (open) udom = "{S. udom.ideal S}"
       
   249 by (fast intro: udom.ideal_principal)
       
   250 
       
   251 instantiation udom :: sq_ord
       
   252 begin
       
   253 
       
   254 definition
       
   255   "x \<sqsubseteq> y \<longleftrightarrow> Rep_udom x \<subseteq> Rep_udom y"
       
   256 
       
   257 instance ..
       
   258 end
       
   259 
       
   260 instance udom :: po
       
   261 by (rule udom.typedef_ideal_po
       
   262     [OF type_definition_udom sq_le_udom_def])
       
   263 
       
   264 instance udom :: cpo
       
   265 by (rule udom.typedef_ideal_cpo
       
   266     [OF type_definition_udom sq_le_udom_def])
       
   267 
       
   268 lemma Rep_udom_lub:
       
   269   "chain Y \<Longrightarrow> Rep_udom (\<Squnion>i. Y i) = (\<Union>i. Rep_udom (Y i))"
       
   270 by (rule udom.typedef_ideal_rep_contlub
       
   271     [OF type_definition_udom sq_le_udom_def])
       
   272 
       
   273 lemma ideal_Rep_udom: "udom.ideal (Rep_udom xs)"
       
   274 by (rule Rep_udom [unfolded mem_Collect_eq])
       
   275 
       
   276 definition
       
   277   udom_principal :: "nat \<Rightarrow> udom" where
       
   278   "udom_principal t = Abs_udom {u. ubasis_le u t}"
       
   279 
       
   280 lemma Rep_udom_principal:
       
   281   "Rep_udom (udom_principal t) = {u. ubasis_le u t}"
       
   282 unfolding udom_principal_def
       
   283 by (simp add: Abs_udom_inverse udom.ideal_principal)
       
   284 
       
   285 interpretation udom:
       
   286   ideal_completion [ubasis_le ubasis_take udom_principal Rep_udom]
       
   287 apply unfold_locales
       
   288 apply (rule ideal_Rep_udom)
       
   289 apply (erule Rep_udom_lub)
       
   290 apply (rule Rep_udom_principal)
       
   291 apply (simp only: sq_le_udom_def)
       
   292 done
       
   293 
       
   294 text {* Universal domain is pointed *}
       
   295 
       
   296 lemma udom_minimal: "udom_principal 0 \<sqsubseteq> x"
       
   297 apply (induct x rule: udom.principal_induct)
       
   298 apply (simp, simp add: ubasis_le_minimal)
       
   299 done
       
   300 
       
   301 instance udom :: pcpo
       
   302 by intro_classes (fast intro: udom_minimal)
       
   303 
       
   304 lemma inst_udom_pcpo: "\<bottom> = udom_principal 0"
       
   305 by (rule udom_minimal [THEN UU_I, symmetric])
       
   306 
       
   307 text {* Universal domain is bifinite *}
       
   308 
       
   309 instantiation udom :: bifinite
       
   310 begin
       
   311 
       
   312 definition
       
   313   approx_udom_def: "approx = udom.completion_approx"
       
   314 
       
   315 instance
       
   316 apply (intro_classes, unfold approx_udom_def)
       
   317 apply (rule udom.chain_completion_approx)
       
   318 apply (rule udom.lub_completion_approx)
       
   319 apply (rule udom.completion_approx_idem)
       
   320 apply (rule udom.finite_fixes_completion_approx)
       
   321 done
       
   322 
       
   323 end
       
   324 
       
   325 lemma approx_udom_principal [simp]:
       
   326   "approx n\<cdot>(udom_principal x) = udom_principal (ubasis_take n x)"
       
   327 unfolding approx_udom_def
       
   328 by (rule udom.completion_approx_principal)
       
   329 
       
   330 lemma approx_eq_udom_principal:
       
   331   "\<exists>a\<in>Rep_udom x. approx n\<cdot>x = udom_principal (ubasis_take n a)"
       
   332 unfolding approx_udom_def
       
   333 by (rule udom.completion_approx_eq_principal)
       
   334 
       
   335 
       
   336 subsection {* Universality of @{typ udom} *}
       
   337 
       
   338 defaultsort bifinite
       
   339 
       
   340 subsubsection {* Choosing a maximal element from a finite set *}
       
   341 
       
   342 lemma finite_has_maximal:
       
   343   fixes A :: "'a::po set"
       
   344   shows "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y"
       
   345 proof (induct rule: finite_ne_induct)
       
   346   case (singleton x)
       
   347     show ?case by simp
       
   348 next
       
   349   case (insert a A)
       
   350   from `\<exists>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y`
       
   351   obtain x where x: "x \<in> A"
       
   352            and x_eq: "\<And>y. \<lbrakk>y \<in> A; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> x = y" by fast
       
   353   show ?case
       
   354   proof (intro bexI ballI impI)
       
   355     fix y
       
   356     assume "y \<in> insert a A" and "(if x \<sqsubseteq> a then a else x) \<sqsubseteq> y"
       
   357     thus "(if x \<sqsubseteq> a then a else x) = y"
       
   358       apply auto
       
   359       apply (frule (1) trans_less)
       
   360       apply (frule (1) x_eq)
       
   361       apply (rule antisym_less, assumption)
       
   362       apply simp
       
   363       apply (erule (1) x_eq)
       
   364       done
       
   365   next
       
   366     show "(if x \<sqsubseteq> a then a else x) \<in> insert a A"
       
   367       by (simp add: x)
       
   368   qed
       
   369 qed
       
   370 
       
   371 definition
       
   372   choose :: "'a compact_basis set \<Rightarrow> 'a compact_basis"
       
   373 where
       
   374   "choose A = (SOME x. x \<in> {x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y})"
       
   375 
       
   376 lemma choose_lemma:
       
   377   "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> choose A \<in> {x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y}"
       
   378 unfolding choose_def
       
   379 apply (rule someI_ex)
       
   380 apply (frule (1) finite_has_maximal, fast)
       
   381 done
       
   382 
       
   383 lemma maximal_choose:
       
   384   "\<lbrakk>finite A; y \<in> A; choose A \<sqsubseteq> y\<rbrakk> \<Longrightarrow> choose A = y"
       
   385 apply (cases "A = {}", simp)
       
   386 apply (frule (1) choose_lemma, simp)
       
   387 done
       
   388 
       
   389 lemma choose_in: "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> choose A \<in> A"
       
   390 by (frule (1) choose_lemma, simp)
       
   391 
       
   392 function
       
   393   choose_pos :: "'a compact_basis set \<Rightarrow> 'a compact_basis \<Rightarrow> nat"
       
   394 where
       
   395   "choose_pos A x =
       
   396     (if finite A \<and> x \<in> A \<and> x \<noteq> choose A
       
   397       then Suc (choose_pos (A - {choose A}) x) else 0)"
       
   398 by auto
       
   399 
       
   400 termination choose_pos
       
   401 apply (relation "measure (card \<circ> fst)", simp)
       
   402 apply clarsimp
       
   403 apply (rule card_Diff1_less)
       
   404 apply assumption
       
   405 apply (erule choose_in)
       
   406 apply clarsimp
       
   407 done
       
   408 
       
   409 declare choose_pos.simps [simp del]
       
   410 
       
   411 lemma choose_pos_choose: "finite A \<Longrightarrow> choose_pos A (choose A) = 0"
       
   412 by (simp add: choose_pos.simps)
       
   413 
       
   414 lemma inj_on_choose_pos [OF refl]:
       
   415   "\<lbrakk>card A = n; finite A\<rbrakk> \<Longrightarrow> inj_on (choose_pos A) A"
       
   416  apply (induct n arbitrary: A)
       
   417   apply simp
       
   418  apply (case_tac "A = {}", simp)
       
   419  apply (frule (1) choose_in)
       
   420  apply (rule inj_onI)
       
   421  apply (drule_tac x="A - {choose A}" in meta_spec, simp)
       
   422  apply (simp add: choose_pos.simps)
       
   423  apply (simp split: split_if_asm)
       
   424  apply (erule (1) inj_onD, simp, simp)
       
   425 done
       
   426 
       
   427 lemma choose_pos_bounded [OF refl]:
       
   428   "\<lbrakk>card A = n; finite A; x \<in> A\<rbrakk> \<Longrightarrow> choose_pos A x < n"
       
   429 apply (induct n arbitrary: A)
       
   430 apply simp
       
   431  apply (case_tac "A = {}", simp)
       
   432  apply (frule (1) choose_in)
       
   433 apply (subst choose_pos.simps)
       
   434 apply simp
       
   435 done
       
   436 
       
   437 lemma choose_pos_lessD:
       
   438   "\<lbrakk>choose_pos A x < choose_pos A y; finite A; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> \<not> x \<sqsubseteq> y"
       
   439  apply (induct A x arbitrary: y rule: choose_pos.induct)
       
   440  apply simp
       
   441  apply (case_tac "x = choose A")
       
   442   apply simp
       
   443   apply (rule notI)
       
   444   apply (frule (2) maximal_choose)
       
   445   apply simp
       
   446  apply (case_tac "y = choose A")
       
   447   apply (simp add: choose_pos_choose)
       
   448  apply (drule_tac x=y in meta_spec)
       
   449  apply simp
       
   450  apply (erule meta_mp)
       
   451  apply (simp add: choose_pos.simps)
       
   452 done
       
   453 
       
   454 subsubsection {* Rank of basis elements *}
       
   455 
       
   456 primrec
       
   457   cb_take :: "nat \<Rightarrow> 'a compact_basis \<Rightarrow> 'a compact_basis"
       
   458 where
       
   459   "cb_take 0 = (\<lambda>x. compact_bot)"
       
   460 | "cb_take (Suc n) = compact_take n"
       
   461 
       
   462 lemma cb_take_covers: "\<exists>n. cb_take n x = x"
       
   463 apply (rule exE [OF compact_basis.take_covers [where a=x]])
       
   464 apply (rename_tac n, rule_tac x="Suc n" in exI, simp)
       
   465 done
       
   466 
       
   467 lemma cb_take_less: "cb_take n x \<sqsubseteq> x"
       
   468 by (cases n, simp, simp add: compact_basis.take_less)
       
   469 
       
   470 lemma cb_take_idem: "cb_take n (cb_take n x) = cb_take n x"
       
   471 by (cases n, simp, simp add: compact_basis.take_take)
       
   472 
       
   473 lemma cb_take_mono: "x \<sqsubseteq> y \<Longrightarrow> cb_take n x \<sqsubseteq> cb_take n y"
       
   474 by (cases n, simp, simp add: compact_basis.take_mono)
       
   475 
       
   476 lemma cb_take_chain_le: "m \<le> n \<Longrightarrow> cb_take m x \<sqsubseteq> cb_take n x"
       
   477 apply (cases m, simp)
       
   478 apply (cases n, simp)
       
   479 apply (simp add: compact_basis.take_chain_le)
       
   480 done
       
   481 
       
   482 lemma range_const: "range (\<lambda>x. c) = {c}"
       
   483 by auto
       
   484 
       
   485 lemma finite_range_cb_take: "finite (range (cb_take n))"
       
   486 apply (cases n)
       
   487 apply (simp add: range_const)
       
   488 apply (simp add: compact_basis.finite_range_take)
       
   489 done
       
   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 antisym_less [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 definition
       
   515   rank_le :: "'a compact_basis \<Rightarrow> 'a compact_basis set"
       
   516 where
       
   517   "rank_le x = {y. rank y \<le> rank x}"
       
   518 
       
   519 definition
       
   520   rank_lt :: "'a compact_basis \<Rightarrow> 'a compact_basis set"
       
   521 where
       
   522   "rank_lt x = {y. rank y < rank x}"
       
   523 
       
   524 definition
       
   525   rank_eq :: "'a compact_basis \<Rightarrow> 'a compact_basis set"
       
   526 where
       
   527   "rank_eq x = {y. rank y = rank x}"
       
   528 
       
   529 lemma rank_eq_cong: "rank x = rank y \<Longrightarrow> rank_eq x = rank_eq y"
       
   530 unfolding rank_eq_def by simp
       
   531 
       
   532 lemma rank_lt_cong: "rank x = rank y \<Longrightarrow> rank_lt x = rank_lt y"
       
   533 unfolding rank_lt_def by simp
       
   534 
       
   535 lemma rank_eq_subset: "rank_eq x \<subseteq> rank_le x"
       
   536 unfolding rank_eq_def rank_le_def by auto
       
   537 
       
   538 lemma rank_lt_subset: "rank_lt x \<subseteq> rank_le x"
       
   539 unfolding rank_lt_def rank_le_def by auto
       
   540 
       
   541 lemma finite_rank_le: "finite (rank_le x)"
       
   542 unfolding rank_le_def
       
   543 apply (rule finite_subset [where B="range (cb_take (rank x))"])
       
   544 apply clarify
       
   545 apply (rule range_eqI)
       
   546 apply (erule rank_leD [symmetric])
       
   547 apply (rule finite_range_cb_take)
       
   548 done
       
   549 
       
   550 lemma finite_rank_eq: "finite (rank_eq x)"
       
   551 by (rule finite_subset [OF rank_eq_subset finite_rank_le])
       
   552 
       
   553 lemma finite_rank_lt: "finite (rank_lt x)"
       
   554 by (rule finite_subset [OF rank_lt_subset finite_rank_le])
       
   555 
       
   556 lemma rank_lt_Int_rank_eq: "rank_lt x \<inter> rank_eq x = {}"
       
   557 unfolding rank_lt_def rank_eq_def rank_le_def by auto
       
   558 
       
   559 lemma rank_lt_Un_rank_eq: "rank_lt x \<union> rank_eq x = rank_le x"
       
   560 unfolding rank_lt_def rank_eq_def rank_le_def by auto
       
   561 
       
   562 subsubsection {* Reordering of basis elements *}
       
   563 
       
   564 definition
       
   565   reorder :: "'a compact_basis \<Rightarrow> nat"
       
   566 where
       
   567   "reorder x = card (rank_lt x) + choose_pos (rank_eq x) x"
       
   568 
       
   569 lemma reorder_bounded: "reorder x < card (rank_le x)"
       
   570 unfolding reorder_def
       
   571  apply (rule ord_less_eq_trans)
       
   572   apply (rule add_strict_left_mono)
       
   573   apply (rule choose_pos_bounded)
       
   574    apply (rule finite_rank_eq)
       
   575   apply (simp add: rank_eq_def)
       
   576  apply (subst card_Un_disjoint [symmetric])
       
   577     apply (rule finite_rank_lt)
       
   578    apply (rule finite_rank_eq)
       
   579   apply (rule rank_lt_Int_rank_eq)
       
   580  apply (simp add: rank_lt_Un_rank_eq)
       
   581 done
       
   582 
       
   583 lemma reorder_ge: "card (rank_lt x) \<le> reorder x"
       
   584 unfolding reorder_def by simp
       
   585 
       
   586 lemma reorder_rank_mono:
       
   587   fixes x y :: "'a compact_basis"
       
   588   shows "rank x < rank y \<Longrightarrow> reorder x < reorder y"
       
   589 apply (rule less_le_trans [OF reorder_bounded])
       
   590 apply (rule order_trans [OF _ reorder_ge])
       
   591 apply (rule card_mono)
       
   592 apply (rule finite_rank_lt)
       
   593 apply (simp add: rank_le_def rank_lt_def subset_eq)
       
   594 done
       
   595 
       
   596 lemma reorder_eqD: "reorder x = reorder y \<Longrightarrow> x = y"
       
   597  apply (rule linorder_cases [where x="rank x" and y="rank y"])
       
   598    apply (drule reorder_rank_mono, simp)
       
   599   apply (simp add: reorder_def)
       
   600   apply (rule inj_on_choose_pos [where A="rank_eq x", THEN inj_onD])
       
   601      apply (rule finite_rank_eq)
       
   602     apply (simp cong: rank_lt_cong rank_eq_cong)
       
   603    apply (simp add: rank_eq_def)
       
   604   apply (simp add: rank_eq_def)
       
   605  apply (drule reorder_rank_mono, simp)
       
   606 done
       
   607 
       
   608 lemma inj_reorder: "inj reorder"
       
   609 by (rule inj_onI, erule reorder_eqD)
       
   610 
       
   611 subsubsection {* Embedding and projection on basis elements *}
       
   612 
       
   613 function
       
   614   basis_emb :: "'a compact_basis \<Rightarrow> ubasis"
       
   615 where
       
   616   "basis_emb x = (if x = compact_bot then 0 else
       
   617     node
       
   618       (reorder x)
       
   619       (case rank x of 0 \<Rightarrow> 0 | Suc k \<Rightarrow> basis_emb (cb_take k x))
       
   620       (basis_emb ` {y. reorder y < reorder x \<and> x \<sqsubseteq> y}))"
       
   621 by auto
       
   622 
       
   623 termination basis_emb
       
   624 apply (relation "measure reorder", simp)
       
   625 apply simp
       
   626 apply (rule reorder_rank_mono)
       
   627 apply (simp add: less_Suc_eq_le)
       
   628 apply (rule rank_leI)
       
   629 apply (rule cb_take_idem)
       
   630 apply simp
       
   631 done
       
   632 
       
   633 declare basis_emb.simps [simp del]
       
   634 
       
   635 lemma basis_emb_compact_bot [simp]: "basis_emb compact_bot = 0"
       
   636 by (simp add: basis_emb.simps)
       
   637 
       
   638 lemma fin1: "finite {y. reorder y < reorder x \<and> x \<sqsubseteq> y}"
       
   639 apply (subst Collect_conj_eq)
       
   640 apply (rule finite_Int)
       
   641 apply (rule disjI1)
       
   642 apply (subgoal_tac "finite (reorder -` {n. n < reorder x})", simp)
       
   643 apply (rule finite_vimageI [OF _ inj_reorder])
       
   644 apply (simp add: lessThan_def [symmetric])
       
   645 done
       
   646 
       
   647 lemma fin2: "finite (basis_emb ` {y. reorder y < reorder x \<and> x \<sqsubseteq> y})"
       
   648 by (rule finite_imageI [OF fin1])
       
   649 
       
   650 lemma basis_emb_mono [OF refl]:
       
   651   "\<lbrakk>n = max (reorder x) (reorder y); x \<sqsubseteq> y\<rbrakk>
       
   652     \<Longrightarrow> ubasis_le (basis_emb x) (basis_emb y)"
       
   653 proof (induct n arbitrary: x y rule: less_induct)
       
   654   case (less n)
       
   655   assume IH:
       
   656     "\<And>(m::nat) (x::'a compact_basis) y.
       
   657       \<lbrakk>m < n; m = max (reorder x) (reorder y); x \<sqsubseteq> y\<rbrakk>
       
   658         \<Longrightarrow> ubasis_le (basis_emb x) (basis_emb y)"
       
   659   assume n: "n = max (reorder x) (reorder y)"
       
   660   assume less: "x \<sqsubseteq> y"
       
   661   show ?case
       
   662   proof (cases)
       
   663     assume "x = compact_bot"
       
   664     thus ?case by (simp add: ubasis_le_minimal)
       
   665   next
       
   666     assume x_neq [simp]: "x \<noteq> compact_bot"
       
   667     with less have y_neq [simp]: "y \<noteq> compact_bot"
       
   668       apply clarify
       
   669       apply (drule antisym_less [OF compact_minimal])
       
   670       apply simp
       
   671       done
       
   672     show ?case
       
   673     proof (rule linorder_cases)
       
   674       assume 1: "reorder x < reorder y"
       
   675       show ?case
       
   676       proof (rule linorder_cases)
       
   677         assume "rank x < rank y"
       
   678         with 1 show ?case
       
   679           apply (case_tac "rank y", simp)
       
   680           apply (subst basis_emb.simps [where x=y])
       
   681           apply simp
       
   682           apply (rule ubasis_le_trans [OF _ ubasis_le_lower [OF fin2]])
       
   683           apply (rule IH [OF _ refl, unfolded n])
       
   684            apply (simp add: less_max_iff_disj)
       
   685            apply (rule reorder_rank_mono)
       
   686            apply (simp add: less_Suc_eq_le)
       
   687            apply (rule rank_leI)
       
   688            apply (rule cb_take_idem)
       
   689           apply (simp add: less_Suc_eq_le)
       
   690           apply (subgoal_tac "cb_take nat x \<sqsubseteq> cb_take nat y")
       
   691            apply (simp add: rank_leD)
       
   692           apply (rule cb_take_mono [OF less])
       
   693           done
       
   694       next
       
   695         assume "rank x = rank y"
       
   696         with 1 show ?case
       
   697           apply (simp add: reorder_def)
       
   698           apply (simp cong: rank_lt_cong rank_eq_cong)
       
   699           apply (drule choose_pos_lessD)
       
   700              apply (rule finite_rank_eq)
       
   701             apply (simp add: rank_eq_def)
       
   702            apply (simp add: rank_eq_def)
       
   703           apply (simp add: less)
       
   704           done
       
   705       next
       
   706         assume "rank x > rank y"
       
   707         hence "reorder x > reorder y"
       
   708           by (rule reorder_rank_mono)
       
   709         with 1 show ?case by simp
       
   710       qed
       
   711     next
       
   712       assume "reorder x = reorder y"
       
   713       hence "x = y" by (rule reorder_eqD)
       
   714       thus ?case by (simp add: ubasis_le_refl)
       
   715     next
       
   716       assume "reorder x > reorder y"
       
   717       with less show ?case
       
   718         apply (simp add: basis_emb.simps [where x=x])
       
   719         apply (rule ubasis_le_upper [OF fin2], simp)
       
   720         apply (cases "rank x")
       
   721          apply (simp add: ubasis_le_minimal)
       
   722         apply simp
       
   723         apply (rule IH [OF _ refl, unfolded n])
       
   724          apply (simp add: less_max_iff_disj)
       
   725          apply (rule reorder_rank_mono)
       
   726          apply (simp add: less_Suc_eq_le)
       
   727          apply (rule rank_leI)
       
   728          apply (rule cb_take_idem)
       
   729         apply (erule rev_trans_less)
       
   730         apply (rule cb_take_less)
       
   731        done
       
   732     qed
       
   733   qed
       
   734 qed
       
   735 
       
   736 lemma inj_basis_emb: "inj basis_emb"
       
   737  apply (rule inj_onI)
       
   738  apply (case_tac "x = compact_bot")
       
   739   apply (case_tac [!] "y = compact_bot")
       
   740     apply simp
       
   741    apply (simp add: basis_emb.simps)
       
   742   apply (simp add: basis_emb.simps)
       
   743  apply (simp add: basis_emb.simps)
       
   744  apply (simp add: fin2 inj_eq [OF inj_reorder])
       
   745 done
       
   746 
       
   747 definition
       
   748   basis_prj :: "nat \<Rightarrow> 'a compact_basis"
       
   749 where
       
   750   "basis_prj x = inv basis_emb
       
   751     (ubasis_until (\<lambda>x. x \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)) x)"
       
   752 
       
   753 lemma basis_prj_basis_emb: "\<And>x. basis_prj (basis_emb x) = x"
       
   754 unfolding basis_prj_def
       
   755  apply (subst ubasis_until_same)
       
   756   apply (rule rangeI)
       
   757  apply (rule inv_f_f)
       
   758  apply (rule inj_basis_emb)
       
   759 done
       
   760 
       
   761 lemma basis_prj_node:
       
   762   "\<lbrakk>finite A; node i x A \<notin> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)\<rbrakk>
       
   763     \<Longrightarrow> basis_prj (node i x A) = (basis_prj x :: 'a compact_basis)"
       
   764 unfolding basis_prj_def by simp
       
   765 
       
   766 lemma basis_prj_0: "basis_prj 0 = compact_bot"
       
   767 apply (subst basis_emb_compact_bot [symmetric])
       
   768 apply (rule basis_prj_basis_emb)
       
   769 done
       
   770 
       
   771 lemma basis_prj_mono: "ubasis_le x y \<Longrightarrow> basis_prj x \<sqsubseteq> basis_prj y"
       
   772  apply (erule ubasis_le.induct)
       
   773     apply (rule refl_less)
       
   774    apply (erule (1) trans_less)
       
   775   apply (case_tac "node i x A \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)")
       
   776    apply (erule rangeE, rename_tac a)
       
   777    apply (case_tac "a = compact_bot", simp)
       
   778    apply (simp add: basis_prj_basis_emb)
       
   779    apply (simp add: basis_emb.simps)
       
   780    apply (clarsimp simp add: fin2)
       
   781    apply (case_tac "rank a", simp)
       
   782     apply (simp add: basis_prj_0)
       
   783    apply (simp add: basis_prj_basis_emb)
       
   784    apply (rule cb_take_less)
       
   785   apply (simp add: basis_prj_node)
       
   786  apply (case_tac "node i x A \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)")
       
   787   apply (erule rangeE, rename_tac a)
       
   788   apply (case_tac "a = compact_bot", simp)
       
   789   apply (simp add: basis_prj_basis_emb)
       
   790   apply (simp add: basis_emb.simps)
       
   791   apply (clarsimp simp add: fin2)
       
   792   apply (case_tac "rank a", simp add: basis_prj_basis_emb)
       
   793   apply (simp add: basis_prj_basis_emb)
       
   794  apply (simp add: basis_prj_node)
       
   795 done
       
   796 
       
   797 lemma basis_emb_prj_less: "ubasis_le (basis_emb (basis_prj x)) x"
       
   798 unfolding basis_prj_def
       
   799  apply (subst f_inv_f [where f=basis_emb])
       
   800   apply (rule ubasis_until)
       
   801   apply (rule range_eqI [where x=compact_bot])
       
   802   apply simp
       
   803  apply (rule ubasis_until_less)
       
   804 done
       
   805 
       
   806 hide (open) const
       
   807   node
       
   808   choose
       
   809   choose_pos
       
   810   reorder
       
   811 
       
   812 subsubsection {* EP-pair from any bifinite domain into @{typ udom} *}
       
   813 
       
   814 definition
       
   815   udom_emb :: "'a::bifinite \<rightarrow> udom"
       
   816 where
       
   817   "udom_emb = compact_basis.basis_fun (\<lambda>x. udom_principal (basis_emb x))"
       
   818 
       
   819 definition
       
   820   udom_prj :: "udom \<rightarrow> 'a::bifinite"
       
   821 where
       
   822   "udom_prj = udom.basis_fun (\<lambda>x. Rep_compact_basis (basis_prj x))"
       
   823 
       
   824 lemma udom_emb_principal:
       
   825   "udom_emb\<cdot>(Rep_compact_basis x) = udom_principal (basis_emb x)"
       
   826 unfolding udom_emb_def
       
   827 apply (rule compact_basis.basis_fun_principal)
       
   828 apply (rule udom.principal_mono)
       
   829 apply (erule basis_emb_mono)
       
   830 done
       
   831 
       
   832 lemma udom_prj_principal:
       
   833   "udom_prj\<cdot>(udom_principal x) = Rep_compact_basis (basis_prj x)"
       
   834 unfolding udom_prj_def
       
   835 apply (rule udom.basis_fun_principal)
       
   836 apply (rule compact_basis.principal_mono)
       
   837 apply (erule basis_prj_mono)
       
   838 done
       
   839 
       
   840 lemma ep_pair_udom: "ep_pair udom_emb udom_prj"
       
   841  apply default
       
   842   apply (rule compact_basis.principal_induct, simp)
       
   843   apply (simp add: udom_emb_principal udom_prj_principal)
       
   844   apply (simp add: basis_prj_basis_emb)
       
   845  apply (rule udom.principal_induct, simp)
       
   846  apply (simp add: udom_emb_principal udom_prj_principal)
       
   847  apply (rule basis_emb_prj_less)
       
   848 done
       
   849 
       
   850 end