Theory Cont
section ‹Continuity and monotonicity›
theory Cont
imports Pcpo
begin
text ‹
Now we change the default class! Form now on all untyped type variables are
of default class po
›
default_sort po
subsection ‹Definitions›
definition monofun :: "('a ⇒ 'b) ⇒ bool"
where "monofun f ⟷ (∀x y. x ⊑ y ⟶ f x ⊑ f y)"
definition cont :: "('a::cpo ⇒ 'b::cpo) ⇒ bool"
where "cont f = (∀Y. chain Y ⟶ range (λi. f (Y i)) <<| f (⨆i. Y i))"
lemma contI: "(⋀Y. chain Y ⟹ range (λi. f (Y i)) <<| f (⨆i. Y i)) ⟹ cont f"
by (simp add: cont_def)
lemma contE: "cont f ⟹ chain Y ⟹ range (λi. f (Y i)) <<| f (⨆i. Y i)"
by (simp add: cont_def)
lemma monofunI: "(⋀x y. x ⊑ y ⟹ f x ⊑ f y) ⟹ monofun f"
by (simp add: monofun_def)
lemma monofunE: "monofun f ⟹ x ⊑ y ⟹ f x ⊑ f y"
by (simp add: monofun_def)
subsection ‹Equivalence of alternate definition›
text ‹monotone functions map chains to chains›
lemma ch2ch_monofun: "monofun f ⟹ chain Y ⟹ chain (λi. f (Y i))"
apply (rule chainI)
apply (erule monofunE)
apply (erule chainE)
done
text ‹monotone functions map upper bound to upper bounds›
lemma ub2ub_monofun: "monofun f ⟹ range Y <| u ⟹ range (λi. f (Y i)) <| f u"
apply (rule ub_rangeI)
apply (erule monofunE)
apply (erule ub_rangeD)
done
text ‹a lemma about binary chains›
lemma binchain_cont: "cont f ⟹ x ⊑ y ⟹ range (λi::nat. f (if i = 0 then x else y)) <<| f y"
apply (subgoal_tac "f (⨆i::nat. if i = 0 then x else y) = f y")
apply (erule subst)
apply (erule contE)
apply (erule bin_chain)
apply (rule_tac f=f in arg_cong)
apply (erule is_lub_bin_chain [THEN lub_eqI])
done
text ‹continuity implies monotonicity›
lemma cont2mono: "cont f ⟹ monofun f"
apply (rule monofunI)
apply (drule (1) binchain_cont)
apply (drule_tac i=0 in is_lub_rangeD1)
apply simp
done
lemmas cont2monofunE = cont2mono [THEN monofunE]
lemmas ch2ch_cont = cont2mono [THEN ch2ch_monofun]
text ‹continuity implies preservation of lubs›
lemma cont2contlubE: "cont f ⟹ chain Y ⟹ f (⨆i. Y i) = (⨆i. f (Y i))"
apply (rule lub_eqI [symmetric])
apply (erule (1) contE)
done
lemma contI2:
fixes f :: "'a::cpo ⇒ 'b::cpo"
assumes mono: "monofun f"
assumes below: "⋀Y. ⟦chain Y; chain (λi. f (Y i))⟧ ⟹ f (⨆i. Y i) ⊑ (⨆i. f (Y i))"
shows "cont f"
proof (rule contI)
fix Y :: "nat ⇒ 'a"
assume Y: "chain Y"
with mono have fY: "chain (λi. f (Y i))"
by (rule ch2ch_monofun)
have "(⨆i. f (Y i)) = f (⨆i. Y i)"
apply (rule below_antisym)
apply (rule lub_below [OF fY])
apply (rule monofunE [OF mono])
apply (rule is_ub_thelub [OF Y])
apply (rule below [OF Y fY])
done
with fY show "range (λi. f (Y i)) <<| f (⨆i. Y i)"
by (rule thelubE)
qed
subsection ‹Collection of continuity rules›
named_theorems cont2cont "continuity intro rule"
subsection ‹Continuity of basic functions›
text ‹The identity function is continuous›
lemma cont_id [simp, cont2cont]: "cont (λx. x)"
apply (rule contI)
apply (erule cpo_lubI)
done
text ‹constant functions are continuous›
lemma cont_const [simp, cont2cont]: "cont (λx. c)"
using is_lub_const by (rule contI)
text ‹application of functions is continuous›
lemma cont_apply:
fixes f :: "'a::cpo ⇒ 'b::cpo ⇒ 'c::cpo" and t :: "'a ⇒ 'b"
assumes 1: "cont (λx. t x)"
assumes 2: "⋀x. cont (λy. f x y)"
assumes 3: "⋀y. cont (λx. f x y)"
shows "cont (λx. (f x) (t x))"
proof (rule contI2 [OF monofunI])
fix x y :: "'a"
assume "x ⊑ y"
then show "f x (t x) ⊑ f y (t y)"
by (auto intro: cont2monofunE [OF 1]
cont2monofunE [OF 2]
cont2monofunE [OF 3]
below_trans)
next
fix Y :: "nat ⇒ 'a"
assume "chain Y"
then show "f (⨆i. Y i) (t (⨆i. Y i)) ⊑ (⨆i. f (Y i) (t (Y i)))"
by (simp only: cont2contlubE [OF 1] ch2ch_cont [OF 1]
cont2contlubE [OF 2] ch2ch_cont [OF 2]
cont2contlubE [OF 3] ch2ch_cont [OF 3]
diag_lub below_refl)
qed
lemma cont_compose: "cont c ⟹ cont (λx. f x) ⟹ cont (λx. c (f x))"
by (rule cont_apply [OF _ _ cont_const])
text ‹Least upper bounds preserve continuity›
lemma cont2cont_lub [simp]:
assumes chain: "⋀x. chain (λi. F i x)"
and cont: "⋀i. cont (λx. F i x)"
shows "cont (λx. ⨆i. F i x)"
apply (rule contI2)
apply (simp add: monofunI cont2monofunE [OF cont] lub_mono chain)
apply (simp add: cont2contlubE [OF cont])
apply (simp add: diag_lub ch2ch_cont [OF cont] chain)
done
text ‹if-then-else is continuous›
lemma cont_if [simp, cont2cont]: "cont f ⟹ cont g ⟹ cont (λx. if b then f x else g x)"
by (induct b) simp_all
subsection ‹Finite chains and flat pcpos›
text ‹Monotone functions map finite chains to finite chains.›
lemma monofun_finch2finch: "monofun f ⟹ finite_chain Y ⟹ finite_chain (λn. f (Y n))"
by (force simp add: finite_chain_def ch2ch_monofun max_in_chain_def)
text ‹The same holds for continuous functions.›
lemma cont_finch2finch: "cont f ⟹ finite_chain Y ⟹ finite_chain (λn. f (Y n))"
by (rule cont2mono [THEN monofun_finch2finch])
text ‹All monotone functions with chain-finite domain are continuous.›
lemma chfindom_monofun2cont: "monofun f ⟹ cont f"
for f :: "'a::chfin ⇒ 'b::cpo"
apply (erule contI2)
apply (frule chfin2finch)
apply (clarsimp simp add: finite_chain_def)
apply (subgoal_tac "max_in_chain i (λi. f (Y i))")
apply (simp add: maxinch_is_thelub ch2ch_monofun)
apply (force simp add: max_in_chain_def)
done
text ‹All strict functions with flat domain are continuous.›
lemma flatdom_strict2mono: "f ⊥ = ⊥ ⟹ monofun f"
for f :: "'a::flat ⇒ 'b::pcpo"
apply (rule monofunI)
apply (drule ax_flat)
apply auto
done
lemma flatdom_strict2cont: "f ⊥ = ⊥ ⟹ cont f"
for f :: "'a::flat ⇒ 'b::pcpo"
by (rule flatdom_strict2mono [THEN chfindom_monofun2cont])
text ‹All functions with discrete domain are continuous.›
lemma cont_discrete_cpo [simp, cont2cont]: "cont f"
for f :: "'a::discrete_cpo ⇒ 'b::cpo"
apply (rule contI)
apply (drule discrete_chain_const, clarify)
apply simp
done
end