(* Title: HOL/HOLCF/Product_Cpo.thy
Author: Franz Regensburger
*)
section \<open>The cpo of cartesian products\<close>
theory Product_Cpo
imports Adm
begin
default_sort cpo
subsection \<open>Unit type is a pcpo\<close>
instantiation unit :: discrete_cpo
begin
definition below_unit_def [simp]: "x \ (y::unit) \ True"
instance
by standard simp
end
instance unit :: pcpo
by standard simp
subsection \<open>Product type is a partial order\<close>
instantiation prod :: (below, below) below
begin
definition below_prod_def: "(\) \ \p1 p2. (fst p1 \ fst p2 \ snd p1 \ snd p2)"
instance ..
end
instance prod :: (po, po) po
proof
fix x :: "'a \ 'b"
show "x \ x"
by (simp add: below_prod_def)
next
fix x y :: "'a \ 'b"
assume "x \ y" "y \ x"
then show "x = y"
unfolding below_prod_def prod_eq_iff
by (fast intro: below_antisym)
next
fix x y z :: "'a \ 'b"
assume "x \ y" "y \ z"
then show "x \ z"
unfolding below_prod_def
by (fast intro: below_trans)
qed
subsection \<open>Monotonicity of \emph{Pair}, \emph{fst}, \emph{snd}\<close>
lemma prod_belowI: "fst p \ fst q \ snd p \ snd q \ p \ q"
by (simp add: below_prod_def)
lemma Pair_below_iff [simp]: "(a, b) \ (c, d) \ a \ c \ b \ d"
by (simp add: below_prod_def)
text \<open>Pair \<open>(_,_)\<close> is monotone in both arguments\<close>
lemma monofun_pair1: "monofun (\x. (x, y))"
by (simp add: monofun_def)
lemma monofun_pair2: "monofun (\y. (x, y))"
by (simp add: monofun_def)
lemma monofun_pair: "x1 \ x2 \ y1 \ y2 \ (x1, y1) \ (x2, y2)"
by simp
lemma ch2ch_Pair [simp]: "chain X \ chain Y \ chain (\i. (X i, Y i))"
by (rule chainI, simp add: chainE)
text \<open>\<^term>\<open>fst\<close> and \<^term>\<open>snd\<close> are monotone\<close>
lemma fst_monofun: "x \ y \ fst x \ fst y"
by (simp add: below_prod_def)
lemma snd_monofun: "x \ y \ snd x \ snd y"
by (simp add: below_prod_def)
lemma monofun_fst: "monofun fst"
by (simp add: monofun_def below_prod_def)
lemma monofun_snd: "monofun snd"
by (simp add: monofun_def below_prod_def)
lemmas ch2ch_fst [simp] = ch2ch_monofun [OF monofun_fst]
lemmas ch2ch_snd [simp] = ch2ch_monofun [OF monofun_snd]
lemma prod_chain_cases:
assumes chain: "chain Y"
obtains A B
where "chain A" and "chain B" and "Y = (\i. (A i, B i))"
proof
from chain show "chain (\i. fst (Y i))"
by (rule ch2ch_fst)
from chain show "chain (\i. snd (Y i))"
by (rule ch2ch_snd)
show "Y = (\i. (fst (Y i), snd (Y i)))"
by simp
qed
subsection \<open>Product type is a cpo\<close>
lemma is_lub_Pair: "range A <<| x \ range B <<| y \ range (\i. (A i, B i)) <<| (x, y)"
by (simp add: is_lub_def is_ub_def below_prod_def)
lemma lub_Pair: "chain A \ chain B \ (\i. (A i, B i)) = (\i. A i, \i. B i)"
for A :: "nat \ 'a::cpo" and B :: "nat \ 'b::cpo"
by (fast intro: lub_eqI is_lub_Pair elim: thelubE)
lemma is_lub_prod:
fixes S :: "nat \ ('a::cpo \ 'b::cpo)"
assumes "chain S"
shows "range S <<| (\i. fst (S i), \i. snd (S i))"
using assms by (auto elim: prod_chain_cases simp: is_lub_Pair cpo_lubI)
lemma lub_prod: "chain S \ (\i. S i) = (\i. fst (S i), \i. snd (S i))"
for S :: "nat \ 'a::cpo \ 'b::cpo"
by (rule is_lub_prod [THEN lub_eqI])
instance prod :: (cpo, cpo) cpo
proof
fix S :: "nat \ ('a \ 'b)"
assume "chain S"
then have "range S <<| (\i. fst (S i), \i. snd (S i))"
by (rule is_lub_prod)
then show "\x. range S <<| x" ..
qed
instance prod :: (discrete_cpo, discrete_cpo) discrete_cpo
proof
fix x y :: "'a \ 'b"
show "x \ y \ x = y"
by (simp add: below_prod_def prod_eq_iff)
qed
subsection \<open>Product type is pointed\<close>
lemma minimal_prod: "(\, \) \ p"
by (simp add: below_prod_def)
instance prod :: (pcpo, pcpo) pcpo
by intro_classes (fast intro: minimal_prod)
lemma inst_prod_pcpo: "\ = (\, \)"
by (rule minimal_prod [THEN bottomI, symmetric])
lemma Pair_bottom_iff [simp]: "(x, y) = \ \ x = \ \ y = \"
by (simp add: inst_prod_pcpo)
lemma fst_strict [simp]: "fst \ = \"
unfolding inst_prod_pcpo by (rule fst_conv)
lemma snd_strict [simp]: "snd \ = \"
unfolding inst_prod_pcpo by (rule snd_conv)
lemma Pair_strict [simp]: "(\, \) = \"
by simp
lemma split_strict [simp]: "case_prod f \ = f \ \"
by (simp add: split_def)
subsection \<open>Continuity of \emph{Pair}, \emph{fst}, \emph{snd}\<close>
lemma cont_pair1: "cont (\x. (x, y))"
apply (rule contI)
apply (rule is_lub_Pair)
apply (erule cpo_lubI)
apply (rule is_lub_const)
done
lemma cont_pair2: "cont (\y. (x, y))"
apply (rule contI)
apply (rule is_lub_Pair)
apply (rule is_lub_const)
apply (erule cpo_lubI)
done
lemma cont_fst: "cont fst"
apply (rule contI)
apply (simp add: lub_prod)
apply (erule cpo_lubI [OF ch2ch_fst])
done
lemma cont_snd: "cont snd"
apply (rule contI)
apply (simp add: lub_prod)
apply (erule cpo_lubI [OF ch2ch_snd])
done
lemma cont2cont_Pair [simp, cont2cont]:
assumes f: "cont (\x. f x)"
assumes g: "cont (\x. g x)"
shows "cont (\x. (f x, g x))"
apply (rule cont_apply [OF f cont_pair1])
apply (rule cont_apply [OF g cont_pair2])
apply (rule cont_const)
done
lemmas cont2cont_fst [simp, cont2cont] = cont_compose [OF cont_fst]
lemmas cont2cont_snd [simp, cont2cont] = cont_compose [OF cont_snd]
lemma cont2cont_case_prod:
assumes f1: "\a b. cont (\x. f x a b)"
assumes f2: "\x b. cont (\a. f x a b)"
assumes f3: "\x a. cont (\b. f x a b)"
assumes g: "cont (\x. g x)"
shows "cont (\x. case g x of (a, b) \ f x a b)"
unfolding split_def
apply (rule cont_apply [OF g])
apply (rule cont_apply [OF cont_fst f2])
apply (rule cont_apply [OF cont_snd f3])
apply (rule cont_const)
apply (rule f1)
done
lemma prod_contI:
assumes f1: "\y. cont (\x. f (x, y))"
assumes f2: "\x. cont (\y. f (x, y))"
shows "cont f"
proof -
have "cont (\(x, y). f (x, y))"
by (intro cont2cont_case_prod f1 f2 cont2cont)
then show "cont f"
by (simp only: case_prod_eta)
qed
lemma prod_cont_iff: "cont f \ (\y. cont (\x. f (x, y))) \ (\x. cont (\y. f (x, y)))"
apply safe
apply (erule cont_compose [OF _ cont_pair1])
apply (erule cont_compose [OF _ cont_pair2])
apply (simp only: prod_contI)
done
lemma cont2cont_case_prod' [simp, cont2cont]:
assumes f: "cont (\p. f (fst p) (fst (snd p)) (snd (snd p)))"
assumes g: "cont (\x. g x)"
shows "cont (\x. case_prod (f x) (g x))"
using assms by (simp add: cont2cont_case_prod prod_cont_iff)
text \<open>The simple version (due to Joachim Breitner) is needed if
either element type of the pair is not a cpo.\<close>
lemma cont2cont_split_simple [simp, cont2cont]:
assumes "\a b. cont (\x. f x a b)"
shows "cont (\x. case p of (a, b) \ f x a b)"
using assms by (cases p) auto
text \<open>Admissibility of predicates on product types.\<close>
lemma adm_case_prod [simp]:
assumes "adm (\x. P x (fst (f x)) (snd (f x)))"
shows "adm (\x. case f x of (a, b) \ P x a b)"
unfolding case_prod_beta using assms .
subsection \<open>Compactness and chain-finiteness\<close>
lemma fst_below_iff: "fst x \ y \ x \ (y, snd x)"
for x :: "'a \ 'b"
by (simp add: below_prod_def)
lemma snd_below_iff: "snd x \ y \ x \ (fst x, y)"
for x :: "'a \ 'b"
by (simp add: below_prod_def)
lemma compact_fst: "compact x \ compact (fst x)"
by (rule compactI) (simp add: fst_below_iff)
lemma compact_snd: "compact x \ compact (snd x)"
by (rule compactI) (simp add: snd_below_iff)
lemma compact_Pair: "compact x \ compact y \ compact (x, y)"
by (rule compactI) (simp add: below_prod_def)
lemma compact_Pair_iff [simp]: "compact (x, y) \ compact x \ compact y"
apply (safe intro!: compact_Pair)
apply (drule compact_fst, simp)
apply (drule compact_snd, simp)
done
instance prod :: (chfin, chfin) chfin
apply intro_classes
apply (erule compact_imp_max_in_chain)
apply (case_tac "\i. Y i", simp)
done
end
¤ Dauer der Verarbeitung: 0.1 Sekunden
(vorverarbeitet)
¤
|
Haftungshinweis
Die Informationen auf dieser Webseite wurden
nach bestem Wissen sorgfältig zusammengestellt. Es wird jedoch weder Vollständigkeit, noch Richtigkeit,
noch Qualität der bereit gestellten Informationen zugesichert.
Bemerkung:
Die farbliche Syntaxdarstellung ist noch experimentell.
|