merged.
--- a/src/HOLCF/Cfun.thy Thu Jan 15 15:51:50 2009 +0100
+++ b/src/HOLCF/Cfun.thy Thu Jan 15 09:10:42 2009 -0800
@@ -7,8 +7,7 @@
header {* The type of continuous functions *}
theory Cfun
-imports Pcpodef Ffun
-uses ("Tools/cont_proc.ML")
+imports Pcpodef Ffun Product_Cpo
begin
defaultsort cpo
@@ -301,7 +300,7 @@
text {* cont2cont lemma for @{term Rep_CFun} *}
-lemma cont2cont_Rep_CFun:
+lemma cont2cont_Rep_CFun [cont2cont]:
assumes f: "cont (\<lambda>x. f x)"
assumes t: "cont (\<lambda>x. t x)"
shows "cont (\<lambda>x. (f x)\<cdot>(t x))"
@@ -321,6 +320,11 @@
text {* cont2cont Lemma for @{term "%x. LAM y. f x y"} *}
+text {*
+ Not suitable as a cont2cont rule, because on nested lambdas
+ it causes exponential blow-up in the number of subgoals.
+*}
+
lemma cont2cont_LAM:
assumes f1: "\<And>x. cont (\<lambda>y. f x y)"
assumes f2: "\<And>y. cont (\<lambda>x. f x y)"
@@ -331,17 +335,40 @@
from f2 show "cont f" by (rule cont2cont_lambda)
qed
-text {* continuity simplification procedure *}
+text {*
+ This version does work as a cont2cont rule, since it
+ has only a single subgoal.
+*}
+
+lemma cont2cont_LAM' [cont2cont]:
+ fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo"
+ assumes f: "cont (\<lambda>p. f (fst p) (snd p))"
+ shows "cont (\<lambda>x. \<Lambda> y. f x y)"
+proof (rule cont2cont_LAM)
+ fix x :: 'a
+ have "cont (\<lambda>y. (x, y))"
+ by (rule cont_pair2)
+ with f have "cont (\<lambda>y. f (fst (x, y)) (snd (x, y)))"
+ by (rule cont2cont_app3)
+ thus "cont (\<lambda>y. f x y)"
+ by (simp only: fst_conv snd_conv)
+next
+ fix y :: 'b
+ have "cont (\<lambda>x. (x, y))"
+ by (rule cont_pair1)
+ with f have "cont (\<lambda>x. f (fst (x, y)) (snd (x, y)))"
+ by (rule cont2cont_app3)
+ thus "cont (\<lambda>x. f x y)"
+ by (simp only: fst_conv snd_conv)
+qed
+
+lemma cont2cont_LAM_discrete [cont2cont]:
+ "(\<And>y::'a::discrete_cpo. cont (\<lambda>x. f x y)) \<Longrightarrow> cont (\<lambda>x. \<Lambda> y. f x y)"
+by (simp add: cont2cont_LAM)
lemmas cont_lemmas1 =
cont_const cont_id cont_Rep_CFun2 cont2cont_Rep_CFun cont2cont_LAM
-use "Tools/cont_proc.ML";
-setup ContProc.setup;
-
-(*val cont_tac = (fn i => (resolve_tac cont_lemmas i));*)
-(*val cont_tacR = (fn i => (REPEAT (cont_tac i)));*)
-
subsection {* Miscellaneous *}
text {* Monotonicity of @{term Abs_CFun} *}
@@ -530,7 +557,8 @@
monocontlub2cont [OF monofun_strictify2 contlub_strictify2, standard]
lemma strictify_conv_if: "strictify\<cdot>f\<cdot>x = (if x = \<bottom> then \<bottom> else f\<cdot>x)"
-by (unfold strictify_def, simp add: cont_strictify1 cont_strictify2)
+ unfolding strictify_def
+ by (simp add: cont_strictify1 cont_strictify2 cont2cont_LAM)
lemma strictify1 [simp]: "strictify\<cdot>f\<cdot>\<bottom> = \<bottom>"
by (simp add: strictify_conv_if)
--- a/src/HOLCF/Cont.thy Thu Jan 15 15:51:50 2009 +0100
+++ b/src/HOLCF/Cont.thy Thu Jan 15 09:10:42 2009 -0800
@@ -104,6 +104,8 @@
apply simp
done
+lemmas cont2monofunE = cont2mono [THEN monofunE]
+
lemmas ch2ch_cont = cont2mono [THEN ch2ch_monofun]
text {* right to left: @{prop "cont f \<Longrightarrow> monofun f \<and> contlub f"} *}
@@ -135,22 +137,82 @@
apply (erule cpo_lubI)
done
+subsection {* Continuity simproc *}
+
+ML {*
+structure Cont2ContData = NamedThmsFun
+ ( val name = "cont2cont" val description = "continuity intro rule" )
+*}
+
+setup {* Cont2ContData.setup *}
+
+text {*
+ Given the term @{term "cont f"}, the procedure tries to construct the
+ theorem @{term "cont f == True"}. If this theorem cannot be completely
+ solved by the introduction rules, then the procedure returns a
+ conditional rewrite rule with the unsolved subgoals as premises.
+*}
+
+setup {*
+let
+ fun solve_cont thy ss t =
+ let
+ val tr = instantiate' [] [SOME (cterm_of thy t)] Eq_TrueI;
+ val rules = Cont2ContData.get (Simplifier.the_context ss);
+ val tac = REPEAT_ALL_NEW (resolve_tac rules);
+ in Option.map fst (Seq.pull (tac 1 tr)) end
+
+ val proc =
+ Simplifier.simproc @{theory} "cont_proc" ["cont f"] solve_cont;
+in
+ Simplifier.map_simpset (fn ss => ss addsimprocs [proc])
+end
+*}
+
subsection {* Continuity of basic functions *}
text {* The identity function is continuous *}
-lemma cont_id: "cont (\<lambda>x. x)"
+lemma cont_id [cont2cont]: "cont (\<lambda>x. x)"
apply (rule contI)
apply (erule cpo_lubI)
done
text {* constant functions are continuous *}
-lemma cont_const: "cont (\<lambda>x. c)"
+lemma cont_const [cont2cont]: "cont (\<lambda>x. c)"
apply (rule contI)
apply (rule lub_const)
done
+text {* application of functions is continuous *}
+
+lemma cont2cont_apply:
+ fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo" and t :: "'a \<Rightarrow> 'b"
+ assumes f1: "\<And>y. cont (\<lambda>x. f x y)"
+ assumes f2: "\<And>x. cont (\<lambda>y. f x y)"
+ assumes t: "cont (\<lambda>x. t x)"
+ shows "cont (\<lambda>x. (f x) (t x))"
+proof (rule monocontlub2cont [OF monofunI contlubI])
+ fix x y :: "'a" assume "x \<sqsubseteq> y"
+ then show "f x (t x) \<sqsubseteq> f y (t y)"
+ by (auto intro: cont2monofunE [OF f1]
+ cont2monofunE [OF f2]
+ cont2monofunE [OF t]
+ trans_less)
+next
+ fix Y :: "nat \<Rightarrow> 'a" assume "chain Y"
+ then show "f (\<Squnion>i. Y i) (t (\<Squnion>i. Y i)) = (\<Squnion>i. f (Y i) (t (Y i)))"
+ by (simp only: cont2contlubE [OF t] ch2ch_cont [OF t]
+ cont2contlubE [OF f1] ch2ch_cont [OF f1]
+ cont2contlubE [OF f2] ch2ch_cont [OF f2]
+ diag_lub)
+qed
+
+lemma cont2cont_compose:
+ "\<lbrakk>cont c; cont (\<lambda>x. f x)\<rbrakk> \<Longrightarrow> cont (\<lambda>x. c (f x))"
+by (rule cont2cont_apply [OF cont_const])
+
text {* if-then-else is continuous *}
lemma cont_if [simp]:
--- a/src/HOLCF/Cprod.thy Thu Jan 15 15:51:50 2009 +0100
+++ b/src/HOLCF/Cprod.thy Thu Jan 15 09:10:42 2009 -0800
@@ -12,23 +12,6 @@
subsection {* Type @{typ unit} is a pcpo *}
-instantiation unit :: sq_ord
-begin
-
-definition
- less_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<equiv> 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 \<rightarrow> unit \<rightarrow> 'a" where
"unit_when = (\<Lambda> a _. a)"
@@ -39,165 +22,6 @@
lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>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>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)"
-
-instance ..
-end
-
-instance "*" :: (po, po) po
-proof
- fix x :: "'a \<times> 'b"
- show "x \<sqsubseteq> x"
- unfolding less_cprod_def by simp
-next
- fix x y :: "'a \<times> '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 \<times> '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: "\<lbrakk>fst p \<sqsubseteq> fst q; snd p \<sqsubseteq> snd q\<rbrakk> \<Longrightarrow> p \<sqsubseteq> q"
-unfolding less_cprod_def by simp
-
-lemma Pair_less_iff [simp]: "(a, b) \<sqsubseteq> (c, d) = (a \<sqsubseteq> c \<and> b \<sqsubseteq> d)"
-unfolding less_cprod_def by simp
-
-text {* Pair @{text "(_,_)"} is monotone in both arguments *}
-
-lemma monofun_pair1: "monofun (\<lambda>x. (x, y))"
-by (simp add: monofun_def)
-
-lemma monofun_pair2: "monofun (\<lambda>y. (x, y))"
-by (simp add: monofun_def)
-
-lemma monofun_pair:
- "\<lbrakk>x1 \<sqsubseteq> x2; y1 \<sqsubseteq> y2\<rbrakk> \<Longrightarrow> (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:
- "\<lbrakk>range X <<| x; range Y <<| y\<rbrakk> \<Longrightarrow> range (\<lambda>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 \<Rightarrow> ('a::cpo \<times> 'b::cpo)"
- assumes S: "chain S"
- shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
-proof -
- have "chain (\<lambda>i. fst (S i))"
- using monofun_fst S by (rule ch2ch_monofun)
- hence 1: "range (\<lambda>i. fst (S i)) <<| (\<Squnion>i. fst (S i))"
- by (rule cpo_lubI)
- have "chain (\<lambda>i. snd (S i))"
- using monofun_snd S by (rule ch2ch_monofun)
- hence 2: "range (\<lambda>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 \<Rightarrow> 'a::cpo \<times> 'b::cpo)
- \<Longrightarrow> (\<Squnion>i. S i) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
-by (rule lub_cprod [THEN thelubI])
-
-instance "*" :: (cpo, cpo) cpo
-proof
- fix S :: "nat \<Rightarrow> ('a \<times> 'b)"
- assume "chain S"
- hence "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
- by (rule lub_cprod)
- thus "\<exists>x. range S <<| x" ..
-qed
-
-instance "*" :: (finite_po, finite_po) finite_po ..
-
-instance "*" :: (discrete_cpo, discrete_cpo) discrete_cpo
-proof
- fix x y :: "'a \<times> 'b"
- show "x \<sqsubseteq> y \<longleftrightarrow> x = y"
- unfolding less_cprod_def Pair_fst_snd_eq
- by simp
-qed
-
-subsection {* Product type is pointed *}
-
-lemma minimal_cprod: "(\<bottom>, \<bottom>) \<sqsubseteq> p"
-by (simp add: less_cprod_def)
-
-instance "*" :: (pcpo, pcpo) pcpo
-by intro_classes (fast intro: minimal_cprod)
-
-lemma inst_cprod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)"
-by (rule minimal_cprod [THEN UU_I, symmetric])
-
-
-subsection {* Continuity of @{text "(_,_)"}, @{term fst}, @{term snd} *}
-
-lemma cont_pair1: "cont (\<lambda>x. (x, y))"
-apply (rule contI)
-apply (rule is_lub_Pair)
-apply (erule cpo_lubI)
-apply (rule lub_const)
-done
-
-lemma cont_pair2: "cont (\<lambda>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
@@ -245,10 +69,10 @@
by (simp add: cpair_eq_pair)
lemma cpair_less [iff]: "(<a, b> \<sqsubseteq> <a', b'>) = (a \<sqsubseteq> a' \<and> b \<sqsubseteq> b')"
-by (simp add: cpair_eq_pair less_cprod_def)
+by (simp add: cpair_eq_pair)
lemma cpair_defined_iff [iff]: "(<x, y> = \<bottom>) = (x = \<bottom> \<and> y = \<bottom>)"
-by (simp add: inst_cprod_pcpo cpair_eq_pair)
+by (simp add: cpair_eq_pair)
lemma cpair_strict [simp]: "\<langle>\<bottom>, \<bottom>\<rangle> = \<bottom>"
by simp
@@ -273,10 +97,10 @@
by (simp add: cpair_eq_pair csnd_def cont_snd)
lemma cfst_strict [simp]: "cfst\<cdot>\<bottom> = \<bottom>"
-unfolding inst_cprod_pcpo2 by (rule cfst_cpair)
+by (simp add: cfst_def)
lemma csnd_strict [simp]: "csnd\<cdot>\<bottom> = \<bottom>"
-unfolding inst_cprod_pcpo2 by (rule csnd_cpair)
+by (simp add: csnd_def)
lemma cpair_cfst_csnd: "\<langle>cfst\<cdot>p, csnd\<cdot>p\<rangle> = p"
by (cases p rule: cprodE, simp)
@@ -302,19 +126,10 @@
by (rule compactI, simp add: csnd_less_iff)
lemma compact_cpair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact <x, y>"
-by (rule compactI, simp add: less_cprod)
+by (simp add: cpair_eq_pair)
lemma compact_cpair_iff [simp]: "compact <x, y> = (compact x \<and> 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
+by (simp add: cpair_eq_pair)
lemma lub_cprod2:
"chain S \<Longrightarrow> range S <<| <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>"
--- a/src/HOLCF/Dsum.thy Thu Jan 15 15:51:50 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,251 +0,0 @@
-(* Title: HOLCF/Dsum.thy
- Author: Brian Huffman
-*)
-
-header {* The cpo of disjoint sums *}
-
-theory Dsum
-imports Bifinite
-begin
-
-subsection {* Ordering on type @{typ "'a + 'b"} *}
-
-instantiation "+" :: (sq_ord, sq_ord) sq_ord
-begin
-
-definition
- less_sum_def: "x \<sqsubseteq> y \<equiv> case x of
- Inl a \<Rightarrow> (case y of Inl b \<Rightarrow> a \<sqsubseteq> b | Inr b \<Rightarrow> False) |
- Inr a \<Rightarrow> (case y of Inl b \<Rightarrow> False | Inr b \<Rightarrow> a \<sqsubseteq> b)"
-
-instance ..
-end
-
-lemma Inl_less_iff [simp]: "Inl x \<sqsubseteq> Inl y = x \<sqsubseteq> y"
-unfolding less_sum_def by simp
-
-lemma Inr_less_iff [simp]: "Inr x \<sqsubseteq> Inr y = x \<sqsubseteq> y"
-unfolding less_sum_def by simp
-
-lemma Inl_less_Inr [simp]: "\<not> Inl x \<sqsubseteq> Inr y"
-unfolding less_sum_def by simp
-
-lemma Inr_less_Inl [simp]: "\<not> Inr x \<sqsubseteq> Inl y"
-unfolding less_sum_def by simp
-
-lemma Inl_mono: "x \<sqsubseteq> y \<Longrightarrow> Inl x \<sqsubseteq> Inl y"
-by simp
-
-lemma Inr_mono: "x \<sqsubseteq> y \<Longrightarrow> Inr x \<sqsubseteq> Inr y"
-by simp
-
-lemma Inl_lessE: "\<lbrakk>Inl a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
-by (cases x, simp_all)
-
-lemma Inr_lessE: "\<lbrakk>Inr a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
-by (cases x, simp_all)
-
-lemmas sum_less_elims = Inl_lessE Inr_lessE
-
-lemma sum_less_cases:
- "\<lbrakk>x \<sqsubseteq> y;
- \<And>a b. \<lbrakk>x = Inl a; y = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R;
- \<And>a b. \<lbrakk>x = Inr a; y = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk>
- \<Longrightarrow> R"
-by (cases x, safe elim!: sum_less_elims, auto)
-
-subsection {* Sum type is a complete partial order *}
-
-instance "+" :: (po, po) po
-proof
- fix x :: "'a + 'b"
- show "x \<sqsubseteq> x"
- by (induct x, simp_all)
-next
- fix x y :: "'a + 'b"
- assume "x \<sqsubseteq> y" and "y \<sqsubseteq> x" thus "x = y"
- by (induct x, auto elim!: sum_less_elims intro: antisym_less)
-next
- fix x y z :: "'a + 'b"
- assume "x \<sqsubseteq> y" and "y \<sqsubseteq> z" thus "x \<sqsubseteq> z"
- by (induct x, auto elim!: sum_less_elims intro: trans_less)
-qed
-
-lemma monofun_inv_Inl: "monofun (\<lambda>p. THE a. p = Inl a)"
-by (rule monofunI, erule sum_less_cases, simp_all)
-
-lemma monofun_inv_Inr: "monofun (\<lambda>p. THE b. p = Inr b)"
-by (rule monofunI, erule sum_less_cases, simp_all)
-
-lemma sum_chain_cases:
- assumes Y: "chain Y"
- assumes A: "\<And>A. \<lbrakk>chain A; Y = (\<lambda>i. Inl (A i))\<rbrakk> \<Longrightarrow> R"
- assumes B: "\<And>B. \<lbrakk>chain B; Y = (\<lambda>i. Inr (B i))\<rbrakk> \<Longrightarrow> R"
- shows "R"
- apply (cases "Y 0")
- apply (rule A)
- apply (rule ch2ch_monofun [OF monofun_inv_Inl Y])
- apply (rule ext)
- apply (cut_tac j=i in chain_mono [OF Y le0], simp)
- apply (erule Inl_lessE, simp)
- apply (rule B)
- apply (rule ch2ch_monofun [OF monofun_inv_Inr Y])
- apply (rule ext)
- apply (cut_tac j=i in chain_mono [OF Y le0], simp)
- apply (erule Inr_lessE, simp)
-done
-
-lemma is_lub_Inl: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inl (S i)) <<| Inl x"
- apply (rule is_lubI)
- apply (rule ub_rangeI)
- apply (simp add: is_ub_lub)
- apply (frule ub_rangeD [where i=arbitrary])
- apply (erule Inl_lessE, simp)
- apply (erule is_lub_lub)
- apply (rule ub_rangeI)
- apply (drule ub_rangeD, simp)
-done
-
-lemma is_lub_Inr: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inr (S i)) <<| Inr x"
- apply (rule is_lubI)
- apply (rule ub_rangeI)
- apply (simp add: is_ub_lub)
- apply (frule ub_rangeD [where i=arbitrary])
- apply (erule Inr_lessE, simp)
- apply (erule is_lub_lub)
- apply (rule ub_rangeI)
- apply (drule ub_rangeD, simp)
-done
-
-instance "+" :: (cpo, cpo) cpo
- apply intro_classes
- apply (erule sum_chain_cases, safe)
- apply (rule exI)
- apply (rule is_lub_Inl)
- apply (erule cpo_lubI)
- apply (rule exI)
- apply (rule is_lub_Inr)
- apply (erule cpo_lubI)
-done
-
-subsection {* Continuity of @{term Inl}, @{term Inr}, @{term sum_case} *}
-
-lemma cont2cont_Inl [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inl (f x))"
-by (fast intro: contI is_lub_Inl elim: contE)
-
-lemma cont2cont_Inr [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inr (f x))"
-by (fast intro: contI is_lub_Inr elim: contE)
-
-lemma cont_Inl: "cont Inl"
-by (rule cont2cont_Inl [OF cont_id])
-
-lemma cont_Inr: "cont Inr"
-by (rule cont2cont_Inr [OF cont_id])
-
-lemmas ch2ch_Inl [simp] = ch2ch_cont [OF cont_Inl]
-lemmas ch2ch_Inr [simp] = ch2ch_cont [OF cont_Inr]
-
-lemmas lub_Inl = cont2contlubE [OF cont_Inl, symmetric]
-lemmas lub_Inr = cont2contlubE [OF cont_Inr, symmetric]
-
-lemma cont_sum_case1:
- assumes f: "\<And>a. cont (\<lambda>x. f x a)"
- assumes g: "\<And>b. cont (\<lambda>x. g x b)"
- shows "cont (\<lambda>x. case y of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
-by (induct y, simp add: f, simp add: g)
-
-lemma cont_sum_case2: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (sum_case f g)"
-apply (rule contI)
-apply (erule sum_chain_cases)
-apply (simp add: cont2contlubE [OF cont_Inl, symmetric] contE)
-apply (simp add: cont2contlubE [OF cont_Inr, symmetric] contE)
-done
-
-lemma cont2cont_sum_case [simp]:
- assumes f1: "\<And>a. cont (\<lambda>x. f x a)" and f2: "\<And>x. cont (\<lambda>a. f x a)"
- assumes g1: "\<And>b. cont (\<lambda>x. g x b)" and g2: "\<And>x. cont (\<lambda>b. g x b)"
- assumes h: "cont (\<lambda>x. h x)"
- shows "cont (\<lambda>x. case h x of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
-apply (rule cont2cont_app2 [OF cont2cont_lambda _ h])
-apply (rule cont_sum_case1 [OF f1 g1])
-apply (rule cont_sum_case2 [OF f2 g2])
-done
-
-subsection {* Compactness and chain-finiteness *}
-
-lemma compact_Inl: "compact a \<Longrightarrow> compact (Inl a)"
-apply (rule compactI2)
-apply (erule sum_chain_cases, safe)
-apply (simp add: lub_Inl)
-apply (erule (2) compactD2)
-apply (simp add: lub_Inr)
-done
-
-lemma compact_Inr: "compact a \<Longrightarrow> compact (Inr a)"
-apply (rule compactI2)
-apply (erule sum_chain_cases, safe)
-apply (simp add: lub_Inl)
-apply (simp add: lub_Inr)
-apply (erule (2) compactD2)
-done
-
-lemma compact_Inl_rev: "compact (Inl a) \<Longrightarrow> compact a"
-unfolding compact_def
-by (drule adm_subst [OF cont_Inl], simp)
-
-lemma compact_Inr_rev: "compact (Inr a) \<Longrightarrow> compact a"
-unfolding compact_def
-by (drule adm_subst [OF cont_Inr], simp)
-
-lemma compact_Inl_iff [simp]: "compact (Inl a) = compact a"
-by (safe elim!: compact_Inl compact_Inl_rev)
-
-lemma compact_Inr_iff [simp]: "compact (Inr a) = compact a"
-by (safe elim!: compact_Inr compact_Inr_rev)
-
-instance "+" :: (chfin, chfin) chfin
-apply intro_classes
-apply (erule compact_imp_max_in_chain)
-apply (case_tac "\<Squnion>i. Y i", simp_all)
-done
-
-instance "+" :: (finite_po, finite_po) finite_po ..
-
-instance "+" :: (discrete_cpo, discrete_cpo) discrete_cpo
-by intro_classes (simp add: less_sum_def split: sum.split)
-
-subsection {* Sum type is a bifinite domain *}
-
-instantiation "+" :: (profinite, profinite) profinite
-begin
-
-definition
- approx_sum_def: "approx =
- (\<lambda>n. \<Lambda> x. case x of Inl a \<Rightarrow> Inl (approx n\<cdot>a) | Inr b \<Rightarrow> Inr (approx n\<cdot>b))"
-
-lemma approx_Inl [simp]: "approx n\<cdot>(Inl x) = Inl (approx n\<cdot>x)"
- unfolding approx_sum_def by simp
-
-lemma approx_Inr [simp]: "approx n\<cdot>(Inr x) = Inr (approx n\<cdot>x)"
- unfolding approx_sum_def by simp
-
-instance proof
- fix i :: nat and x :: "'a + 'b"
- show "chain (approx :: nat \<Rightarrow> 'a + 'b \<rightarrow> 'a + 'b)"
- unfolding approx_sum_def
- by (rule ch2ch_LAM, case_tac x, simp_all)
- show "(\<Squnion>i. approx i\<cdot>x) = x"
- by (induct x, simp_all add: lub_Inl lub_Inr)
- show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
- by (induct x, simp_all)
- have "{x::'a + 'b. approx i\<cdot>x = x} \<subseteq>
- {x::'a. approx i\<cdot>x = x} <+> {x::'b. approx i\<cdot>x = x}"
- by (rule subsetI, case_tac x, simp_all add: InlI InrI)
- thus "finite {x::'a + 'b. approx i\<cdot>x = x}"
- by (rule finite_subset,
- intro finite_Plus finite_fixes_approx)
-qed
-
-end
-
-end
--- a/src/HOLCF/Fixrec.thy Thu Jan 15 15:51:50 2009 +0100
+++ b/src/HOLCF/Fixrec.thy Thu Jan 15 09:10:42 2009 -0800
@@ -55,6 +55,7 @@
"maybe_when\<cdot>f\<cdot>r\<cdot>fail = f"
"maybe_when\<cdot>f\<cdot>r\<cdot>(return\<cdot>x) = r\<cdot>x"
by (simp_all add: return_def fail_def maybe_when_def cont_Rep_maybe
+ cont2cont_LAM
cont_Abs_maybe Abs_maybe_inverse Rep_maybe_strict)
translations
--- a/src/HOLCF/HOLCF.thy Thu Jan 15 15:51:50 2009 +0100
+++ b/src/HOLCF/HOLCF.thy Thu Jan 15 09:10:42 2009 -0800
@@ -6,10 +6,11 @@
theory HOLCF
imports
- Domain ConvexPD Algebraic Universal Dsum Main
+ Domain ConvexPD Algebraic Universal Sum_Cpo Main
uses
"holcf_logic.ML"
"Tools/cont_consts.ML"
+ "Tools/cont_proc.ML"
"Tools/domain/domain_library.ML"
"Tools/domain/domain_syntax.ML"
"Tools/domain/domain_axioms.ML"
--- a/src/HOLCF/IsaMakefile Thu Jan 15 15:51:50 2009 +0100
+++ b/src/HOLCF/IsaMakefile Thu Jan 15 09:10:42 2009 -0800
@@ -1,5 +1,3 @@
-#
-# $Id$
#
# IsaMakefile for HOLCF
#
@@ -41,7 +39,6 @@
Discrete.thy \
Deflation.thy \
Domain.thy \
- Dsum.thy \
Eventual.thy \
Ffun.thy \
Fixrec.thy \
@@ -54,8 +51,10 @@
Pcpodef.thy \
Pcpo.thy \
Porder.thy \
+ Product_Cpo.thy \
Sprod.thy \
Ssum.thy \
+ Sum_Cpo.thy \
Tr.thy \
Universal.thy \
UpperPD.thy \
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Product_Cpo.thy Thu Jan 15 09:10:42 2009 -0800
@@ -0,0 +1,247 @@
+(* Title: HOLCF/Product_Cpo.thy
+ Author: Franz Regensburger
+*)
+
+header {* The cpo of cartesian products *}
+
+theory Product_Cpo
+imports Adm
+begin
+
+defaultsort cpo
+
+subsection {* Type @{typ unit} is a pcpo *}
+
+instantiation unit :: sq_ord
+begin
+
+definition
+ less_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<equiv> True"
+
+instance ..
+end
+
+instance unit :: discrete_cpo
+by intro_classes simp
+
+instance unit :: finite_po ..
+
+instance unit :: pcpo
+by intro_classes simp
+
+
+subsection {* Product type is a partial order *}
+
+instantiation "*" :: (sq_ord, sq_ord) sq_ord
+begin
+
+definition
+ less_cprod_def: "(op \<sqsubseteq>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)"
+
+instance ..
+end
+
+instance "*" :: (po, po) po
+proof
+ fix x :: "'a \<times> 'b"
+ show "x \<sqsubseteq> x"
+ unfolding less_cprod_def by simp
+next
+ fix x y :: "'a \<times> '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 \<times> '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: "\<lbrakk>fst p \<sqsubseteq> fst q; snd p \<sqsubseteq> snd q\<rbrakk> \<Longrightarrow> p \<sqsubseteq> q"
+unfolding less_cprod_def by simp
+
+lemma Pair_less_iff [simp]: "(a, b) \<sqsubseteq> (c, d) \<longleftrightarrow> a \<sqsubseteq> c \<and> b \<sqsubseteq> d"
+unfolding less_cprod_def by simp
+
+text {* Pair @{text "(_,_)"} is monotone in both arguments *}
+
+lemma monofun_pair1: "monofun (\<lambda>x. (x, y))"
+by (simp add: monofun_def)
+
+lemma monofun_pair2: "monofun (\<lambda>y. (x, y))"
+by (simp add: monofun_def)
+
+lemma monofun_pair:
+ "\<lbrakk>x1 \<sqsubseteq> x2; y1 \<sqsubseteq> y2\<rbrakk> \<Longrightarrow> (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:
+ "\<lbrakk>range X <<| x; range Y <<| y\<rbrakk> \<Longrightarrow> range (\<lambda>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 \<Rightarrow> ('a::cpo \<times> 'b::cpo)"
+ assumes S: "chain S"
+ shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
+proof -
+ have "chain (\<lambda>i. fst (S i))"
+ using monofun_fst S by (rule ch2ch_monofun)
+ hence 1: "range (\<lambda>i. fst (S i)) <<| (\<Squnion>i. fst (S i))"
+ by (rule cpo_lubI)
+ have "chain (\<lambda>i. snd (S i))"
+ using monofun_snd S by (rule ch2ch_monofun)
+ hence 2: "range (\<lambda>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 \<Rightarrow> 'a::cpo \<times> 'b::cpo)
+ \<Longrightarrow> (\<Squnion>i. S i) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
+by (rule lub_cprod [THEN thelubI])
+
+instance "*" :: (cpo, cpo) cpo
+proof
+ fix S :: "nat \<Rightarrow> ('a \<times> 'b)"
+ assume "chain S"
+ hence "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
+ by (rule lub_cprod)
+ thus "\<exists>x. range S <<| x" ..
+qed
+
+instance "*" :: (finite_po, finite_po) finite_po ..
+
+instance "*" :: (discrete_cpo, discrete_cpo) discrete_cpo
+proof
+ fix x y :: "'a \<times> 'b"
+ show "x \<sqsubseteq> y \<longleftrightarrow> x = y"
+ unfolding less_cprod_def Pair_fst_snd_eq
+ by simp
+qed
+
+subsection {* Product type is pointed *}
+
+lemma minimal_cprod: "(\<bottom>, \<bottom>) \<sqsubseteq> p"
+by (simp add: less_cprod_def)
+
+instance "*" :: (pcpo, pcpo) pcpo
+by intro_classes (fast intro: minimal_cprod)
+
+lemma inst_cprod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)"
+by (rule minimal_cprod [THEN UU_I, symmetric])
+
+lemma Pair_defined_iff [simp]: "(x, y) = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>"
+unfolding inst_cprod_pcpo by simp
+
+lemma fst_strict [simp]: "fst \<bottom> = \<bottom>"
+unfolding inst_cprod_pcpo by (rule fst_conv)
+
+lemma csnd_strict [simp]: "snd \<bottom> = \<bottom>"
+unfolding inst_cprod_pcpo by (rule snd_conv)
+
+lemma Pair_strict [simp]: "(\<bottom>, \<bottom>) = \<bottom>"
+by simp
+
+lemma split_strict [simp]: "split f \<bottom> = f \<bottom> \<bottom>"
+unfolding split_def by simp
+
+subsection {* Continuity of @{text "(_,_)"}, @{term fst}, @{term snd} *}
+
+lemma cont_pair1: "cont (\<lambda>x. (x, y))"
+apply (rule contI)
+apply (rule is_lub_Pair)
+apply (erule cpo_lubI)
+apply (rule lub_const)
+done
+
+lemma cont_pair2: "cont (\<lambda>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
+
+lemma cont2cont_Pair [cont2cont]:
+ assumes f: "cont (\<lambda>x. f x)"
+ assumes g: "cont (\<lambda>x. g x)"
+ shows "cont (\<lambda>x. (f x, g x))"
+apply (rule cont2cont_apply [OF _ cont_pair1 f])
+apply (rule cont2cont_apply [OF _ cont_pair2 g])
+apply (rule cont_const)
+done
+
+lemmas cont2cont_fst [cont2cont] = cont2cont_compose [OF cont_fst]
+
+lemmas cont2cont_snd [cont2cont] = cont2cont_compose [OF cont_snd]
+
+subsection {* Compactness and chain-finiteness *}
+
+lemma fst_less_iff: "fst (x::'a \<times> 'b) \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (y, snd x)"
+unfolding less_cprod_def by simp
+
+lemma snd_less_iff: "snd (x::'a \<times> 'b) \<sqsubseteq> y = x \<sqsubseteq> (fst x, y)"
+unfolding less_cprod_def by simp
+
+lemma compact_fst: "compact x \<Longrightarrow> compact (fst x)"
+by (rule compactI, simp add: fst_less_iff)
+
+lemma compact_snd: "compact x \<Longrightarrow> compact (snd x)"
+by (rule compactI, simp add: snd_less_iff)
+
+lemma compact_Pair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact (x, y)"
+by (rule compactI, simp add: less_cprod_def)
+
+lemma compact_Pair_iff [simp]: "compact (x, y) \<longleftrightarrow> compact x \<and> compact y"
+apply (safe intro!: compact_Pair)
+apply (drule compact_fst, simp)
+apply (drule compact_snd, simp)
+done
+
+instance "*" :: (chfin, chfin) chfin
+apply intro_classes
+apply (erule compact_imp_max_in_chain)
+apply (case_tac "\<Squnion>i. Y i", simp)
+done
+
+end
--- a/src/HOLCF/Ssum.thy Thu Jan 15 15:51:50 2009 +0100
+++ b/src/HOLCF/Ssum.thy Thu Jan 15 09:10:42 2009 -0800
@@ -188,7 +188,7 @@
lemma beta_sscase:
"sscase\<cdot>f\<cdot>g\<cdot>s = (\<Lambda><t, x, y>. If t then f\<cdot>x else g\<cdot>y fi)\<cdot>(Rep_Ssum s)"
-unfolding sscase_def by (simp add: cont_Rep_Ssum)
+unfolding sscase_def by (simp add: cont_Rep_Ssum cont2cont_LAM)
lemma sscase1 [simp]: "sscase\<cdot>f\<cdot>g\<cdot>\<bottom> = \<bottom>"
unfolding beta_sscase by (simp add: Rep_Ssum_strict)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Sum_Cpo.thy Thu Jan 15 09:10:42 2009 -0800
@@ -0,0 +1,251 @@
+(* Title: HOLCF/Sum_Cpo.thy
+ Author: Brian Huffman
+*)
+
+header {* The cpo of disjoint sums *}
+
+theory Sum_Cpo
+imports Bifinite
+begin
+
+subsection {* Ordering on type @{typ "'a + 'b"} *}
+
+instantiation "+" :: (sq_ord, sq_ord) sq_ord
+begin
+
+definition
+ less_sum_def: "x \<sqsubseteq> y \<equiv> case x of
+ Inl a \<Rightarrow> (case y of Inl b \<Rightarrow> a \<sqsubseteq> b | Inr b \<Rightarrow> False) |
+ Inr a \<Rightarrow> (case y of Inl b \<Rightarrow> False | Inr b \<Rightarrow> a \<sqsubseteq> b)"
+
+instance ..
+end
+
+lemma Inl_less_iff [simp]: "Inl x \<sqsubseteq> Inl y = x \<sqsubseteq> y"
+unfolding less_sum_def by simp
+
+lemma Inr_less_iff [simp]: "Inr x \<sqsubseteq> Inr y = x \<sqsubseteq> y"
+unfolding less_sum_def by simp
+
+lemma Inl_less_Inr [simp]: "\<not> Inl x \<sqsubseteq> Inr y"
+unfolding less_sum_def by simp
+
+lemma Inr_less_Inl [simp]: "\<not> Inr x \<sqsubseteq> Inl y"
+unfolding less_sum_def by simp
+
+lemma Inl_mono: "x \<sqsubseteq> y \<Longrightarrow> Inl x \<sqsubseteq> Inl y"
+by simp
+
+lemma Inr_mono: "x \<sqsubseteq> y \<Longrightarrow> Inr x \<sqsubseteq> Inr y"
+by simp
+
+lemma Inl_lessE: "\<lbrakk>Inl a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by (cases x, simp_all)
+
+lemma Inr_lessE: "\<lbrakk>Inr a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by (cases x, simp_all)
+
+lemmas sum_less_elims = Inl_lessE Inr_lessE
+
+lemma sum_less_cases:
+ "\<lbrakk>x \<sqsubseteq> y;
+ \<And>a b. \<lbrakk>x = Inl a; y = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R;
+ \<And>a b. \<lbrakk>x = Inr a; y = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk>
+ \<Longrightarrow> R"
+by (cases x, safe elim!: sum_less_elims, auto)
+
+subsection {* Sum type is a complete partial order *}
+
+instance "+" :: (po, po) po
+proof
+ fix x :: "'a + 'b"
+ show "x \<sqsubseteq> x"
+ by (induct x, simp_all)
+next
+ fix x y :: "'a + 'b"
+ assume "x \<sqsubseteq> y" and "y \<sqsubseteq> x" thus "x = y"
+ by (induct x, auto elim!: sum_less_elims intro: antisym_less)
+next
+ fix x y z :: "'a + 'b"
+ assume "x \<sqsubseteq> y" and "y \<sqsubseteq> z" thus "x \<sqsubseteq> z"
+ by (induct x, auto elim!: sum_less_elims intro: trans_less)
+qed
+
+lemma monofun_inv_Inl: "monofun (\<lambda>p. THE a. p = Inl a)"
+by (rule monofunI, erule sum_less_cases, simp_all)
+
+lemma monofun_inv_Inr: "monofun (\<lambda>p. THE b. p = Inr b)"
+by (rule monofunI, erule sum_less_cases, simp_all)
+
+lemma sum_chain_cases:
+ assumes Y: "chain Y"
+ assumes A: "\<And>A. \<lbrakk>chain A; Y = (\<lambda>i. Inl (A i))\<rbrakk> \<Longrightarrow> R"
+ assumes B: "\<And>B. \<lbrakk>chain B; Y = (\<lambda>i. Inr (B i))\<rbrakk> \<Longrightarrow> R"
+ shows "R"
+ apply (cases "Y 0")
+ apply (rule A)
+ apply (rule ch2ch_monofun [OF monofun_inv_Inl Y])
+ apply (rule ext)
+ apply (cut_tac j=i in chain_mono [OF Y le0], simp)
+ apply (erule Inl_lessE, simp)
+ apply (rule B)
+ apply (rule ch2ch_monofun [OF monofun_inv_Inr Y])
+ apply (rule ext)
+ apply (cut_tac j=i in chain_mono [OF Y le0], simp)
+ apply (erule Inr_lessE, simp)
+done
+
+lemma is_lub_Inl: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inl (S i)) <<| Inl x"
+ apply (rule is_lubI)
+ apply (rule ub_rangeI)
+ apply (simp add: is_ub_lub)
+ apply (frule ub_rangeD [where i=arbitrary])
+ apply (erule Inl_lessE, simp)
+ apply (erule is_lub_lub)
+ apply (rule ub_rangeI)
+ apply (drule ub_rangeD, simp)
+done
+
+lemma is_lub_Inr: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inr (S i)) <<| Inr x"
+ apply (rule is_lubI)
+ apply (rule ub_rangeI)
+ apply (simp add: is_ub_lub)
+ apply (frule ub_rangeD [where i=arbitrary])
+ apply (erule Inr_lessE, simp)
+ apply (erule is_lub_lub)
+ apply (rule ub_rangeI)
+ apply (drule ub_rangeD, simp)
+done
+
+instance "+" :: (cpo, cpo) cpo
+ apply intro_classes
+ apply (erule sum_chain_cases, safe)
+ apply (rule exI)
+ apply (rule is_lub_Inl)
+ apply (erule cpo_lubI)
+ apply (rule exI)
+ apply (rule is_lub_Inr)
+ apply (erule cpo_lubI)
+done
+
+subsection {* Continuity of @{term Inl}, @{term Inr}, @{term sum_case} *}
+
+lemma cont2cont_Inl [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inl (f x))"
+by (fast intro: contI is_lub_Inl elim: contE)
+
+lemma cont2cont_Inr [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inr (f x))"
+by (fast intro: contI is_lub_Inr elim: contE)
+
+lemma cont_Inl: "cont Inl"
+by (rule cont2cont_Inl [OF cont_id])
+
+lemma cont_Inr: "cont Inr"
+by (rule cont2cont_Inr [OF cont_id])
+
+lemmas ch2ch_Inl [simp] = ch2ch_cont [OF cont_Inl]
+lemmas ch2ch_Inr [simp] = ch2ch_cont [OF cont_Inr]
+
+lemmas lub_Inl = cont2contlubE [OF cont_Inl, symmetric]
+lemmas lub_Inr = cont2contlubE [OF cont_Inr, symmetric]
+
+lemma cont_sum_case1:
+ assumes f: "\<And>a. cont (\<lambda>x. f x a)"
+ assumes g: "\<And>b. cont (\<lambda>x. g x b)"
+ shows "cont (\<lambda>x. case y of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
+by (induct y, simp add: f, simp add: g)
+
+lemma cont_sum_case2: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (sum_case f g)"
+apply (rule contI)
+apply (erule sum_chain_cases)
+apply (simp add: cont2contlubE [OF cont_Inl, symmetric] contE)
+apply (simp add: cont2contlubE [OF cont_Inr, symmetric] contE)
+done
+
+lemma cont2cont_sum_case [simp]:
+ assumes f1: "\<And>a. cont (\<lambda>x. f x a)" and f2: "\<And>x. cont (\<lambda>a. f x a)"
+ assumes g1: "\<And>b. cont (\<lambda>x. g x b)" and g2: "\<And>x. cont (\<lambda>b. g x b)"
+ assumes h: "cont (\<lambda>x. h x)"
+ shows "cont (\<lambda>x. case h x of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
+apply (rule cont2cont_app2 [OF cont2cont_lambda _ h])
+apply (rule cont_sum_case1 [OF f1 g1])
+apply (rule cont_sum_case2 [OF f2 g2])
+done
+
+subsection {* Compactness and chain-finiteness *}
+
+lemma compact_Inl: "compact a \<Longrightarrow> compact (Inl a)"
+apply (rule compactI2)
+apply (erule sum_chain_cases, safe)
+apply (simp add: lub_Inl)
+apply (erule (2) compactD2)
+apply (simp add: lub_Inr)
+done
+
+lemma compact_Inr: "compact a \<Longrightarrow> compact (Inr a)"
+apply (rule compactI2)
+apply (erule sum_chain_cases, safe)
+apply (simp add: lub_Inl)
+apply (simp add: lub_Inr)
+apply (erule (2) compactD2)
+done
+
+lemma compact_Inl_rev: "compact (Inl a) \<Longrightarrow> compact a"
+unfolding compact_def
+by (drule adm_subst [OF cont_Inl], simp)
+
+lemma compact_Inr_rev: "compact (Inr a) \<Longrightarrow> compact a"
+unfolding compact_def
+by (drule adm_subst [OF cont_Inr], simp)
+
+lemma compact_Inl_iff [simp]: "compact (Inl a) = compact a"
+by (safe elim!: compact_Inl compact_Inl_rev)
+
+lemma compact_Inr_iff [simp]: "compact (Inr a) = compact a"
+by (safe elim!: compact_Inr compact_Inr_rev)
+
+instance "+" :: (chfin, chfin) chfin
+apply intro_classes
+apply (erule compact_imp_max_in_chain)
+apply (case_tac "\<Squnion>i. Y i", simp_all)
+done
+
+instance "+" :: (finite_po, finite_po) finite_po ..
+
+instance "+" :: (discrete_cpo, discrete_cpo) discrete_cpo
+by intro_classes (simp add: less_sum_def split: sum.split)
+
+subsection {* Sum type is a bifinite domain *}
+
+instantiation "+" :: (profinite, profinite) profinite
+begin
+
+definition
+ approx_sum_def: "approx =
+ (\<lambda>n. \<Lambda> x. case x of Inl a \<Rightarrow> Inl (approx n\<cdot>a) | Inr b \<Rightarrow> Inr (approx n\<cdot>b))"
+
+lemma approx_Inl [simp]: "approx n\<cdot>(Inl x) = Inl (approx n\<cdot>x)"
+ unfolding approx_sum_def by simp
+
+lemma approx_Inr [simp]: "approx n\<cdot>(Inr x) = Inr (approx n\<cdot>x)"
+ unfolding approx_sum_def by simp
+
+instance proof
+ fix i :: nat and x :: "'a + 'b"
+ show "chain (approx :: nat \<Rightarrow> 'a + 'b \<rightarrow> 'a + 'b)"
+ unfolding approx_sum_def
+ by (rule ch2ch_LAM, case_tac x, simp_all)
+ show "(\<Squnion>i. approx i\<cdot>x) = x"
+ by (induct x, simp_all add: lub_Inl lub_Inr)
+ show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
+ by (induct x, simp_all)
+ have "{x::'a + 'b. approx i\<cdot>x = x} \<subseteq>
+ {x::'a. approx i\<cdot>x = x} <+> {x::'b. approx i\<cdot>x = x}"
+ by (rule subsetI, case_tac x, simp_all add: InlI InrI)
+ thus "finite {x::'a + 'b. approx i\<cdot>x = x}"
+ by (rule finite_subset,
+ intro finite_Plus finite_fixes_approx)
+qed
+
+end
+
+end
--- a/src/HOLCF/Up.thy Thu Jan 15 15:51:50 2009 +0100
+++ b/src/HOLCF/Up.thy Thu Jan 15 09:10:42 2009 -0800
@@ -282,10 +282,10 @@
text {* properties of fup *}
lemma fup1 [simp]: "fup\<cdot>f\<cdot>\<bottom> = \<bottom>"
-by (simp add: fup_def cont_Ifup1 cont_Ifup2 inst_up_pcpo)
+by (simp add: fup_def cont_Ifup1 cont_Ifup2 inst_up_pcpo cont2cont_LAM)
lemma fup2 [simp]: "fup\<cdot>f\<cdot>(up\<cdot>x) = f\<cdot>x"
-by (simp add: up_def fup_def cont_Iup cont_Ifup1 cont_Ifup2)
+by (simp add: up_def fup_def cont_Iup cont_Ifup1 cont_Ifup2 cont2cont_LAM)
lemma fup3 [simp]: "fup\<cdot>up\<cdot>x = x"
by (cases x, simp_all)
--- a/src/HOLCF/ex/Stream.thy Thu Jan 15 15:51:50 2009 +0100
+++ b/src/HOLCF/ex/Stream.thy Thu Jan 15 09:10:42 2009 -0800
@@ -521,7 +521,7 @@
section "smap"
lemma smap_unfold: "smap = (\<Lambda> f t. case t of x&&xs \<Rightarrow> f$x && smap$f$xs)"
-by (insert smap_def [THEN eq_reflection, THEN fix_eq2], auto)
+by (insert smap_def [where 'a='a and 'b='b, THEN eq_reflection, THEN fix_eq2], auto)
lemma smap_empty [simp]: "smap\<cdot>f\<cdot>\<bottom> = \<bottom>"
by (subst smap_unfold, simp)
@@ -540,7 +540,7 @@
lemma sfilter_unfold:
"sfilter = (\<Lambda> p s. case s of x && xs \<Rightarrow>
If p\<cdot>x then x && sfilter\<cdot>p\<cdot>xs else sfilter\<cdot>p\<cdot>xs fi)"
-by (insert sfilter_def [THEN eq_reflection, THEN fix_eq2], auto)
+by (insert sfilter_def [where 'a='a, THEN eq_reflection, THEN fix_eq2], auto)
lemma strict_sfilter: "sfilter\<cdot>\<bottom> = \<bottom>"
apply (rule ext_cfun)