(* Title: HOLCF/FunCpo.thy ID: $Id: Ffun.thy,v 1.14 2008/03/27 18:49:24 huffman Exp $ Author: Franz Regensburger Definition of the partial ordering for the type of all functions => (fun) Class instance of => (fun) for class pcpo. *) header {* Class instances for the full function space *} theory Ffun imports Cont begin subsection {* Full function space is a partial order *} instantiation "fun" :: (type, sq_ord) sq_ord begin definition less_fun_def: "(op \<sqsubseteq>) ≡ (λf g. ∀x. f x \<sqsubseteq> g x)" instance .. end instance "fun" :: (type, po) po proof fix f :: "'a => 'b" show "f \<sqsubseteq> f" by (simp add: less_fun_def) next fix f g :: "'a => 'b" assume "f \<sqsubseteq> g" and "g \<sqsubseteq> f" thus "f = g" by (simp add: less_fun_def expand_fun_eq antisym_less) next fix f g h :: "'a => 'b" assume "f \<sqsubseteq> g" and "g \<sqsubseteq> h" thus "f \<sqsubseteq> h" unfolding less_fun_def by (fast elim: trans_less) qed text {* make the symbol @{text "<<"} accessible for type fun *} lemma expand_fun_less: "(f \<sqsubseteq> g) = (∀x. f x \<sqsubseteq> g x)" by (simp add: less_fun_def) lemma less_fun_ext: "(!!x. f x \<sqsubseteq> g x) ==> f \<sqsubseteq> g" by (simp add: less_fun_def) subsection {* Full function space is chain complete *} text {* function application is monotone *} lemma monofun_app: "monofun (λf. f x)" by (rule monofunI, simp add: less_fun_def) text {* chains of functions yield chains in the po range *} lemma ch2ch_fun: "chain S ==> chain (λi. S i x)" by (simp add: chain_def less_fun_def) lemma ch2ch_lambda: "(!!x. chain (λi. S i x)) ==> chain S" by (simp add: chain_def less_fun_def) text {* upper bounds of function chains yield upper bound in the po range *} lemma ub2ub_fun: "range S <| u ==> range (λi. S i x) <| u x" by (auto simp add: is_ub_def less_fun_def) text {* Type @{typ "'a::type => 'b::cpo"} is chain complete *} lemma is_lub_lambda: assumes f: "!!x. range (λi. Y i x) <<| f x" shows "range Y <<| f" apply (rule is_lubI) apply (rule ub_rangeI) apply (rule less_fun_ext) apply (rule is_ub_lub [OF f]) apply (rule less_fun_ext) apply (rule is_lub_lub [OF f]) apply (erule ub2ub_fun) done lemma lub_fun: "chain (S::nat => 'a::type => 'b::cpo) ==> range S <<| (λx. \<Squnion>i. S i x)" apply (rule is_lub_lambda) apply (rule cpo_lubI) apply (erule ch2ch_fun) done lemma thelub_fun: "chain (S::nat => 'a::type => 'b::cpo) ==> lub (range S) = (λx. \<Squnion>i. S i x)" by (rule lub_fun [THEN thelubI]) lemma cpo_fun: "chain (S::nat => 'a::type => 'b::cpo) ==> ∃x. range S <<| x" by (rule exI, erule lub_fun) instance "fun" :: (type, cpo) cpo by intro_classes (rule cpo_fun) instance "fun" :: (finite, finite_po) finite_po .. instance "fun" :: (type, discrete_cpo) discrete_cpo proof fix f g :: "'a => 'b" show "f \<sqsubseteq> g <-> f = g" unfolding expand_fun_less expand_fun_eq by simp qed text {* chain-finite function spaces *} lemma maxinch2maxinch_lambda: "(!!x. max_in_chain n (λi. S i x)) ==> max_in_chain n S" unfolding max_in_chain_def expand_fun_eq by simp lemma maxinch_mono: "[|max_in_chain i Y; i ≤ j|] ==> max_in_chain j Y" unfolding max_in_chain_def proof (intro allI impI) fix k assume Y: "∀n≥i. Y i = Y n" assume ij: "i ≤ j" assume jk: "j ≤ k" from ij jk have ik: "i ≤ k" by simp from Y ij have Yij: "Y i = Y j" by simp from Y ik have Yik: "Y i = Y k" by simp from Yij Yik show "Y j = Y k" by auto qed instance "fun" :: (finite, chfin) chfin proof fix Y :: "nat => 'a => 'b" let ?n = "λx. LEAST n. max_in_chain n (λi. Y i x)" assume "chain Y" hence "!!x. chain (λi. Y i x)" by (rule ch2ch_fun) hence "!!x. ∃n. max_in_chain n (λi. Y i x)" by (rule chfin) hence "!!x. max_in_chain (?n x) (λi. Y i x)" by (rule LeastI_ex) hence "!!x. max_in_chain (Max (range ?n)) (λi. Y i x)" by (rule maxinch_mono [OF _ Max_ge], simp_all) hence "max_in_chain (Max (range ?n)) Y" by (rule maxinch2maxinch_lambda) thus "∃n. max_in_chain n Y" .. qed subsection {* Full function space is pointed *} lemma minimal_fun: "(λx. ⊥) \<sqsubseteq> f" by (simp add: less_fun_def) lemma least_fun: "∃x::'a::type => 'b::pcpo. ∀y. x \<sqsubseteq> y" apply (rule_tac x = "λx. ⊥" in exI) apply (rule minimal_fun [THEN allI]) done instance "fun" :: (type, pcpo) pcpo by intro_classes (rule least_fun) text {* for compatibility with old HOLCF-Version *} lemma inst_fun_pcpo: "⊥ = (λx. ⊥)" by (rule minimal_fun [THEN UU_I, symmetric]) text {* function application is strict in the left argument *} lemma app_strict [simp]: "⊥ x = ⊥" by (simp add: inst_fun_pcpo) text {* The following results are about application for functions in @{typ "'a=>'b"} *} lemma monofun_fun_fun: "f \<sqsubseteq> g ==> f x \<sqsubseteq> g x" by (simp add: less_fun_def) lemma monofun_fun_arg: "[|monofun f; x \<sqsubseteq> y|] ==> f x \<sqsubseteq> f y" by (rule monofunE) lemma monofun_fun: "[|monofun f; monofun g; f \<sqsubseteq> g; x \<sqsubseteq> y|] ==> f x \<sqsubseteq> g y" by (rule trans_less [OF monofun_fun_arg monofun_fun_fun]) subsection {* Propagation of monotonicity and continuity *} text {* the lub of a chain of monotone functions is monotone *} lemma monofun_lub_fun: "[|chain (F::nat => 'a => 'b::cpo); ∀i. monofun (F i)|] ==> monofun (\<Squnion>i. F i)" apply (rule monofunI) apply (simp add: thelub_fun) apply (rule lub_mono) apply (erule ch2ch_fun) apply (erule ch2ch_fun) apply (simp add: monofunE) done text {* the lub of a chain of continuous functions is continuous *} declare range_composition [simp del] lemma contlub_lub_fun: "[|chain F; ∀i. cont (F i)|] ==> contlub (\<Squnion>i. F i)" apply (rule contlubI) apply (simp add: thelub_fun) apply (simp add: cont2contlubE) apply (rule ex_lub) apply (erule ch2ch_fun) apply (simp add: ch2ch_cont) done lemma cont_lub_fun: "[|chain F; ∀i. cont (F i)|] ==> cont (\<Squnion>i. F i)" apply (rule monocontlub2cont) apply (erule monofun_lub_fun) apply (simp add: cont2mono) apply (erule (1) contlub_lub_fun) done lemma cont2cont_lub: "[|chain F; !!i. cont (F i)|] ==> cont (λx. \<Squnion>i. F i x)" by (simp add: thelub_fun [symmetric] cont_lub_fun) lemma mono2mono_fun: "monofun f ==> monofun (λx. f x y)" apply (rule monofunI) apply (erule (1) monofun_fun_arg [THEN monofun_fun_fun]) done lemma cont2cont_fun: "cont f ==> cont (λx. f x y)" apply (rule monocontlub2cont) apply (erule cont2mono [THEN mono2mono_fun]) apply (rule contlubI) apply (simp add: cont2contlubE) apply (simp add: thelub_fun ch2ch_cont) done text {* Note @{text "(λx. λy. f x y) = f"} *} lemma mono2mono_lambda: assumes f: "!!y. monofun (λx. f x y)" shows "monofun f" apply (rule monofunI) apply (rule less_fun_ext) apply (erule monofunE [OF f]) done lemma cont2cont_lambda [simp]: assumes f: "!!y. cont (λx. f x y)" shows "cont f" apply (subgoal_tac "monofun f") apply (rule monocontlub2cont) apply assumption apply (rule contlubI) apply (rule ext) apply (simp add: thelub_fun ch2ch_monofun) apply (erule cont2contlubE [OF f]) apply (simp add: mono2mono_lambda cont2mono f) done text {* What D.A.Schmidt calls continuity of abstraction; never used here *} lemma contlub_lambda: "(!!x::'a::type. chain (λi. S i x::'b::cpo)) ==> (λx. \<Squnion>i. S i x) = (\<Squnion>i. (λx. S i x))" by (simp add: thelub_fun ch2ch_lambda) lemma contlub_abstraction: "[|chain Y; ∀y. cont (λx.(c::'a::cpo=>'b::type=>'c::cpo) x y)|] ==> (λy. \<Squnion>i. c (Y i) y) = (\<Squnion>i. (λy. c (Y i) y))" apply (rule thelub_fun [symmetric]) apply (simp add: ch2ch_cont) done lemma mono2mono_app: "[|monofun f; ∀x. monofun (f x); monofun t|] ==> monofun (λx. (f x) (t x))" apply (rule monofunI) apply (simp add: monofun_fun monofunE) done lemma cont2contlub_app: "[|cont f; ∀x. cont (f x); cont t|] ==> contlub (λx. (f x) (t x))" apply (rule contlubI) apply (subgoal_tac "chain (λi. f (Y i))") apply (subgoal_tac "chain (λi. t (Y i))") apply (simp add: cont2contlubE thelub_fun) apply (rule diag_lub) apply (erule ch2ch_fun) apply (drule spec) apply (erule (1) ch2ch_cont) apply (erule (1) ch2ch_cont) apply (erule (1) ch2ch_cont) done lemma cont2cont_app: "[|cont f; ∀x. cont (f x); cont t|] ==> cont (λx. (f x) (t x))" by (blast intro: monocontlub2cont mono2mono_app cont2mono cont2contlub_app) lemmas cont2cont_app2 = cont2cont_app [rule_format] lemma cont2cont_app3: "[|cont f; cont t|] ==> cont (λx. f (t x))" by (rule cont2cont_app2 [OF cont_const]) end
lemma expand_fun_less:
f << g = (∀x. f x << g x)
lemma less_fun_ext:
(!!x. f x << g x) ==> f << g
lemma monofun_app:
monofun (λf. f x)
lemma ch2ch_fun:
chain S ==> chain (λi. S i x)
lemma ch2ch_lambda:
(!!x. chain (λi. S i x)) ==> chain S
lemma ub2ub_fun:
range S <| u ==> range (λi. S i x) <| u x
lemma is_lub_lambda:
(!!x. range (λi. Y i x) <<| f x) ==> range Y <<| f
lemma lub_fun:
chain S ==> range S <<| (λx. LUB i. S i x)
lemma thelub_fun:
chain S ==> Lub S = (λx. LUB i. S i x)
lemma cpo_fun:
chain S ==> ∃x. range S <<| x
lemma maxinch2maxinch_lambda:
(!!x. max_in_chain n (λi. S i x)) ==> max_in_chain n S
lemma maxinch_mono:
[| max_in_chain i Y; i ≤ j |] ==> max_in_chain j Y
lemma minimal_fun:
(λx. UU) << f
lemma least_fun:
∃x. ∀y. x << y
lemma inst_fun_pcpo:
UU = (λx. UU)
lemma app_strict:
UU x = UU
lemma monofun_fun_fun:
f << g ==> f x << g x
lemma monofun_fun_arg:
[| monofun f; x << y |] ==> f x << f y
lemma monofun_fun:
[| monofun f; monofun g; f << g; x << y |] ==> f x << g y
lemma monofun_lub_fun:
[| chain F; ∀i. monofun (F i) |] ==> monofun (LUB i. F i)
lemma contlub_lub_fun:
[| chain F; ∀i. cont (F i) |] ==> contlub (LUB i. F i)
lemma cont_lub_fun:
[| chain F; ∀i. cont (F i) |] ==> cont (LUB i. F i)
lemma cont2cont_lub:
[| chain F; !!i. cont (F i) |] ==> cont (λx. LUB i. F i x)
lemma mono2mono_fun:
monofun f ==> monofun (λx. f x y)
lemma cont2cont_fun:
cont f ==> cont (λx. f x y)
lemma mono2mono_lambda:
(!!y. monofun (λx. f x y)) ==> monofun f
lemma cont2cont_lambda:
(!!y. cont (λx. f x y)) ==> cont f
lemma contlub_lambda:
(!!x. chain (λi. S i x)) ==> (λx. LUB i. S i x) = (LUB i. S i)
lemma contlub_abstraction:
[| chain Y; ∀y. cont (λx. c x y) |]
==> (λy. LUB i. c (Y i) y) = (LUB i. c (Y i))
lemma mono2mono_app:
[| monofun f; ∀x. monofun (f x); monofun t |] ==> monofun (λx. f x (t x))
lemma cont2contlub_app:
[| cont f; ∀x. cont (f x); cont t |] ==> contlub (λx. f x (t x))
lemma cont2cont_app:
[| cont f; ∀x. cont (f x); cont t |] ==> cont (λx. f x (t x))
lemma cont2cont_app2:
[| cont f; !!x. cont (f x); cont t |] ==> cont (λx. f x (t x))
lemma cont2cont_app3:
[| cont f; cont t |] ==> cont (λx. f (t x))