completely remove constants cpair, cfst, csnd
authorhuffman
Mon, 22 Mar 2010 23:34:23 -0700
changeset 35926 e6aec5d665f0
parent 35925 3da8db225c93
child 35927 343d5b0df29a
completely remove constants cpair, cfst, csnd
src/HOLCF/Cprod.thy
src/HOLCF/Fixrec.thy
src/HOLCF/HOLCF.thy
--- a/src/HOLCF/Cprod.thy	Mon Mar 22 23:33:23 2010 -0700
+++ b/src/HOLCF/Cprod.thy	Mon Mar 22 23:34:23 2010 -0700
@@ -22,150 +22,22 @@
 lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a"
 by (simp add: unit_when_def)
 
-subsection {* Continuous versions of constants *}
-
-definition
-  cpair :: "'a \<rightarrow> 'b \<rightarrow> ('a * 'b)"  -- {* continuous pairing *}  where
-  "cpair = (\<Lambda> x y. (x, y))"
-
-definition
-  cfst :: "('a * 'b) \<rightarrow> 'a" where
-  "cfst = (\<Lambda> p. fst p)"
-
-definition
-  csnd :: "('a * 'b) \<rightarrow> 'b" where
-  "csnd = (\<Lambda> p. snd p)"
+subsection {* Continuous version of split function *}
 
 definition
   csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a * 'b) \<rightarrow> 'c" where
   "csplit = (\<Lambda> f p. f\<cdot>(fst p)\<cdot>(snd p))"
 
-syntax
-  "_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b"  ("(1<_,/ _>)")
-
-syntax (xsymbols)
-  "_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b"  ("(1\<langle>_,/ _\<rangle>)")
-
 translations
-  "\<langle>x, y, z\<rangle>" == "\<langle>x, \<langle>y, z\<rangle>\<rangle>"
-  "\<langle>x, y\<rangle>"    == "CONST cpair\<cdot>x\<cdot>y"
-
-translations
-  "\<Lambda>(CONST cpair\<cdot>x\<cdot>y). t" == "CONST csplit\<cdot>(\<Lambda> x y. t)"
-  "\<Lambda>(CONST Pair x y). t" => "CONST csplit\<cdot>(\<Lambda> x y. t)"
+  "\<Lambda>(CONST Pair x y). t" == "CONST csplit\<cdot>(\<Lambda> 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> \<Longrightarrow> a = aa \<and> b = ba"
-by (simp add: cpair_eq_pair)
-
-lemma cpair_eq [iff]: "(<a, b> = <a', b'>) = (a = a' \<and> b = b')"
-by (simp add: cpair_eq_pair)
-
-lemma cpair_below [iff]: "(<a, b> \<sqsubseteq> <a', b'>) = (a \<sqsubseteq> a' \<and> b \<sqsubseteq> b')"
-by (simp add: cpair_eq_pair)
-
-lemma cpair_defined_iff [iff]: "(<x, y> = \<bottom>) = (x = \<bottom> \<and> y = \<bottom>)"
-by (simp add: cpair_eq_pair)
-
-lemma cpair_strict [simp]: "\<langle>\<bottom>, \<bottom>\<rangle> = \<bottom>"
-by simp
-
-lemma inst_cprod_pcpo2: "\<bottom> = <\<bottom>, \<bottom>>"
-by (rule cpair_strict [symmetric])
-
-lemma defined_cpair_rev: 
- "<a,b> = \<bottom> \<Longrightarrow> a = \<bottom> \<and> b = \<bottom>"
-by simp
-
-lemma Exh_Cprod2: "\<exists>a b. z = <a, b>"
-by (simp add: cpair_eq_pair)
-
-lemma cprodE: "\<lbrakk>\<And>x y. p = <x, y> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-by (cut_tac Exh_Cprod2, auto)
-
-lemma cfst_cpair [simp]: "cfst\<cdot><x, y> = x"
-by (simp add: cpair_eq_pair cfst_def)
-
-lemma csnd_cpair [simp]: "csnd\<cdot><x, y> = y"
-by (simp add: cpair_eq_pair csnd_def)
-
-lemma cfst_strict [simp]: "cfst\<cdot>\<bottom> = \<bottom>"
-by (simp add: cfst_def)
-
-lemma csnd_strict [simp]: "csnd\<cdot>\<bottom> = \<bottom>"
-by (simp add: csnd_def)
-
-lemma cpair_cfst_csnd: "\<langle>cfst\<cdot>p, csnd\<cdot>p\<rangle> = p"
-by (cases p rule: cprodE, simp)
-
-lemmas surjective_pairing_Cprod2 = cpair_cfst_csnd
-
-lemma below_cprod: "x \<sqsubseteq> y = (cfst\<cdot>x \<sqsubseteq> cfst\<cdot>y \<and> csnd\<cdot>x \<sqsubseteq> csnd\<cdot>y)"
-by (simp add: below_prod_def cfst_def csnd_def)
-
-lemma eq_cprod: "(x = y) = (cfst\<cdot>x = cfst\<cdot>y \<and> csnd\<cdot>x = csnd\<cdot>y)"
-by (auto simp add: po_eq_conv below_cprod)
-
-lemma cfst_below_iff: "cfst\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <y, csnd\<cdot>x>"
-by (simp add: below_cprod)
-
-lemma csnd_below_iff: "csnd\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <cfst\<cdot>x, y>"
-by (simp add: below_cprod)
-
-lemma compact_cfst: "compact x \<Longrightarrow> compact (cfst\<cdot>x)"
-by (rule compactI, simp add: cfst_below_iff)
-
-lemma compact_csnd: "compact x \<Longrightarrow> compact (csnd\<cdot>x)"
-by (rule compactI, simp add: csnd_below_iff)
-
-lemma compact_cpair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact <x, y>"
-by (simp add: cpair_eq_pair)
-
-lemma compact_cpair_iff [simp]: "compact <x, y> = (compact x \<and> compact y)"
-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)>"
-apply (simp add: cpair_eq_pair cfst_def csnd_def)
-apply (erule lub_cprod)
-done
-
-lemma thelub_cprod2:
-  "chain S \<Longrightarrow> (\<Squnion>i. S i) = <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>"
-by (rule lub_cprod2 [THEN thelubI])
-
 lemma csplit1 [simp]: "csplit\<cdot>f\<cdot>\<bottom> = f\<cdot>\<bottom>\<cdot>\<bottom>"
 by (simp add: csplit_def)
 
