Theory Cfun

Up to index of Isabelle/HOLCF

theory Cfun
imports Pcpodef
uses (Tools/cont_proc.ML)
begin

(*  Title:      HOLCF/Cfun.thy
    ID:         $Id: Cfun.thy,v 1.32 2007/10/21 14:27:42 wenzelm Exp $
    Author:     Franz Regensburger

Definition of the type ->  of continuous functions.
*)

header {* The type of continuous functions *}

theory Cfun
imports Pcpodef
uses ("Tools/cont_proc.ML")
begin

defaultsort cpo

subsection {* Definition of continuous function type *}

lemma Ex_cont: "∃f. cont f"
by (rule exI, rule cont_const)

lemma adm_cont: "adm cont"
by (rule admI, rule cont_lub_fun)

cpodef (CFun)  ('a, 'b) "->" (infixr "->" 0) = "{f::'a => 'b. cont f}"
by (simp add: Ex_cont adm_cont)

syntax (xsymbols)
  "->"     :: "[type, type] => type"      ("(_ ->/ _)" [1,0]0)

notation
  Rep_CFun  ("(_$/_)" [999,1000] 999)

notation (xsymbols)
  Rep_CFun  ("(_·/_)" [999,1000] 999)

notation (HTML output)
  Rep_CFun  ("(_·/_)" [999,1000] 999)

subsection {* Syntax for continuous lambda abstraction *}

syntax "_cabs" :: "'a"

parse_translation {*
(* rewrites (_cabs x t) => (Abs_CFun (%x. t)) *)
  [mk_binder_tr ("_cabs", @{const_syntax Abs_CFun})];
*}

