Theory Cont

Up to index of Isabelle/HOLCF

theory Cont
imports Pcpo

(*  Title:      HOLCF/Cont.thy
    Author:     Franz Regensburger
*)

header {* 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
*}

defaultsort po

subsection {* Definitions *}

definition
  monofun :: "('a => 'b) => bool"  -- "monotonicity"  where
  "monofun f = (∀x y. x \<sqsubseteq> y --> f x \<sqsubseteq> f y)"

definition
  contlub :: "('a::cpo => 'b::cpo) => bool"  -- "first cont. def" where
  "contlub f = (∀Y. chain Y --> f (\<Squnion>i. Y i) = (\<Squnion>i. f (Y i)))"

definition
  cont :: "('a::cpo => 'b::cpo) => bool"  -- "secnd cont. def" where
  "cont f = (∀Y. chain Y --> range (λi. f (Y i)) <<| f (\<Squnion>i. Y i))"

lemma contlubI:
  "[|!!Y. chain Y ==> f (\<Squnion>i. Y i) = (\<Squnion>i. f (Y i))|] ==> contlub f"
by (simp add: contlub_def)

lemma contlubE: 
  "[|contlub f; chain Y|] ==> f (\<Squnion>i. Y i) = (\<Squnion>i. f (Y i))" 
by (simp add: contlub_def)

lemma contI:
  "[|!!Y. chain Y ==> range (λi. f (Y i)) <<| f (\<Squnion>i. Y i)|] ==> cont f"
by (simp add: cont_def)

lemma contE:
  "[|cont f; chain Y|] ==> range (λi. f (Y i)) <<| f (\<Squnion>i. Y i)"
by (simp add: cont_def)

lemma monofunI: 
  "[|!!x y. x \<sqsubseteq> y ==> f x \<sqsubseteq> f y|] ==> monofun f"
by (simp add: monofun_def)

lemma monofunE: 
  "[|monofun f; x \<sqsubseteq> y|] ==> f x \<sqsubseteq> f y"
by (simp add: monofun_def)


subsection {* @{prop "monofun f ∧ contlub f ≡ cont f"} *}

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 {* left to right: @{prop "monofun f ∧ contlub f ==> cont f"} *}

lemma monocontlub2cont: "[|monofun f; contlub f|] ==> cont f"
apply (rule contI)
apply (rule thelubE)
apply (erule (1) ch2ch_monofun)
apply (erule (1) contlubE [symmetric])
done

text {* first a lemma about binary chains *}

lemma binchain_cont:
  "[|cont f; x \<sqsubseteq> y|] ==> range (λi::nat. f (if i = 0 then x else y)) <<| f y"
