src/HOLCF/Up.thy
changeset 16753 fb6801c926d2
parent 16553 aa36d41e4263
child 16933 91ded127f5f7
--- a/src/HOLCF/Up.thy	Fri Jul 08 02:39:53 2005 +0200
+++ b/src/HOLCF/Up.thy	Fri Jul 08 02:41:19 2005 +0200
@@ -15,95 +15,46 @@
 
 subsection {* Definition of new type for lifting *}
 
-typedef (Up) 'a u = "UNIV :: 'a option set" ..
+datatype 'a u = Ibottom | Iup 'a
 
 consts
-  Iup         :: "'a \<Rightarrow> 'a u"
-  Ifup        :: "('a \<rightarrow> 'b::pcpo) \<Rightarrow> 'a u \<Rightarrow> 'b"
-
-defs
-  Iup_def:     "Iup x \<equiv> Abs_Up (Some x)"
-  Ifup_def:    "Ifup f x \<equiv> case Rep_Up x of None \<Rightarrow> \<bottom> | Some z \<Rightarrow> f\<cdot>z"
-
-lemma Abs_Up_inverse2: "Rep_Up (Abs_Up y) = y"
-by (simp add: Up_def Abs_Up_inverse)
-
-lemma Exh_Up: "z = Abs_Up None \<or> (\<exists>x. z = Iup x)"
-apply (unfold Iup_def)
-apply (rule Rep_Up_inverse [THEN subst])
-apply (case_tac "Rep_Up z")
-apply auto
-done
+  Ifup :: "('a \<rightarrow> 'b::pcpo) \<Rightarrow> 'a u \<Rightarrow> 'b"
 
-lemma inj_Abs_Up: "inj Abs_Up" (* worthless *)
-apply (rule inj_on_inverseI)
-apply (rule Abs_Up_inverse2)
-done
-
-lemma inj_Rep_Up: "inj Rep_Up" (* worthless *)
-apply (rule inj_on_inverseI)
-apply (rule Rep_Up_inverse)
-done
-
-lemma Iup_eq [simp]: "(Iup x = Iup y) = (x = y)"
-by (simp add: Iup_def Abs_Up_inject Up_def)
-
-lemma Iup_defined [simp]: "Iup x \<noteq> Abs_Up None"
-by (simp add: Iup_def Abs_Up_inject Up_def)
-
-lemma upE: "\<lbrakk>p = Abs_Up None \<Longrightarrow> Q; \<And>x. p = Iup x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-by (rule Exh_Up [THEN disjE], auto)
-
-lemma Ifup1 [simp]: "Ifup f (Abs_Up None) = \<bottom>"
-by (simp add: Ifup_def Abs_Up_inverse2)
-
-lemma Ifup2 [simp]: "Ifup f (Iup x) = f\<cdot>x"
-by (simp add: Ifup_def Iup_def Abs_Up_inverse2)
+primrec
+  "Ifup f Ibottom = \<bottom>"
+  "Ifup f (Iup x) = f\<cdot>x"
 
 subsection {* Ordering on type @{typ "'a u"} *}
 
 instance u :: (sq_ord) sq_ord ..
 
 defs (overloaded)
-  less_up_def: "(op \<sqsubseteq>) \<equiv> (\<lambda>x1 x2. case Rep_Up x1 of
-               None \<Rightarrow> True
-             | Some y1 \<Rightarrow> (case Rep_Up x2 of None \<Rightarrow> False
-                                           | Some y2 \<Rightarrow> y1 \<sqsubseteq> y2))"
+  less_up_def:
+    "(op \<sqsubseteq>) \<equiv> (\<lambda>x y. case x of Ibottom \<Rightarrow> True | Iup a \<Rightarrow>
+      (case y of Ibottom \<Rightarrow> False | Iup b \<Rightarrow> a \<sqsubseteq> b))"
 
-lemma minimal_up [iff]: "Abs_Up None \<sqsubseteq> z"
-by (simp add: less_up_def Abs_Up_inverse2)
+lemma minimal_up [iff]: "Ibottom \<sqsubseteq> z"
+by (simp add: less_up_def)
 
-lemma not_Iup_less [iff]: "\<not> Iup x \<sqsubseteq> Abs_Up None"
-by (simp add: Iup_def less_up_def Abs_Up_inverse2)
+lemma not_Iup_less [iff]: "\<not> Iup x \<sqsubseteq> Ibottom"
+by (simp add: less_up_def)
 
 lemma Iup_less [iff]: "(Iup x \<sqsubseteq> Iup y) = (x \<sqsubseteq> y)"
-by (simp add: Iup_def less_up_def Abs_Up_inverse2)
+by (simp add: less_up_def)
 
 subsection {* Type @{typ "'a u"} is a partial order *}
 
-lemma refl_less_up: "(p::'a u) \<sqsubseteq> p"
-by (rule_tac p = "p" in upE, auto)
+lemma refl_less_up: "(x::'a u) \<sqsubseteq> x"
+by (simp add: less_up_def split: u.split)
 