text {* To avoid eta-contraction of body: *}
typed_print_translation {*
  let
    fun cabs_tr' _ _ [Abs abs] = let
          val (x,t) = atomic_abs_tr' abs
        in Syntax.const "_cabs" $ x $ t end

      | cabs_tr' _ T [t] = let
          val xT = domain_type (domain_type T);
          val abs' = ("x",xT,(incr_boundvars 1 t)$Bound 0);
          val (x,t') = atomic_abs_tr' abs';
        in Syntax.const "_cabs" $ x $ t' end;

  in [(@{const_syntax Abs_CFun}, cabs_tr')] end;
*}

text {* Syntax for nested abstractions *}

syntax
  "_Lambda" :: "[cargs, 'a] => logic"  ("(3LAM _./ _)" [1000, 10] 10)

syntax (xsymbols)
  "_Lambda" :: "[cargs, 'a] => logic" ("(3Λ_./ _)" [1000, 10] 10)

parse_ast_translation {*
(* rewrites (LAM x y z. t) => (_cabs x (_cabs y (_cabs z t))) *)
(* cf. Syntax.lambda_ast_tr from Syntax/syn_trans.ML *)
  let
    fun Lambda_ast_tr [pats, body] =
          Syntax.fold_ast_p "_cabs" (Syntax.unfold_ast "_cargs" pats, body)
      | Lambda_ast_tr asts = raise Syntax.AST ("Lambda_ast_tr", asts);
  in [("_Lambda", Lambda_ast_tr)] end;
*}

print_ast_translation {*
(* rewrites (_cabs x (_cabs y (_cabs z t))) => (LAM x y z. t) *)
(* cf. Syntax.abs_ast_tr' from Syntax/syn_trans.ML *)
  let
    fun cabs_ast_tr' asts =
      (case Syntax.unfold_ast_p "_cabs"
          (Syntax.Appl (Syntax.Constant "_cabs" :: asts)) of
        ([], _) => raise Syntax.AST ("cabs_ast_tr'", asts)
      | (xs, body) => Syntax.Appl
          [Syntax.Constant "_Lambda", Syntax.fold_ast "_cargs" xs, body]);
  in [("_cabs", cabs_ast_tr')] end;
*}

text {* Dummy patterns for continuous abstraction *}
translations
  "Λ _. t" => "CONST Abs_CFun (λ _. t)"


subsection {* Continuous function space is pointed *}

lemma UU_CFun: "⊥ ∈ CFun"
by (simp add: CFun_def inst_fun_pcpo cont_const)

instance "->" :: (cpo, pcpo) pcpo
by (rule typedef_pcpo [OF type_definition_CFun less_CFun_def UU_CFun])

lemmas Rep_CFun_strict =
  typedef_Rep_strict [OF type_definition_CFun less_CFun_def UU_CFun]

lemmas Abs_CFun_strict =
  typedef_Abs_strict [OF type_definition_CFun less_CFun_def UU_CFun]

text {* function application is strict in its first argument *}

lemma Rep_CFun_strict1 [simp]: "⊥·x = ⊥"
by (simp add: Rep_CFun_strict)

text {* for compatibility with old HOLCF-Version *}
lemma inst_cfun_pcpo: "⊥ = (Λ x. ⊥)"
by (simp add: inst_fun_pcpo [symmetric] Abs_CFun_strict)

subsection {* Basic properties of continuous functions *}

text {* Beta-equality for continuous functions *}

lemma Abs_CFun_inverse2: "cont f ==> Rep_CFun (Abs_CFun f) = f"
by (simp add: Abs_CFun_inverse CFun_def)

lemma beta_cfun [simp]: "cont f ==> (Λ x. f x)·u = f u"
by (simp add: Abs_CFun_inverse2)

text {* Eta-equality for continuous functions *}

lemma eta_cfun: "(Λ x. f·x) = f"
by (rule Rep_CFun_inverse)

text {* Extensionality for continuous functions *}

lemma expand_cfun_eq: "(f = g) = (∀x. f·x = g·x)"
by (simp add: Rep_CFun_inject [symmetric] expand_fun_eq)

lemma ext_cfun: "(!!x. f·x = g·x) ==> f = g"
by (simp add: expand_cfun_eq)

text {* Extensionality wrt. ordering for continuous functions *}

lemma expand_cfun_less: "f \<sqsubseteq> g = (∀x. f·x \<sqsubseteq> g·x)" 
by (simp add: less_CFun_def expand_fun_less)

lemma less_cfun_ext: "(!!x. f·x \<sqsubseteq> g·x) ==> f \<sqsubseteq> g"
by (simp add: expand_cfun_less)

text {* Congruence for continuous function application *}

lemma cfun_cong: "[|f = g; x = y|] ==> f·x = g·y"
by simp

lemma cfun_fun_cong: "f = g ==> f·x = g·x"
by simp

lemma cfun_arg_cong: "x = y ==> f·x = f·y"
by simp

subsection {* Continuity of application *}

lemma cont_Rep_CFun1: "cont (λf. f·x)"
by (rule cont_Rep_CFun [THEN cont2cont_fun])

lemma cont_Rep_CFun2: "cont (λx. f·x)"
apply (cut_tac x=f in Rep_CFun)
apply (simp add: CFun_def)
done

lemmas monofun_Rep_CFun = cont_Rep_CFun [THEN cont2mono]
lemmas contlub_Rep_CFun = cont_Rep_CFun [THEN cont2contlub]

lemmas monofun_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2mono, standard]
lemmas contlub_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2contlub, standard]
lemmas monofun_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2mono, standard]
lemmas contlub_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2contlub, standard]

text {* contlub, cont properties of @{term Rep_CFun} in each argument *}

lemma contlub_cfun_arg: "chain Y ==> f·(lub (range Y)) = (\<Squnion>i. f·(Y i))"
by (rule contlub_Rep_CFun2 [THEN contlubE])

lemma cont_cfun_arg: "chain Y ==> range (λi. f·(Y i)) <<| f·(lub (range Y))"
by (rule cont_Rep_CFun2 [THEN contE])

lemma contlub_cfun_fun: "chain F ==> lub (range F)·x = (\<Squnion>i. F i·x)"
by (rule contlub_Rep_CFun1 [THEN contlubE])

lemma cont_cfun_fun: "chain F ==> range (λi. F i·x) <<| lub (range F)·x"
by (rule cont_Rep_CFun1 [THEN contE])

text {* monotonicity of application *}

lemma monofun_cfun_fun: "f \<sqsubseteq> g ==> f·x \<sqsubseteq> g·x"
by (simp add: expand_cfun_less)

lemma monofun_cfun_arg: "x \<sqsubseteq> y ==> f·x \<sqsubseteq> f·y"
by (rule monofun_Rep_CFun2 [THEN monofunE])

lemma monofun_cfun: "[|f \<sqsubseteq> g; x \<sqsubseteq> y|] ==> f·x \<sqsubseteq> g·y"
by (rule trans_less [OF monofun_cfun_fun monofun_cfun_arg])

text {* ch2ch - rules for the type @{typ "'a -> 'b"} *}

lemma chain_monofun: "chain Y ==> chain (λi. f·(Y i))"
by (erule monofun_Rep_CFun2 [THEN ch2ch_monofun])

lemma ch2ch_Rep_CFunR: "chain Y ==> chain (λi. f·(Y i))"
by (rule monofun_Rep_CFun2 [THEN ch2ch_monofun])

lemma ch2ch_Rep_CFunL: "chain F ==> chain (λi. (F i)·x)"
by (rule monofun_Rep_CFun1 [THEN ch2ch_monofun])

lemma ch2ch_Rep_CFun [simp]:
  "[|chain F; chain Y|] ==> chain (λi. (F i)·(Y i))"
apply (rule chainI)
apply (rule monofun_cfun)
apply (erule chainE)
apply (erule chainE)
done

lemma ch2ch_LAM: "[|!!x. chain (λi. S i x); !!i. cont (λx. S i x)|]
    ==> chain (λi. Λ x. S i x)"
by (simp add: chain_def expand_cfun_less)

text {* contlub, cont properties of @{term Rep_CFun} in both arguments *}

lemma contlub_cfun: 
  "[|chain F; chain Y|] ==> (\<Squnion>i. F i)·(\<Squnion>i. Y i) = (\<Squnion>i. F i·(Y i))"
by (simp add: contlub_cfun_fun contlub_cfun_arg diag_lub)

lemma cont_cfun: 
  "[|chain F; chain Y|] ==> range (λi. F i·(Y i)) <<| (\<Squnion>i. F i)·(\<Squnion>i. Y i)"
apply (rule thelubE)
apply (simp only: ch2ch_Rep_CFun)
apply (simp only: contlub_cfun)
done

lemma contlub_LAM:
  "[|!!x. chain (λi. F i x); !!i. cont (λx. F i x)|]
    ==> (Λ x. \<Squnion>i. F i x) = (\<Squnion>i. Λ x. F i x)"
apply (simp add: thelub_CFun ch2ch_LAM)
apply (simp add: Abs_CFun_inverse2)
apply (simp add: thelub_fun ch2ch_lambda)
done

text {* strictness *}

lemma strictI: "f·x = ⊥ ==> f·⊥ = ⊥"
apply (rule UU_I)
apply (erule subst)
apply (rule minimal [THEN monofun_cfun_arg])
done

text {* the lub of a chain of continous functions is monotone *}

lemma lub_cfun_mono: "chain F ==> monofun (λx. \<Squnion>i. F i·x)"
apply (drule ch2ch_monofun [OF monofun_Rep_CFun])
apply (simp add: thelub_fun [symmetric])
apply (erule monofun_lub_fun)
apply (simp add: monofun_Rep_CFun2)
done

text {* a lemma about the exchange of lubs for type @{typ "'a -> 'b"} *}

lemma ex_lub_cfun:
  "[|chain F; chain Y|] ==> (\<Squnion>j. \<Squnion>i. F j·(Y i)) = (\<Squnion>i. \<Squnion>j. F j·(Y i))"
by (simp add: diag_lub)

text {* the lub of a chain of cont. functions is continuous *}

lemma cont_lub_cfun: "chain F ==> cont (λx. \<Squnion>i. F i·x)"
apply (rule cont2cont_lub)
apply (erule monofun_Rep_CFun [THEN ch2ch_monofun])
apply (rule cont_Rep_CFun2)
done

text {* type @{typ "'a -> 'b"} is chain complete *}

lemma lub_cfun: "chain F ==> range F <<| (Λ x. \<Squnion>i. F i·x)"
by (simp only: contlub_cfun_fun [symmetric] eta_cfun thelubE)

lemma thelub_cfun: "chain F ==> lub (range F) = (Λ x. \<Squnion>i. F i·x)"
by (rule lub_cfun [THEN thelubI])

subsection {* Continuity simplification procedure *}

text {* cont2cont lemma for @{term Rep_CFun} *}

lemma cont2cont_Rep_CFun:
  "[|cont f; cont t|] ==> cont (λx. (f x)·(t x))"
by (best intro: cont2cont_app2 cont_const cont_Rep_CFun cont_Rep_CFun2)

text {* cont2mono Lemma for @{term "%x. LAM y. c1(x)(y)"} *}

lemma cont2mono_LAM:
assumes p1: "!!x. cont(c1 x)"
assumes p2: "!!y. monofun(%x. c1 x y)"
shows "monofun(%x. LAM y. c1 x y)"
apply (rule monofunI)
apply (rule less_cfun_ext)
apply (simp add: p1)
apply (erule p2 [THEN monofunE])
done

text {* cont2cont Lemma for @{term "%x. LAM y. c1 x y"} *}

lemma cont2cont_LAM:
assumes p1: "!!x. cont(c1 x)"
assumes p2: "!!y. cont(%x. c1 x y)"
shows "cont(%x. LAM y. c1 x y)"
apply (rule cont_Abs_CFun)
apply (simp add: p1 CFun_def)
apply (simp add: p2 cont2cont_lambda)
done

text {* continuity simplification procedure *}

lemmas cont_lemmas1 =
  cont_const cont_id cont_Rep_CFun2 cont2cont_Rep_CFun cont2cont_LAM

use "Tools/cont_proc.ML";
setup ContProc.setup;

(*val cont_tac = (fn i => (resolve_tac cont_lemmas i));*)
(*val cont_tacR = (fn i => (REPEAT (cont_tac i)));*)

subsection {* Miscellaneous *}

text {* Monotonicity of @{term Abs_CFun} *}

lemma semi_monofun_Abs_CFun:
  "[|cont f; cont g; f \<sqsubseteq> g|] ==> Abs_CFun f \<sqsubseteq> Abs_CFun g"
by (simp add: less_CFun_def Abs_CFun_inverse2)

text {* some lemmata for functions with flat/chfin domain/range types *}

lemma chfin_Rep_CFunR: "chain (Y::nat => 'a::cpo->'b::chfin)  
      ==> !s. ? n. lub(range(Y))$s = Y n$s"
apply (rule allI)
apply (subst contlub_cfun_fun)
apply assumption
apply (fast intro!: thelubI chfin lub_finch2 chfin2finch ch2ch_Rep_CFunL)
done

lemma adm_chfindom: "adm (λ(u::'a::cpo -> 'b::chfin). P(u·s))"
by (rule adm_subst, simp, rule adm_chfin)

subsection {* Continuous injection-retraction pairs *}

text {* Continuous retractions are strict. *}

lemma retraction_strict:
  "∀x. f·(g·x) = x ==> f·⊥ = ⊥"
apply (rule UU_I)
apply (drule_tac x="⊥" in spec)
apply (erule subst)
apply (rule monofun_cfun_arg)
apply (rule minimal)
done

lemma injection_eq:
  "∀x. f·(g·x) = x ==> (g·x = g·y) = (x = y)"
apply (rule iffI)
apply (drule_tac f=f in cfun_arg_cong)
apply simp
apply simp
done

lemma injection_less:
  "∀x. f·(g·x) = x ==> (g·x \<sqsubseteq> g·y) = (x \<sqsubseteq> y)"
apply (rule iffI)
apply (drule_tac f=f in monofun_cfun_arg)
apply simp
apply (erule monofun_cfun_arg)
done

lemma injection_defined_rev:
  "[|∀x. f·(g·x) = x; g·z = ⊥|] ==> z = ⊥"
apply (drule_tac f=f in cfun_arg_cong)
apply (simp add: retraction_strict)
done

lemma injection_defined:
  "[|∀x. f·(g·x) = x; z ≠ ⊥|] ==> g·z ≠ ⊥"
by (erule contrapos_nn, rule injection_defined_rev)

text {* propagation of flatness and chain-finiteness by retractions *}

lemma chfin2chfin:
  "∀y. (f::'a::chfin -> 'b)·(g·y) = y
    ==> ∀Y::nat => 'b. chain Y --> (∃n. max_in_chain n Y)"
apply clarify
apply (drule_tac f=g in chain_monofun)
apply (drule chfin [rule_format])
apply (unfold max_in_chain_def)
apply (simp add: injection_eq)
done

lemma flat2flat:
  "∀y. (f::'a::flat -> 'b::pcpo)·(g·y) = y
    ==> ∀x y::'b. x \<sqsubseteq> y --> x = ⊥ ∨ x = y"
apply clarify
apply (drule_tac f=g in monofun_cfun_arg)
apply (drule ax_flat [rule_format])
apply (erule disjE)
apply (simp add: injection_defined_rev)
apply (simp add: injection_eq)
done

text {* a result about functions with flat codomain *}

lemma flat_eqI: "[|(x::'a::flat) \<sqsubseteq> y; x ≠ ⊥|] ==> x = y"
by (drule ax_flat [rule_format], simp)

lemma flat_codom:
  "f·x = (c::'b::flat) ==> f·⊥ = ⊥ ∨ (∀z. f·z = c)"
apply (case_tac "f·x = ⊥")
apply (rule disjI1)
apply (rule UU_I)
apply (erule_tac t="⊥" in subst)
apply (rule minimal [THEN monofun_cfun_arg])
apply clarify
apply (rule_tac a = "f·⊥" in refl [THEN box_equals])
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
done


subsection {* Identity and composition *}

definition
  ID :: "'a -> 'a" where
  "ID = (Λ x. x)"

definition
  cfcomp  :: "('b -> 'c) -> ('a -> 'b) -> 'a -> 'c" where
  oo_def: "cfcomp = (Λ f g x. f·(g·x))"

abbreviation
  cfcomp_syn :: "['b -> 'c, 'a -> 'b] => 'a -> 'c"  (infixr "oo" 100)  where
  "f oo g == cfcomp·f·g"

lemma ID1 [simp]: "ID·x = x"
by (simp add: ID_def)

lemma cfcomp1: "(f oo g) = (Λ x. f·(g·x))"
by (simp add: oo_def)

lemma cfcomp2 [simp]: "(f oo g)·x = f·(g·x)"
by (simp add: cfcomp1)

lemma cfcomp_strict [simp]: "⊥ oo f = ⊥"
by (simp add: expand_cfun_eq)

text {*
  Show that interpretation of (pcpo,@{text "_->_"}) is a category.
  The class of objects is interpretation of syntactical class pcpo.
  The class of arrows  between objects @{typ 'a} and @{typ 'b} is interpret. of @{typ "'a -> 'b"}.
  The identity arrow is interpretation of @{term ID}.
  The composition of f and g is interpretation of @{text "oo"}.
*}

lemma ID2 [simp]: "f oo ID = f"
by (rule ext_cfun, simp)

lemma ID3 [simp]: "ID oo f = f"
by (rule ext_cfun, simp)

lemma assoc_oo: "f oo (g oo h) = (f oo g) oo h"
by (rule ext_cfun, simp)


subsection {* Strictified functions *}

defaultsort pcpo

definition
  strictify  :: "('a -> 'b) -> 'a -> 'b" where
  "strictify = (Λ f x. if x = ⊥ then ⊥ else f·x)"

text {* results about strictify *}

lemma cont_strictify1: "cont (λf. if x = ⊥ then ⊥ else f·x)"
by (simp add: cont_if)

lemma monofun_strictify2: "monofun (λx. if x = ⊥ then ⊥ else f·x)"
apply (rule monofunI)
apply (auto simp add: monofun_cfun_arg eq_UU_iff [symmetric])
done

(*FIXME: long proof*)
lemma contlub_strictify2: "contlub (λx. if x = ⊥ then ⊥ else f·x)"
apply (rule contlubI)
apply (case_tac "lub (range Y) = ⊥")
apply (drule (1) chain_UU_I)
apply simp
apply (simp del: if_image_distrib)
apply (simp only: contlub_cfun_arg)
apply (rule lub_equal2)
apply (rule chain_mono2 [THEN exE])
apply (erule chain_UU_I_inverse2)
apply (assumption)
apply (rule_tac x=x in exI, clarsimp)
apply (erule chain_monofun)
apply (erule monofun_strictify2 [THEN ch2ch_monofun])
done

lemmas cont_strictify2 =
  monocontlub2cont [OF monofun_strictify2 contlub_strictify2, standard]

lemma strictify_conv_if: "strictify·f·x = (if x = ⊥ then ⊥ else f·x)"
by (unfold strictify_def, simp add: cont_strictify1 cont_strictify2)

lemma strictify1 [simp]: "strictify·f·⊥ = ⊥"
by (simp add: strictify_conv_if)

lemma strictify2 [simp]: "x ≠ ⊥ ==> strictify·f·x = f·x"
by (simp add: strictify_conv_if)

subsection {* Continuous let-bindings *}

definition
  CLet :: "'a -> ('a -> 'b) -> 'b" where
  "CLet = (Λ s f. f·s)"

syntax
  "_CLet" :: "[letbinds, 'a] => 'a" ("(Let (_)/ in (_))" 10)

translations
  "_CLet (_binds b bs) e" == "_CLet b (_CLet bs e)"
  "Let x = a in e" == "CONST CLet·a·(Λ x. e)"

end

Definition of continuous function type

lemma Ex_cont:

  f. cont f

lemma adm_cont:

  adm cont

Syntax for continuous lambda abstraction

Continuous function space is pointed

lemma UU_CFun:

  UU ∈ CFun

lemma Rep_CFun_strict:

  Rep_CFun UU = UU

lemma Abs_CFun_strict:

  (LAM x. UU x) = UU

lemma Rep_CFun_strict1:

  UU·x = UU

lemma inst_cfun_pcpo:

  UU = (LAM x. UU)

Basic properties of continuous functions

lemma Abs_CFun_inverse2:

  cont f ==> Rep_CFun (LAM x. f x) = f

lemma beta_cfun:

  cont f ==> (LAM x. f xu = f u

lemma eta_cfun:

  (LAM x. f·x) = f

lemma expand_cfun_eq:

  (f = g) = (∀x. f·x = g·x)

lemma ext_cfun:

  (!!x. f·x = g·x) ==> f = g

lemma expand_cfun_less:

  f << g = (∀x. f·x << g·x)

lemma less_cfun_ext:

  (!!x. f·x << g·x) ==> f << g

lemma cfun_cong:

  [| f = g; x = y |] ==> f·x = g·y

lemma cfun_fun_cong:

  f = g ==> f·x = g·x

lemma cfun_arg_cong:

  x = y ==> f·x = f·y

Continuity of application

lemma cont_Rep_CFun1:

  contf. f·x)

lemma cont_Rep_CFun2:

  cont (Rep_CFun f)

lemma monofun_Rep_CFun:

  monofun Rep_CFun

lemma contlub_Rep_CFun:

  contlub Rep_CFun

lemma monofun_Rep_CFun1:

  monofunf. f·x)

lemma contlub_Rep_CFun1:

  contlubf. f·x)

