src/HOL/HOLCF/UpperPD.thy
changeset 40774 0437dbc127b3
parent 40734 a292fc5157f8
child 40888 28cd51cff70c
equal deleted inserted replaced
40773:6c12f5e24e34 40774:0437dbc127b3
       
     1 (*  Title:      HOLCF/UpperPD.thy
       
     2     Author:     Brian Huffman
       
     3 *)
       
     4 
       
     5 header {* Upper powerdomain *}
       
     6 
       
     7 theory UpperPD
       
     8 imports CompactBasis
       
     9 begin
       
    10 
       
    11 subsection {* Basis preorder *}
       
    12 
       
    13 definition
       
    14   upper_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix "\<le>\<sharp>" 50) where
       
    15   "upper_le = (\<lambda>u v. \<forall>y\<in>Rep_pd_basis v. \<exists>x\<in>Rep_pd_basis u. x \<sqsubseteq> y)"
       
    16 
       
    17 lemma upper_le_refl [simp]: "t \<le>\<sharp> t"
       
    18 unfolding upper_le_def by fast
       
    19 
       
    20 lemma upper_le_trans: "\<lbrakk>t \<le>\<sharp> u; u \<le>\<sharp> v\<rbrakk> \<Longrightarrow> t \<le>\<sharp> v"
       
    21 unfolding upper_le_def
       
    22 apply (rule ballI)
       
    23 apply (drule (1) bspec, erule bexE)
       
    24 apply (drule (1) bspec, erule bexE)
       
    25 apply (erule rev_bexI)
       
    26 apply (erule (1) below_trans)
       
    27 done
       
    28 
       
    29 interpretation upper_le: preorder upper_le
       
    30 by (rule preorder.intro, rule upper_le_refl, rule upper_le_trans)
       
    31 
       
    32 lemma upper_le_minimal [simp]: "PDUnit compact_bot \<le>\<sharp> t"
       
    33 unfolding upper_le_def Rep_PDUnit by simp
       
    34 
       
    35 lemma PDUnit_upper_mono: "x \<sqsubseteq> y \<Longrightarrow> PDUnit x \<le>\<sharp> PDUnit y"
       
    36 unfolding upper_le_def Rep_PDUnit by simp
       
    37 
       
    38 lemma PDPlus_upper_mono: "\<lbrakk>s \<le>\<sharp> t; u \<le>\<sharp> v\<rbrakk> \<Longrightarrow> PDPlus s u \<le>\<sharp> PDPlus t v"
       
    39 unfolding upper_le_def Rep_PDPlus by fast
       
    40 
       
    41 lemma PDPlus_upper_le: "PDPlus t u \<le>\<sharp> t"
       
    42 unfolding upper_le_def Rep_PDPlus by fast
       
    43 
       
    44 lemma upper_le_PDUnit_PDUnit_iff [simp]:
       
    45   "(PDUnit a \<le>\<sharp> PDUnit b) = (a \<sqsubseteq> b)"
       
    46 unfolding upper_le_def Rep_PDUnit by fast
       
    47 
       
    48 lemma upper_le_PDPlus_PDUnit_iff:
       
    49   "(PDPlus t u \<le>\<sharp> PDUnit a) = (t \<le>\<sharp> PDUnit a \<or> u \<le>\<sharp> PDUnit a)"
       
    50 unfolding upper_le_def Rep_PDPlus Rep_PDUnit by fast
       
    51 
       
    52 lemma upper_le_PDPlus_iff: "(t \<le>\<sharp> PDPlus u v) = (t \<le>\<sharp> u \<and> t \<le>\<sharp> v)"
       
    53 unfolding upper_le_def Rep_PDPlus by fast
       
    54 
       
    55 lemma upper_le_induct [induct set: upper_le]:
       
    56   assumes le: "t \<le>\<sharp> u"
       
    57   assumes 1: "\<And>a b. a \<sqsubseteq> b \<Longrightarrow> P (PDUnit a) (PDUnit b)"
       
    58   assumes 2: "\<And>t u a. P t (PDUnit a) \<Longrightarrow> P (PDPlus t u) (PDUnit a)"
       
    59   assumes 3: "\<And>t u v. \<lbrakk>P t u; P t v\<rbrakk> \<Longrightarrow> P t (PDPlus u v)"
       
    60   shows "P t u"
       
    61 using le apply (induct u arbitrary: t rule: pd_basis_induct)
       
    62 apply (erule rev_mp)
       
    63 apply (induct_tac t rule: pd_basis_induct)
       
    64 apply (simp add: 1)
       
    65 apply (simp add: upper_le_PDPlus_PDUnit_iff)
       
    66 apply (simp add: 2)
       
    67 apply (subst PDPlus_commute)
       
    68 apply (simp add: 2)
       
    69 apply (simp add: upper_le_PDPlus_iff 3)
       
    70 done
       
    71 
       
    72 
       
    73 subsection {* Type definition *}
       
    74 
       
    75 typedef (open) 'a upper_pd =
       
    76   "{S::'a pd_basis set. upper_le.ideal S}"
       
    77 by (fast intro: upper_le.ideal_principal)
       
    78 
       
    79 instantiation upper_pd :: ("domain") below
       
    80 begin
       
    81 
       
    82 definition
       
    83   "x \<sqsubseteq> y \<longleftrightarrow> Rep_upper_pd x \<subseteq> Rep_upper_pd y"
       
    84 
       
    85 instance ..
       
    86 end
       
    87 
       
    88 instance upper_pd :: ("domain") po
       
    89 using type_definition_upper_pd below_upper_pd_def
       
    90 by (rule upper_le.typedef_ideal_po)
       
    91 
       
    92 instance upper_pd :: ("domain") cpo
       
    93 using type_definition_upper_pd below_upper_pd_def
       
    94 by (rule upper_le.typedef_ideal_cpo)
       
    95 
       
    96 definition
       
    97   upper_principal :: "'a pd_basis \<Rightarrow> 'a upper_pd" where
       
    98   "upper_principal t = Abs_upper_pd {u. u \<le>\<sharp> t}"
       
    99 
       
   100 interpretation upper_pd:
       
   101   ideal_completion upper_le upper_principal Rep_upper_pd
       
   102 using type_definition_upper_pd below_upper_pd_def
       
   103 using upper_principal_def pd_basis_countable
       
   104 by (rule upper_le.typedef_ideal_completion)
       
   105 
       
   106 text {* Upper powerdomain is pointed *}
       
   107 
       
   108 lemma upper_pd_minimal: "upper_principal (PDUnit compact_bot) \<sqsubseteq> ys"
       
   109 by (induct ys rule: upper_pd.principal_induct, simp, simp)
       
   110 
       
   111 instance upper_pd :: ("domain") pcpo
       
   112 by intro_classes (fast intro: upper_pd_minimal)
       
   113 
       
   114 lemma inst_upper_pd_pcpo: "\<bottom> = upper_principal (PDUnit compact_bot)"
       
   115 by (rule upper_pd_minimal [THEN UU_I, symmetric])
       
   116 
       
   117 
       
   118 subsection {* Monadic unit and plus *}
       
   119 
       
   120 definition
       
   121   upper_unit :: "'a \<rightarrow> 'a upper_pd" where
       
   122   "upper_unit = compact_basis.basis_fun (\<lambda>a. upper_principal (PDUnit a))"
       
   123 
       
   124 definition
       
   125   upper_plus :: "'a upper_pd \<rightarrow> 'a upper_pd \<rightarrow> 'a upper_pd" where
       
   126   "upper_plus = upper_pd.basis_fun (\<lambda>t. upper_pd.basis_fun (\<lambda>u.
       
   127       upper_principal (PDPlus t u)))"
       
   128 
       
   129 abbreviation
       
   130   upper_add :: "'a upper_pd \<Rightarrow> 'a upper_pd \<Rightarrow> 'a upper_pd"
       
   131     (infixl "+\<sharp>" 65) where
       
   132   "xs +\<sharp> ys == upper_plus\<cdot>xs\<cdot>ys"
       
   133 
       
   134 syntax
       
   135   "_upper_pd" :: "args \<Rightarrow> 'a upper_pd" ("{_}\<sharp>")
       
   136 
       
   137 translations
       
   138   "{x,xs}\<sharp>" == "{x}\<sharp> +\<sharp> {xs}\<sharp>"
       
   139   "{x}\<sharp>" == "CONST upper_unit\<cdot>x"
       
   140 
       
   141 lemma upper_unit_Rep_compact_basis [simp]:
       
   142   "{Rep_compact_basis a}\<sharp> = upper_principal (PDUnit a)"
       
   143 unfolding upper_unit_def
       
   144 by (simp add: compact_basis.basis_fun_principal PDUnit_upper_mono)
       
   145 
       
   146 lemma upper_plus_principal [simp]:
       
   147   "upper_principal t +\<sharp> upper_principal u = upper_principal (PDPlus t u)"
       
   148 unfolding upper_plus_def
       
   149 by (simp add: upper_pd.basis_fun_principal
       
   150     upper_pd.basis_fun_mono PDPlus_upper_mono)
       
   151 
       
   152 interpretation upper_add: semilattice upper_add proof
       
   153   fix xs ys zs :: "'a upper_pd"
       
   154   show "(xs +\<sharp> ys) +\<sharp> zs = xs +\<sharp> (ys +\<sharp> zs)"
       
   155     apply (induct xs ys arbitrary: zs rule: upper_pd.principal_induct2, simp, simp)
       
   156     apply (rule_tac x=zs in upper_pd.principal_induct, simp)
       
   157     apply (simp add: PDPlus_assoc)
       
   158     done
       
   159   show "xs +\<sharp> ys = ys +\<sharp> xs"
       
   160     apply (induct xs ys rule: upper_pd.principal_induct2, simp, simp)
       
   161     apply (simp add: PDPlus_commute)
       
   162     done
       
   163   show "xs +\<sharp> xs = xs"
       
   164     apply (induct xs rule: upper_pd.principal_induct, simp)
       
   165     apply (simp add: PDPlus_absorb)
       
   166     done
       
   167 qed
       
   168 
       
   169 lemmas upper_plus_assoc = upper_add.assoc
       
   170 lemmas upper_plus_commute = upper_add.commute
       
   171 lemmas upper_plus_absorb = upper_add.idem
       
   172 lemmas upper_plus_left_commute = upper_add.left_commute
       
   173 lemmas upper_plus_left_absorb = upper_add.left_idem
       
   174 
       
   175 text {* Useful for @{text "simp add: upper_plus_ac"} *}
       
   176 lemmas upper_plus_ac =
       
   177   upper_plus_assoc upper_plus_commute upper_plus_left_commute
       
   178 
       
   179 text {* Useful for @{text "simp only: upper_plus_aci"} *}
       
   180 lemmas upper_plus_aci =
       
   181   upper_plus_ac upper_plus_absorb upper_plus_left_absorb
       
   182 
       
   183 lemma upper_plus_below1: "xs +\<sharp> ys \<sqsubseteq> xs"
       
   184 apply (induct xs ys rule: upper_pd.principal_induct2, simp, simp)
       
   185 apply (simp add: PDPlus_upper_le)
       
   186 done
       
   187 
       
   188 lemma upper_plus_below2: "xs +\<sharp> ys \<sqsubseteq> ys"
       
   189 by (subst upper_plus_commute, rule upper_plus_below1)
       
   190 
       
   191 lemma upper_plus_greatest: "\<lbrakk>xs \<sqsubseteq> ys; xs \<sqsubseteq> zs\<rbrakk> \<Longrightarrow> xs \<sqsubseteq> ys +\<sharp> zs"
       
   192 apply (subst upper_plus_absorb [of xs, symmetric])
       
   193 apply (erule (1) monofun_cfun [OF monofun_cfun_arg])
       
   194 done
       
   195 
       
   196 lemma upper_below_plus_iff [simp]:
       
   197   "xs \<sqsubseteq> ys +\<sharp> zs \<longleftrightarrow> xs \<sqsubseteq> ys \<and> xs \<sqsubseteq> zs"
       
   198 apply safe
       
   199 apply (erule below_trans [OF _ upper_plus_below1])
       
   200 apply (erule below_trans [OF _ upper_plus_below2])
       
   201 apply (erule (1) upper_plus_greatest)
       
   202 done
       
   203 
       
   204 lemma upper_plus_below_unit_iff [simp]:
       
   205   "xs +\<sharp> ys \<sqsubseteq> {z}\<sharp> \<longleftrightarrow> xs \<sqsubseteq> {z}\<sharp> \<or> ys \<sqsubseteq> {z}\<sharp>"
       
   206 apply (induct xs rule: upper_pd.principal_induct, simp)
       
   207 apply (induct ys rule: upper_pd.principal_induct, simp)
       
   208 apply (induct z rule: compact_basis.principal_induct, simp)
       
   209 apply (simp add: upper_le_PDPlus_PDUnit_iff)
       
   210 done
       
   211 
       
   212 lemma upper_unit_below_iff [simp]: "{x}\<sharp> \<sqsubseteq> {y}\<sharp> \<longleftrightarrow> x \<sqsubseteq> y"
       
   213 apply (induct x rule: compact_basis.principal_induct, simp)
       
   214 apply (induct y rule: compact_basis.principal_induct, simp)
       
   215 apply simp
       
   216 done
       
   217 
       
   218 lemmas upper_pd_below_simps =
       
   219   upper_unit_below_iff
       
   220   upper_below_plus_iff
       
   221   upper_plus_below_unit_iff
       
   222 
       
   223 lemma upper_unit_eq_iff [simp]: "{x}\<sharp> = {y}\<sharp> \<longleftrightarrow> x = y"
       
   224 unfolding po_eq_conv by simp
       
   225 
       
   226 lemma upper_unit_strict [simp]: "{\<bottom>}\<sharp> = \<bottom>"
       
   227 using upper_unit_Rep_compact_basis [of compact_bot]
       
   228 by (simp add: inst_upper_pd_pcpo)
       
   229 
       
   230 lemma upper_plus_strict1 [simp]: "\<bottom> +\<sharp> ys = \<bottom>"
       
   231 by (rule UU_I, rule upper_plus_below1)
       
   232 
       
   233 lemma upper_plus_strict2 [simp]: "xs +\<sharp> \<bottom> = \<bottom>"
       
   234 by (rule UU_I, rule upper_plus_below2)
       
   235 
       
   236 lemma upper_unit_bottom_iff [simp]: "{x}\<sharp> = \<bottom> \<longleftrightarrow> x = \<bottom>"
       
   237 unfolding upper_unit_strict [symmetric] by (rule upper_unit_eq_iff)
       
   238 
       
   239 lemma upper_plus_bottom_iff [simp]:
       
   240   "xs +\<sharp> ys = \<bottom> \<longleftrightarrow> xs = \<bottom> \<or> ys = \<bottom>"
       
   241 apply (rule iffI)
       
   242 apply (erule rev_mp)
       
   243 apply (rule upper_pd.principal_induct2 [where x=xs and y=ys], simp, simp)
       
   244 apply (simp add: inst_upper_pd_pcpo upper_pd.principal_eq_iff
       
   245                  upper_le_PDPlus_PDUnit_iff)
       
   246 apply auto
       
   247 done
       
   248 
       
   249 lemma compact_upper_unit: "compact x \<Longrightarrow> compact {x}\<sharp>"
       
   250 by (auto dest!: compact_basis.compact_imp_principal)
       
   251 
       
   252 lemma compact_upper_unit_iff [simp]: "compact {x}\<sharp> \<longleftrightarrow> compact x"
       
   253 apply (safe elim!: compact_upper_unit)
       
   254 apply (simp only: compact_def upper_unit_below_iff [symmetric])
       
   255 apply (erule adm_subst [OF cont_Rep_cfun2])
       
   256 done
       
   257 
       
   258 lemma compact_upper_plus [simp]:
       
   259   "\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs +\<sharp> ys)"
       
   260 by (auto dest!: upper_pd.compact_imp_principal)
       
   261 
       
   262 
       
   263 subsection {* Induction rules *}
       
   264 
       
   265 lemma upper_pd_induct1:
       
   266   assumes P: "adm P"
       
   267   assumes unit: "\<And>x. P {x}\<sharp>"
       
   268   assumes insert: "\<And>x ys. \<lbrakk>P {x}\<sharp>; P ys\<rbrakk> \<Longrightarrow> P ({x}\<sharp> +\<sharp> ys)"
       
   269   shows "P (xs::'a upper_pd)"
       
   270 apply (induct xs rule: upper_pd.principal_induct, rule P)
       
   271 apply (induct_tac a rule: pd_basis_induct1)
       
   272 apply (simp only: upper_unit_Rep_compact_basis [symmetric])
       
   273 apply (rule unit)
       
   274 apply (simp only: upper_unit_Rep_compact_basis [symmetric]
       
   275                   upper_plus_principal [symmetric])
       
   276 apply (erule insert [OF unit])
       
   277 done
       
   278 
       
   279 lemma upper_pd_induct
       
   280   [case_names adm upper_unit upper_plus, induct type: upper_pd]:
       
   281   assumes P: "adm P"
       
   282   assumes unit: "\<And>x. P {x}\<sharp>"
       
   283   assumes plus: "\<And>xs ys. \<lbrakk>P xs; P ys\<rbrakk> \<Longrightarrow> P (xs +\<sharp> ys)"
       
   284   shows "P (xs::'a upper_pd)"
       
   285 apply (induct xs rule: upper_pd.principal_induct, rule P)
       
   286 apply (induct_tac a rule: pd_basis_induct)
       
   287 apply (simp only: upper_unit_Rep_compact_basis [symmetric] unit)
       
   288 apply (simp only: upper_plus_principal [symmetric] plus)
       
   289 done
       
   290 
       
   291 
       
   292 subsection {* Monadic bind *}
       
   293 
       
   294 definition
       
   295   upper_bind_basis ::
       
   296   "'a pd_basis \<Rightarrow> ('a \<rightarrow> 'b upper_pd) \<rightarrow> 'b upper_pd" where
       
   297   "upper_bind_basis = fold_pd
       
   298     (\<lambda>a. \<Lambda> f. f\<cdot>(Rep_compact_basis a))
       
   299     (\<lambda>x y. \<Lambda> f. x\<cdot>f +\<sharp> y\<cdot>f)"
       
   300 
       
   301 lemma ACI_upper_bind:
       
   302   "class.ab_semigroup_idem_mult (\<lambda>x y. \<Lambda> f. x\<cdot>f +\<sharp> y\<cdot>f)"
       
   303 apply unfold_locales
       
   304 apply (simp add: upper_plus_assoc)
       
   305 apply (simp add: upper_plus_commute)
       
   306 apply (simp add: eta_cfun)
       
   307 done
       
   308 
       
   309 lemma upper_bind_basis_simps [simp]:
       
   310   "upper_bind_basis (PDUnit a) =
       
   311     (\<Lambda> f. f\<cdot>(Rep_compact_basis a))"
       
   312   "upper_bind_basis (PDPlus t u) =
       
   313     (\<Lambda> f. upper_bind_basis t\<cdot>f +\<sharp> upper_bind_basis u\<cdot>f)"
       
   314 unfolding upper_bind_basis_def
       
   315 apply -
       
   316 apply (rule fold_pd_PDUnit [OF ACI_upper_bind])
       
   317 apply (rule fold_pd_PDPlus [OF ACI_upper_bind])
       
   318 done
       
   319 
       
   320 lemma upper_bind_basis_mono:
       
   321   "t \<le>\<sharp> u \<Longrightarrow> upper_bind_basis t \<sqsubseteq> upper_bind_basis u"
       
   322 unfolding cfun_below_iff
       
   323 apply (erule upper_le_induct, safe)
       
   324 apply (simp add: monofun_cfun)
       
   325 apply (simp add: below_trans [OF upper_plus_below1])
       
   326 apply simp
       
   327 done
       
   328 
       
   329 definition
       
   330   upper_bind :: "'a upper_pd \<rightarrow> ('a \<rightarrow> 'b upper_pd) \<rightarrow> 'b upper_pd" where
       
   331   "upper_bind = upper_pd.basis_fun upper_bind_basis"
       
   332 
       
   333 lemma upper_bind_principal [simp]:
       
   334   "upper_bind\<cdot>(upper_principal t) = upper_bind_basis t"
       
   335 unfolding upper_bind_def
       
   336 apply (rule upper_pd.basis_fun_principal)
       
   337 apply (erule upper_bind_basis_mono)
       
   338 done
       
   339 
       
   340 lemma upper_bind_unit [simp]:
       
   341   "upper_bind\<cdot>{x}\<sharp>\<cdot>f = f\<cdot>x"
       
   342 by (induct x rule: compact_basis.principal_induct, simp, simp)
       
   343 
       
   344 lemma upper_bind_plus [simp]:
       
   345   "upper_bind\<cdot>(xs +\<sharp> ys)\<cdot>f = upper_bind\<cdot>xs\<cdot>f +\<sharp> upper_bind\<cdot>ys\<cdot>f"
       
   346 by (induct xs ys rule: upper_pd.principal_induct2, simp, simp, simp)
       
   347 
       
   348 lemma upper_bind_strict [simp]: "upper_bind\<cdot>\<bottom>\<cdot>f = f\<cdot>\<bottom>"
       
   349 unfolding upper_unit_strict [symmetric] by (rule upper_bind_unit)
       
   350 
       
   351 lemma upper_bind_bind:
       
   352   "upper_bind\<cdot>(upper_bind\<cdot>xs\<cdot>f)\<cdot>g = upper_bind\<cdot>xs\<cdot>(\<Lambda> x. upper_bind\<cdot>(f\<cdot>x)\<cdot>g)"
       
   353 by (induct xs, simp_all)
       
   354 
       
   355 
       
   356 subsection {* Map *}
       
   357 
       
   358 definition
       
   359   upper_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a upper_pd \<rightarrow> 'b upper_pd" where
       
   360   "upper_map = (\<Lambda> f xs. upper_bind\<cdot>xs\<cdot>(\<Lambda> x. {f\<cdot>x}\<sharp>))"
       
   361 
       
   362 lemma upper_map_unit [simp]:
       
   363   "upper_map\<cdot>f\<cdot>{x}\<sharp> = {f\<cdot>x}\<sharp>"
       
   364 unfolding upper_map_def by simp
       
   365 
       
   366 lemma upper_map_plus [simp]:
       
   367   "upper_map\<cdot>f\<cdot>(xs +\<sharp> ys) = upper_map\<cdot>f\<cdot>xs +\<sharp> upper_map\<cdot>f\<cdot>ys"
       
   368 unfolding upper_map_def by simp
       
   369 
       
   370 lemma upper_map_bottom [simp]: "upper_map\<cdot>f\<cdot>\<bottom> = {f\<cdot>\<bottom>}\<sharp>"
       
   371 unfolding upper_map_def by simp
       
   372 
       
   373 lemma upper_map_ident: "upper_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs"
       
   374 by (induct xs rule: upper_pd_induct, simp_all)
       
   375 
       
   376 lemma upper_map_ID: "upper_map\<cdot>ID = ID"
       
   377 by (simp add: cfun_eq_iff ID_def upper_map_ident)
       
   378 
       
   379 lemma upper_map_map:
       
   380   "upper_map\<cdot>f\<cdot>(upper_map\<cdot>g\<cdot>xs) = upper_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs"
       
   381 by (induct xs rule: upper_pd_induct, simp_all)
       
   382 
       
   383 lemma ep_pair_upper_map: "ep_pair e p \<Longrightarrow> ep_pair (upper_map\<cdot>e) (upper_map\<cdot>p)"
       
   384 apply default
       
   385 apply (induct_tac x rule: upper_pd_induct, simp_all add: ep_pair.e_inverse)
       
   386 apply (induct_tac y rule: upper_pd_induct)
       
   387 apply (simp_all add: ep_pair.e_p_below monofun_cfun del: upper_below_plus_iff)
       
   388 done
       
   389 
       
   390 lemma deflation_upper_map: "deflation d \<Longrightarrow> deflation (upper_map\<cdot>d)"
       
   391 apply default
       
   392 apply (induct_tac x rule: upper_pd_induct, simp_all add: deflation.idem)
       
   393 apply (induct_tac x rule: upper_pd_induct)
       
   394 apply (simp_all add: deflation.below monofun_cfun del: upper_below_plus_iff)
       
   395 done
       
   396 
       
   397 (* FIXME: long proof! *)
       
   398 lemma finite_deflation_upper_map:
       
   399   assumes "finite_deflation d" shows "finite_deflation (upper_map\<cdot>d)"
       
   400 proof (rule finite_deflation_intro)
       
   401   interpret d: finite_deflation d by fact
       
   402   have "deflation d" by fact
       
   403   thus "deflation (upper_map\<cdot>d)" by (rule deflation_upper_map)
       
   404   have "finite (range (\<lambda>x. d\<cdot>x))" by (rule d.finite_range)
       
   405   hence "finite (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))"
       
   406     by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject)
       
   407   hence "finite (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x)))" by simp
       
   408   hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))"
       
   409     by (rule finite_vimageI, simp add: inj_on_def Rep_pd_basis_inject)
       
   410   hence *: "finite (upper_principal ` Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))" by simp
       
   411   hence "finite (range (\<lambda>xs. upper_map\<cdot>d\<cdot>xs))"
       
   412     apply (rule rev_finite_subset)
       
   413     apply clarsimp
       
   414     apply (induct_tac xs rule: upper_pd.principal_induct)
       
   415     apply (simp add: adm_mem_finite *)
       
   416     apply (rename_tac t, induct_tac t rule: pd_basis_induct)
       
   417     apply (simp only: upper_unit_Rep_compact_basis [symmetric] upper_map_unit)
       
   418     apply simp
       
   419     apply (subgoal_tac "\<exists>b. d\<cdot>(Rep_compact_basis a) = Rep_compact_basis b")
       
   420     apply clarsimp
       
   421     apply (rule imageI)
       
   422     apply (rule vimageI2)
       
   423     apply (simp add: Rep_PDUnit)
       
   424     apply (rule range_eqI)
       
   425     apply (erule sym)
       
   426     apply (rule exI)
       
   427     apply (rule Abs_compact_basis_inverse [symmetric])
       
   428     apply (simp add: d.compact)
       
   429     apply (simp only: upper_plus_principal [symmetric] upper_map_plus)
       
   430     apply clarsimp
       
   431     apply (rule imageI)
       
   432     apply (rule vimageI2)
       
   433     apply (simp add: Rep_PDPlus)
       
   434     done
       
   435   thus "finite {xs. upper_map\<cdot>d\<cdot>xs = xs}"
       
   436     by (rule finite_range_imp_finite_fixes)
       
   437 qed
       
   438 
       
   439 subsection {* Upper powerdomain is a domain *}
       
   440 
       
   441 definition
       
   442   upper_approx :: "nat \<Rightarrow> udom upper_pd \<rightarrow> udom upper_pd"
       
   443 where
       
   444   "upper_approx = (\<lambda>i. upper_map\<cdot>(udom_approx i))"
       
   445 
       
   446 lemma upper_approx: "approx_chain upper_approx"
       
   447 using upper_map_ID finite_deflation_upper_map
       
   448 unfolding upper_approx_def by (rule approx_chain_lemma1)
       
   449 
       
   450 definition upper_defl :: "defl \<rightarrow> defl"
       
   451 where "upper_defl = defl_fun1 upper_approx upper_map"
       
   452 
       
   453 lemma cast_upper_defl:
       
   454   "cast\<cdot>(upper_defl\<cdot>A) =
       
   455     udom_emb upper_approx oo upper_map\<cdot>(cast\<cdot>A) oo udom_prj upper_approx"
       
   456 using upper_approx finite_deflation_upper_map
       
   457 unfolding upper_defl_def by (rule cast_defl_fun1)
       
   458 
       
   459 instantiation upper_pd :: ("domain") liftdomain
       
   460 begin
       
   461 
       
   462 definition
       
   463   "emb = udom_emb upper_approx oo upper_map\<cdot>emb"
       
   464 
       
   465 definition
       
   466   "prj = upper_map\<cdot>prj oo udom_prj upper_approx"
       
   467 
       
   468 definition
       
   469   "defl (t::'a upper_pd itself) = upper_defl\<cdot>DEFL('a)"
       
   470 
       
   471 definition
       
   472   "(liftemb :: 'a upper_pd u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   473 
       
   474 definition
       
   475   "(liftprj :: udom \<rightarrow> 'a upper_pd u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   476 
       
   477 definition
       
   478   "liftdefl (t::'a upper_pd itself) = u_defl\<cdot>DEFL('a upper_pd)"
       
   479 
       
   480 instance
       
   481 using liftemb_upper_pd_def liftprj_upper_pd_def liftdefl_upper_pd_def
       
   482 proof (rule liftdomain_class_intro)
       
   483   show "ep_pair emb (prj :: udom \<rightarrow> 'a upper_pd)"
       
   484     unfolding emb_upper_pd_def prj_upper_pd_def
       
   485     using ep_pair_udom [OF upper_approx]
       
   486     by (intro ep_pair_comp ep_pair_upper_map ep_pair_emb_prj)
       
   487 next
       
   488   show "cast\<cdot>DEFL('a upper_pd) = emb oo (prj :: udom \<rightarrow> 'a upper_pd)"
       
   489     unfolding emb_upper_pd_def prj_upper_pd_def defl_upper_pd_def cast_upper_defl
       
   490     by (simp add: cast_DEFL oo_def cfun_eq_iff upper_map_map)
       
   491 qed
       
   492 
       
   493 end
       
   494 
       
   495 lemma DEFL_upper: "DEFL('a upper_pd) = upper_defl\<cdot>DEFL('a)"
       
   496 by (rule defl_upper_pd_def)
       
   497 
       
   498 
       
   499 subsection {* Join *}
       
   500 
       
   501 definition
       
   502   upper_join :: "'a upper_pd upper_pd \<rightarrow> 'a upper_pd" where
       
   503   "upper_join = (\<Lambda> xss. upper_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))"
       
   504 
       
   505 lemma upper_join_unit [simp]:
       
   506   "upper_join\<cdot>{xs}\<sharp> = xs"
       
   507 unfolding upper_join_def by simp
       
   508 
       
   509 lemma upper_join_plus [simp]:
       
   510   "upper_join\<cdot>(xss +\<sharp> yss) = upper_join\<cdot>xss +\<sharp> upper_join\<cdot>yss"
       
   511 unfolding upper_join_def by simp
       
   512 
       
   513 lemma upper_join_bottom [simp]: "upper_join\<cdot>\<bottom> = \<bottom>"
       
   514 unfolding upper_join_def by simp
       
   515 
       
   516 lemma upper_join_map_unit:
       
   517   "upper_join\<cdot>(upper_map\<cdot>upper_unit\<cdot>xs) = xs"
       
   518 by (induct xs rule: upper_pd_induct, simp_all)
       
   519 
       
   520 lemma upper_join_map_join:
       
   521   "upper_join\<cdot>(upper_map\<cdot>upper_join\<cdot>xsss) = upper_join\<cdot>(upper_join\<cdot>xsss)"
       
   522 by (induct xsss rule: upper_pd_induct, simp_all)
       
   523 
       
   524 lemma upper_join_map_map:
       
   525   "upper_join\<cdot>(upper_map\<cdot>(upper_map\<cdot>f)\<cdot>xss) =
       
   526    upper_map\<cdot>f\<cdot>(upper_join\<cdot>xss)"
       
   527 by (induct xss rule: upper_pd_induct, simp_all)
       
   528 
       
   529 end