use default names sprod/Rep_sprod/Abs_sprod from pcpodef instead of Sprod/Rep_Sprod/Abs_Sprod; similarly for ssum
authorhuffman
Sun, 24 Oct 2010 03:43:12 -0700
changeset 40098 9dbb01456031
parent 40097 429cd74f795f
child 40099 0fb7891f0c7c
use default names sprod/Rep_sprod/Abs_sprod from pcpodef instead of Sprod/Rep_Sprod/Abs_Sprod; similarly for ssum
src/HOLCF/Sprod.thy
src/HOLCF/Ssum.thy
--- a/src/HOLCF/Sprod.thy	Sat Oct 23 19:56:33 2010 -0700
+++ b/src/HOLCF/Sprod.thy	Sun Oct 24 03:43:12 2010 -0700
@@ -12,12 +12,12 @@
 
 subsection {* Definition of strict product type *}
 
-pcpodef (Sprod)  ('a, 'b) sprod (infixr "**" 20) =
+pcpodef ('a, 'b) sprod (infixr "**" 20) =
         "{p::'a \<times> 'b. p = \<bottom> \<or> (fst p \<noteq> \<bottom> \<and> snd p \<noteq> \<bottom>)}"
 by simp_all
 
 instance sprod :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin
-by (rule typedef_chfin [OF type_definition_Sprod below_Sprod_def])
+by (rule typedef_chfin [OF type_definition_sprod below_sprod_def])
 
 type_notation (xsymbols)
   sprod  ("(_ \<otimes>/ _)" [21,20] 20)
@@ -28,15 +28,15 @@
 
 definition
   sfst :: "('a ** 'b) \<rightarrow> 'a" where
-  "sfst = (\<Lambda> p. fst (Rep_Sprod p))"
+  "sfst = (\<Lambda> p. fst (Rep_sprod p))"
 
 definition
   ssnd :: "('a ** 'b) \<rightarrow> 'b" where
-  "ssnd = (\<Lambda> p. snd (Rep_Sprod p))"
+  "ssnd = (\<Lambda> p. snd (Rep_sprod p))"
 
 definition
   spair :: "'a \<rightarrow> 'b \<rightarrow> ('a ** 'b)" where
-  "spair = (\<Lambda> a b. Abs_Sprod (strict\<cdot>b\<cdot>a, strict\<cdot>a\<cdot>b))"
+  "spair = (\<Lambda> a b. Abs_sprod (strict\<cdot>b\<cdot>a, strict\<cdot>a\<cdot>b))"
 
 definition
   ssplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a ** 'b) \<rightarrow> 'c" where
@@ -53,20 +53,20 @@
 
 subsection {* Case analysis *}
 
-lemma spair_Sprod: "(strict\<cdot>b\<cdot>a, strict\<cdot>a\<cdot>b) \<in> Sprod"
-by (simp add: Sprod_def strict_conv_if)
+lemma spair_sprod: "(strict\<cdot>b\<cdot>a, strict\<cdot>a\<cdot>b) \<in> sprod"
+by (simp add: sprod_def strict_conv_if)
 
-lemma Rep_Sprod_spair: "Rep_Sprod (:a, b:) = (strict\<cdot>b\<cdot>a, strict\<cdot>a\<cdot>b)"
-by (simp add: spair_def cont_Abs_Sprod Abs_Sprod_inverse spair_Sprod)
+lemma Rep_sprod_spair: "Rep_sprod (:a, b:) = (strict\<cdot>b\<cdot>a, strict\<cdot>a\<cdot>b)"
+by (simp add: spair_def cont_Abs_sprod Abs_sprod_inverse spair_sprod)
 
-lemmas Rep_Sprod_simps =
-  Rep_Sprod_inject [symmetric] below_Sprod_def
+lemmas Rep_sprod_simps =
+  Rep_sprod_inject [symmetric] below_sprod_def
   Pair_fst_snd_eq below_prod_def
-  Rep_Sprod_strict Rep_Sprod_spair
+  Rep_sprod_strict Rep_sprod_spair
 
 lemma sprodE [case_names bottom spair, cases type: sprod]:
   obtains "p = \<bottom>" | x y where "p = (:x, y:)" and "x \<noteq> \<bottom>" and "y \<noteq> \<bottom>"
-using Rep_Sprod [of p] by (auto simp add: Sprod_def Rep_Sprod_simps)
+using Rep_sprod [of p] by (auto simp add: sprod_def Rep_sprod_simps)
 
 lemma sprod_induct [case_names bottom spair, induct type: sprod]:
   "\<lbrakk>P \<bottom>; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> P (:x, y:)\<rbrakk> \<Longrightarrow> P x"