lemma monofun_Rep_CFun2:

  monofun (Rep_CFun f)

lemma contlub_Rep_CFun2:

  contlub (Rep_CFun f)

lemma contlub_cfun_arg:

  chain Y ==> f·(Lub Y) = (LUB i. f·(Y i))

lemma cont_cfun_arg:

  chain Y ==> rangei. f·(Y i)) <<| f·(Lub Y)

lemma contlub_cfun_fun:

  chain F ==> Lub F·x = (LUB i. F i·x)

lemma cont_cfun_fun:

  chain F ==> rangei. F i·x) <<| Lub F·x

lemma monofun_cfun_fun:

  f << g ==> f·x << g·x

lemma monofun_cfun_arg:

  x << y ==> f·x << f·y

lemma monofun_cfun:

  [| f << g; x << y |] ==> f·x << g·y

lemma chain_monofun:

  chain Y ==> chaini. f·(Y i))

lemma ch2ch_Rep_CFunR:

  chain Y ==> chaini. f·(Y i))

lemma ch2ch_Rep_CFunL:

  chain F ==> chaini. F i·x)

lemma ch2ch_Rep_CFun:

  [| chain F; chain Y |] ==> chaini. F i·(Y i))

lemma ch2ch_LAM:

  [| !!x. chaini. S i x); !!i. cont (S i) |] ==> chaini. LAM x. S i x)

