src/HOLCF/Pcpo.thy
changeset 31071 845a6acd3bf3
parent 31024 0fdf666e08bf
child 31076 99fe356cbbc2
--- a/src/HOLCF/Pcpo.thy	Fri May 08 10:59:11 2009 +0200
+++ b/src/HOLCF/Pcpo.thy	Fri May 08 13:34:27 2009 +0200
@@ -13,28 +13,28 @@
 text {* The class cpo of chain complete partial orders *}
 
 class cpo = po +
-        -- {* class axiom: *}
-  assumes cpo:   "chain S \<Longrightarrow> \<exists>x :: 'a::po. range S <<| x"
+  assumes cpo: "chain S \<Longrightarrow> \<exists>x. range S <<| x"
+begin
 
 text {* in cpo's everthing equal to THE lub has lub properties for every chain *}
 
-lemma cpo_lubI: "chain (S::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> range S <<| (\<Squnion>i. S i)"
-by (fast dest: cpo elim: lubI)
+lemma cpo_lubI: "chain S \<Longrightarrow> range S <<| (\<Squnion>i. S i)"
+  by (fast dest: cpo elim: lubI)
 
-lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = (l::'a::cpo)\<rbrakk> \<Longrightarrow> range S <<| l"
-by (blast dest: cpo intro: lubI)
+lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = l\<rbrakk> \<Longrightarrow> range S <<| l"
+  by (blast dest: cpo intro: lubI)
 
 text {* Properties of the lub *}
 
-lemma is_ub_thelub: "chain (S::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)"
-by (blast dest: cpo intro: lubI [THEN is_ub_lub])
+lemma is_ub_thelub: "chain S \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)"
+  by (blast dest: cpo intro: lubI [THEN is_ub_lub])
 
 lemma is_lub_thelub:
-  "\<lbrakk>chain (S::nat \<Rightarrow> 'a::cpo); range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x"
-by (blast dest: cpo intro: lubI [THEN is_lub_lub])
+  "\<lbrakk>chain S; range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x"
+  by (blast dest: cpo intro: lubI [THEN is_lub_lub])
 
 lemma lub_range_mono:
-  "\<lbrakk>range X \<subseteq> range Y; chain Y; chain (X::nat \<Rightarrow> 'a::cpo)\<rbrakk>
+  "\<lbrakk>range X \<subseteq> range Y; chain Y; chain X\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)"
 apply (erule is_lub_thelub)
 apply (rule ub_rangeI)
@@ -45,7 +45,7 @@
 done
 
 lemma lub_range_shift:
-  "chain (Y::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)"
+  "chain Y \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)"
 apply (rule antisym_less)
 apply (rule lub_range_mono)
 apply    fast
@@ -62,7 +62,7 @@
 done
 
 lemma maxinch_is_thelub:
-  "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = ((Y i)::'a::cpo))"
+  "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = Y i)"
 apply (rule iffI)
 apply (fast intro!: thelubI lub_finch1)
 apply (unfold max_in_chain_def)
@@ -75,7 +75,7 @@
 text {* the @{text "\<sqsubseteq>"} relation between two chains is preserved by their lubs *}
 
 lemma lub_mono:
-  "\<lbrakk>chain (X::nat \<Rightarrow> 'a::cpo); chain Y; \<And>i. X i \<sqsubseteq> Y i\<rbrakk> 
+  "\<lbrakk>chain X; chain Y; \<And>i. X i \<sqsubseteq> Y i\<rbrakk> 
     \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)"
 apply (erule is_lub_thelub)
 apply (rule ub_rangeI)
@@ -87,14 +87,14 @@
 text {* the = relation between two chains is preserved by their lubs *}
 
 lemma lub_equal:
-  "\<lbrakk>chain (X::nat \<Rightarrow> 'a::cpo); chain Y; \<forall>k. X k = Y k\<rbrakk>
+  "\<lbrakk>chain X; chain Y; \<forall>k. X k = Y k\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)"
-by (simp only: expand_fun_eq [symmetric])
+  by (simp only: expand_fun_eq [symmetric])
 
 text {* more results about mono and = of lubs of chains *}
 
 lemma lub_mono2:
-  "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain (X::nat \<Rightarrow> 'a::cpo); chain Y\<rbrakk>
+  "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain X; chain Y\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)"
 apply (erule exE)
 apply (subgoal_tac "(\<Squnion>i. X (i + Suc j)) \<sqsubseteq> (\<Squnion>i. Y (i + Suc j))")
@@ -104,12 +104,12 @@
 done
 
 lemma lub_equal2:
-  "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain (X::nat \<Rightarrow> 'a::cpo); chain Y\<rbrakk>
+  "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain X; chain Y\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)"
-by (blast intro: antisym_less lub_mono2 sym)
+  by (blast intro: antisym_less lub_mono2 sym)
 
 lemma lub_mono3:
-  "\<lbrakk>chain (Y::nat \<Rightarrow> 'a::cpo); chain X; \<forall>i. \<exists>j. Y i \<sqsubseteq> X j\<rbrakk>
+  "\<lbrakk>chain Y; chain X; \<forall>i. \<exists>j. Y i \<sqsubseteq> X j\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. X i)"
 apply (erule is_lub_thelub)
 apply (rule ub_rangeI)
@@ -120,7 +120,6 @@
 done
 
 lemma ch2ch_lub:
-  fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo"
   assumes 1: "\<And>j. chain (\<lambda>i. Y i j)"
   assumes 2: "\<And>i. chain (\<lambda>j. Y i j)"
   shows "chain (\<lambda>i. \<Squnion>j. Y i j)"
@@ -130,7 +129,6 @@
 done
 
 lemma diag_lub:
-  fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo"
   assumes 1: "\<And>j. chain (\<lambda>i. Y i j)"
   assumes 2: "\<And>i. chain (\<lambda>j. Y i j)"
   shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>i. Y i i)"
