src/HOLCF/Bifinite.thy
changeset 27186 416d66c36d8f
parent 26962 c8b20f615d6c
child 27309 c74270fd72a8
     1.1 --- a/src/HOLCF/Bifinite.thy	Thu Jun 12 22:30:00 2008 +0200
     1.2 +++ b/src/HOLCF/Bifinite.thy	Thu Jun 12 22:41:03 2008 +0200
     1.3 @@ -27,18 +27,17 @@
     1.4  apply (clarify, erule subst, rule exI, rule refl)
     1.5  done
     1.6  
     1.7 -lemma chain_approx [simp]:
     1.8 -  "chain (approx :: nat \<Rightarrow> 'a::profinite \<rightarrow> 'a)"
     1.9 +lemma chain_approx [simp]: "chain approx"
    1.10  apply (rule chainI)
    1.11  apply (rule less_cfun_ext)
    1.12  apply (rule chainE)
    1.13  apply (rule chain_approx_app)
    1.14  done
    1.15  
    1.16 -lemma lub_approx [simp]: "(\<Squnion>i. approx i) = (\<Lambda>(x::'a::profinite). x)"
    1.17 +lemma lub_approx [simp]: "(\<Squnion>i. approx i) = (\<Lambda> x. x)"
    1.18  by (rule ext_cfun, simp add: contlub_cfun_fun)
    1.19  
    1.20 -lemma approx_less: "approx i\<cdot>x \<sqsubseteq> (x::'a::profinite)"
    1.21 +lemma approx_less: "approx i\<cdot>x \<sqsubseteq> x"
    1.22  apply (subgoal_tac "approx i\<cdot>x \<sqsubseteq> (\<Squnion>i. approx i\<cdot>x)", simp)
    1.23  apply (rule is_ub_thelub, simp)
    1.24  done
    1.25 @@ -47,7 +46,7 @@
    1.26  by (rule UU_I, rule approx_less)
    1.27  
    1.28  lemma approx_approx1:
    1.29 -  "i \<le> j \<Longrightarrow> approx i\<cdot>(approx j\<cdot>x) = approx i\<cdot>(x::'a::profinite)"
    1.30 +  "i \<le> j \<Longrightarrow> approx i\<cdot>(approx j\<cdot>x) = approx i\<cdot>x"
    1.31  apply (rule antisym_less)
    1.32  apply (rule monofun_cfun_arg [OF approx_less])
    1.33  apply (rule sq_ord_eq_less_trans [OF approx_idem [symmetric]])
    1.34 @@ -57,7 +56,7 @@
    1.35  done
    1.36  
    1.37  lemma approx_approx2:
    1.38 -  "j \<le> i \<Longrightarrow> approx i\<cdot>(approx j\<cdot>x) = approx j\<cdot>(x::'a::profinite)"
    1.39 +  "j \<le> i \<Longrightarrow> approx i\<cdot>(approx j\<cdot>x) = approx j\<cdot>x"
    1.40  apply (rule antisym_less)
    1.41  apply (rule approx_less)
    1.42  apply (rule sq_ord_eq_less_trans [OF approx_idem [symmetric]])
    1.43 @@ -66,7 +65,7 @@
    1.44  done
    1.45  
    1.46  lemma approx_approx [simp]:
    1.47 -  "approx i\<cdot>(approx j\<cdot>x) = approx (min i j)\<cdot>(x::'a::profinite)"
    1.48 +  "approx i\<cdot>(approx j\<cdot>x) = approx (min i j)\<cdot>x"
    1.49  apply (rule_tac x=i and y=j in linorder_le_cases)
    1.50  apply (simp add: approx_approx1 min_def)
    1.51  apply (simp add: approx_approx2 min_def)
    1.52 @@ -76,16 +75,16 @@
    1.53    "\<forall>x. f (f x) = f x \<Longrightarrow> {x. f x = x} = {y. \<exists>x. y = f x}"
    1.54  by (auto simp add: eq_sym_conv)
    1.55  
    1.56 -lemma finite_approx: "finite {y::'a::profinite. \<exists>x. y = approx n\<cdot>x}"
    1.57 +lemma finite_approx: "finite {y. \<exists>x. y = approx n\<cdot>x}"
    1.58  using finite_fixes_approx by (simp add: idem_fixes_eq_range)
    1.59  
    1.60 -lemma finite_range_approx:
    1.61 -  "finite (range (\<lambda>x::'a::profinite. approx n\<cdot>x))"
    1.62 -by (simp add: image_def finite_approx)
    1.63 +lemma finite_image_approx: "finite ((\<lambda>x. approx n\<cdot>x) ` A)"
    1.64 +by (rule finite_subset [OF _ finite_fixes_approx [where i=n]]) auto
    1.65  
    1.66 -lemma compact_approx [simp]:
    1.67 -  fixes x :: "'a::profinite"
    1.68 -  shows "compact (approx n\<cdot>x)"
    1.69 +lemma finite_range_approx: "finite (range (\<lambda>x. approx n\<cdot>x))"
    1.70 +by (rule finite_image_approx)
    1.71 +
    1.72 +lemma compact_approx [simp]: "compact (approx n\<cdot>x)"
    1.73  proof (rule compactI2)
    1.74    fix Y::"nat \<Rightarrow> 'a"
    1.75    assume Y: "chain Y"
    1.76 @@ -115,7 +114,6 @@
    1.77  qed
    1.78  
    1.79  lemma bifinite_compact_eq_approx:
    1.80 -  fixes x :: "'a::profinite"
    1.81    assumes x: "compact x"
    1.82    shows "\<exists>i. approx i\<cdot>x = x"
    1.83  proof -
    1.84 @@ -129,7 +127,7 @@
    1.85  qed
    1.86  
    1.87  lemma bifinite_compact_iff:
    1.88 -  "compact (x::'a::profinite) = (\<exists>n. approx n\<cdot>x = x)"
    1.89 +  "compact x \<longleftrightarrow> (\<exists>n. approx n\<cdot>x = x)"
    1.90   apply (rule iffI)
    1.91    apply (erule bifinite_compact_eq_approx)
    1.92   apply (erule exE)
    1.93 @@ -139,16 +137,14 @@
    1.94  
    1.95  lemma approx_induct:
    1.96    assumes adm: "adm P" and P: "\<And>n x. P (approx n\<cdot>x)"
    1.97 -  shows "P (x::'a::profinite)"
    1.98 +  shows "P x"
    1.99  proof -
   1.100    have "P (\<Squnion>n. approx n\<cdot>x)"
   1.101      by (rule admD [OF adm], simp, simp add: P)
   1.102    thus "P x" by simp
   1.103  qed
   1.104  
   1.105 -lemma bifinite_less_ext:
   1.106 -  fixes x y :: "'a::profinite"
   1.107 -  shows "(\<And>i. approx i\<cdot>x \<sqsubseteq> approx i\<cdot>y) \<Longrightarrow> x \<sqsubseteq> y"
   1.108 +lemma bifinite_less_ext: "(\<And>i. approx i\<cdot>x \<sqsubseteq> approx i\<cdot>y) \<Longrightarrow> x \<sqsubseteq> y"
   1.109  apply (subgoal_tac "(\<Squnion>i. approx i\<cdot>x) \<sqsubseteq> (\<Squnion>i. approx i\<cdot>y)", simp)
   1.110  apply (rule lub_mono, simp, simp, simp)
   1.111  done