lemma contlub_cfun:

  [| chain F; chain Y |] ==> (LUB i. F i)·(LUB i. Y i) = (LUB i. F i·(Y i))

lemma cont_cfun:

  [| chain F; chain Y |] ==> rangei. F i·(Y i)) <<| (LUB i. F i)·(LUB i. Y i)

lemma contlub_LAM:

  [| !!x. chaini. F i x); !!i. cont (F i) |]
  ==> (LAM x. LUB i. F i x) = (LUB i. LAM x. F i x)

lemma strictI:

  f·x = UU ==> f·UU = UU

lemma lub_cfun_mono:

  chain F ==> monofunx. LUB i. F i·x)

lemma ex_lub_cfun:

  [| chain F; chain Y |] ==> (LUB j i. F j·(Y i)) = (LUB i j. F j·(Y i))

lemma cont_lub_cfun:

  chain F ==> contx. LUB i. F i·x)

lemma lub_cfun:

  chain F ==> range F <<| (LAM x. LUB i. F i·x)

lemma thelub_cfun:

  chain F ==> Lub F = (LAM x. LUB i. F i·x)

Continuity simplification procedure

lemma cont2cont_Rep_CFun:

  [| cont f; cont t |] ==> contx. f x·(t x))

lemma cont2mono_LAM:

  [| !!x. cont (c1.0 x); !!y. monofunx. c1.0 x y) |]
  ==> monofunx. LAM y. c1.0 x y)

