Theory Sprod

Up to index of Isabelle/HOLCF

theory Sprod
imports Bifinite

(*  Title:      HOLCF/Sprod.thy
Author: Franz Regensburger and Brian Huffman
*)


header {* The type of strict products *}

theory Sprod
imports Bifinite
begin


default_sort pcpo

subsection {* Definition of strict product type *}

pcpodef (Sprod) ('a, 'b) sprod (infixr "**" 20) =
"{p::'a × 'b. p = ⊥ ∨ (fst p ≠ ⊥ ∧ snd p ≠ ⊥)}"

by simp_all

instance sprod :: ("{finite_po,pcpo}", "{finite_po,pcpo}") finite_po
by (rule typedef_finite_po [OF type_definition_Sprod])

instance sprod :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin
by (rule typedef_chfin [OF type_definition_Sprod below_Sprod_def])

type_notation (xsymbols)
sprod ("(_ ⊗/ _)" [21,20] 20)

type_notation (HTML output)
sprod ("(_ ⊗/ _)" [21,20] 20)


lemma spair_lemma:
"(strictify·(Λ b. a)·b, strictify·(Λ a. b)·a) ∈ Sprod"

by (simp add: Sprod_def strictify_conv_if)

subsection {* Definitions of constants *}

definition
sfst :: "('a ** 'b) -> 'a" where
"sfst = (Λ p. fst (Rep_Sprod p))"


definition
ssnd :: "('a ** 'b) -> 'b" where
"ssnd = (Λ p. snd (Rep_Sprod p))"


definition
spair :: "'a -> 'b -> ('a ** 'b)" where
"spair = (Λ a b. Abs_Sprod
(strictify·(Λ b. a)·b, strictify·(Λ a. b)·a))"


definition
ssplit :: "('a -> 'b -> 'c) -> ('a ** 'b) -> 'c" where
"ssplit = (Λ f. strictify·(Λ p. f·(sfst·p)·(ssnd·p)))"


syntax
"_stuple" :: "['a, args] => 'a ** 'b" ("(1'(:_,/ _:'))")

translations
"(:x, y, z:)" == "(:x, (:y, z:):)"
"(:x, y:)" == "CONST spair·x·y"


translations
"Λ(CONST spair·x·y). t" == "CONST ssplit·(Λ x y. t)"


subsection {* Case analysis *}

lemma Rep_Sprod_spair:
"Rep_Sprod (:a, b:) = (strictify·(Λ b. a)·b, strictify·(Λ a. b)·a)"

unfolding spair_def
by (simp add: cont_Abs_Sprod Abs_Sprod_inverse spair_lemma)

lemmas Rep_Sprod_simps =
Rep_Sprod_inject [symmetric] below_Sprod_def
Rep_Sprod_strict Rep_Sprod_spair


lemma Exh_Sprod:
"z = ⊥ ∨ (∃a b. z = (:a, b:) ∧ a ≠ ⊥ ∧ b ≠ ⊥)"

apply (insert Rep_Sprod [of z])
apply (simp add: Rep_Sprod_simps Pair_fst_snd_eq)
apply (simp add: Sprod_def)
apply (erule disjE, simp)
apply (simp add: strictify_conv_if)
apply fast
done

lemma sprodE [case_names bottom spair, cases type: sprod]:
"[|p = ⊥ ==> Q; !!x y. [|p = (:x, y:); x ≠ ⊥; y ≠ ⊥|] ==> Q|] ==> Q"

using Exh_Sprod [of p] by auto

lemma sprod_induct [case_names bottom spair, induct type: sprod]:
"[|P ⊥; !!x y. [|x ≠ ⊥; y ≠ ⊥|] ==> P (:x, y:)|] ==> P x"

by (cases x, simp_all)

subsection {* Properties of \emph{spair} *}

lemma spair_strict1 [simp]: "(:⊥, y:) = ⊥"
by (simp add: Rep_Sprod_simps strictify_conv_if)

lemma spair_strict2 [simp]: "(:x, ⊥:) = ⊥"
by (simp add: Rep_Sprod_simps strictify_conv_if)

lemma spair_strict_iff [simp]: "((:x, y:) = ⊥) = (x = ⊥ ∨ y = ⊥)"
by (simp add: Rep_Sprod_simps strictify_conv_if)

lemma spair_below_iff:
"((:a, b:) \<sqsubseteq> (:c, d:)) = (a = ⊥ ∨ b = ⊥ ∨ (a \<sqsubseteq> c ∧ b \<sqsubseteq> d))"

by (simp add: Rep_Sprod_simps strictify_conv_if)

lemma spair_eq_iff:
"((:a, b:) = (:c, d:)) =
(a = c ∧ b = d ∨ (a = ⊥ ∨ b = ⊥) ∧ (c = ⊥ ∨ d = ⊥))"

by (simp add: Rep_Sprod_simps strictify_conv_if)

lemma spair_strict: "x = ⊥ ∨ y = ⊥ ==> (:x, y:) = ⊥"
by simp