@@ -75,22 +75,22 @@
 subsection {* Properties of \emph{spair} *}
 
 lemma spair_strict1 [simp]: "(:\<bottom>, y:) = \<bottom>"
-by (simp add: Rep_Sprod_simps)
+by (simp add: Rep_sprod_simps)
 
 lemma spair_strict2 [simp]: "(:x, \<bottom>:) = \<bottom>"
-by (simp add: Rep_Sprod_simps)
+by (simp add: Rep_sprod_simps)
 
 lemma spair_strict_iff [simp]: "((:x, y:) = \<bottom>) = (x = \<bottom> \<or> y = \<bottom>)"
-by (simp add: Rep_Sprod_simps strict_conv_if)
+by (simp add: Rep_sprod_simps strict_conv_if)
 
 lemma spair_below_iff:
   "((:a, b:) \<sqsubseteq> (:c, d:)) = (a = \<bottom> \<or> b = \<bottom> \<or> (a \<sqsubseteq> c \<and> b \<sqsubseteq> d))"
-by (simp add: Rep_Sprod_simps strict_conv_if)
+by (simp add: Rep_sprod_simps strict_conv_if)
 
 lemma spair_eq_iff:
   "((:a, b:) = (:c, d:)) =
     (a = c \<and> b = d \<or> (a = \<bottom> \<or> b = \<bottom>) \<and> (c = \<bottom> \<or> d = \<bottom>))"
-by (simp add: Rep_Sprod_simps strict_conv_if)
+by (simp add: Rep_sprod_simps strict_conv_if)
 
 lemma spair_strict: "x = \<bottom> \<or> y = \<bottom> \<Longrightarrow> (:x, y:) = \<bottom>"
 by simp
@@ -125,16 +125,16 @@
 subsection {* Properties of \emph{sfst} and \emph{ssnd} *}
 
 lemma sfst_strict [simp]: "sfst\<cdot>\<bottom> = \<bottom>"
-by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_strict)
+by (simp add: sfst_def cont_Rep_sprod Rep_sprod_strict)
 
 lemma ssnd_strict [simp]: "ssnd\<cdot>\<bottom> = \<bottom>"
-by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_strict)
+by (simp add: ssnd_def cont_Rep_sprod Rep_sprod_strict)
 
 lemma sfst_spair [simp]: "y \<noteq> \<bottom> \<Longrightarrow> sfst\<cdot>(:x, y:) = x"
-by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_spair)
+by (simp add: sfst_def cont_Rep_sprod Rep_sprod_spair)
 
 lemma ssnd_spair [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssnd\<cdot>(:x, y:) = y"
-by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_spair)
+by (simp add: ssnd_def cont_Rep_sprod Rep_sprod_spair)
 
 lemma sfst_defined_iff [simp]: "(sfst\<cdot>p = \<bottom>) = (p = \<bottom>)"
 by (cases p, simp_all)
@@ -152,7 +152,7 @@
 by (cases p, simp_all)
 
 lemma below_sprod: "x \<sqsubseteq> y = (sfst\<cdot>x \<sqsubseteq> sfst\<cdot>y \<and> ssnd\<cdot>x \<sqsubseteq> ssnd\<cdot>y)"
-by (simp add: Rep_Sprod_simps sfst_def ssnd_def cont_Rep_Sprod)
+by (simp add: Rep_sprod_simps sfst_def ssnd_def cont_Rep_sprod)
 
 lemma eq_sprod: "(x = y) = (sfst\<cdot>x = sfst\<cdot>y \<and> ssnd\<cdot>x = ssnd\<cdot>y)"
 by (auto simp add: po_eq_conv below_sprod)
@@ -176,7 +176,7 @@
 by (rule compactI, simp add: ssnd_below_iff)
 
 lemma compact_spair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact (:x, y:)"
-by (rule compact_Sprod, simp add: Rep_Sprod_spair strict_conv_if)
+by (rule compact_sprod, simp add: Rep_sprod_spair strict_conv_if)
 
 lemma compact_spair_iff:
   "compact (:x, y:) = (x = \<bottom> \<or> y = \<bottom> \<or> (compact x \<and> compact y))"
--- a/src/HOLCF/Ssum.thy	Sat Oct 23 19:56:33 2010 -0700
+++ b/src/HOLCF/Ssum.thy	Sun Oct 24 03:43:12 2010 -0700
@@ -12,14 +12,14 @@
 
 subsection {* Definition of strict sum type *}
 
-pcpodef (Ssum)  ('a, 'b) ssum (infixr "++" 10) = 
+pcpodef ('a, 'b) ssum (infixr "++" 10) = 
   "{p :: tr \<times> ('a \<times> 'b). p = \<bottom> \<or>
     (fst p = TT \<and> fst (snd p) \<noteq> \<bottom> \<and> snd (snd p) = \<bottom>) \<or>
     (fst p = FF \<and> fst (snd p) = \<bottom> \<and> snd (snd p) \<noteq> \<bottom>) }"
 by simp_all
 
 instance ssum :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin
