(* Title: HOLCF/Pcpo.thy ID: $Id: Pcpo.thy,v 1.48 2008/03/29 18:14:01 wenzelm Exp $ Author: Franz Regensburger *) header {* Classes cpo and pcpo *} theory Pcpo imports Porder begin subsection {* Complete partial orders *} text {* The class cpo of chain complete partial orders *} axclass cpo < po -- {* class axiom: *} cpo: "chain S ==> ∃x. range S <<| x" text {* in cpo's everthing equal to THE lub has lub properties for every chain *} lemma cpo_lubI: "chain (S::nat => 'a::cpo) ==> range S <<| lub (range S)" by (fast dest: cpo elim: lubI) lemma thelubE: "[|chain S; (\<Squnion>i. S i) = (l::'a::cpo)|] ==> range S <<| l" by (blast dest: cpo intro: lubI) text {* Properties of the lub *} lemma is_ub_thelub: "chain (S::nat => 'a::cpo) ==> S x \<sqsubseteq> (\<Squnion>i. S i)" by (blast dest: cpo intro: lubI [THEN is_ub_lub]) lemma is_lub_thelub: "[|chain (S::nat => 'a::cpo); range S <| x|] ==> (\<Squnion>i. S i) \<sqsubseteq> x" by (blast dest: cpo intro: lubI [THEN is_lub_lub]) lemma lub_range_mono: "[|range X ⊆ range Y; chain Y; chain (X::nat => 'a::cpo)|] ==> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" apply (erule is_lub_thelub) apply (rule ub_rangeI) apply (subgoal_tac "∃j. X i = Y j") apply clarsimp apply (erule is_ub_thelub) apply auto done lemma lub_range_shift: "chain (Y::nat => 'a::cpo) ==> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)" apply (rule antisym_less) apply (rule lub_range_mono) apply fast apply assumption apply (erule chain_shift) apply (rule is_lub_thelub) apply assumption apply (rule ub_rangeI) apply (rule_tac y="Y (i + j)" in trans_less) apply (erule chain_mono) apply (rule le_add1) apply (rule is_ub_thelub) apply (erule chain_shift) done lemma maxinch_is_thelub: "chain Y ==> max_in_chain i Y = ((\<Squnion>i. Y i) = ((Y i)::'a::cpo))" apply (rule iffI) apply (fast intro!: thelubI lub_finch1) apply (unfold max_in_chain_def) apply (safe intro!: antisym_less) apply (fast elim!: chain_mono) apply (drule sym) apply (force elim!: is_ub_thelub) done text {* the @{text "\<sqsubseteq>"} relation between two chains is preserved by their lubs *} lemma lub_mono: "[|chain (X::nat => 'a::cpo); chain Y; !!i. X i \<sqsubseteq> Y i|] ==> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" apply (erule is_lub_thelub) apply (rule ub_rangeI) apply (rule trans_less) apply (erule meta_spec) apply (erule is_ub_thelub) done text {* the = relation between two chains is preserved by their lubs *} lemma lub_equal: "[|chain (X::nat => 'a::cpo); chain Y; ∀k. X k = Y k|] ==> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" by (simp only: expand_fun_eq [symmetric]) text {* more results about mono and = of lubs of chains *} lemma lub_mono2: "[|∃j. ∀i>j. X i = Y i; chain (X::nat => 'a::cpo); chain Y|] ==> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" apply (erule exE) apply (subgoal_tac "(\<Squnion>i. X (i + Suc j)) \<sqsubseteq> (\<Squnion>i. Y (i + Suc j))") apply (thin_tac "∀i>j. X i = Y i") apply (simp only: lub_range_shift) apply simp done lemma lub_equal2: "[|∃j. ∀i>j. X i = Y i; chain (X::nat => 'a::cpo); chain Y|] ==> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" by (blast intro: antisym_less lub_mono2 sym) lemma lub_mono3: "[|chain (Y::nat => 'a::cpo); chain X; ∀i. ∃j. Y i \<sqsubseteq> X j|] ==> (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. X i)" apply (erule is_lub_thelub) apply (rule ub_rangeI) apply (erule allE) apply (erule exE) apply (erule trans_less) apply (erule is_ub_thelub) done lemma ch2ch_lub: fixes Y :: "nat => nat => 'a::cpo" assumes 1: "!!j. chain (λi. Y i j)" assumes 2: "!!i. chain (λj. Y i j)" shows "chain (λi. \<Squnion>j. Y i j)" apply (rule chainI) apply (rule lub_mono [OF 2 2]) apply (rule chainE [OF 1]) done lemma diag_lub: fixes Y :: "nat => nat => 'a::cpo" assumes 1: "!!j. chain (λi. Y i j)" assumes 2: "!!i. chain (λj. Y i j)" shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>i. Y i i)" proof (rule antisym_less) have 3: "chain (λi. Y i i)" apply (rule chainI) apply (rule trans_less) apply (rule chainE [OF 1]) apply (rule chainE [OF 2]) done have 4: "chain (λi. \<Squnion>j. Y i j)" by (rule ch2ch_lub [OF 1 2]) show "(\<Squnion>i. \<Squnion>j. Y i j) \<sqsubseteq> (\<Squnion>i. Y i i)" apply (rule is_lub_thelub [OF 4]) apply (rule ub_rangeI) apply (rule lub_mono3 [rule_format, OF 2 3]) apply (rule exI) apply (rule trans_less) apply (rule chain_mono [OF 1 le_maxI1]) apply (rule chain_mono [OF 2 le_maxI2]) done show "(\<Squnion>i. Y i i) \<sqsubseteq> (\<Squnion>i. \<Squnion>j. Y i j)" apply (rule lub_mono [OF 3 4]) apply (rule is_ub_thelub [OF 2]) done qed lemma ex_lub: fixes Y :: "nat => nat => 'a::cpo" assumes 1: "!!j. chain (λi. Y i j)" assumes 2: "!!i. chain (λj. Y i j)" shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)" by (simp add: diag_lub 1 2) subsection {* Pointed cpos *} text {* The class pcpo of pointed cpos *} axclass pcpo < cpo least: "∃x. ∀y. x \<sqsubseteq> y" definition UU :: "'a::pcpo" where "UU = (THE x. ∀y. x \<sqsubseteq> y)" notation (xsymbols) UU ("⊥") text {* derive the old rule minimal *} lemma UU_least: "∀z. ⊥ \<sqsubseteq> z" apply (unfold UU_def) apply (rule theI') apply (rule ex_ex1I) apply (rule least) apply (blast intro: antisym_less) done lemma minimal [iff]: "⊥ \<sqsubseteq> x" by (rule UU_least [THEN spec]) lemma UU_reorient: "(⊥ = x) = (x = ⊥)" by auto ML {* local val meta_UU_reorient = thm "UU_reorient" RS eq_reflection; fun reorient_proc sg _ (_ $ t $ u) = case u of Const("Pcpo.UU",_) => NONE | Const("HOL.zero", _) => NONE | Const("HOL.one", _) => NONE | Const("Numeral.number_of", _) $ _ => NONE | _ => SOME meta_UU_reorient; in val UU_reorient_simproc = Simplifier.simproc @{theory} "UU_reorient_simproc" ["UU=x"] reorient_proc end; Addsimprocs [UU_reorient_simproc]; *} text {* useful lemmas about @{term ⊥} *} lemma less_UU_iff [simp]: "(x \<sqsubseteq> ⊥) = (x = ⊥)" by (simp add: po_eq_conv) lemma eq_UU_iff: "(x = ⊥) = (x \<sqsubseteq> ⊥)" by simp lemma UU_I: "x \<sqsubseteq> ⊥ ==> x = ⊥" by (subst eq_UU_iff) lemma not_less2not_eq: "¬ (x::'a::po) \<sqsubseteq> y ==> x ≠ y" by auto lemma chain_UU_I: "[|chain Y; (\<Squnion>i. Y i) = ⊥|] ==> ∀i. Y i = ⊥" apply (rule allI) apply (rule UU_I) apply (erule subst) apply (erule is_ub_thelub) done lemma chain_UU_I_inverse: "∀i::nat. Y i = ⊥ ==> (\<Squnion>i. Y i) = ⊥" apply (rule lub_chain_maxelem) apply (erule spec) apply simp done lemma chain_UU_I_inverse2: "(\<Squnion>i. Y i) ≠ ⊥ ==> ∃i::nat. Y i ≠ ⊥" by (blast intro: chain_UU_I_inverse) lemma notUU_I: "[|x \<sqsubseteq> y; x ≠ ⊥|] ==> y ≠ ⊥" by (blast intro: UU_I) lemma chain_mono2: "[|∃j. Y j ≠ ⊥; chain Y|] ==> ∃j. ∀i>j. Y i ≠ ⊥" by (blast dest: notUU_I chain_mono_less) subsection {* Chain-finite and flat cpos *} text {* further useful classes for HOLCF domains *} axclass finite_po < finite, po axclass chfin < po chfin: "chain Y ==> ∃n. max_in_chain n Y" axclass flat < pcpo ax_flat: "x \<sqsubseteq> y ==> (x = ⊥) ∨ (x = y)" text {* finite partial orders are chain-finite and directed-complete *} instance finite_po < chfin apply intro_classes apply (drule finite_range_imp_finch) apply (rule finite) apply (simp add: finite_chain_def) done instance finite_po < cpo apply intro_classes apply (drule directed_chain) apply (drule directed_finiteD [OF _ finite subset_refl]) apply (erule bexE, rule exI) apply (erule (1) is_lub_maximal) done text {* some properties for chfin and flat *} text {* chfin types are cpo *} instance chfin < cpo apply intro_classes apply (frule chfin) apply (blast intro: lub_finch1) done text {* flat types are chfin *} instance flat < chfin apply intro_classes apply (unfold max_in_chain_def) apply (case_tac "∀i. Y i = ⊥") apply simp apply simp apply (erule exE) apply (rule_tac x="i" in exI) apply clarify apply (blast dest: chain_mono ax_flat) done text {* flat subclass of chfin; @{text adm_flat} not needed *} lemma flat_less_iff: fixes x y :: "'a::flat" shows "(x \<sqsubseteq> y) = (x = ⊥ ∨ x = y)" by (safe dest!: ax_flat) lemma flat_eq: "(a::'a::flat) ≠ ⊥ ==> a \<sqsubseteq> b = (a = b)" by (safe dest!: ax_flat) lemma chfin2finch: "chain (Y::nat => 'a::chfin) ==> finite_chain Y" by (simp add: chfin finite_chain_def) text {* Discrete cpos *} axclass discrete_cpo < sq_ord discrete_cpo [simp]: "x \<sqsubseteq> y <-> x = y" instance discrete_cpo < po by (intro_classes, simp_all) text {* In a discrete cpo, every chain is constant *} lemma discrete_chain_const: assumes S: "chain (S::nat => 'a::discrete_cpo)" shows "∃x. S = (λi. x)" proof (intro exI ext) fix i :: nat have "S 0 \<sqsubseteq> S i" using S le0 by (rule chain_mono) hence "S 0 = S i" by simp thus "S i = S 0" by (rule sym) qed instance discrete_cpo < cpo proof fix S :: "nat => 'a" assume S: "chain S" hence "∃x. S = (λi. x)" by (rule discrete_chain_const) thus "∃x. range S <<| x" by (fast intro: lub_const) qed text {* lemmata for improved admissibility introdution rule *} lemma infinite_chain_adm_lemma: "[|chain Y; ∀i. P (Y i); !!Y. [|chain Y; ∀i. P (Y i); ¬ finite_chain Y|] ==> P (\<Squnion>i. Y i)|] ==> P (\<Squnion>i. Y i)" apply (case_tac "finite_chain Y") prefer 2 apply fast apply (unfold finite_chain_def) apply safe apply (erule lub_finch1 [THEN thelubI, THEN ssubst]) apply assumption apply (erule spec) done lemma increasing_chain_adm_lemma: "[|chain Y; ∀i. P (Y i); !!Y. [|chain Y; ∀i. P (Y i); ∀i. ∃j>i. Y i ≠ Y j ∧ Y i \<sqsubseteq> Y j|] ==> P (\<Squnion>i. Y i)|] ==> P (\<Squnion>i. Y i)" apply (erule infinite_chain_adm_lemma) apply assumption apply (erule thin_rl) apply (unfold finite_chain_def) apply (unfold max_in_chain_def) apply (fast dest: le_imp_less_or_eq elim: chain_mono_less) done end
lemma cpo_lubI:
chain S ==> range S <<| Lub S
lemma thelubE:
[| chain S; (LUB i. S i) = l |] ==> range S <<| l
lemma is_ub_thelub:
chain S ==> S x << (LUB i. S i)
lemma is_lub_thelub:
[| chain S; range S <| x |] ==> (LUB i. S i) << x
lemma lub_range_mono:
[| range X ⊆ range Y; chain Y; chain X |] ==> (LUB i. X i) << (LUB i. Y i)
lemma lub_range_shift:
chain Y ==> (LUB i. Y (i + j)) = (LUB i. Y i)
lemma maxinch_is_thelub:
chain Y ==> max_in_chain i Y = ((LUB i. Y i) = Y i)
lemma lub_mono:
[| chain X; chain Y; !!i. X i << Y i |] ==> (LUB i. X i) << (LUB i. Y i)
lemma lub_equal:
[| chain X; chain Y; ∀k. X k = Y k |] ==> (LUB i. X i) = (LUB i. Y i)
lemma lub_mono2:
[| ∃j. ∀i>j. X i = Y i; chain X; chain Y |] ==> (LUB i. X i) << (LUB i. Y i)
lemma lub_equal2:
[| ∃j. ∀i>j. X i = Y i; chain X; chain Y |] ==> (LUB i. X i) = (LUB i. Y i)
lemma lub_mono3:
[| chain Y; chain X; ∀i. ∃j. Y i << X j |] ==> (LUB i. Y i) << (LUB i. X i)
lemma ch2ch_lub:
[| !!j. chain (λi. Y i j); !!i. chain (Y i) |] ==> chain (λi. LUB j. Y i j)
lemma diag_lub:
[| !!j. chain (λi. Y i j); !!i. chain (Y i) |]
==> (LUB i j. Y i j) = (LUB i. Y i i)
lemma ex_lub:
[| !!j. chain (λi. Y i j); !!i. chain (Y i) |]
==> (LUB i j. Y i j) = (LUB j i. Y i j)
lemma UU_least:
∀z. UU << z
lemma minimal:
UU << x
lemma UU_reorient:
(UU = x) = (x = UU)
lemma less_UU_iff:
x << UU = (x = UU)
lemma eq_UU_iff:
(x = UU) = x << UU
lemma UU_I:
x << UU ==> x = UU
lemma not_less2not_eq:
¬ x << y ==> x ≠ y
lemma chain_UU_I:
[| chain Y; (LUB i. Y i) = UU |] ==> ∀i. Y i = UU
lemma chain_UU_I_inverse:
∀i. Y i = UU ==> (LUB i. Y i) = UU
lemma chain_UU_I_inverse2:
(LUB i. Y i) ≠ UU ==> ∃i. Y i ≠ UU
lemma notUU_I:
[| x << y; x ≠ UU |] ==> y ≠ UU
lemma chain_mono2:
[| ∃j. Y j ≠ UU; chain Y |] ==> ∃j. ∀i>j. Y i ≠ UU
lemma flat_less_iff:
x << y = (x = UU ∨ x = y)
lemma flat_eq:
a ≠ UU ==> a << b = (a = b)
lemma chfin2finch:
chain Y ==> finite_chain Y
lemma discrete_chain_const:
chain S ==> ∃x. S = (λi. x)
lemma infinite_chain_adm_lemma:
[| chain Y; ∀i. P (Y i);
!!Y. [| chain Y; ∀i. P (Y i); ¬ finite_chain Y |] ==> P (LUB i. Y i) |]
==> P (LUB i. Y i)
lemma increasing_chain_adm_lemma:
[| chain Y; ∀i. P (Y i);
!!Y. [| chain Y; ∀i. P (Y i); ∀i. ∃j>i. Y i ≠ Y j ∧ Y i << Y j |]
==> P (LUB i. Y i) |]
==> P (LUB i. Y i)