added theorems less_sprod, spair_less, spair_eq, spair_inject
authorhuffman
Wed, 08 Jun 2005 00:18:26 +0200
changeset 16317 868eddbcaf6e
parent 16316 17db5df51a35
child 16318 45b12a01382f
added theorems less_sprod, spair_less, spair_eq, spair_inject
src/HOLCF/Sprod.ML
src/HOLCF/Sprod.thy
--- a/src/HOLCF/Sprod.ML	Wed Jun 08 00:16:28 2005 +0200
+++ b/src/HOLCF/Sprod.ML	Wed Jun 08 00:18:26 2005 +0200
@@ -6,7 +6,6 @@
 val sfst_def = thm "sfst_def";
 val ssnd_def = thm "ssnd_def";
 val ssplit_def = thm "ssplit_def";
-val spair_inject = thm "spair_inject";
 val inst_sprod_pcpo2 = thm "inst_sprod_pcpo2";
 val spair_strict = thm "spair_strict";
 val spair_strict1 = thm "spair_strict1";
@@ -14,6 +13,8 @@
 val spair_strict_rev = thm "spair_strict_rev";
 val spair_defined_rev = thm "spair_defined_rev";
 val spair_defined = thm "spair_defined";
+val spair_eq = thm "spair_eq";
+val spair_inject = thm "spair_inject";
 val Exh_Sprod2 = thm "Exh_Sprod2";
 val sprodE = thm "sprodE";
 val sfst_strict = thm "sfst_strict";
@@ -22,6 +23,8 @@
 val ssnd_spair = thm "ssnd_spair";
 val sfstssnd_defined = thm "sfstssnd_defined";
 val surjective_pairing_Sprod2 = thm "surjective_pairing_Sprod2";
+val less_sprod = thm "less_sprod";
+val spair_less = thm "spair_less";
 val ssplit1 = thm "ssplit1";
 val ssplit2 = thm "ssplit2";
 val ssplit3 = thm "ssplit3";
--- a/src/HOLCF/Sprod.thy	Wed Jun 08 00:16:28 2005 +0200
+++ b/src/HOLCF/Sprod.thy	Wed Jun 08 00:18:26 2005 +0200
@@ -113,20 +113,20 @@
 
 subsection {* Properties of @{term spair} *}
 
-lemma spair_strict1 [simp]: "(:\<bottom>, b:) = \<bottom>"
+lemma spair_strict1 [simp]: "(:\<bottom>, y:) = \<bottom>"
 by (simp add: spair_Abs_Sprod UU_Abs_Sprod strictify_conv_if)
 
-lemma spair_strict2 [simp]: "(:a, \<bottom>:) = \<bottom>"
+lemma spair_strict2 [simp]: "(:x, \<bottom>:) = \<bottom>"
 by (simp add: spair_Abs_Sprod UU_Abs_Sprod strictify_conv_if)
 
-lemma spair_strict: "a = \<bottom> \<or> b = \<bottom> \<Longrightarrow> (:a, b:) = \<bottom>"
+lemma spair_strict: "x = \<bottom> \<or> y = \<bottom> \<Longrightarrow> (:x, y:) = \<bottom>"
 by auto
 
 lemma spair_strict_rev: "(:x, y:) \<noteq> \<bottom> \<Longrightarrow> x \<noteq> \<bottom> \<and> y \<noteq> \<bottom>"
 by (erule contrapos_np, auto)
 
 lemma spair_defined [simp]: 
-  "\<lbrakk>a \<noteq> \<bottom>; b \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:a, b:) \<noteq> \<bottom>"
+  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<noteq> \<bottom>"
 apply (simp add: spair_Abs_Sprod UU_Abs_Sprod)
 apply (subst Abs_Sprod_inject)
 apply (simp add: Sprod_def)
@@ -134,15 +134,19 @@
 apply simp
 done
 
-lemma spair_defined_rev: "(:a, b:) = \<bottom> \<Longrightarrow> a = \<bottom> \<or> b = \<bottom>"
+lemma spair_defined_rev: "(:x, y:) = \<bottom> \<Longrightarrow> x = \<bottom> \<or> y = \<bottom>"
 by (erule contrapos_pp, simp)
 
+lemma spair_eq:
+  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> ((:x, y:) = (:a, b:)) = (x = a \<and> y = b)"
+apply (simp add: spair_Abs_Sprod)
+apply (simp add: Abs_Sprod_inject [OF _ spair_lemma] Sprod_def)
+apply (simp add: strictify_conv_if)
+done
+
 lemma spair_inject:
-  "\<lbrakk>aa \<noteq> \<bottom>; ba \<noteq> \<bottom>; (:a,b:) = (:aa,ba:)\<rbrakk> \<Longrightarrow> a = aa \<and> b = ba"
-apply (simp add: spair_Abs_Sprod)
-apply (simp add: Abs_Sprod_inject [OF spair_lemma] Sprod_def)
-apply (simp add: strictify_conv_if split: split_if_asm)
-done
+  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>; (:x, y:) = (:a, b:)\<rbrakk> \<Longrightarrow> x = a \<and> y = b"
+by (rule spair_eq [THEN iffD1])
 
 lemma inst_sprod_pcpo2: "UU = (:UU,UU:)"
 by simp
@@ -169,10 +173,25 @@
 
 lemma sfstssnd_defined: "p \<noteq> \<bottom> \<Longrightarrow> sfst\<cdot>p \<noteq> \<bottom> \<and> ssnd\<cdot>p \<noteq> \<bottom>"
 by (rule_tac p=p in sprodE, simp_all)
- 
+
 lemma surjective_pairing_Sprod2: "(:sfst\<cdot>p, ssnd\<cdot>p:) = p"
 by (rule_tac p=p in sprodE, simp_all)
 
+lemma less_sprod: "p1 \<sqsubseteq> p2 = (sfst\<cdot>p1 \<sqsubseteq> sfst\<cdot>p2 \<and> ssnd\<cdot>p1 \<sqsubseteq> ssnd\<cdot>p2)"
+apply (simp add: less_sprod_def sfst_def ssnd_def cont_Rep_Sprod)
+apply (rule less_cprod)
+done
+
+lemma spair_less:
+  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<sqsubseteq> (:a, b:) = (x \<sqsubseteq> a \<and> y \<sqsubseteq> b)"
+apply (case_tac "a = \<bottom>")
+apply (simp add: eq_UU_iff [symmetric])
+apply (case_tac "b = \<bottom>")
+apply (simp add: eq_UU_iff [symmetric])
+apply (simp add: less_sprod)
+done
+
+
 subsection {* Properties of @{term ssplit} *}
 
 lemma ssplit1 [simp]: "ssplit\<cdot>f\<cdot>\<bottom> = \<bottom>"