Up to index of Isabelle/HOLCF
theory Cfun(* Title: HOLCF/Cfun.thy
ID: $Id: Cfun.thy,v 1.17 2005/07/26 16:28:52 huffman Exp $
Author: Franz Regensburger
Definition of the type -> of continuous functions.
*)
header {* The type of continuous functions *}
theory Cfun
imports Pcpodef
uses ("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
Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("_$_" [999,1000] 999)
(* application *)
Abs_CFun :: "('a => 'b) => ('a -> 'b)" (binder "LAM " 10)
(* abstraction *)
syntax (xsymbols)
"->" :: "[type, type] => type" ("(_ ->/ _)" [1,0]0)
"LAM " :: "[idts, 'a => 'b] => ('a -> 'b)"
("(3Λ_./ _)" [0, 10] 10)
Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_·_)" [999,1000] 999)
syntax (HTML output)
Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_·_)" [999,1000] 999)
subsection {* Class instances *}
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 {* Additional lemma about the isomorphism between
@{typ "'a -> 'b"} and @{term CFun} *}
lemma Abs_CFun_inverse2: "cont f ==> Rep_CFun (Abs_CFun f) = f"
by (simp add: Abs_CFun_inverse CFun_def)
text {* Beta-equality for continuous functions *}
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 ext_cfun: "(!!x. f·x = g·x) ==> f = g"
by (simp add: Rep_CFun_inject [symmetric] ext)
text {* lemmas about application of continuous functions *}
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_CF1L])
lemma cont_Rep_CFun2: "cont (λx. f·x)"
apply (rule_tac P = "cont" in CollectD)
apply (fold CFun_def)
apply (rule Rep_CFun)
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 {* Extensionality wrt. @{term "op <<"} in @{typ "'a -> 'b"} *}
lemma less_cfun_ext: "(!!x. f·x \<sqsubseteq> g·x) ==> f \<sqsubseteq> g"
by (simp add: less_CFun_def less_fun_def)
text {* monotonicity of application *}
lemma monofun_cfun_fun: "f \<sqsubseteq> g ==> f·x \<sqsubseteq> g·x"
by (simp add: less_CFun_def less_fun_def)
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: "[|chain F; chain Y|] ==> chain (λi. (F i)·(Y i))"
apply (rule chainI)
apply (rule monofun_cfun)
apply (erule chainE)
apply (erule chainE)
done
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))"
apply (simp only: contlub_cfun_fun)
apply (simp only: contlub_cfun_arg)
apply (rule diag_lub)
apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun])
apply (erule monofun_Rep_CFun2 [THEN ch2ch_monofun])
done
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
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 ch2ch_Rep_CFunL ch2ch_Rep_CFunR)
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 {* 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 {* for compatibility with old HOLCF-Version *}
lemma inst_cfun_pcpo: "⊥ = (Λ x. ⊥)"
by (simp add: inst_fun_pcpo [symmetric] Abs_CFun_strict)
subsection {* Continuity of application *}
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_CF1L_rev)
done
text {* continuity simplification procedure *}
lemmas cont_lemmas1 =
cont_const cont_id cont_Rep_CFun2 cont2cont_Rep_CFun cont2cont_LAM
use "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)));*)
text {* function application is strict in its first argument *}
lemma Rep_CFun_strict1 [simp]: "⊥·x = ⊥"
by (simp add: Rep_CFun_strict)
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
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 *}
consts
ID :: "'a -> 'a"
cfcomp :: "('b -> 'c) -> ('a -> 'b) -> 'a -> 'c"
syntax "@oo" :: "['b -> 'c, 'a -> 'b] => 'a -> 'c" (infixr "oo" 100)
translations "f1 oo f2" == "cfcomp$f1$f2"
defs
ID_def: "ID ≡ (Λ x. x)"
oo_def: "cfcomp ≡ (Λ f g x. f·(g·x))"
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)
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
consts
Istrictify :: "('a -> 'b) => 'a => 'b"
strictify :: "('a -> 'b) -> 'a -> 'b"
defs
Istrictify_def: "Istrictify f x ≡ if x = ⊥ then ⊥ else f·x"
strictify_def: "strictify ≡ (Λ f x. Istrictify f x)"
text {* results about strictify *}
lemma Istrictify1: "Istrictify f ⊥ = ⊥"
by (simp add: Istrictify_def)
lemma Istrictify2: "x ≠ ⊥ ==> Istrictify f x = f·x"
by (simp add: Istrictify_def)
lemma cont_Istrictify1: "cont (λf. Istrictify f x)"
apply (case_tac "x = ⊥")
apply (simp add: Istrictify1)
apply (simp add: Istrictify2)
done
lemma monofun_Istrictify2: "monofun (λx. Istrictify f x)"
apply (rule monofunI)
apply (simp add: Istrictify_def monofun_cfun_arg)
apply clarify
apply (simp add: eq_UU_iff)
done
lemma contlub_Istrictify2: "contlub (λx. Istrictify f x)"
apply (rule contlubI)
apply (case_tac "lub (range Y) = ⊥")
apply (drule (1) chain_UU_I)
apply (simp add: Istrictify1 thelub_const)
apply (simp add: Istrictify2)
apply (simp add: contlub_cfun_arg)
apply (rule lub_equal2)
apply (rule chain_mono2 [THEN exE])
apply (erule chain_UU_I_inverse2)
apply (assumption)
apply (blast intro: Istrictify2 [symmetric])
apply (erule chain_monofun)
apply (erule monofun_Istrictify2 [THEN ch2ch_monofun])
done
lemmas cont_Istrictify2 =
monocontlub2cont [OF monofun_Istrictify2 contlub_Istrictify2, standard]
lemma strictify1 [simp]: "strictify·f·⊥ = ⊥"
apply (unfold strictify_def)
apply (simp add: cont_Istrictify1 cont_Istrictify2)
apply (rule Istrictify1)
done
lemma strictify2 [simp]: "x ≠ ⊥ ==> strictify·f·x = f·x"
apply (unfold strictify_def)
apply (simp add: cont_Istrictify1 cont_Istrictify2)
apply (erule Istrictify2)
done
lemma strictify_conv_if: "strictify·f·x = (if x = ⊥ then ⊥ else f·x)"
by simp
end
lemma Ex_cont:
∃f. cont f
lemma adm_cont:
adm cont
lemma UU_CFun:
UU ∈ CFun
lemmas Rep_CFun_strict:
Rep_CFun UU = UU
lemmas Rep_CFun_strict:
Rep_CFun UU = UU
lemmas Abs_CFun_strict:
Abs_CFun UU = UU
lemmas Abs_CFun_strict:
Abs_CFun UU = UU
lemma Abs_CFun_inverse2:
cont f ==> Rep_CFun (Abs_CFun f) = f
lemma beta_cfun:
cont f ==> (LAM x. f x)·u = f u
lemma eta_cfun:
(LAM x. f·x) = f
lemma ext_cfun:
(!!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
lemma cont_Rep_CFun1:
cont (%f. f·x)
lemma cont_Rep_CFun2:
cont (Rep_CFun f)
lemmas monofun_Rep_CFun:
monofun Rep_CFun
lemmas monofun_Rep_CFun:
monofun Rep_CFun
lemmas contlub_Rep_CFun:
contlub Rep_CFun
lemmas contlub_Rep_CFun:
contlub Rep_CFun
lemmas monofun_Rep_CFun1:
monofun (%f. f·x)
lemmas monofun_Rep_CFun1:
monofun (%f. f·x)
lemmas contlub_Rep_CFun1:
contlub (%f. f·x)
lemmas contlub_Rep_CFun1:
contlub (%f. f·x)
lemmas monofun_Rep_CFun2:
monofun (Rep_CFun f)
lemmas monofun_Rep_CFun2:
monofun (Rep_CFun f)
lemmas contlub_Rep_CFun2:
contlub (Rep_CFun f)
lemmas contlub_Rep_CFun2:
contlub (Rep_CFun f)
lemma contlub_cfun_arg:
chain Y ==> f·(lub (range Y)) = (LUB i. f·(Y i))
lemma cont_cfun_arg:
chain Y ==> range (%i. f·(Y i)) <<| f·(lub (range Y))
lemma contlub_cfun_fun:
chain F ==> lub (range F)·x = (LUB i. F i·x)
lemma cont_cfun_fun:
chain F ==> range (%i. F i·x) <<| lub (range F)·x
lemma less_cfun_ext:
(!!x. f·x << g·x) ==> f << g
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 ==> chain (%i. f·(Y i))
lemma ch2ch_Rep_CFunR:
chain Y ==> chain (%i. f·(Y i))
lemma ch2ch_Rep_CFunL:
chain F ==> chain (%i. F i·x)
lemma ch2ch_Rep_CFun:
[| chain F; chain Y |] ==> chain (%i. F i·(Y i))
lemma contlub_cfun:
[| chain F; chain Y |] ==> lub (range F)·(lub (range Y)) = (LUB i. F i·(Y i))
lemma cont_cfun:
[| chain F; chain Y |] ==> range (%i. F i·(Y i)) <<| lub (range F)·(lub (range Y))
lemma strictI:
f·x = UU ==> f·UU = UU
lemma lub_cfun_mono:
chain F ==> monofun (%x. LUB i. F i·x)
lemma ex_lub_cfun:
[| chain F; chain Y |] ==> (LUB j. LUB i. F j·(Y i)) = (LUB i. LUB j. F j·(Y i))
lemma cont_lub_cfun:
chain F ==> cont (%x. 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 (range F) = (LAM x. LUB i. F i·x)
lemma semi_monofun_Abs_CFun:
[| cont f; cont g; f << g |] ==> Abs_CFun f << Abs_CFun g
lemma inst_cfun_pcpo:
UU = (LAM x. UU)
lemma cont2cont_Rep_CFun:
[| cont f; cont t |] ==> cont (%x. f x·(t x))
lemma cont2mono_LAM:
[| !!x. cont (c1.0 x); !!y. monofun (%x. c1.0 x y) |] ==> monofun (%x. LAM y. c1.0 x y)
lemma cont2cont_LAM:
[| !!x. cont (c1.0 x); !!y. cont (%x. c1.0 x y) |] ==> cont (%x. LAM y. c1.0 x y)
lemmas cont_lemmas1:
cont (%x. c)
cont (%x. x)
cont (Rep_CFun f)
[| cont f; cont t |] ==> cont (%x. f x·(t x))
[| !!x. cont (c1.0 x); !!y. cont (%x. c1.0 x y) |] ==> cont (%x. LAM y. c1.0 x y)
lemmas cont_lemmas1:
cont (%x. c)
cont (%x. x)
cont (Rep_CFun f)
[| cont f; cont t |] ==> cont (%x. f x·(t x))
[| !!x. cont (c1.0 x); !!y. cont (%x. c1.0 x y) |] ==> cont (%x. LAM y. c1.0 x y)
lemma Rep_CFun_strict1:
UU·x = UU
lemma chfin_Rep_CFunR:
chain Y ==> ∀s. ∃n. lub (range Y)·s = Y n·s
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 = UU ∨ x = y
lemma flat_eqI:
[| x << y; x ≠ UU |] ==> x = y
lemma flat_codom:
f·x = c ==> f·UU = UU ∨ (∀z. f·z = c)
lemma ID1:
ID·x = x
lemma cfcomp1:
f oo g = (LAM x. f·(g·x))
lemma cfcomp2:
(f oo g)·x = f·(g·x)
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
lemma Istrictify1:
Istrictify f UU = UU
lemma Istrictify2:
x ≠ UU ==> Istrictify f x = f·x
lemma cont_Istrictify1:
cont (%f. Istrictify f x)
lemma monofun_Istrictify2:
monofun (Istrictify f)
lemma contlub_Istrictify2:
contlub (Istrictify f)
lemmas cont_Istrictify2:
cont (Istrictify f)
lemmas cont_Istrictify2:
cont (Istrictify f)
lemma strictify1:
strictify·f·UU = UU
lemma strictify2:
x ≠ UU ==> strictify·f·x = f·x
lemma strictify_conv_if:
strictify·f·x = (if x = UU then UU else f·x)