apply (subgoal_tac "f (\<Squnion>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 lub_bin_chain [THEN thelubI])
done

text {* right to left: @{prop "cont f ==> monofun f ∧ contlub f"} *}
text {* part1: @{prop "cont f ==> monofun f"} *}

lemma cont2mono: "cont f ==> monofun f"
apply (rule monofunI)
apply (drule (1) binchain_cont)
apply (drule_tac i=0 in is_ub_lub)
apply simp
done

lemmas cont2monofunE = cont2mono [THEN monofunE]

lemmas ch2ch_cont = cont2mono [THEN ch2ch_monofun]

text {* right to left: @{prop "cont f ==> monofun f ∧ contlub f"} *}
text {* part2: @{prop "cont f ==> contlub f"} *}

lemma cont2contlub: "cont f ==> contlub f"
apply (rule contlubI)
apply (rule thelubI [symmetric])
apply (erule (1) contE)
done

lemmas cont2contlubE = cont2contlub [THEN contlubE]

lemma contI2:
  assumes mono: "monofun f"
  assumes less: "!!Y. [|chain Y; chain (λi. f (Y i))|]
     ==> f (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. f (Y i))"
  shows "cont f"
apply (rule monocontlub2cont)
apply (rule mono)
apply (rule contlubI)
apply (rule antisym_less)
apply (rule less, assumption)
apply (erule ch2ch_monofun [OF mono])
apply (rule is_lub_thelub)
apply (erule ch2ch_monofun [OF mono])
apply (rule ub2ub_monofun [OF mono])
apply (rule is_lubD1)
apply (erule cpo_lubI)
done

subsection {* Continuity simproc *}

ML {*
structure Cont2ContData = NamedThmsFun
  ( val name = "cont2cont" val description = "continuity intro rule" )
*}

setup {* Cont2ContData.setup *}

text {*
  Given the term @{term "cont f"}, the procedure tries to construct the
  theorem @{term "cont f == True"}. If this theorem cannot be completely
  solved by the introduction rules, then the procedure returns a
  conditional rewrite rule with the unsolved subgoals as premises.
*}

setup {*
let
  fun solve_cont thy ss t =
    let
      val tr = instantiate' [] [SOME (cterm_of thy t)] Eq_TrueI;
      val rules = Cont2ContData.get (Simplifier.the_context ss);
      val tac = REPEAT_ALL_NEW (match_tac rules);
    in Option.map fst (Seq.pull (tac 1 tr)) end

  val proc =
    Simplifier.simproc @{theory} "cont_proc" ["cont f"] solve_cont;
in
  Simplifier.map_simpset (fn ss => ss addsimprocs [proc])
end
*}

subsection {* Continuity of basic functions *}

text {* The identity function is continuous *}

lemma cont_id [cont2cont]: "cont (λx. x)"
apply (rule contI)
apply (erule cpo_lubI)
done

text {* constant functions are continuous *}

lemma cont_const [cont2cont]: "cont (λx. c)"
apply (rule contI)
apply (rule lub_const)
done

text {* application of functions is continuous *}

lemma cont2cont_apply:
  fixes f :: "'a::cpo => 'b::cpo => 'c::cpo" and t :: "'a => 'b"
  assumes f1: "!!y. cont (λx. f x y)"
  assumes f2: "!!x. cont (λy. f x y)"
  assumes t: "cont (λx. t x)"
  shows "cont (λx. (f x) (t x))"
proof (rule monocontlub2cont [OF monofunI contlubI])
  fix x y :: "'a" assume "x \<sqsubseteq> y"
  then show "f x (t x) \<sqsubseteq> f y (t y)"
    by (auto intro: cont2monofunE [OF f1]
                    cont2monofunE [OF f2]
                    cont2monofunE [OF t]
                    trans_less)
next
  fix Y :: "nat => 'a" assume "chain Y"
  then show "f (\<Squnion>i. Y i) (t (\<Squnion>i. Y i)) = (\<Squnion>i. f (Y i) (t (Y i)))"
    by (simp only: cont2contlubE [OF t]  ch2ch_cont [OF t]
                   cont2contlubE [OF f1] ch2ch_cont [OF f1]
                   cont2contlubE [OF f2] ch2ch_cont [OF f2]
                   diag_lub)
qed

lemma cont2cont_compose:
  "[|cont c; cont (λx. f x)|] ==> cont (λx. c (f x))"
by (rule cont2cont_apply [OF cont_const])

text {* if-then-else is continuous *}

lemma cont_if [simp]:
  "[|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))"
apply (unfold finite_chain_def)
apply (simp add: ch2ch_monofun)
apply (force simp add: max_in_chain_def)
done

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])

lemma chfindom_monofun2cont: "monofun f ==> cont (f::'a::chfin => 'b::cpo)"
apply (rule monocontlub2cont)
apply assumption
apply (rule contlubI)
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 {* some properties of flat *}

lemma flatdom_strict2mono: "f ⊥ = ⊥ ==> monofun (f::'a::flat => 'b::pcpo)"
apply (rule monofunI)
apply (drule ax_flat)
apply auto
done

lemma flatdom_strict2cont: "f ⊥ = ⊥ ==> cont (f::'a::flat => 'b::pcpo)"
by (rule flatdom_strict2mono [THEN chfindom_monofun2cont])

text {* functions with discrete domain *}

lemma cont_discrete_cpo [simp]: "cont (f::'a::discrete_cpo => 'b::cpo)"
apply (rule contI)
apply (drule discrete_chain_const, clarify)
apply (simp add: lub_const)
done

end