-lemma antisym_less_up: "\<lbrakk>(p1::'a u) \<sqsubseteq> p2; p2 \<sqsubseteq> p1\<rbrakk> \<Longrightarrow> p1 = p2"
-apply (rule_tac p = "p1" in upE)
-apply (rule_tac p = "p2" in upE)
-apply simp
-apply simp
-apply (rule_tac p = "p2" in upE)
-apply simp
-apply simp
-apply (drule antisym_less, assumption)
-apply simp
+lemma antisym_less_up: "\<lbrakk>(x::'a u) \<sqsubseteq> y; y \<sqsubseteq> x\<rbrakk> \<Longrightarrow> x = y"
+apply (simp add: less_up_def split: u.split_asm)
+apply (erule (1) antisym_less)
 done
 
-lemma trans_less_up: "\<lbrakk>(p1::'a u) \<sqsubseteq> p2; p2 \<sqsubseteq> p3\<rbrakk> \<Longrightarrow> p1 \<sqsubseteq> p3"
-apply (rule_tac p = "p1" in upE)
-apply simp
-apply (rule_tac p = "p2" in upE)
-apply simp
-apply (rule_tac p = "p3" in upE)
-apply simp
-apply (auto elim: trans_less)
+lemma trans_less_up: "\<lbrakk>(x::'a u) \<sqsubseteq> y; y \<sqsubseteq> z\<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
+apply (simp add: less_up_def split: u.split_asm)
+apply (erule (1) trans_less)
 done
 
 instance u :: (cpo) po
@@ -118,7 +69,7 @@
 apply (rule ub_rangeI)
 apply (subst Iup_less)
 apply (erule is_ub_lub)
-apply (rule_tac p="u" in upE)
+apply (case_tac u)
 apply (drule ub_rangeD)
 apply simp
 apply simp
@@ -130,25 +81,25 @@
 
 text {* Now some lemmas about chains of @{typ "'a u"} elements *}
 
-lemma up_lemma1: "z \<noteq> Abs_Up None \<Longrightarrow> Iup (THE a. Iup a = z) = z"
-by (rule_tac p="z" in upE, simp_all)
+lemma up_lemma1: "z \<noteq> Ibottom \<Longrightarrow> Iup (THE a. Iup a = z) = z"
+by (case_tac z, simp_all)
 
 lemma up_lemma2:
-  "\<lbrakk>chain Y; Y j \<noteq> Abs_Up None\<rbrakk> \<Longrightarrow> Y (i + j) \<noteq> Abs_Up None"
+  "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk> \<Longrightarrow> Y (i + j) \<noteq> Ibottom"
 apply (erule contrapos_nn)
 apply (drule_tac x="j" and y="i + j" in chain_mono3)
 apply (rule le_add2)
-apply (rule_tac p="Y j" in upE)
+apply (case_tac "Y j")
 apply assumption
 apply simp
 done
 
 lemma up_lemma3:
-  "\<lbrakk>chain Y; Y j \<noteq> Abs_Up None\<rbrakk> \<Longrightarrow> Iup (THE a. Iup a = Y (i + j)) = Y (i + j)"
+  "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk> \<Longrightarrow> Iup (THE a. Iup a = Y (i + j)) = Y (i + j)"
 by (rule up_lemma1 [OF up_lemma2])
 
 lemma up_lemma4:
-  "\<lbrakk>chain Y; Y j \<noteq> Abs_Up None\<rbrakk> \<Longrightarrow> chain (\<lambda>i. THE a. Iup a = Y (i + j))"
+  "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk> \<Longrightarrow> chain (\<lambda>i. THE a. Iup a = Y (i + j))"
 apply (rule chainI)
 apply (rule Iup_less [THEN iffD1])
 apply (subst up_lemma3, assumption+)+
@@ -156,25 +107,25 @@
 done
 
 lemma up_lemma5:
-  "\<lbrakk>chain Y; Y j \<noteq> Abs_Up None\<rbrakk> \<Longrightarrow>
+  "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk> \<Longrightarrow>
     (\<lambda>i. Y (i + j)) = (\<lambda>i. Iup (THE a. Iup a = Y (i + j)))"
 by (rule ext, rule up_lemma3 [symmetric])
 
 lemma up_lemma6:
-  "\<lbrakk>chain Y; Y j \<noteq> Abs_Up None\<rbrakk>  
+  "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk>  
       \<Longrightarrow> range Y <<| Iup (\<Squnion>i. THE a. Iup a = Y(i + j))"
 apply (rule_tac j1="j" in is_lub_range_shift [THEN iffD1])
 apply assumption
 apply (subst up_lemma5, assumption+)
 apply (rule is_lub_Iup)
 apply (rule thelubE [OF _ refl])
-apply (rule up_lemma4, assumption+)
+apply (erule (1) up_lemma4)
 done
 
 lemma up_chain_cases:
   "chain Y \<Longrightarrow>
    (\<exists>A. chain A \<and> lub (range Y) = Iup (lub (range A)) \<and>
-   (\<exists>j. \<forall>i. Y (i + j) = Iup (A i))) \<or> (Y = (\<lambda>i. Abs_Up None))"
+   (\<exists>j. \<forall>i. Y (i + j) = Iup (A i))) \<or> (Y = (\<lambda>i. Ibottom))"
 apply (rule disjCI)
 apply (simp add: expand_fun_eq)
 apply (erule exE, rename_tac j)