-by (rule typedef_chfin [OF type_definition_Ssum below_Ssum_def])
+by (rule typedef_chfin [OF type_definition_ssum below_ssum_def])
 
 type_notation (xsymbols)
   ssum  ("(_ \<oplus>/ _)" [21, 20] 20)
@@ -31,44 +31,44 @@
 
 definition
   sinl :: "'a \<rightarrow> ('a ++ 'b)" where
-  "sinl = (\<Lambda> a. Abs_Ssum (strict\<cdot>a\<cdot>TT, a, \<bottom>))"
+  "sinl = (\<Lambda> a. Abs_ssum (strict\<cdot>a\<cdot>TT, a, \<bottom>))"
 
 definition
   sinr :: "'b \<rightarrow> ('a ++ 'b)" where
-  "sinr = (\<Lambda> b. Abs_Ssum (strict\<cdot>b\<cdot>FF, \<bottom>, b))"
+  "sinr = (\<Lambda> b. Abs_ssum (strict\<cdot>b\<cdot>FF, \<bottom>, b))"
 
-lemma sinl_Ssum: "(strict\<cdot>a\<cdot>TT, a, \<bottom>) \<in> Ssum"
-by (simp add: Ssum_def strict_conv_if)
+lemma sinl_ssum: "(strict\<cdot>a\<cdot>TT, a, \<bottom>) \<in> ssum"
+by (simp add: ssum_def strict_conv_if)
 
-lemma sinr_Ssum: "(strict\<cdot>b\<cdot>FF, \<bottom>, b) \<in> Ssum"
-by (simp add: Ssum_def strict_conv_if)
+lemma sinr_ssum: "(strict\<cdot>b\<cdot>FF, \<bottom>, b) \<in> ssum"
+by (simp add: ssum_def strict_conv_if)
 
-lemma Rep_Ssum_sinl: "Rep_Ssum (sinl\<cdot>a) = (strict\<cdot>a\<cdot>TT, a, \<bottom>)"
-by (simp add: sinl_def cont_Abs_Ssum Abs_Ssum_inverse sinl_Ssum)
+lemma Rep_ssum_sinl: "Rep_ssum (sinl\<cdot>a) = (strict\<cdot>a\<cdot>TT, a, \<bottom>)"
+by (simp add: sinl_def cont_Abs_ssum Abs_ssum_inverse sinl_ssum)
 
-lemma Rep_Ssum_sinr: "Rep_Ssum (sinr\<cdot>b) = (strict\<cdot>b\<cdot>FF, \<bottom>, b)"
-by (simp add: sinr_def cont_Abs_Ssum Abs_Ssum_inverse sinr_Ssum)
+lemma Rep_ssum_sinr: "Rep_ssum (sinr\<cdot>b) = (strict\<cdot>b\<cdot>FF, \<bottom>, b)"
+by (simp add: sinr_def cont_Abs_ssum Abs_ssum_inverse sinr_ssum)
 
-lemmas Rep_Ssum_simps =
-  Rep_Ssum_inject [symmetric] below_Ssum_def
+lemmas Rep_ssum_simps =
+  Rep_ssum_inject [symmetric] below_ssum_def
   Pair_fst_snd_eq below_prod_def
-  Rep_Ssum_strict Rep_Ssum_sinl Rep_Ssum_sinr
+  Rep_ssum_strict Rep_ssum_sinl Rep_ssum_sinr
 
 subsection {* Properties of \emph{sinl} and \emph{sinr} *}
 
 text {* Ordering *}
 
 lemma sinl_below [simp]: "(sinl\<cdot>x \<sqsubseteq> sinl\<cdot>y) = (x \<sqsubseteq> y)"
-by (simp add: Rep_Ssum_simps strict_conv_if)
+by (simp add: Rep_ssum_simps strict_conv_if)
 
 lemma sinr_below [simp]: "(sinr\<cdot>x \<sqsubseteq> sinr\<cdot>y) = (x \<sqsubseteq> y)"
-by (simp add: Rep_Ssum_simps strict_conv_if)
+by (simp add: Rep_ssum_simps strict_conv_if)
 
 lemma sinl_below_sinr [simp]: "(sinl\<cdot>x \<sqsubseteq> sinr\<cdot>y) = (x = \<bottom>)"
-by (simp add: Rep_Ssum_simps strict_conv_if)
+by (simp add: Rep_ssum_simps strict_conv_if)
 
 lemma sinr_below_sinl [simp]: "(sinr\<cdot>x \<sqsubseteq> sinl\<cdot>y) = (x = \<bottom>)"