lemma spair_strict_rev: "(:x, y:) ≠ ⊥ ==> x ≠ ⊥ ∧ y ≠ ⊥"
by simp

lemma spair_defined: "[|x ≠ ⊥; y ≠ ⊥|] ==> (:x, y:) ≠ ⊥"
by simp

lemma spair_defined_rev: "(:x, y:) = ⊥ ==> x = ⊥ ∨ y = ⊥"
by simp

lemma spair_eq:
"[|x ≠ ⊥; y ≠ ⊥|] ==> ((:x, y:) = (:a, b:)) = (x = a ∧ y = b)"

by (simp add: spair_eq_iff)

lemma spair_inject:
"[|x ≠ ⊥; y ≠ ⊥; (:x, y:) = (:a, b:)|] ==> x = a ∧ y = b"

by (rule spair_eq [THEN iffD1])

lemma inst_sprod_pcpo2: "UU = (:UU,UU:)"
by simp

lemma sprodE2: "(!!x y. p = (:x, y:) ==> Q) ==> Q"
by (cases p, simp only: inst_sprod_pcpo2, simp)

subsection {* Properties of \emph{sfst} and \emph{ssnd} *}

lemma sfst_strict [simp]: "sfst·⊥ = ⊥"
by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_strict)

lemma ssnd_strict [simp]: "ssnd·⊥ = ⊥"
by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_strict)

lemma sfst_spair [simp]: "y ≠ ⊥ ==> sfst·(:x, y:) = x"
by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_spair)

lemma ssnd_spair [simp]: "x ≠ ⊥ ==> ssnd·(:x, y:) = y"
by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_spair)

lemma sfst_defined_iff [simp]: "(sfst·p = ⊥) = (p = ⊥)"
by (cases p, simp_all)

lemma ssnd_defined_iff [simp]: "(ssnd·p = ⊥) = (p = ⊥)"
by (cases p, simp_all)

lemma sfst_defined: "p ≠ ⊥ ==> sfst·p ≠ ⊥"
by simp

lemma ssnd_defined: "p ≠ ⊥ ==> ssnd·p ≠ ⊥"
by simp

lemma surjective_pairing_Sprod2: "(:sfst·p, ssnd·p:) = p"
by (cases p, simp_all)

lemma below_sprod: "x \<sqsubseteq> y = (sfst·x \<sqsubseteq> sfst·y ∧ ssnd·x \<sqsubseteq> ssnd·y)"
apply (simp add: below_Sprod_def sfst_def ssnd_def cont_Rep_Sprod)
apply (simp only: below_prod_def)
done

lemma eq_sprod: "(x = y) = (sfst·x = sfst·y ∧ ssnd·x = ssnd·y)"
by (auto simp add: po_eq_conv below_sprod)

lemma spair_below:
"[|x ≠ ⊥; y ≠ ⊥|] ==> (:x, y:) \<sqsubseteq> (:a, b:) = (x \<sqsubseteq> a ∧ y \<sqsubseteq> b)"

apply (cases "a = ⊥", simp)
apply (cases "b = ⊥", simp)
apply (simp add: below_sprod)
done

lemma sfst_below_iff: "sfst·x \<sqsubseteq> y = x \<sqsubseteq> (:y, ssnd·x:)"
apply (cases "x = ⊥", simp, cases "y = ⊥", simp)
apply (simp add: below_sprod)
done

lemma ssnd_below_iff: "ssnd·x \<sqsubseteq> y = x \<sqsubseteq> (:sfst·x, y:)"
apply (cases "x = ⊥", simp, cases "y = ⊥", simp)
apply (simp add: below_sprod)
done

subsection {* Compactness *}

lemma compact_sfst: "compact x ==> compact (sfst·x)"
by (rule compactI, simp add: sfst_below_iff)

lemma compact_ssnd: "compact x ==> compact (ssnd·x)"
by (rule compactI, simp add: ssnd_below_iff)

lemma compact_spair: "[|compact x; compact y|] ==> compact (:x, y:)"
by (rule compact_Sprod, simp add: Rep_Sprod_spair strictify_conv_if)

lemma compact_spair_iff:
"compact (:x, y:) = (x = ⊥ ∨ y = ⊥ ∨ (compact x ∧ compact y))"

apply (safe elim!: compact_spair)
apply (drule compact_sfst, simp)
apply (drule compact_ssnd, simp)
apply simp
apply simp
done

subsection {* Properties of \emph{ssplit} *}

lemma ssplit1 [simp]: "ssplit·f·⊥ = ⊥"
by (simp add: ssplit_def)

lemma ssplit2 [simp]: "[|x ≠ ⊥; y ≠ ⊥|] ==> ssplit·f·(:x, y:) = f·x·y"
by (simp add: ssplit_def)

lemma ssplit3 [simp]: "ssplit·spair·z = z"
by (cases z, simp_all)

subsection {* Strict product preserves flatness *}

instance sprod :: (flat, flat) flat
proof
fix x y :: "'a ⊗ 'b"
assume "x \<sqsubseteq> y" thus "x = ⊥ ∨ x = y"
apply (induct x, simp)
apply (induct y, simp)
apply (simp add: spair_below_iff flat_below_iff)
done
qed