lemma cont2cont_LAM:

  [| !!x. cont (c1.0 x); !!y. contx. c1.0 x y) |]
  ==> contx. LAM y. c1.0 x y)

lemma cont_lemmas1:

  contx. c)
  contx. x)
  cont (Rep_CFun f)
  [| cont f; cont t |] ==> contx. f x·(t x))
  [| !!x. cont (c1.0 x); !!y. contx. c1.0 x y) |]
  ==> contx. LAM y. c1.0 x y)

Miscellaneous

lemma semi_monofun_Abs_CFun:

  [| cont f; cont g; f << g |] ==> (LAM x. f x) << (LAM x. g x)

lemma chfin_Rep_CFunR:

  chain Y ==> ∀s. ∃n. Lub Y·s = Y n·s

lemma adm_chfindom:

  admu. P (u·s))

Continuous injection-retraction pairs

lemma retraction_strict:

  x. f·(g·x) = x ==> f·UU = UU

lemma injection_eq:

  x. f·(g·x) = x ==> (g·x = g·y) = (x = y)

lemma injection_less:

  x. f·(g·x) = x ==> g·x << g·y = x << y

lemma injection_defined_rev:

  [| ∀x. f·(g·x) = x; g·z = UU |] ==> z = UU

lemma injection_defined:

  [| ∀x. f·(g·x) = x; z  UU |] ==> g·z  UU