-by (simp add: Rep_Ssum_simps strict_conv_if)
+by (simp add: Rep_ssum_simps strict_conv_if)
 
 text {* Equality *}
 
@@ -93,10 +93,10 @@
 text {* Strictness *}
 
 lemma sinl_strict [simp]: "sinl\<cdot>\<bottom> = \<bottom>"
-by (simp add: Rep_Ssum_simps)
+by (simp add: Rep_ssum_simps)
 
 lemma sinr_strict [simp]: "sinr\<cdot>\<bottom> = \<bottom>"
-by (simp add: Rep_Ssum_simps)
+by (simp add: Rep_ssum_simps)
 
 lemma sinl_defined_iff [simp]: "(sinl\<cdot>x = \<bottom>) = (x = \<bottom>)"
 using sinl_eq [of "x" "\<bottom>"] by simp
@@ -113,10 +113,10 @@
 text {* Compactness *}
 
 lemma compact_sinl: "compact x \<Longrightarrow> compact (sinl\<cdot>x)"
-by (rule compact_Ssum, simp add: Rep_Ssum_sinl)
+by (rule compact_ssum, simp add: Rep_ssum_sinl)
 
 lemma compact_sinr: "compact x \<Longrightarrow> compact (sinr\<cdot>x)"
-by (rule compact_Ssum, simp add: Rep_Ssum_sinr)
+by (rule compact_ssum, simp add: Rep_ssum_sinr)
 
 lemma compact_sinlD: "compact (sinl\<cdot>x) \<Longrightarrow> compact x"
 unfolding compact_def
@@ -138,7 +138,7 @@
   obtains "p = \<bottom>"
   | x where "p = sinl\<cdot>x" and "x \<noteq> \<bottom>"
   | y where "p = sinr\<cdot>y" and "y \<noteq> \<bottom>"
-using Rep_Ssum [of p] by (auto simp add: Ssum_def Rep_Ssum_simps)
+using Rep_ssum [of p] by (auto simp add: ssum_def Rep_ssum_simps)
 
 lemma ssum_induct [case_names bottom sinl sinr, induct type: ssum]:
   "\<lbrakk>P \<bottom>;
@@ -160,7 +160,7 @@
 
 definition
   sscase :: "('a \<rightarrow> 'c) \<rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a ++ 'b) \<rightarrow> 'c" where
-  "sscase = (\<Lambda> f g s. (\<lambda>(t, x, y). If t then f\<cdot>x else g\<cdot>y fi) (Rep_Ssum s))"
+  "sscase = (\<Lambda> f g s. (\<lambda>(t, x, y). If t then f\<cdot>x else g\<cdot>y fi) (Rep_ssum s))"
 
 translations
   "case s of XCONST sinl\<cdot>x \<Rightarrow> t1 | XCONST sinr\<cdot>y \<Rightarrow> t2" == "CONST sscase\<cdot>(\<Lambda> x. t1)\<cdot>(\<Lambda> y. t2)\<cdot>s"
@@ -170,17 +170,17 @@
   "\<Lambda>(XCONST sinr\<cdot>y). t" == "CONST sscase\<cdot>\<bottom>\<cdot>(\<Lambda> y. t)"
 
 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) (Rep_Ssum s)"
-unfolding sscase_def by (simp add: cont_Rep_Ssum [THEN cont_compose])
+  "sscase\<cdot>f\<cdot>g\<cdot>s = (\<lambda>(t, x, y). If t then f\<cdot>x else g\<cdot>y fi) (Rep_ssum s)"
+unfolding sscase_def by (simp add: cont_Rep_ssum [THEN cont_compose])
 
 lemma sscase1 [simp]: "sscase\<cdot>f\<cdot>g\<cdot>\<bottom> = \<bottom>"
-unfolding beta_sscase by (simp add: Rep_Ssum_strict)
+unfolding beta_sscase by (simp add: Rep_ssum_strict)
 
 lemma sscase2 [simp]: "x \<noteq> \<bottom> \<Longrightarrow> sscase\<cdot>f\<cdot>g\<cdot>(sinl\<cdot>x) = f\<cdot>x"
-unfolding beta_sscase by (simp add: Rep_Ssum_sinl)
+unfolding beta_sscase by (simp add: Rep_ssum_sinl)
 
 lemma sscase3 [simp]: "y \<noteq> \<bottom> \<Longrightarrow> sscase\<cdot>f\<cdot>g\<cdot>(sinr\<cdot>y) = g\<cdot>y"
-unfolding beta_sscase by (simp add: Rep_Ssum_sinr)
+unfolding beta_sscase by (simp add: Rep_ssum_sinr)
 
 lemma sscase4 [simp]: "sscase\<cdot>sinl\<cdot>sinr\<cdot>z = z"
 by (cases z, simp_all)