subsection {* Map function for strict products *}

definition
sprod_map :: "('a -> 'b) -> ('c -> 'd) -> 'a ⊗ 'c -> 'b ⊗ 'd"
where
"sprod_map = (Λ f g. ssplit·(Λ x y. (:f·x, g·y:)))"


lemma sprod_map_strict [simp]: "sprod_map·a·b·⊥ = ⊥"
unfolding sprod_map_def by simp

lemma sprod_map_spair [simp]:
"x ≠ ⊥ ==> y ≠ ⊥ ==> sprod_map·f·g·(:x, y:) = (:f·x, g·y:)"

by (simp add: sprod_map_def)

lemma sprod_map_spair':
"f·⊥ = ⊥ ==> g·⊥ = ⊥ ==> sprod_map·f·g·(:x, y:) = (:f·x, g·y:)"

by (cases "x = ⊥ ∨ y = ⊥") auto

lemma sprod_map_ID: "sprod_map·ID·ID = ID"
unfolding sprod_map_def by (simp add: expand_cfun_eq eta_cfun)

lemma sprod_map_map:
"[|f1·⊥ = ⊥; g1·⊥ = ⊥|] ==>
sprod_map·f1·g1·(sprod_map·f2·g2·p) =
sprod_map·(Λ x. f1·(f2·x))·(Λ x. g1·(g2·x))·p"

apply (induct p, simp)
apply (case_tac "f2·x = ⊥", simp)
apply (case_tac "g2·y = ⊥", simp)
apply simp
done

lemma ep_pair_sprod_map:
assumes "ep_pair e1 p1" and "ep_pair e2 p2"
shows "ep_pair (sprod_map·e1·e2) (sprod_map·p1·p2)"

proof
interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
fix x show "sprod_map·p1·p2·(sprod_map·e1·e2·x) = x"
by (induct x) simp_all
fix y show "sprod_map·e1·e2·(sprod_map·p1·p2·y) \<sqsubseteq> y"
apply (induct y, simp)
apply (case_tac "p1·x = ⊥", simp, case_tac "p2·y = ⊥", simp)
apply (simp add: monofun_cfun e1p1.e_p_below e2p2.e_p_below)
done
qed

lemma deflation_sprod_map:
assumes "deflation d1" and "deflation d2"
shows "deflation (sprod_map·d1·d2)"

proof
interpret d1: deflation d1 by fact
interpret d2: deflation d2 by fact
fix x
show "sprod_map·d1·d2·(sprod_map·d1·d2·x) = sprod_map·d1·d2·x"
apply (induct x, simp)
apply (case_tac "d1·x = ⊥", simp, case_tac "d2·y = ⊥", simp)
apply (simp add: d1.idem d2.idem)
done
show "sprod_map·d1·d2·x \<sqsubseteq> x"
apply (induct x, simp)
apply (simp add: monofun_cfun d1.below d2.below)
done
qed

lemma finite_deflation_sprod_map:
assumes "finite_deflation d1" and "finite_deflation d2"
shows "finite_deflation (sprod_map·d1·d2)"

proof (intro finite_deflation.intro finite_deflation_axioms.intro)
interpret d1: finite_deflation d1 by fact
interpret d2: finite_deflation d2 by fact
have "deflation d1" and "deflation d2" by fact+
thus "deflation (sprod_map·d1·d2)" by (rule deflation_sprod_map)
have "{x. sprod_map·d1·d2·x = x} ⊆ insert ⊥
((λ(x, y). (:x, y:)) ` ({x. d1·x = x} × {y. d2·y = y}))"

by (rule subsetI, case_tac x, auto simp add: spair_eq_iff)
thus "finite {x. sprod_map·d1·d2·x = x}"
by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
qed

subsection {* Strict product is a bifinite domain *}

instantiation sprod :: (bifinite, bifinite) bifinite
begin


definition
approx_sprod_def:
"approx = (λn. sprod_map·(approx n)·(approx n))"


instance proof
fix i :: nat and x :: "'a ⊗ 'b"
show "chain (approx :: nat => 'a ⊗ 'b -> 'a ⊗ 'b)"
unfolding approx_sprod_def by simp
show "(\<Squnion>i. approx i·x) = x"
unfolding approx_sprod_def sprod_map_def
by (simp add: lub_distribs eta_cfun)
show "approx i·(approx i·x) = approx i·x"
unfolding approx_sprod_def sprod_map_def
by (simp add: ssplit_def strictify_conv_if)
show "finite {x::'a ⊗ 'b. approx i·x = x}"
unfolding approx_sprod_def
by (intro finite_deflation.finite_fixes
finite_deflation_sprod_map
finite_deflation_approx)

qed

end

lemma approx_spair [simp]:
"approx i·(:x, y:) = (:approx i·x, approx i·y:)"

unfolding approx_sprod_def sprod_map_def
by (simp add: ssplit_def strictify_conv_if)

end