Up to index of Isabelle/HOLCF
theory Pcpodef(* Title: HOLCF/Pcpodef.thy
ID: $Id: Pcpodef.thy,v 1.3 2005/07/26 16:24:29 huffman Exp $
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 \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
shows "OFCLASS('b, po_class)"
apply (intro_classes, unfold less)
apply (rule refl_less)
apply (rule type_definition.Rep_inject [OF type, THEN iffD1])
apply (erule (1) antisym_less)
apply (erule (1) trans_less)
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 monofun_Rep:
assumes less: "op \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
shows "monofun Rep"
by (rule monofunI, unfold less)
lemmas ch2ch_Rep = ch2ch_monofun [OF monofun_Rep]
lemmas ub2ub_Rep = ub2ub_monofun [OF monofun_Rep]
lemma Abs_inverse_lub_Rep:
fixes Abs :: "'a::cpo => 'b::po"
assumes type: "type_definition Rep Abs A"
and less: "op \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
and adm: "adm (λx. x ∈ A)"
shows "chain S ==> Rep (Abs (\<Squnion>i. Rep (S i))) = (\<Squnion>i. Rep (S i))"
apply (rule type_definition.Abs_inverse [OF type])
apply (erule admD [OF adm ch2ch_Rep [OF less], rule_format])
apply (rule type_definition.Rep [OF type])
done
theorem typedef_lub:
fixes Abs :: "'a::cpo => 'b::po"
assumes type: "type_definition Rep Abs A"
and less: "op \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
and adm: "adm (λx. x ∈ A)"
shows "chain S ==> range S <<| Abs (\<Squnion>i. Rep (S i))"
apply (frule ch2ch_Rep [OF less])
apply (rule is_lubI)
apply (rule ub_rangeI)
apply (simp only: less Abs_inverse_lub_Rep [OF type less adm])
apply (erule is_ub_thelub)
apply (simp only: less Abs_inverse_lub_Rep [OF type less adm])
apply (erule is_lub_thelub)
apply (erule ub2ub_Rep [OF less])
done
lemmas typedef_thelub = typedef_lub [THEN thelubI, standard]
theorem typedef_cpo:
fixes Abs :: "'a::cpo => 'b::po"
assumes type: "type_definition Rep Abs A"
and less: "op \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
and adm: "adm (λx. x ∈ A)"
shows "OFCLASS('b, cpo_class)"
proof
fix S::"nat => 'b" assume "chain S"
hence "range S <<| Abs (\<Squnion>i. Rep (S i))"
by (rule typedef_lub [OF type less adm])
thus "∃x. range S <<| x" ..
qed
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 \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
and adm: "adm (λx. x ∈ A)"
shows "cont Rep"
apply (rule contI)
apply (simp only: typedef_thelub [OF type less adm])
apply (simp only: Abs_inverse_lub_Rep [OF type less adm])
apply (rule thelubE [OF _ refl])
apply (erule ch2ch_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_is_lubI:
assumes less: "op \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
shows "range (λi. Rep (S i)) <<| Rep x ==> range S <<| x"
apply (rule is_lubI)
apply (rule ub_rangeI)
apply (subst less)
apply (erule is_ub_lub)
apply (subst less)
apply (erule is_lub_lub)
apply (erule ub2ub_Rep [OF less])
done
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 \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
and adm: "adm (λx. x ∈ A)" (* not used *)
and f_in_A: "!!x. f x ∈ A"
and cont_f: "cont f"
shows "cont (λx. Abs (f x))"
apply (rule contI)
apply (rule typedef_is_lubI [OF less])
apply (simp only: type_definition.Abs_inverse [OF type f_in_A])
apply (erule cont_f [THEN contE])
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_generic:
fixes Abs :: "'a::cpo => 'b::cpo"
assumes type: "type_definition Rep Abs A"
and less: "op \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
and z_in_A: "z ∈ A"
and z_least: "!!x. x ∈ A ==> 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 ⊥}.
*}
theorem typedef_pcpo:
fixes Abs :: "'a::pcpo => 'b::cpo"
assumes type: "type_definition Rep Abs A"
and less: "op \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
and UU_in_A: "⊥ ∈ A"
shows "OFCLASS('b, pcpo_class)"
by (rule typedef_pcpo_generic [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 \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> 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 \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> 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 \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> 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 \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> 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
theorem typedef_po:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y |] ==> OFCLASS('b, po_class)
lemma monofun_Rep:
op << == %x y. Rep x << Rep y ==> monofun Rep
lemmas ch2ch_Rep:
[| op << == %x y. f x << f y; chain Y |] ==> chain (%i. f (Y i))
lemmas ch2ch_Rep:
[| op << == %x y. f x << f y; chain Y |] ==> chain (%i. f (Y i))
lemmas ub2ub_Rep:
[| op << == %x y. f x << f y; range Y <| u |] ==> range (%i. f (Y i)) <| f u
lemmas ub2ub_Rep:
[| op << == %x y. f x << f y; range Y <| u |] ==> range (%i. f (Y i)) <| f u
lemma Abs_inverse_lub_Rep:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; adm (%x. x ∈ A); chain S |] ==> Rep (Abs (LUB i. Rep (S i))) = (LUB i. Rep (S i))
theorem typedef_lub:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; adm (%x. x ∈ A); chain S |] ==> range S <<| Abs (LUB i. Rep (S i))
lemmas typedef_thelub:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; adm (%x. x ∈ A); chain S |] ==> lub (range S) = Abs (LUB i. Rep (S i))
lemmas typedef_thelub:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; adm (%x. x ∈ A); chain S |] ==> lub (range S) = Abs (LUB i. Rep (S i))
theorem typedef_cpo:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; adm (%x. x ∈ A) |] ==> OFCLASS('b, cpo_class)
theorem typedef_cont_Rep:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; adm (%x. x ∈ A) |] ==> cont Rep
theorem typedef_is_lubI:
[| op << == %x y. Rep x << Rep y; range (%i. Rep (S i)) <<| Rep x |] ==> range S <<| x
theorem typedef_cont_Abs:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; adm (%x. x ∈ A); !!x. f x ∈ A; cont f |] ==> cont (%x. Abs (f x))
theorem typedef_pcpo_generic:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; z ∈ A; !!x. x ∈ A ==> z << x |] ==> OFCLASS('b, pcpo_class)
theorem typedef_pcpo:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; UU ∈ A |] ==> OFCLASS('b, pcpo_class)
theorem typedef_Abs_strict:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; UU ∈ A |] ==> Abs UU = UU
theorem typedef_Rep_strict:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; UU ∈ A |] ==> Rep UU = UU
theorem typedef_Abs_defined:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; UU ∈ A; x ≠ UU; x ∈ A |] ==> Abs x ≠ UU
theorem typedef_Rep_defined:
[| type_definition Rep Abs A; op << == %x y. Rep x << Rep y; UU ∈ A; x ≠ UU |] ==> Rep x ≠ UU