@@ -192,7 +143,7 @@
 apply (rule_tac x="Iup (lub (range A))" in exI)
 apply (erule_tac j1="j" in is_lub_range_shift [THEN iffD1])
 apply (simp add: is_lub_Iup thelubE)
-apply (rule_tac x="Abs_Up None" in exI)
+apply (rule_tac x="Ibottom" in exI)
 apply (rule lub_const)
 done
 
@@ -202,7 +153,7 @@
 subsection {* Type @{typ "'a u"} is pointed *}
 
 lemma least_up: "EX x::'a u. ALL y. x\<sqsubseteq>y"
-apply (rule_tac x = "Abs_Up None" in exI)
+apply (rule_tac x = "Ibottom" in exI)
 apply (rule minimal_up [THEN allI])
 done
 
@@ -210,21 +161,9 @@
 by intro_classes (rule least_up)
 
 text {* for compatibility with old HOLCF-Version *}
-lemma inst_up_pcpo: "\<bottom> = Abs_Up None"
+lemma inst_up_pcpo: "\<bottom> = Ibottom"
 by (rule minimal_up [THEN UU_I, symmetric])
 
-text {* some lemmas restated for class pcpo *}
-
-lemma less_up3b: "~ Iup(x) \<sqsubseteq> \<bottom>"
-apply (subst inst_up_pcpo)
-apply simp
-done
-
-lemma defined_Iup2 [iff]: "Iup(x) ~= \<bottom>"
-apply (subst inst_up_pcpo)
-apply (rule Iup_defined)
-done
-
 subsection {* Continuity of @{term Iup} and @{term Ifup} *}
 
 text {* continuity for @{term Iup} *}
@@ -238,18 +177,12 @@
 text {* continuity for @{term Ifup} *}
 
 lemma cont_Ifup1: "cont (\<lambda>f. Ifup f x)"
-apply (rule contI)
-apply (rule_tac p="x" in upE)
-apply (simp add: lub_const)
-apply (simp add: cont_cfun_fun)
-done
+by (induct x, simp_all)
 
 lemma monofun_Ifup2: "monofun (\<lambda>x. Ifup f x)"
 apply (rule monofunI)
-apply (rule_tac p="x" in upE)
-apply simp
-apply (rule_tac p="y" in upE)
-apply simp
+apply (case_tac x, simp)
+apply (case_tac y, simp)
 apply (simp add: monofun_cfun_arg)
 done
 
@@ -272,21 +205,21 @@
   "fup \<equiv> \<Lambda> f p. Ifup f p"
 
 translations
-"case l of up\<cdot>x => t1" == "fup\<cdot>(LAM x. t1)\<cdot>l"
+"case l of up\<cdot>x \<Rightarrow> t" == "fup\<cdot>(LAM x. t)\<cdot>l"
 
 text {* continuous versions of lemmas for @{typ "('a)u"} *}
 
-lemma Exh_Up1: "z = \<bottom> \<or> (\<exists>x. z = up\<cdot>x)"
-apply (rule_tac p="z" in upE)
+lemma Exh_Up: "z = \<bottom> \<or> (\<exists>x. z = up\<cdot>x)"
+apply (induct z)
 apply (simp add: inst_up_pcpo)
 apply (simp add: up_def cont_Iup)
 done
 
-lemma up_inject: "up\<cdot>x = up\<cdot>y \<Longrightarrow> x = y"
+lemma up_eq [simp]: "(up\<cdot>x = up\<cdot>y) = (x = y)"
 by (simp add: up_def cont_Iup)
 
-lemma up_eq [simp]: "(up\<cdot>x = up\<cdot>y) = (x = y)"
-by (rule iffI, erule up_inject, simp)
+lemma up_inject: "up\<cdot>x = up\<cdot>y \<Longrightarrow> x = y"
+by simp
 
 lemma up_defined [simp]: " up\<cdot>x \<noteq> \<bottom>"
 by (simp add: up_def cont_Iup inst_up_pcpo)
@@ -297,8 +230,8 @@
 lemma up_less [simp]: "(up\<cdot>x \<sqsubseteq> up\<cdot>y) = (x \<sqsubseteq> y)"
 by (simp add: up_def cont_Iup)
 
-lemma upE1: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x. p = up\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-apply (rule_tac p="p" in upE)
+lemma upE: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x. p = up\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
+apply (case_tac p)
 apply (simp add: inst_up_pcpo)
 apply (simp add: up_def cont_Iup)
 done
@@ -307,9 +240,9 @@
 by (simp add: fup_def cont_Ifup1 cont_Ifup2 inst_up_pcpo)
 
 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)
 
 lemma fup3 [simp]: "fup\<cdot>up\<cdot>x = x"
-by (rule_tac p=x in upE1, simp_all)
+by (rule_tac p=x in upE, simp_all)
 
 end