(* Title: HOLCF/Cprod.thy ID: $Id: Cprod.thy,v 1.40 2008/05/19 21:49:21 huffman Exp $ Author: Franz Regensburger Partial ordering for cartesian product of HOL products. *) header {* The cpo of cartesian products *} theory Cprod imports Bifinite begin defaultsort cpo subsection {* Type @{typ unit} is a pcpo *} instantiation unit :: sq_ord begin definition less_unit_def [simp]: "x \<sqsubseteq> (y::unit) ≡ True" instance .. end instance unit :: discrete_cpo by intro_classes simp instance unit :: finite_po .. instance unit :: pcpo by intro_classes simp definition unit_when :: "'a -> unit -> 'a" where "unit_when = (Λ a _. a)" translations "Λ(). t" == "CONST unit_when·t" lemma unit_when [simp]: "unit_when·a·u = a" by (simp add: unit_when_def) subsection {* Product type is a partial order *} instantiation "*" :: (sq_ord, sq_ord) sq_ord begin definition less_cprod_def: "(op \<sqsubseteq>) ≡ λp1 p2. (fst p1 \<sqsubseteq> fst p2 ∧ snd p1 \<sqsubseteq> snd p2)" instance .. end instance "*" :: (po, po) po proof fix x :: "'a × 'b" show "x \<sqsubseteq> x" unfolding less_cprod_def by simp next fix x y :: "'a × 'b" assume "x \<sqsubseteq> y" "y \<sqsubseteq> x" thus "x = y" unfolding less_cprod_def Pair_fst_snd_eq by (fast intro: antisym_less) next fix x y z :: "'a × 'b" assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" thus "x \<sqsubseteq> z" unfolding less_cprod_def by (fast intro: trans_less) qed subsection {* Monotonicity of @{text "(_,_)"}, @{term fst}, @{term snd} *} lemma prod_lessI: "[|fst p \<sqsubseteq> fst q; snd p \<sqsubseteq> snd q|] ==> p \<sqsubseteq> q" unfolding less_cprod_def by simp lemma Pair_less_iff [simp]: "(a, b) \<sqsubseteq> (c, d) = (a \<sqsubseteq> c ∧ b \<sqsubseteq> d)" unfolding less_cprod_def by simp text {* Pair @{text "(_,_)"} is monotone in both arguments *} 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 \<sqsubseteq> x2; y1 \<sqsubseteq> y2|] ==> (x1, y1) \<sqsubseteq> (x2, y2)" by simp text {* @{term fst} and @{term snd} are monotone *} lemma monofun_fst: "monofun fst" by (simp add: monofun_def less_cprod_def) lemma monofun_snd: "monofun snd" by (simp add: monofun_def less_cprod_def) subsection {* Product type is a cpo *} lemma is_lub_Pair: "[|range X <<| x; range Y <<| y|] ==> range (λi. (X i, Y i)) <<| (x, y)" apply (rule is_lubI [OF ub_rangeI]) apply (simp add: less_cprod_def is_ub_lub) apply (frule ub2ub_monofun [OF monofun_fst]) apply (drule ub2ub_monofun [OF monofun_snd]) apply (simp add: less_cprod_def is_lub_lub) done lemma lub_cprod: fixes S :: "nat => ('a::cpo × 'b::cpo)" assumes S: "chain S" shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" proof - have "chain (λi. fst (S i))" using monofun_fst S by (rule ch2ch_monofun) hence 1: "range (λi. fst (S i)) <<| (\<Squnion>i. fst (S i))" by (rule cpo_lubI) have "chain (λi. snd (S i))" using monofun_snd S by (rule ch2ch_monofun) hence 2: "range (λi. snd (S i)) <<| (\<Squnion>i. snd (S i))" by (rule cpo_lubI) show "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" using is_lub_Pair [OF 1 2] by simp qed lemma thelub_cprod: "chain (S::nat => 'a::cpo × 'b::cpo) ==> lub (range S) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" by (rule lub_cprod [THEN thelubI]) instance "*" :: (cpo, cpo) cpo proof fix S :: "nat => ('a × 'b)" assume "chain S" hence "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" by (rule lub_cprod) thus "∃x. range S <<| x" .. qed instance "*" :: (finite_po, finite_po) finite_po .. instance "*" :: (discrete_cpo, discrete_cpo) discrete_cpo proof fix x y :: "'a × 'b" show "x \<sqsubseteq> y <-> x = y" unfolding less_cprod_def Pair_fst_snd_eq by simp qed subsection {* Product type is pointed *} lemma minimal_cprod: "(⊥, ⊥) \<sqsubseteq> p" by (simp add: less_cprod_def) instance "*" :: (pcpo, pcpo) pcpo by intro_classes (fast intro: minimal_cprod) lemma inst_cprod_pcpo: "⊥ = (⊥, ⊥)" by (rule minimal_cprod [THEN UU_I, symmetric]) subsection {* Continuity of @{text "(_,_)"}, @{term fst}, @{term snd} *} lemma cont_pair1: "cont (λx. (x, y))" apply (rule contI) apply (rule is_lub_Pair) apply (erule cpo_lubI) apply (rule lub_const) done lemma cont_pair2: "cont (λy. (x, y))" apply (rule contI) apply (rule is_lub_Pair) apply (rule lub_const) apply (erule cpo_lubI) done lemma contlub_fst: "contlub fst" apply (rule contlubI) apply (simp add: thelub_cprod) done lemma contlub_snd: "contlub snd" apply (rule contlubI) apply (simp add: thelub_cprod) done lemma cont_fst: "cont fst" apply (rule monocontlub2cont) apply (rule monofun_fst) apply (rule contlub_fst) done lemma cont_snd: "cont snd" apply (rule monocontlub2cont) apply (rule monofun_snd) apply (rule contlub_snd) done subsection {* Continuous versions of constants *} definition cpair :: "'a -> 'b -> ('a * 'b)" -- {* continuous pairing *} where "cpair = (Λ x y. (x, y))" definition cfst :: "('a * 'b) -> 'a" where "cfst = (Λ p. fst p)" definition csnd :: "('a * 'b) -> 'b" where "csnd = (Λ p. snd p)" definition csplit :: "('a -> 'b -> 'c) -> ('a * 'b) -> 'c" where "csplit = (Λ f p. f·(cfst·p)·(csnd·p))" syntax "_ctuple" :: "['a, args] => 'a * 'b" ("(1<_,/ _>)") syntax (xsymbols) "_ctuple" :: "['a, args] => 'a * 'b" ("(1〈_,/ _〉)") translations "〈x, y, z〉" == "〈x, 〈y, z〉〉" "〈x, y〉" == "CONST cpair·x·y" translations "Λ(CONST cpair·x·y). t" == "CONST csplit·(Λ x y. t)" subsection {* Convert all lemmas to the continuous versions *} lemma cpair_eq_pair: "<x, y> = (x, y)" by (simp add: cpair_def cont_pair1 cont_pair2) lemma pair_eq_cpair: "(x, y) = <x, y>" by (simp add: cpair_def cont_pair1 cont_pair2) lemma inject_cpair: "<a,b> = <aa,ba> ==> a = aa ∧ b = ba" by (simp add: cpair_eq_pair) lemma cpair_eq [iff]: "(<a, b> = <a', b'>) = (a = a' ∧ b = b')" by (simp add: cpair_eq_pair) lemma cpair_less [iff]: "(<a, b> \<sqsubseteq> <a', b'>) = (a \<sqsubseteq> a' ∧ b \<sqsubseteq> b')" by (simp add: cpair_eq_pair less_cprod_def) lemma cpair_defined_iff [iff]: "(<x, y> = ⊥) = (x = ⊥ ∧ y = ⊥)" by (simp add: inst_cprod_pcpo cpair_eq_pair) lemma cpair_strict [simp]: "〈⊥, ⊥〉 = ⊥" by simp lemma inst_cprod_pcpo2: "⊥ = <⊥, ⊥>" by (rule cpair_strict [symmetric]) lemma defined_cpair_rev: "<a,b> = ⊥ ==> a = ⊥ ∧ b = ⊥" by simp lemma Exh_Cprod2: "∃a b. z = <a, b>" by (simp add: cpair_eq_pair) lemma cprodE: "[|!!x y. p = <x, y> ==> Q|] ==> Q" by (cut_tac Exh_Cprod2, auto) lemma cfst_cpair [simp]: "cfst·<x, y> = x" by (simp add: cpair_eq_pair cfst_def cont_fst) lemma csnd_cpair [simp]: "csnd·<x, y> = y" by (simp add: cpair_eq_pair csnd_def cont_snd) lemma cfst_strict [simp]: "cfst·⊥ = ⊥" unfolding inst_cprod_pcpo2 by (rule cfst_cpair) lemma csnd_strict [simp]: "csnd·⊥ = ⊥" unfolding inst_cprod_pcpo2 by (rule csnd_cpair) lemma cpair_cfst_csnd: "〈cfst·p, csnd·p〉 = p" by (cases p rule: cprodE, simp) lemmas surjective_pairing_Cprod2 = cpair_cfst_csnd lemma less_cprod: "x \<sqsubseteq> y = (cfst·x \<sqsubseteq> cfst·y ∧ csnd·x \<sqsubseteq> csnd·y)" by (simp add: less_cprod_def cfst_def csnd_def cont_fst cont_snd) lemma eq_cprod: "(x = y) = (cfst·x = cfst·y ∧ csnd·x = csnd·y)" by (auto simp add: po_eq_conv less_cprod) lemma cfst_less_iff: "cfst·x \<sqsubseteq> y = x \<sqsubseteq> <y, csnd·x>" by (simp add: less_cprod) lemma csnd_less_iff: "csnd·x \<sqsubseteq> y = x \<sqsubseteq> <cfst·x, y>" by (simp add: less_cprod) lemma compact_cfst: "compact x ==> compact (cfst·x)" by (rule compactI, simp add: cfst_less_iff) lemma compact_csnd: "compact x ==> compact (csnd·x)" by (rule compactI, simp add: csnd_less_iff) lemma compact_cpair: "[|compact x; compact y|] ==> compact <x, y>" by (rule compactI, simp add: less_cprod) lemma compact_cpair_iff [simp]: "compact <x, y> = (compact x ∧ compact y)" apply (safe intro!: compact_cpair) apply (drule compact_cfst, simp) apply (drule compact_csnd, simp) done instance "*" :: (chfin, chfin) chfin apply intro_classes apply (erule compact_imp_max_in_chain) apply (rule_tac p="\<Squnion>i. Y i" in cprodE, simp) done lemma lub_cprod2: "chain S ==> range S <<| <\<Squnion>i. cfst·(S i), \<Squnion>i. csnd·(S i)>" apply (simp add: cpair_eq_pair cfst_def csnd_def cont_fst cont_snd) apply (erule lub_cprod) done lemma thelub_cprod2: "chain S ==> lub (range S) = <\<Squnion>i. cfst·(S i), \<Squnion>i. csnd·(S i)>" by (rule lub_cprod2 [THEN thelubI]) lemma csplit1 [simp]: "csplit·f·⊥ = f·⊥·⊥" by (simp add: csplit_def) lemma csplit2 [simp]: "csplit·f·<x,y> = f·x·y" by (simp add: csplit_def) lemma csplit3 [simp]: "csplit·cpair·z = z" by (simp add: csplit_def cpair_cfst_csnd) lemmas Cprod_rews = cfst_cpair csnd_cpair csplit2 subsection {* Product type is a bifinite domain *} instantiation "*" :: (profinite, profinite) profinite begin definition approx_cprod_def: "approx = (λn. Λ〈x, y〉. 〈approx n·x, approx n·y〉)" instance proof fix i :: nat and x :: "'a × 'b" show "chain (λi. approx i·x)" unfolding approx_cprod_def by simp show "(\<Squnion>i. approx i·x) = x" unfolding approx_cprod_def by (simp add: lub_distribs eta_cfun) show "approx i·(approx i·x) = approx i·x" unfolding approx_cprod_def csplit_def by simp have "{x::'a × 'b. approx i·x = x} ⊆ {x::'a. approx i·x = x} × {x::'b. approx i·x = x}" unfolding approx_cprod_def by (clarsimp simp add: pair_eq_cpair) thus "finite {x::'a × 'b. approx i·x = x}" by (rule finite_subset, intro finite_cartesian_product finite_fixes_approx) qed end instance "*" :: (bifinite, bifinite) bifinite .. lemma approx_cpair [simp]: "approx i·〈x, y〉 = 〈approx i·x, approx i·y〉" unfolding approx_cprod_def by simp lemma cfst_approx: "cfst·(approx i·p) = approx i·(cfst·p)" by (cases p rule: cprodE, simp) lemma csnd_approx: "csnd·(approx i·p) = approx i·(csnd·p)" by (cases p rule: cprodE, simp) end
lemma unit_when:
(LAM (). a)·u = a
lemma prod_lessI:
[| fst p << fst q; snd p << snd q |] ==> p << q
lemma Pair_less_iff:
(a, b) << (c, d) = (a << c ∧ b << d)
lemma monofun_pair1:
monofun (λx. (x, y))
lemma monofun_pair2:
monofun (Pair x)
lemma monofun_pair:
[| x1.0 << x2.0; y1.0 << y2.0 |] ==> (x1.0, y1.0) << (x2.0, y2.0)
lemma monofun_fst:
monofun fst
lemma monofun_snd:
monofun snd
lemma is_lub_Pair:
[| range X <<| x; range Y <<| y |] ==> range (λi. (X i, Y i)) <<| (x, y)
lemma lub_cprod:
chain S ==> range S <<| (LUB i. fst (S i), LUB i. snd (S i))
lemma thelub_cprod:
chain S ==> Lub S = (LUB i. fst (S i), LUB i. snd (S i))
lemma minimal_cprod:
(UU, UU) << p
lemma inst_cprod_pcpo:
UU = (UU, UU)
lemma cont_pair1:
cont (λx. (x, y))
lemma cont_pair2:
cont (Pair x)
lemma contlub_fst:
contlub fst
lemma contlub_snd:
contlub snd
lemma cont_fst:
cont fst
lemma cont_snd:
cont snd
lemma cpair_eq_pair:
<x, y> = (x, y)
lemma pair_eq_cpair:
(x, y) = <x, y>
lemma inject_cpair:
<a, b> = <aa, ba> ==> a = aa ∧ b = ba
lemma cpair_eq:
(<a, b> = <a', b'>) = (a = a' ∧ b = b')
lemma cpair_less:
<a, b> << <a', b'> = (a << a' ∧ b << b')
lemma cpair_defined_iff:
(<x, y> = UU) = (x = UU ∧ y = UU)
lemma cpair_strict:
<UU, UU> = UU
lemma inst_cprod_pcpo2:
UU = <UU, UU>
lemma defined_cpair_rev:
<a, b> = UU ==> a = UU ∧ b = UU
lemma Exh_Cprod2:
∃a b. z = <a, b>
lemma cprodE:
(!!x y. p = <x, y> ==> Q) ==> Q
lemma cfst_cpair:
cfst·<x, y> = x
lemma csnd_cpair:
csnd·<x, y> = y
lemma cfst_strict:
cfst·UU = UU
lemma csnd_strict:
csnd·UU = UU
lemma cpair_cfst_csnd:
<cfst·p, csnd·p> = p
lemma surjective_pairing_Cprod2:
<cfst·p, csnd·p> = p
lemma less_cprod:
x << y = (cfst·x << cfst·y ∧ csnd·x << csnd·y)
lemma eq_cprod:
(x = y) = (cfst·x = cfst·y ∧ csnd·x = csnd·y)
lemma cfst_less_iff:
cfst·x << y = x << <y, csnd·x>
lemma csnd_less_iff:
csnd·x << y = x << <cfst·x, y>
lemma compact_cfst:
compact x ==> compact (cfst·x)
lemma compact_csnd:
compact x ==> compact (csnd·x)
lemma compact_cpair:
[| compact x; compact y |] ==> compact <x, y>
lemma compact_cpair_iff:
compact <x, y> = (compact x ∧ compact y)
lemma lub_cprod2:
chain S ==> range S <<| <LUB i. cfst·(S i), LUB i. csnd·(S i)>
lemma thelub_cprod2:
chain S ==> Lub S = <LUB i. cfst·(S i), LUB i. csnd·(S i)>
lemma csplit1:
csplit·f·UU = f·UU·UU
lemma csplit2:
csplit·f·<x, y> = f·x·y
lemma csplit3:
csplit·cpair·z = z
lemma Cprod_rews:
cfst·<x, y> = x
csnd·<x, y> = y
csplit·f·<x, y> = f·x·y
lemma approx_cpair:
approx i·<x, y> = <approx i·x, approx i·y>
lemma cfst_approx:
cfst·(approx i·p) = approx i·(cfst·p)
lemma csnd_approx:
csnd·(approx i·p) = approx i·(csnd·p)