-lemma csplit2 [simp]: "csplit\<cdot>f\<cdot><x,y> = f\<cdot>x\<cdot>y"
-by (simp add: csplit_def cpair_def)
-
-lemma csplit3 [simp]: "csplit\<cdot>cpair\<cdot>z = z"
-by (simp add: csplit_def cpair_def)
-
 lemma csplit_Pair [simp]: "csplit\<cdot>f\<cdot>(x, y) = f\<cdot>x\<cdot>y"
 by (simp add: csplit_def)
 
-lemmas Cprod_rews = cfst_cpair csnd_cpair csplit2
-
-subsection {* Product type is a bifinite domain *}
-
-lemma approx_cpair [simp]:
-  "approx i\<cdot>\<langle>x, y\<rangle> = \<langle>approx i\<cdot>x, approx i\<cdot>y\<rangle>"
-by (simp add: cpair_eq_pair)
-
-lemma cfst_approx: "cfst\<cdot>(approx i\<cdot>p) = approx i\<cdot>(cfst\<cdot>p)"
-by (cases p rule: cprodE, simp)
-
-lemma csnd_approx: "csnd\<cdot>(approx i\<cdot>p) = approx i\<cdot>(csnd\<cdot>p)"
-by (cases p rule: cprodE, simp)
-
 end
--- a/src/HOLCF/Fixrec.thy	Mon Mar 22 23:33:23 2010 -0700
+++ b/src/HOLCF/Fixrec.thy	Mon Mar 22 23:34:23 2010 -0700
@@ -276,13 +276,13 @@
 
 definition
   cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
-  "cpair_pat p1 p2 = (\<Lambda>\<langle>x, y\<rangle>.
-    bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. return\<cdot>\<langle>a, b\<rangle>)))"
+  "cpair_pat p1 p2 = (\<Lambda>(x, y).
+    bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. return\<cdot>(a, b))))"
 
 definition
   spair_pat ::
   "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a::pcpo \<otimes> 'b::pcpo, 'c \<times> 'd) pat" where
-  "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>\<langle>x, y\<rangle>)"
+  "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>(x, y))"
 
 definition
   sinl_pat :: "('a, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
@@ -310,7 +310,7 @@
 
 text {* Parse translations (patterns) *}
 translations
-  "_pat (XCONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)"
+  "_pat (XCONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
   "_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
   "_pat (XCONST sinl\<cdot>x)" => "CONST sinl_pat (_pat x)"
   "_pat (XCONST sinr\<cdot>x)" => "CONST sinr_pat (_pat x)"
