src/HOLCF/ConvexPD.thy
changeset 27309 c74270fd72a8
parent 27297 2c42b1505f25
child 27310 d0229bc6c461
     1.1 --- a/src/HOLCF/ConvexPD.thy	Fri Jun 20 22:41:41 2008 +0200
     1.2 +++ b/src/HOLCF/ConvexPD.thy	Fri Jun 20 22:51:50 2008 +0200
     1.3 @@ -327,7 +327,7 @@
     1.4  
     1.5  lemma convex_unit_less_iff [simp]: "{x}\<natural> \<sqsubseteq> {y}\<natural> \<longleftrightarrow> x \<sqsubseteq> y"
     1.6   apply (rule iffI)
     1.7 -  apply (rule bifinite_less_ext)
     1.8 +  apply (rule profinite_less_ext)
     1.9    apply (drule_tac f="approx i" in monofun_cfun_arg, simp)
    1.10    apply (cut_tac x="approx i\<cdot>x" in compact_basis.compact_imp_principal, simp)
    1.11    apply (cut_tac x="approx i\<cdot>y" in compact_basis.compact_imp_principal, simp)
    1.12 @@ -346,7 +346,7 @@
    1.13  
    1.14  lemma compact_convex_unit_iff [simp]:
    1.15    "compact {x}\<natural> \<longleftrightarrow> compact x"
    1.16 -unfolding bifinite_compact_iff by simp
    1.17 +unfolding profinite_compact_iff by simp
    1.18  
    1.19  lemma compact_convex_plus [simp]:
    1.20    "\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs +\<natural> ys)"
    1.21 @@ -589,7 +589,7 @@
    1.22      (convex_to_upper\<cdot>xs \<sqsubseteq> convex_to_upper\<cdot>ys \<and>
    1.23       convex_to_lower\<cdot>xs \<sqsubseteq> convex_to_lower\<cdot>ys)"
    1.24   apply (safe elim!: monofun_cfun_arg)
    1.25 - apply (rule bifinite_less_ext)
    1.26 + apply (rule profinite_less_ext)
    1.27   apply (drule_tac f="approx i" in monofun_cfun_arg)
    1.28   apply (drule_tac f="approx i" in monofun_cfun_arg)
    1.29   apply (cut_tac x="approx i\<cdot>xs" in convex_pd.compact_imp_principal, simp)