(* Title: HOLCF/Ssum.thy ID: $Id: Ssum.thy,v 1.27 2008/05/19 21:49:22 huffman Exp $ Author: Franz Regensburger and Brian Huffman Strict sum with typedef. *) header {* The type of strict sums *} theory Ssum imports Cprod Tr begin defaultsort pcpo subsection {* Definition of strict sum type *} pcpodef (Ssum) ('a, 'b) "++" (infixr "++" 10) = "{p :: tr × ('a × 'b). (cfst·p \<sqsubseteq> TT <-> csnd·(csnd·p) = ⊥) ∧ (cfst·p \<sqsubseteq> FF <-> cfst·(csnd·p) = ⊥)}" by simp instance "++" :: ("{finite_po,pcpo}", "{finite_po,pcpo}") finite_po by (rule typedef_finite_po [OF type_definition_Ssum]) instance "++" :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin by (rule typedef_chfin [OF type_definition_Ssum less_Ssum_def]) syntax (xsymbols) "++" :: "[type, type] => type" ("(_ ⊕/ _)" [21, 20] 20) syntax (HTML output) "++" :: "[type, type] => type" ("(_ ⊕/ _)" [21, 20] 20) subsection {* Definitions of constructors *} definition sinl :: "'a -> ('a ++ 'b)" where "sinl = (Λ a. Abs_Ssum <strictify·(Λ _. TT)·a, a, ⊥>)" definition sinr :: "'b -> ('a ++ 'b)" where "sinr = (Λ b. Abs_Ssum <strictify·(Λ _. FF)·b, ⊥, b>)" lemma sinl_Ssum: "<strictify·(Λ _. TT)·a, a, ⊥> ∈ Ssum" by (simp add: Ssum_def strictify_conv_if) lemma sinr_Ssum: "<strictify·(Λ _. FF)·b, ⊥, b> ∈ Ssum" by (simp add: Ssum_def strictify_conv_if) lemma sinl_Abs_Ssum: "sinl·a = Abs_Ssum <strictify·(Λ _. TT)·a, a, ⊥>" by (unfold sinl_def, simp add: cont_Abs_Ssum sinl_Ssum) lemma sinr_Abs_Ssum: "sinr·b = Abs_Ssum <strictify·(Λ _. FF)·b, ⊥, b>" by (unfold sinr_def, simp add: cont_Abs_Ssum sinr_Ssum) lemma Rep_Ssum_sinl: "Rep_Ssum (sinl·a) = <strictify·(Λ _. TT)·a, a, ⊥>" by (simp add: sinl_Abs_Ssum Abs_Ssum_inverse sinl_Ssum) lemma Rep_Ssum_sinr: "Rep_Ssum (sinr·b) = <strictify·(Λ _. FF)·b, ⊥, b>" by (simp add: sinr_Abs_Ssum Abs_Ssum_inverse sinr_Ssum) subsection {* Properties of @{term sinl} and @{term sinr} *} text {* Ordering *} lemma sinl_less [simp]: "(sinl·x \<sqsubseteq> sinl·y) = (x \<sqsubseteq> y)" by (simp add: less_Ssum_def Rep_Ssum_sinl strictify_conv_if) lemma sinr_less [simp]: "(sinr·x \<sqsubseteq> sinr·y) = (x \<sqsubseteq> y)" by (simp add: less_Ssum_def Rep_Ssum_sinr strictify_conv_if) lemma sinl_less_sinr [simp]: "(sinl·x \<sqsubseteq> sinr·y) = (x = ⊥)" by (simp add: less_Ssum_def Rep_Ssum_sinl Rep_Ssum_sinr strictify_conv_if) lemma sinr_less_sinl [simp]: "(sinr·x \<sqsubseteq> sinl·y) = (x = ⊥)" by (simp add: less_Ssum_def Rep_Ssum_sinl Rep_Ssum_sinr strictify_conv_if) text {* Equality *} lemma sinl_eq [simp]: "(sinl·x = sinl·y) = (x = y)" by (simp add: po_eq_conv) lemma sinr_eq [simp]: "(sinr·x = sinr·y) = (x = y)" by (simp add: po_eq_conv) lemma sinl_eq_sinr [simp]: "(sinl·x = sinr·y) = (x = ⊥ ∧ y = ⊥)" by (subst po_eq_conv, simp) lemma sinr_eq_sinl [simp]: "(sinr·x = sinl·y) = (x = ⊥ ∧ y = ⊥)" by (subst po_eq_conv, simp) lemma sinl_inject: "sinl·x = sinl·y ==> x = y" by (rule sinl_eq [THEN iffD1]) lemma sinr_inject: "sinr·x = sinr·y ==> x = y" by (rule sinr_eq [THEN iffD1]) text {* Strictness *} lemma sinl_strict [simp]: "sinl·⊥ = ⊥" by (simp add: sinl_Abs_Ssum Abs_Ssum_strict) lemma sinr_strict [simp]: "sinr·⊥ = ⊥" by (simp add: sinr_Abs_Ssum Abs_Ssum_strict) lemma sinl_defined_iff [simp]: "(sinl·x = ⊥) = (x = ⊥)" by (cut_tac sinl_eq [of "x" "⊥"], simp) lemma sinr_defined_iff [simp]: "(sinr·x = ⊥) = (x = ⊥)" by (cut_tac sinr_eq [of "x" "⊥"], simp) lemma sinl_defined [intro!]: "x ≠ ⊥ ==> sinl·x ≠ ⊥" by simp lemma sinr_defined [intro!]: "x ≠ ⊥ ==> sinr·x ≠ ⊥" by simp text {* Compactness *} lemma compact_sinl: "compact x ==> compact (sinl·x)" by (rule compact_Ssum, simp add: Rep_Ssum_sinl strictify_conv_if) lemma compact_sinr: "compact x ==> compact (sinr·x)" by (rule compact_Ssum, simp add: Rep_Ssum_sinr strictify_conv_if) lemma compact_sinlD: "compact (sinl·x) ==> compact x" unfolding compact_def by (drule adm_subst [OF cont_Rep_CFun2 [where f=sinl]], simp) lemma compact_sinrD: "compact (sinr·x) ==> compact x" unfolding compact_def by (drule adm_subst [OF cont_Rep_CFun2 [where f=sinr]], simp) lemma compact_sinl_iff [simp]: "compact (sinl·x) = compact x" by (safe elim!: compact_sinl compact_sinlD) lemma compact_sinr_iff [simp]: "compact (sinr·x) = compact x" by (safe elim!: compact_sinr compact_sinrD) subsection {* Case analysis *} lemma Exh_Ssum: "z = ⊥ ∨ (∃a. z = sinl·a ∧ a ≠ ⊥) ∨ (∃b. z = sinr·b ∧ b ≠ ⊥)" apply (rule_tac x=z in Abs_Ssum_induct) apply (rule_tac p=y in cprodE, rename_tac t x) apply (rule_tac p=x in cprodE, rename_tac a b) apply (rule_tac p=t in trE) apply (rule disjI1) apply (simp add: Ssum_def cpair_strict Abs_Ssum_strict) apply (rule disjI2, rule disjI1, rule_tac x=a in exI) apply (simp add: sinl_Abs_Ssum Ssum_def) apply (rule disjI2, rule disjI2, rule_tac x=b in exI) apply (simp add: sinr_Abs_Ssum Ssum_def) done lemma ssumE [cases type: ++]: "[|p = ⊥ ==> Q; !!x. [|p = sinl·x; x ≠ ⊥|] ==> Q; !!y. [|p = sinr·y; y ≠ ⊥|] ==> Q|] ==> Q" by (cut_tac z=p in Exh_Ssum, auto) lemma ssum_induct [induct type: ++]: "[|P ⊥; !!x. x ≠ ⊥ ==> P (sinl·x); !!y. y ≠ ⊥ ==> P (sinr·y)|] ==> P x" by (cases x, simp_all) lemma ssumE2: "[|!!x. p = sinl·x ==> Q; !!y. p = sinr·y ==> Q|] ==> Q" by (cases p, simp only: sinl_strict [symmetric], simp, simp) lemma less_sinlD: "p \<sqsubseteq> sinl·x ==> ∃y. p = sinl·y ∧ y \<sqsubseteq> x" by (cases p, rule_tac x="⊥" in exI, simp_all) lemma less_sinrD: "p \<sqsubseteq> sinr·x ==> ∃y. p = sinr·y ∧ y \<sqsubseteq> x" by (cases p, rule_tac x="⊥" in exI, simp_all) subsection {* Case analysis combinator *} definition sscase :: "('a -> 'c) -> ('b -> 'c) -> ('a ++ 'b) -> 'c" where "sscase = (Λ f g s. (Λ<t, x, y>. If t then f·x else g·y fi)·(Rep_Ssum s))" translations "case s of XCONST sinl·x => t1 | XCONST sinr·y => t2" == "CONST sscase·(Λ x. t1)·(Λ y. t2)·s" translations "Λ(XCONST sinl·x). t" == "CONST sscase·(Λ x. t)·⊥" "Λ(XCONST sinr·y). t" == "CONST sscase·⊥·(Λ y. t)" lemma beta_sscase: "sscase·f·g·s = (Λ<t, x, y>. If t then f·x else g·y fi)·(Rep_Ssum s)" unfolding sscase_def by (simp add: cont_Rep_Ssum) lemma sscase1 [simp]: "sscase·f·g·⊥ = ⊥" unfolding beta_sscase by (simp add: Rep_Ssum_strict) lemma sscase2 [simp]: "x ≠ ⊥ ==> sscase·f·g·(sinl·x) = f·x" unfolding beta_sscase by (simp add: Rep_Ssum_sinl) lemma sscase3 [simp]: "y ≠ ⊥ ==> sscase·f·g·(sinr·y) = g·y" unfolding beta_sscase by (simp add: Rep_Ssum_sinr) lemma sscase4 [simp]: "sscase·sinl·sinr·z = z" by (cases z, simp_all) subsection {* Strict sum preserves flatness *} instance "++" :: (flat, flat) flat apply (intro_classes, clarify) apply (rule_tac p=x in ssumE, simp) apply (rule_tac p=y in ssumE, simp_all add: flat_less_iff) apply (rule_tac p=y in ssumE, simp_all add: flat_less_iff) done subsection {* Strict sum is a bifinite domain *} instantiation "++" :: (bifinite, bifinite) bifinite begin definition approx_ssum_def: "approx = (λn. sscase·(Λ x. sinl·(approx n·x))·(Λ y. sinr·(approx n·y)))" lemma approx_sinl [simp]: "approx i·(sinl·x) = sinl·(approx i·x)" unfolding approx_ssum_def by (cases "x = ⊥") simp_all lemma approx_sinr [simp]: "approx i·(sinr·x) = sinr·(approx i·x)" unfolding approx_ssum_def by (cases "x = ⊥") simp_all instance proof fix i :: nat and x :: "'a ⊕ 'b" show "chain (λi. approx i·x)" unfolding approx_ssum_def by simp show "(\<Squnion>i. approx i·x) = x" unfolding approx_ssum_def by (simp add: lub_distribs eta_cfun) show "approx i·(approx i·x) = approx i·x" by (cases x, simp add: approx_ssum_def, simp, simp) have "{x::'a ⊕ 'b. approx i·x = x} ⊆ (λx. sinl·x) ` {x. approx i·x = x} ∪ (λx. sinr·x) ` {x. approx i·x = x}" by (rule subsetI, rule_tac p=x in ssumE2, simp, simp) thus "finite {x::'a ⊕ 'b. approx i·x = x}" by (rule finite_subset, intro finite_UnI finite_imageI finite_fixes_approx) qed end end
lemma sinl_Ssum:
<strictify·(LAM uu_. TT)·a, a, UU> ∈ Ssum
lemma sinr_Ssum:
<strictify·(LAM uu_. FF)·b, UU, b> ∈ Ssum
lemma sinl_Abs_Ssum:
sinl·a = Abs_Ssum <strictify·(LAM uu_. TT)·a, a, UU>
lemma sinr_Abs_Ssum:
sinr·b = Abs_Ssum <strictify·(LAM uu_. FF)·b, UU, b>
lemma Rep_Ssum_sinl:
Rep_Ssum (sinl·a) = <strictify·(LAM uu_. TT)·a, a, UU>
lemma Rep_Ssum_sinr:
Rep_Ssum (sinr·b) = <strictify·(LAM uu_. FF)·b, UU, b>
lemma sinl_less:
sinl·x << sinl·y = x << y
lemma sinr_less:
sinr·x << sinr·y = x << y
lemma sinl_less_sinr:
sinl·x << sinr·y = (x = UU)
lemma sinr_less_sinl:
sinr·x << sinl·y = (x = UU)
lemma sinl_eq:
(sinl·x = sinl·y) = (x = y)
lemma sinr_eq:
(sinr·x = sinr·y) = (x = y)
lemma sinl_eq_sinr:
(sinl·x = sinr·y) = (x = UU ∧ y = UU)
lemma sinr_eq_sinl:
(sinr·x = sinl·y) = (x = UU ∧ y = UU)
lemma sinl_inject:
sinl·x = sinl·y ==> x = y
lemma sinr_inject:
sinr·x = sinr·y ==> x = y
lemma sinl_strict:
sinl·UU = UU
lemma sinr_strict:
sinr·UU = UU
lemma sinl_defined_iff:
(sinl·x = UU) = (x = UU)
lemma sinr_defined_iff:
(sinr·x = UU) = (x = UU)
lemma sinl_defined:
x ≠ UU ==> sinl·x ≠ UU
lemma sinr_defined:
x ≠ UU ==> sinr·x ≠ UU
lemma compact_sinl:
compact x ==> compact (sinl·x)
lemma compact_sinr:
compact x ==> compact (sinr·x)
lemma compact_sinlD:
compact (sinl·x) ==> compact x
lemma compact_sinrD:
compact (sinr·x) ==> compact x
lemma compact_sinl_iff:
compact (sinl·x) = compact x
lemma compact_sinr_iff:
compact (sinr·x) = compact x
lemma Exh_Ssum:
z = UU ∨ (∃a. z = sinl·a ∧ a ≠ UU) ∨ (∃b. z = sinr·b ∧ b ≠ UU)
lemma ssumE:
[| p = UU ==> Q; !!x. [| p = sinl·x; x ≠ UU |] ==> Q;
!!y. [| p = sinr·y; y ≠ UU |] ==> Q |]
==> Q
lemma ssum_induct:
[| P UU; !!x. x ≠ UU ==> P (sinl·x); !!y. y ≠ UU ==> P (sinr·y) |] ==> P x
lemma ssumE2:
[| !!x. p = sinl·x ==> Q; !!y. p = sinr·y ==> Q |] ==> Q
lemma less_sinlD:
p << sinl·x ==> ∃y. p = sinl·y ∧ y << x
lemma less_sinrD:
p << sinr·x ==> ∃y. p = sinr·y ∧ y << x
lemma beta_sscase:
sscase·f·g·s = (LAM <t, x, y>. If t then f·x else g·y fi)·(Rep_Ssum s)
lemma sscase1:
sscase·f·g·UU = UU
lemma sscase2:
x ≠ UU ==> sscase·f·g·(sinl·x) = f·x
lemma sscase3:
y ≠ UU ==> sscase·f·g·(sinr·y) = g·y
lemma sscase4:
sscase·sinl·sinr·z = z
lemma approx_sinl:
approx i·(sinl·x) = sinl·(approx i·x)
lemma approx_sinr:
approx i·(sinr·x) = sinr·(approx i·x)