# HG changeset patch # User huffman # Date 1291658913 28800 # Node ID ff7d177128ef0da39e48e27747dd333fcd9eac14 # Parent f7d8cfa6e7fc6c713aa8ffe6a2343d457630883f rename lub_fun -> is_lub_fun, thelub_fun -> lub_fun diff -r f7d8cfa6e7fc -r ff7d177128ef NEWS --- a/NEWS Mon Dec 06 08:59:58 2010 -0800 +++ b/NEWS Mon Dec 06 10:08:33 2010 -0800 @@ -524,6 +524,8 @@ unique_lub ~> is_lub_unique is_ub_lub ~> is_lub_rangeD1 lub_bin_chain ~> is_lub_bin_chain + lub_fun ~> is_lub_fun + thelub_fun ~> lub_fun thelub_Pair ~> lub_Pair lub_cprod ~> is_lub_prod thelub_cprod ~> lub_prod diff -r f7d8cfa6e7fc -r ff7d177128ef src/HOL/HOLCF/Cfun.thy --- a/src/HOL/HOLCF/Cfun.thy Mon Dec 06 08:59:58 2010 -0800 +++ b/src/HOL/HOLCF/Cfun.thy Mon Dec 06 10:08:33 2010 -0800 @@ -260,7 +260,7 @@ \ (\i. \ x. F i x) = (\ x. \i. F i x)" apply (simp add: lub_cfun) apply (simp add: Abs_cfun_inverse2) -apply (simp add: thelub_fun ch2ch_lambda) +apply (simp add: lub_fun ch2ch_lambda) done lemmas lub_distribs = lub_APP lub_LAM diff -r f7d8cfa6e7fc -r ff7d177128ef src/HOL/HOLCF/Fun_Cpo.thy --- a/src/HOL/HOLCF/Fun_Cpo.thy Mon Dec 06 08:59:58 2010 -0800 +++ b/src/HOL/HOLCF/Fun_Cpo.thy Mon Dec 06 10:08:33 2010 -0800 @@ -57,19 +57,13 @@ lemma ch2ch_lambda: "(\x. chain (\i. S i x)) \ chain S" by (simp add: chain_def below_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 below_fun_def) - text {* Type @{typ "'a::type => 'b::cpo"} is chain complete *} lemma is_lub_lambda: "(\x. range (\i. Y i x) <<| f x) \ range Y <<| f" unfolding is_lub_def is_ub_def below_fun_def by simp -lemma lub_fun: +lemma is_lub_fun: "chain (S::nat \ 'a::type \ 'b::cpo) \ range S <<| (\x. \i. S i x)" apply (rule is_lub_lambda) @@ -77,13 +71,13 @@ apply (erule ch2ch_fun) done -lemma thelub_fun: +lemma lub_fun: "chain (S::nat \ 'a::type \ 'b::cpo) \ (\i. S i) = (\x. \i. S i x)" -by (rule lub_fun [THEN lub_eqI]) +by (rule is_lub_fun [THEN lub_eqI]) instance "fun" :: (type, cpo) cpo -by intro_classes (rule exI, erule lub_fun) +by intro_classes (rule exI, erule is_lub_fun) subsection {* Chain-finiteness of function space *} @@ -135,12 +129,12 @@ text {* The lub of a chain of monotone functions is monotone. *} lemma adm_monofun: "adm monofun" -by (rule admI, simp add: thelub_fun fun_chain_iff monofun_def lub_mono) +by (rule admI, simp add: lub_fun fun_chain_iff monofun_def lub_mono) text {* The lub of a chain of continuous functions is continuous. *} lemma adm_cont: "adm cont" -by (rule admI, simp add: thelub_fun fun_chain_iff) +by (rule admI, simp add: lub_fun fun_chain_iff) text {* Function application preserves monotonicity and continuity. *} @@ -150,7 +144,7 @@ lemma cont2cont_fun: "cont f \ cont (\x. f x y)" apply (rule contI2) apply (erule cont2mono [THEN mono2mono_fun]) -apply (simp add: cont2contlubE thelub_fun ch2ch_cont) +apply (simp add: cont2contlubE lub_fun ch2ch_cont) done lemma cont_fun: "cont (\f. f x)" @@ -174,6 +168,6 @@ lemma contlub_lambda: "(\x::'a::type. chain (\i. S i x::'b::cpo)) \ (\x. \i. S i x) = (\i. (\x. S i x))" -by (simp add: thelub_fun ch2ch_lambda) +by (simp add: lub_fun ch2ch_lambda) end