@@ -159,12 +157,12 @@
 qed
 
 lemma ex_lub:
-  fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo"
   assumes 1: "\<And>j. chain (\<lambda>i. Y i j)"
   assumes 2: "\<And>i. chain (\<lambda>j. Y i j)"
   shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)"
-by (simp add: diag_lub 1 2)
+  by (simp add: diag_lub 1 2)
 
+end
 
 subsection {* Pointed cpos *}
 
@@ -172,9 +170,9 @@
 
 class pcpo = cpo +
   assumes least: "\<exists>x. \<forall>y. x \<sqsubseteq> y"
+begin
 
-definition
-  UU :: "'a::pcpo" where
+definition UU :: 'a where
   "UU = (THE x. \<forall>y. x \<sqsubseteq> y)"
 
 notation (xsymbols)
@@ -193,6 +191,8 @@
 lemma minimal [iff]: "\<bottom> \<sqsubseteq> x"
 by (rule UU_least [THEN spec])
 
+end
+
 text {* Simproc to rewrite @{term "\<bottom> = x"} to @{term "x = \<bottom>"}. *}
 
 setup {*
@@ -202,6 +202,9 @@
 
 simproc_setup reorient_bottom ("\<bottom> = x") = ReorientProc.proc
 
+context pcpo
+begin
+
 text {* useful lemmas about @{term \<bottom>} *}
 
 lemma less_UU_iff [simp]: "(x \<sqsubseteq> \<bottom>) = (x = \<bottom>)"
@@ -213,9 +216,6 @@
 lemma UU_I: "x \<sqsubseteq> \<bottom> \<Longrightarrow> x = \<bottom>"
 by (subst eq_UU_iff)
 
-lemma not_less2not_eq: "\<not> (x::'a::po) \<sqsubseteq> y \<Longrightarrow> x \<noteq> y"
-by auto
-
 lemma chain_UU_I: "\<lbrakk>chain Y; (\<Squnion>i. Y i) = \<bottom>\<rbrakk> \<Longrightarrow> \<forall>i. Y i = \<bottom>"
 apply (rule allI)
 apply (rule UU_I)
@@ -230,49 +230,53 @@
 done
 
 lemma chain_UU_I_inverse2: "(\<Squnion>i. Y i) \<noteq> \<bottom> \<Longrightarrow> \<exists>i::nat. Y i \<noteq> \<bottom>"
-by (blast intro: chain_UU_I_inverse)
+  by (blast intro: chain_UU_I_inverse)
 
 lemma notUU_I: "\<lbrakk>x \<sqsubseteq> y; x \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> y \<noteq> \<bottom>"
-by (blast intro: UU_I)
+  by (blast intro: UU_I)
 
 lemma chain_mono2: "\<lbrakk>\<exists>j. Y j \<noteq> \<bottom>; chain Y\<rbrakk> \<Longrightarrow> \<exists>j. \<forall>i>j. Y i \<noteq> \<bottom>"
-by (blast dest: notUU_I chain_mono_less)
+  by (blast dest: notUU_I chain_mono_less)
+
+end
 
 subsection {* Chain-finite and flat cpos *}
 
 text {* further useful classes for HOLCF domains *}
 
-class finite_po = finite + po
+class chfin = po +
+  assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n Y"
+begin
 
-class chfin = po +
-  assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n (Y :: nat => 'a::po)"
+subclass cpo
+apply default
+apply (frule chfin)
+apply (blast intro: lub_finch1)
+done
 
-class flat = pcpo +
-  assumes ax_flat: "(x :: 'a::pcpo) \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y"
+lemma chfin2finch: "chain Y \<Longrightarrow> finite_chain Y"
+  by (simp add: chfin finite_chain_def)
+
+end
 
-text {* finite partial orders are chain-finite *}
+class finite_po = finite + po
+begin
 
-instance finite_po < chfin
-apply intro_classes
+subclass chfin
+apply default
 apply (drule finite_range_imp_finch)
 apply (rule finite)
 apply (simp add: finite_chain_def)
 done
 
-text {* some properties for chfin and flat *}
-
-text {* chfin types are cpo *}
+end
 
-instance chfin < cpo
-apply intro_classes
-apply (frule chfin)
-apply (blast intro: lub_finch1)
-done
+class flat = pcpo +
+  assumes ax_flat: "x \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y"
+begin
 
-text {* flat types are chfin *}
-
-instance flat < chfin
-apply intro_classes
+subclass chfin
+apply default
 apply (unfold max_in_chain_def)
 apply (case_tac "\<forall>i. Y i = \<bottom>")
 apply simp
@@ -283,31 +287,28 @@
 apply (blast dest: chain_mono ax_flat)
 done
 
-text {* flat subclass of chfin; @{text adm_flat} not needed *}
-
 lemma flat_less_iff:
-  fixes x y :: "'a::flat"
   shows "(x \<sqsubseteq> y) = (x = \<bottom> \<or> x = y)"
-by (safe dest!: ax_flat)
+  by (safe dest!: ax_flat)
 
-lemma flat_eq: "(a::'a::flat) \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)"
-by (safe dest!: ax_flat)
+lemma flat_eq: "a \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)"
+  by (safe dest!: ax_flat)
 
-lemma chfin2finch: "chain (Y::nat \<Rightarrow> 'a::chfin) \<Longrightarrow> finite_chain Y"
-by (simp add: chfin finite_chain_def)
+end
 
 text {* Discrete cpos *}
 
 class discrete_cpo = sq_ord +
   assumes discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y"
+begin
 
-subclass (in discrete_cpo) po
+subclass po
 proof qed simp_all
 
 text {* In a discrete cpo, every chain is constant *}
 
 lemma discrete_chain_const:
-  assumes S: "chain (S::nat \<Rightarrow> 'a::discrete_cpo)"
+  assumes S: "chain S"
   shows "\<exists>x. S = (\<lambda>i. x)"
 proof (intro exI ext)
   fix i :: nat
@@ -316,7 +317,7 @@
   thus "S i = S 0" by (rule sym)
 qed
 
-instance discrete_cpo < cpo
+subclass cpo
 proof
   fix S :: "nat \<Rightarrow> 'a"
   assume S: "chain S"
@@ -326,31 +327,6 @@
     by (fast intro: lub_const)
 qed
 
-text {* lemmata for improved admissibility introdution rule *}
-
-lemma infinite_chain_adm_lemma:
-  "\<lbrakk>chain Y; \<forall>i. P (Y i);  
-    \<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i); \<not> finite_chain Y\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk>
-      \<Longrightarrow> P (\<Squnion>i. Y i)"
-apply (case_tac "finite_chain Y")
-prefer 2 apply fast
-apply (unfold finite_chain_def)
-apply safe
-apply (erule lub_finch1 [THEN thelubI, THEN ssubst])
-apply assumption
-apply (erule spec)
-done
-
-lemma increasing_chain_adm_lemma:
-  "\<lbrakk>chain Y;  \<forall>i. P (Y i); \<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i);
-    \<forall>i. \<exists>j>i. Y i \<noteq> Y j \<and> Y i \<sqsubseteq> Y j\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk>
-      \<Longrightarrow> P (\<Squnion>i. Y i)"
-apply (erule infinite_chain_adm_lemma)
-apply assumption
-apply (erule thin_rl)
-apply (unfold finite_chain_def)
-apply (unfold max_in_chain_def)
-apply (fast dest: le_imp_less_or_eq elim: chain_mono_less)
-done
+end
 
 end