src/HOLCF/Product_Cpo.thy
changeset 41022 0437dbc127b3
parent 41021 6c12f5e24e34
child 41023 ed7a4eadb2f6
     1.1 --- a/src/HOLCF/Product_Cpo.thy	Sat Nov 27 14:34:54 2010 -0800
     1.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.3 @@ -1,299 +0,0 @@
     1.4 -(*  Title:      HOLCF/Product_Cpo.thy
     1.5 -    Author:     Franz Regensburger
     1.6 -*)
     1.7 -
     1.8 -header {* The cpo of cartesian products *}
     1.9 -
    1.10 -theory Product_Cpo
    1.11 -imports Adm
    1.12 -begin
    1.13 -
    1.14 -default_sort cpo
    1.15 -
    1.16 -subsection {* Unit type is a pcpo *}
    1.17 -
    1.18 -instantiation unit :: discrete_cpo
    1.19 -begin
    1.20 -
    1.21 -definition
    1.22 -  below_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<longleftrightarrow> True"
    1.23 -
    1.24 -instance proof
    1.25 -qed simp
    1.26 -
    1.27 -end
    1.28 -
    1.29 -instance unit :: pcpo
    1.30 -by intro_classes simp
    1.31 -
    1.32 -
    1.33 -subsection {* Product type is a partial order *}
    1.34 -
    1.35 -instantiation prod :: (below, below) below
    1.36 -begin
    1.37 -
    1.38 -definition
    1.39 -  below_prod_def: "(op \<sqsubseteq>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)"
    1.40 -
    1.41 -instance ..
    1.42 -end
    1.43 -
    1.44 -instance prod :: (po, po) po
    1.45 -proof
    1.46 -  fix x :: "'a \<times> 'b"
    1.47 -  show "x \<sqsubseteq> x"
    1.48 -    unfolding below_prod_def by simp
    1.49 -next
    1.50 -  fix x y :: "'a \<times> 'b"
    1.51 -  assume "x \<sqsubseteq> y" "y \<sqsubseteq> x" thus "x = y"
    1.52 -    unfolding below_prod_def Pair_fst_snd_eq
    1.53 -    by (fast intro: below_antisym)
    1.54 -next
    1.55 -  fix x y z :: "'a \<times> 'b"
    1.56 -  assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" thus "x \<sqsubseteq> z"
    1.57 -    unfolding below_prod_def
    1.58 -    by (fast intro: below_trans)
    1.59 -qed
    1.60 -
    1.61 -subsection {* Monotonicity of \emph{Pair}, \emph{fst}, \emph{snd} *}
    1.62 -
    1.63 -lemma prod_belowI: "\<lbrakk>fst p \<sqsubseteq> fst q; snd p \<sqsubseteq> snd q\<rbrakk> \<Longrightarrow> p \<sqsubseteq> q"
    1.64 -unfolding below_prod_def by simp
    1.65 -
    1.66 -lemma Pair_below_iff [simp]: "(a, b) \<sqsubseteq> (c, d) \<longleftrightarrow> a \<sqsubseteq> c \<and> b \<sqsubseteq> d"
    1.67 -unfolding below_prod_def by simp
    1.68 -
    1.69 -text {* Pair @{text "(_,_)"}  is monotone in both arguments *}
    1.70 -
    1.71 -lemma monofun_pair1: "monofun (\<lambda>x. (x, y))"
    1.72 -by (simp add: monofun_def)
    1.73 -
    1.74 -lemma monofun_pair2: "monofun (\<lambda>y. (x, y))"
    1.75 -by (simp add: monofun_def)
    1.76 -
    1.77 -lemma monofun_pair:
    1.78 -  "\<lbrakk>x1 \<sqsubseteq> x2; y1 \<sqsubseteq> y2\<rbrakk> \<Longrightarrow> (x1, y1) \<sqsubseteq> (x2, y2)"
    1.79 -by simp
    1.80 -
    1.81 -lemma ch2ch_Pair [simp]:
    1.82 -  "chain X \<Longrightarrow> chain Y \<Longrightarrow> chain (\<lambda>i. (X i, Y i))"
    1.83 -by (rule chainI, simp add: chainE)
    1.84 -
    1.85 -text {* @{term fst} and @{term snd} are monotone *}
    1.86 -
    1.87 -lemma fst_monofun: "x \<sqsubseteq> y \<Longrightarrow> fst x \<sqsubseteq> fst y"
    1.88 -unfolding below_prod_def by simp
    1.89 -
    1.90 -lemma snd_monofun: "x \<sqsubseteq> y \<Longrightarrow> snd x \<sqsubseteq> snd y"
    1.91 -unfolding below_prod_def by simp
    1.92 -
    1.93 -lemma monofun_fst: "monofun fst"
    1.94 -by (simp add: monofun_def below_prod_def)
    1.95 -
    1.96 -lemma monofun_snd: "monofun snd"
    1.97 -by (simp add: monofun_def below_prod_def)
    1.98 -
    1.99 -lemmas ch2ch_fst [simp] = ch2ch_monofun [OF monofun_fst]
   1.100 -
   1.101 -lemmas ch2ch_snd [simp] = ch2ch_monofun [OF monofun_snd]
   1.102 -
   1.103 -lemma prod_chain_cases:
   1.104 -  assumes "chain Y"
   1.105 -  obtains A B
   1.106 -  where "chain A" and "chain B" and "Y = (\<lambda>i. (A i, B i))"
   1.107 -proof
   1.108 -  from `chain Y` show "chain (\<lambda>i. fst (Y i))" by (rule ch2ch_fst)
   1.109 -  from `chain Y` show "chain (\<lambda>i. snd (Y i))" by (rule ch2ch_snd)
   1.110 -  show "Y = (\<lambda>i. (fst (Y i), snd (Y i)))" by simp
   1.111 -qed
   1.112 -
   1.113 -subsection {* Product type is a cpo *}
   1.114 -
   1.115 -lemma is_lub_Pair:
   1.116 -  "\<lbrakk>range A <<| x; range B <<| y\<rbrakk> \<Longrightarrow> range (\<lambda>i. (A i, B i)) <<| (x, y)"
   1.117 -unfolding is_lub_def is_ub_def ball_simps below_prod_def by simp
   1.118 -
   1.119 -lemma lub_Pair:
   1.120 -  "\<lbrakk>chain (A::nat \<Rightarrow> 'a::cpo); chain (B::nat \<Rightarrow> 'b::cpo)\<rbrakk>
   1.121 -    \<Longrightarrow> (\<Squnion>i. (A i, B i)) = (\<Squnion>i. A i, \<Squnion>i. B i)"
   1.122 -by (fast intro: lub_eqI is_lub_Pair elim: thelubE)
   1.123 -
   1.124 -lemma is_lub_prod:
   1.125 -  fixes S :: "nat \<Rightarrow> ('a::cpo \<times> 'b::cpo)"
   1.126 -  assumes S: "chain S"
   1.127 -  shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
   1.128 -using S by (auto elim: prod_chain_cases simp add: is_lub_Pair cpo_lubI)
   1.129 -
   1.130 -lemma lub_prod:
   1.131 -  "chain (S::nat \<Rightarrow> 'a::cpo \<times> 'b::cpo)
   1.132 -    \<Longrightarrow> (\<Squnion>i. S i) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
   1.133 -by (rule is_lub_prod [THEN lub_eqI])
   1.134 -
   1.135 -instance prod :: (cpo, cpo) cpo
   1.136 -proof
   1.137 -  fix S :: "nat \<Rightarrow> ('a \<times> 'b)"
   1.138 -  assume "chain S"
   1.139 -  hence "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
   1.140 -    by (rule is_lub_prod)
   1.141 -  thus "\<exists>x. range S <<| x" ..
   1.142 -qed
   1.143 -
   1.144 -instance prod :: (discrete_cpo, discrete_cpo) discrete_cpo
   1.145 -proof
   1.146 -  fix x y :: "'a \<times> 'b"
   1.147 -  show "x \<sqsubseteq> y \<longleftrightarrow> x = y"
   1.148 -    unfolding below_prod_def Pair_fst_snd_eq
   1.149 -    by simp
   1.150 -qed
   1.151 -
   1.152 -subsection {* Product type is pointed *}
   1.153 -
   1.154 -lemma minimal_prod: "(\<bottom>, \<bottom>) \<sqsubseteq> p"
   1.155 -by (simp add: below_prod_def)
   1.156 -
   1.157 -instance prod :: (pcpo, pcpo) pcpo
   1.158 -by intro_classes (fast intro: minimal_prod)
   1.159 -
   1.160 -lemma inst_prod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)"
   1.161 -by (rule minimal_prod [THEN UU_I, symmetric])
   1.162 -
   1.163 -lemma Pair_bottom_iff [simp]: "(x, y) = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>"
   1.164 -unfolding inst_prod_pcpo by simp
   1.165 -
   1.166 -lemma fst_strict [simp]: "fst \<bottom> = \<bottom>"
   1.167 -unfolding inst_prod_pcpo by (rule fst_conv)
   1.168 -
   1.169 -lemma snd_strict [simp]: "snd \<bottom> = \<bottom>"
   1.170 -unfolding inst_prod_pcpo by (rule snd_conv)
   1.171 -
   1.172 -lemma Pair_strict [simp]: "(\<bottom>, \<bottom>) = \<bottom>"
   1.173 -by simp
   1.174 -
   1.175 -lemma split_strict [simp]: "split f \<bottom> = f \<bottom> \<bottom>"
   1.176 -unfolding split_def by simp
   1.177 -
   1.178 -subsection {* Continuity of \emph{Pair}, \emph{fst}, \emph{snd} *}
   1.179 -
   1.180 -lemma cont_pair1: "cont (\<lambda>x. (x, y))"
   1.181 -apply (rule contI)
   1.182 -apply (rule is_lub_Pair)
   1.183 -apply (erule cpo_lubI)
   1.184 -apply (rule is_lub_const)
   1.185 -done
   1.186 -
   1.187 -lemma cont_pair2: "cont (\<lambda>y. (x, y))"
   1.188 -apply (rule contI)
   1.189 -apply (rule is_lub_Pair)
   1.190 -apply (rule is_lub_const)
   1.191 -apply (erule cpo_lubI)
   1.192 -done
   1.193 -
   1.194 -lemma cont_fst: "cont fst"
   1.195 -apply (rule contI)
   1.196 -apply (simp add: lub_prod)
   1.197 -apply (erule cpo_lubI [OF ch2ch_fst])
   1.198 -done
   1.199 -
   1.200 -lemma cont_snd: "cont snd"
   1.201 -apply (rule contI)
   1.202 -apply (simp add: lub_prod)
   1.203 -apply (erule cpo_lubI [OF ch2ch_snd])
   1.204 -done
   1.205 -
   1.206 -lemma cont2cont_Pair [simp, cont2cont]:
   1.207 -  assumes f: "cont (\<lambda>x. f x)"
   1.208 -  assumes g: "cont (\<lambda>x. g x)"
   1.209 -  shows "cont (\<lambda>x. (f x, g x))"
   1.210 -apply (rule cont_apply [OF f cont_pair1])
   1.211 -apply (rule cont_apply [OF g cont_pair2])
   1.212 -apply (rule cont_const)
   1.213 -done
   1.214 -
   1.215 -lemmas cont2cont_fst [simp, cont2cont] = cont_compose [OF cont_fst]
   1.216 -
   1.217 -lemmas cont2cont_snd [simp, cont2cont] = cont_compose [OF cont_snd]
   1.218 -
   1.219 -lemma cont2cont_prod_case:
   1.220 -  assumes f1: "\<And>a b. cont (\<lambda>x. f x a b)"
   1.221 -  assumes f2: "\<And>x b. cont (\<lambda>a. f x a b)"
   1.222 -  assumes f3: "\<And>x a. cont (\<lambda>b. f x a b)"
   1.223 -  assumes g: "cont (\<lambda>x. g x)"
   1.224 -  shows "cont (\<lambda>x. case g x of (a, b) \<Rightarrow> f x a b)"
   1.225 -unfolding split_def
   1.226 -apply (rule cont_apply [OF g])
   1.227 -apply (rule cont_apply [OF cont_fst f2])
   1.228 -apply (rule cont_apply [OF cont_snd f3])
   1.229 -apply (rule cont_const)
   1.230 -apply (rule f1)
   1.231 -done
   1.232 -
   1.233 -lemma prod_contI:
   1.234 -  assumes f1: "\<And>y. cont (\<lambda>x. f (x, y))"
   1.235 -  assumes f2: "\<And>x. cont (\<lambda>y. f (x, y))"
   1.236 -  shows "cont f"
   1.237 -proof -
   1.238 -  have "cont (\<lambda>(x, y). f (x, y))"
   1.239 -    by (intro cont2cont_prod_case f1 f2 cont2cont)
   1.240 -  thus "cont f"
   1.241 -    by (simp only: split_eta)
   1.242 -qed
   1.243 -
   1.244 -lemma prod_cont_iff:
   1.245 -  "cont f \<longleftrightarrow> (\<forall>y. cont (\<lambda>x. f (x, y))) \<and> (\<forall>x. cont (\<lambda>y. f (x, y)))"
   1.246 -apply safe
   1.247 -apply (erule cont_compose [OF _ cont_pair1])
   1.248 -apply (erule cont_compose [OF _ cont_pair2])
   1.249 -apply (simp only: prod_contI)
   1.250 -done
   1.251 -
   1.252 -lemma cont2cont_prod_case' [simp, cont2cont]:
   1.253 -  assumes f: "cont (\<lambda>p. f (fst p) (fst (snd p)) (snd (snd p)))"
   1.254 -  assumes g: "cont (\<lambda>x. g x)"
   1.255 -  shows "cont (\<lambda>x. prod_case (f x) (g x))"
   1.256 -using assms by (simp add: cont2cont_prod_case prod_cont_iff)
   1.257 -
   1.258 -text {* The simple version (due to Joachim Breitner) is needed if
   1.259 -  either element type of the pair is not a cpo. *}
   1.260 -
   1.261 -lemma cont2cont_split_simple [simp, cont2cont]:
   1.262 - assumes "\<And>a b. cont (\<lambda>x. f x a b)"
   1.263 - shows "cont (\<lambda>x. case p of (a, b) \<Rightarrow> f x a b)"
   1.264 -using assms by (cases p) auto
   1.265 -
   1.266 -text {* Admissibility of predicates on product types. *}
   1.267 -
   1.268 -lemma adm_prod_case [simp]:
   1.269 -  assumes "adm (\<lambda>x. P x (fst (f x)) (snd (f x)))"
   1.270 -  shows "adm (\<lambda>x. case f x of (a, b) \<Rightarrow> P x a b)"
   1.271 -unfolding prod_case_beta using assms .
   1.272 -
   1.273 -subsection {* Compactness and chain-finiteness *}
   1.274 -
   1.275 -lemma fst_below_iff: "fst (x::'a \<times> 'b) \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (y, snd x)"
   1.276 -unfolding below_prod_def by simp
   1.277 -
   1.278 -lemma snd_below_iff: "snd (x::'a \<times> 'b) \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (fst x, y)"
   1.279 -unfolding below_prod_def by simp
   1.280 -
   1.281 -lemma compact_fst: "compact x \<Longrightarrow> compact (fst x)"
   1.282 -by (rule compactI, simp add: fst_below_iff)
   1.283 -
   1.284 -lemma compact_snd: "compact x \<Longrightarrow> compact (snd x)"
   1.285 -by (rule compactI, simp add: snd_below_iff)
   1.286 -
   1.287 -lemma compact_Pair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact (x, y)"
   1.288 -by (rule compactI, simp add: below_prod_def)
   1.289 -
   1.290 -lemma compact_Pair_iff [simp]: "compact (x, y) \<longleftrightarrow> compact x \<and> compact y"
   1.291 -apply (safe intro!: compact_Pair)
   1.292 -apply (drule compact_fst, simp)
   1.293 -apply (drule compact_snd, simp)
   1.294 -done
   1.295 -
   1.296 -instance prod :: (chfin, chfin) chfin
   1.297 -apply intro_classes
   1.298 -apply (erule compact_imp_max_in_chain)
   1.299 -apply (case_tac "\<Squnion>i. Y i", simp)
   1.300 -done
   1.301 -
   1.302 -end