src/HOLCF/Pcpodef.thy
author huffman
Thu, 07 Jul 2005 18:20:08 +0200
changeset 16738 b70bac29b11d
parent 16697 007f4caab6c1
child 16918 d0fdc7b9a33f
permissions -rw-r--r--
use theorems ch2ch_cont, cont2contlubE

(*  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 \<Rightarrow> 'b::sq_ord"
  assumes type: "type_definition Rep Abs A"
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> 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 \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
  shows "chain S \<Longrightarrow> chain (\<lambda>n. Rep (S n))"
by (rule chainI, drule chainE, unfold less)

lemma lub_Rep_in_A:
  fixes Abs :: "'a::cpo \<Rightarrow> 'b::po"
  assumes type: "type_definition Rep Abs A"
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and adm:  "adm (\<lambda>x. x \<in> A)"
  shows "chain S \<Longrightarrow> (LUB n. Rep (S n)) \<in> 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 \<Rightarrow> 'b::po"
  assumes type: "type_definition Rep Abs A"
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and adm: "adm (\<lambda>x. x \<in> A)"
  shows "chain S \<Longrightarrow> 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 \<Rightarrow> 'b::po"
  assumes type: "type_definition Rep Abs A"
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and adm: "adm (\<lambda>x. x \<in> 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 \<Rightarrow> 'b::cpo"
  assumes type: "type_definition Rep Abs A"
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and adm: "adm (\<lambda>x. x \<in> 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 \<Rightarrow> 'b::cpo"
  fixes f :: "'c::cpo \<Rightarrow> 'a::cpo"
  assumes type: "type_definition Rep Abs A"
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and adm: "adm (\<lambda>x. x \<in> A)"
    and f_in_A: "\<And>x. f x \<in> A"
    and cont_f: "cont f"
  shows "cont (\<lambda>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: cont2contlubE [OF cont_f])
 apply (rule is_lub_thelub)
  apply (erule ch2ch_cont [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 \<Rightarrow> 'b::cpo"
  assumes type: "type_definition Rep Abs A"
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and z_in_A: "z \<in> A"
    and z_least: "\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> 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 \<bottom>}.
*}

theorem typedef_pcpo_UU:
  fixes Abs :: "'a::pcpo \<Rightarrow> 'b::cpo"
  assumes type: "type_definition Rep Abs A"
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and UU_in_A: "\<bottom> \<in> 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 \<bottom>} 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 \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and UU_in_A: "\<bottom> \<in> A"
  shows "Abs \<bottom> = \<bottom>"
 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 \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and UU_in_A: "\<bottom> \<in> A"
  shows "Rep \<bottom> = \<bottom>"
 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 \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and UU_in_A: "\<bottom> \<in> A"
  shows "\<lbrakk>x \<noteq> \<bottom>; x \<in> A\<rbrakk> \<Longrightarrow> Abs x \<noteq> \<bottom>"
 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 \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
    and UU_in_A: "\<bottom> \<in> A"
  shows "x \<noteq> \<bottom> \<Longrightarrow> Rep x \<noteq> \<bottom>"
 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