lemma chfin2chfin:

  y. f·(g·y) = y ==> ∀Y. chain Y --> (∃n. max_in_chain n Y)

lemma flat2flat:

  y. f·(g·y) = y ==> ∀x y. x << y --> x = UUx = y

lemma flat_eqI:

  [| x << y; x  UU |] ==> x = y

lemma flat_codom:

  f·x = c ==> f·UU = UU ∨ (∀z. f·z = c)

Identity and composition

lemma ID1:

  ID·x = x

lemma cfcomp1:

  f oo g = (LAM x. f·(g·x))

lemma cfcomp2:

  (f oo gx = f·(g·x)

lemma cfcomp_strict:

  UU oo f = UU

lemma ID2:

  f oo ID = f

lemma ID3:

  ID oo f = f

lemma assoc_oo:

  f oo g oo h = (f oo g) oo h

Strictified functions

lemma cont_strictify1:

  contf. if x = UU then UU else f·x)

lemma monofun_strictify2:

  monofunx. if x = UU then UU else f·x)

lemma contlub_strictify2:

  contlubx. if x = UU then UU else f·x)

lemma cont_strictify2:

  contx. if x = UU then UU else f·x)

lemma strictify_conv_if:

  strictify·f·x = (if x = UU then UU else f·x)

lemma strictify1:

  strictify·f·UU = UU

lemma strictify2:

  x  UU ==> strictify·f·x = f·x

Continuous let-bindings