@@ -321,12 +321,12 @@
 
 text {* CONST version is also needed for constructors with special syntax *}
 translations
-  "_pat (CONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)"
+  "_pat (CONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
   "_pat (CONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
 
 text {* Parse translations (variables) *}
 translations
-  "_variable (XCONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+  "_variable (XCONST Pair x y) r" => "_variable (_args x y) r"
   "_variable (XCONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
   "_variable (XCONST sinl\<cdot>x) r" => "_variable x r"
   "_variable (XCONST sinr\<cdot>x) r" => "_variable x r"
@@ -336,12 +336,12 @@
   "_variable (XCONST ONE) r" => "_variable _noargs r"
 
 translations
-  "_variable (CONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+  "_variable (CONST Pair x y) r" => "_variable (_args x y) r"
   "_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
 
 text {* Print translations *}
 translations
-  "CONST cpair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)"
+  "CONST Pair (_match p1 v1) (_match p2 v2)"
       <= "_match (CONST cpair_pat p1 p2) (_args v1 v2)"
   "CONST spair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)"
       <= "_match (CONST spair_pat p1 p2) (_args v1 v2)"
@@ -353,20 +353,20 @@
   "CONST ONE" <= "_match (CONST ONE_pat) _noargs"
 
 lemma cpair_pat1:
-  "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = \<bottom>"
+  "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = \<bottom>"
 apply (simp add: branch_def cpair_pat_def)
 apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
 done
 
 lemma cpair_pat2:
-  "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = fail"
+  "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = fail"
 apply (simp add: branch_def cpair_pat_def)
 apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
 done
 
 lemma cpair_pat3:
   "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow>
-   branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = branch q\<cdot>s\<cdot>y"
+   branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = branch q\<cdot>s\<cdot>y"
 apply (simp add: branch_def cpair_pat_def)
 apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
 apply (rule_tac p="q\<cdot>y" in maybeE, simp_all)
@@ -379,7 +379,7 @@
   "branch (spair_pat p1 p2)\<cdot>r\<cdot>\<bottom> = \<bottom>"
   "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk>
      \<Longrightarrow> branch (spair_pat p1 p2)\<cdot>r\<cdot>(:x, y:) =
-         branch (cpair_pat p1 p2)\<cdot>r\<cdot>\<langle>x, y\<rangle>"
+         branch (cpair_pat p1 p2)\<cdot>r\<cdot>(x, y)"
 by (simp_all add: branch_def spair_pat_def)
 
 lemma sinl_pat [simp]:
@@ -425,7 +425,7 @@
 
 definition
   as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where
-  "as_pat p = (\<Lambda> x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> a. return\<cdot>\<langle>x, a\<rangle>))"
+  "as_pat p = (\<Lambda> x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> a. return\<cdot>(x, a)))"
 
 definition
   lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where
@@ -516,7 +516,6 @@
 by (simp_all add: match_UU_def)
 
 lemma match_cpair_simps [simp]:
-  "match_cpair\<cdot>\<langle>x, y\<rangle>\<cdot>k = k\<cdot>x\<cdot>y"
   "match_cpair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y"
 by (simp_all add: match_cpair_def)
 
@@ -602,7 +601,6 @@
       (@{const_name sinl}, @{const_name match_sinl}),
       (@{const_name sinr}, @{const_name match_sinr}),
       (@{const_name spair}, @{const_name match_spair}),
-      (@{const_name cpair}, @{const_name match_cpair}),
       (@{const_name Pair}, @{const_name match_cpair}),
       (@{const_name ONE}, @{const_name match_ONE}),
       (@{const_name TT}, @{const_name match_TT}),
--- a/src/HOLCF/HOLCF.thy	Mon Mar 22 23:33:23 2010 -0700
+++ b/src/HOLCF/HOLCF.thy	Mon Mar 22 23:34:23 2010 -0700
@@ -50,10 +50,6 @@
 lemmas not_Iup_less = not_Iup_below
 lemmas Iup_less = Iup_below
 lemmas up_less = up_below
-lemmas cpair_less = cpair_below
-lemmas less_cprod = below_cprod
-lemmas cfst_less_iff = cfst_below_iff
-lemmas csnd_less_iff = csnd_below_iff
 lemmas Def_inject_less_eq = Def_below_Def
 lemmas Def_less_is_eq = Def_below_iff
 lemmas spair_less_iff = spair_below_iff