# HG changeset patch # User huffman # Date 1120601071 -7200 # Node ID 007f4caab6c13f5b219cb2eb42b891f9f9d5569f # Parent 9486b344235102c7dd580a974ea9a1fe609627db renamed from TypedefPcpo.thy; added theorems Rep_defined, Abs_defined; uses pcpodef_package.ML diff -r 9486b3442351 -r 007f4caab6c1 src/HOLCF/Pcpodef.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Pcpodef.thy Wed Jul 06 00:04:31 2005 +0200 @@ -0,0 +1,223 @@ +(* Title: HOLCF/Pcpodef.thy + ID: $Id$ + Author: Brian Huffman +*) + +header {* Subtypes of pcpos *} + +theory Pcpodef +imports Adm +uses ("pcpodef_package.ML") +begin + +subsection {* Proving a subtype is a partial order *} + +text {* + A subtype of a partial order is itself a partial order, + if the ordering is defined in the standard way. +*} + +theorem typedef_po: + fixes Abs :: "'a::po \ 'b::sq_ord" + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + shows "OFCLASS('b, po_class)" + apply (intro_classes, unfold less) + apply (rule refl_less) + apply (subst type_definition.Rep_inject [OF type, symmetric]) + apply (rule antisym_less, assumption+) + apply (rule trans_less, assumption+) +done + + +subsection {* Proving a subtype is complete *} + +text {* + A subtype of a cpo is itself a cpo if the ordering is + defined in the standard way, and the defining subset + is closed with respect to limits of chains. A set is + closed if and only if membership in the set is an + admissible predicate. +*} + +lemma chain_Rep: + assumes less: "op \ \ \x y. Rep x \ Rep y" + shows "chain S \ chain (\n. Rep (S n))" +by (rule chainI, drule chainE, unfold less) + +lemma lub_Rep_in_A: + fixes Abs :: "'a::cpo \ 'b::po" + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and adm: "adm (\x. x \ A)" + shows "chain S \ (LUB n. Rep (S n)) \ A" + apply (erule admD [OF adm chain_Rep [OF less], rule_format]) + apply (rule type_definition.Rep [OF type]) +done + +theorem typedef_is_lub: + fixes Abs :: "'a::cpo \ 'b::po" + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and adm: "adm (\x. x \ A)" + shows "chain S \ range S <<| Abs (LUB n. Rep (S n))" + apply (rule is_lubI) + apply (rule ub_rangeI) + apply (subst less) + apply (subst type_definition.Abs_inverse [OF type]) + apply (erule lub_Rep_in_A [OF type less adm]) + apply (rule is_ub_thelub) + apply (erule chain_Rep [OF less]) + apply (subst less) + apply (subst type_definition.Abs_inverse [OF type]) + apply (erule lub_Rep_in_A [OF type less adm]) + apply (rule is_lub_thelub) + apply (erule chain_Rep [OF less]) + apply (rule ub_rangeI) + apply (drule ub_rangeD) + apply (unfold less) + apply assumption +done + +theorem typedef_cpo: + fixes Abs :: "'a::cpo \ 'b::po" + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and adm: "adm (\x. x \ A)" + shows "OFCLASS('b, cpo_class)" + apply (intro_classes) + apply (rule_tac x="Abs (LUB n. Rep (S n))" in exI) + apply (erule typedef_is_lub [OF type less adm]) +done + + +subsubsection {* Continuity of @{term Rep} and @{term Abs} *} + +text {* For any sub-cpo, the @{term Rep} function is continuous. *} + +theorem typedef_cont_Rep: + fixes Abs :: "'a::cpo \ 'b::cpo" + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and adm: "adm (\x. x \ A)" + shows "cont Rep" + apply (rule contI) + apply (simp only: typedef_is_lub [OF type less adm, THEN thelubI]) + apply (subst type_definition.Abs_inverse [OF type]) + apply (erule lub_Rep_in_A [OF type less adm]) + apply (rule thelubE [OF _ refl]) + apply (erule chain_Rep [OF less]) +done + +text {* + For a sub-cpo, we can make the @{term Abs} function continuous + only if we restrict its domain to the defining subset by + composing it with another continuous function. +*} + +theorem typedef_cont_Abs: + fixes Abs :: "'a::cpo \ 'b::cpo" + fixes f :: "'c::cpo \ 'a::cpo" + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and adm: "adm (\x. x \ A)" + and f_in_A: "\x. f x \ A" + and cont_f: "cont f" + shows "cont (\x. Abs (f x))" + apply (rule contI) + apply (rule is_lubI) + apply (rule ub_rangeI) + apply (simp only: less type_definition.Abs_inverse [OF type f_in_A]) + apply (rule monofun_fun_arg [OF cont2mono [OF cont_f]]) + apply (erule is_ub_thelub) + apply (simp only: less type_definition.Abs_inverse [OF type f_in_A]) + apply (simp only: contlubE [OF cont2contlub [OF cont_f]]) + apply (rule is_lub_thelub) + apply (erule ch2ch_monofun [OF cont2mono [OF cont_f]]) + apply (rule ub_rangeI) + apply (drule_tac i=i in ub_rangeD) + apply (simp only: less type_definition.Abs_inverse [OF type f_in_A]) +done + +subsection {* Proving a subtype is pointed *} + +text {* + A subtype of a cpo has a least element if and only if + the defining subset has a least element. +*} + +theorem typedef_pcpo: + fixes Abs :: "'a::cpo \ 'b::cpo" + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and z_in_A: "z \ A" + and z_least: "\x. x \ A \ z \ x" + shows "OFCLASS('b, pcpo_class)" + apply (intro_classes) + apply (rule_tac x="Abs z" in exI, rule allI) + apply (unfold less) + apply (subst type_definition.Abs_inverse [OF type z_in_A]) + apply (rule z_least [OF type_definition.Rep [OF type]]) +done + +text {* + As a special case, a subtype of a pcpo has a least element + if the defining subset contains @{term \}. +*} + +theorem typedef_pcpo_UU: + fixes Abs :: "'a::pcpo \ 'b::cpo" + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and UU_in_A: "\ \ A" + shows "OFCLASS('b, pcpo_class)" +by (rule typedef_pcpo [OF type less UU_in_A], rule minimal) + +subsubsection {* Strictness of @{term Rep} and @{term Abs} *} + +text {* + For a sub-pcpo where @{term \} is a member of the defining + subset, @{term Rep} and @{term Abs} are both strict. +*} + +theorem typedef_Abs_strict: + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and UU_in_A: "\ \ A" + shows "Abs \ = \" + apply (rule UU_I, unfold less) + apply (simp add: type_definition.Abs_inverse [OF type UU_in_A]) +done + +theorem typedef_Rep_strict: + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and UU_in_A: "\ \ A" + shows "Rep \ = \" + apply (rule typedef_Abs_strict [OF type less UU_in_A, THEN subst]) + apply (rule type_definition.Abs_inverse [OF type UU_in_A]) +done + +theorem typedef_Abs_defined: + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and UU_in_A: "\ \ A" + shows "\x \ \; x \ A\ \ Abs x \ \" + apply (rule typedef_Abs_strict [OF type less UU_in_A, THEN subst]) + apply (simp add: type_definition.Abs_inject [OF type] UU_in_A) +done + +theorem typedef_Rep_defined: + assumes type: "type_definition Rep Abs A" + and less: "op \ \ \x y. Rep x \ Rep y" + and UU_in_A: "\ \ A" + shows "x \ \ \ Rep x \ \" + apply (rule typedef_Rep_strict [OF type less UU_in_A, THEN subst]) + apply (simp add: type_definition.Rep_inject [OF type]) +done + +subsection {* HOLCF type definition package *} + +use "pcpodef_package.ML" + +end