(* Title: HOLCF/Bifinite.thy ID: $Id: Bifinite.thy,v 1.6 2008/05/19 21:49:20 huffman Exp $ Author: Brian Huffman *) header {* Bifinite domains and approximation *} theory Bifinite imports Cfun begin subsection {* Omega-profinite and bifinite domains *} class profinite = cpo + fixes approx :: "nat => 'a -> 'a" assumes chain_approx_app: "chain (λi. approx i·x)" assumes lub_approx_app [simp]: "(\<Squnion>i. approx i·x) = x" assumes approx_idem: "approx i·(approx i·x) = approx i·x" assumes finite_fixes_approx: "finite {x. approx i·x = x}" class bifinite = profinite + pcpo lemma finite_range_imp_finite_fixes: "finite {x. ∃y. x = f y} ==> finite {x. f x = x}" apply (subgoal_tac "{x. f x = x} ⊆ {x. ∃y. x = f y}") apply (erule (1) finite_subset) apply (clarify, erule subst, rule exI, rule refl) done lemma chain_approx [simp]: "chain (approx :: nat => 'a::profinite -> 'a)" apply (rule chainI) apply (rule less_cfun_ext) apply (rule chainE) apply (rule chain_approx_app) done lemma lub_approx [simp]: "(\<Squnion>i. approx i) = (Λ(x::'a::profinite). x)" by (rule ext_cfun, simp add: contlub_cfun_fun) lemma approx_less: "approx i·x \<sqsubseteq> (x::'a::profinite)" apply (subgoal_tac "approx i·x \<sqsubseteq> (\<Squnion>i. approx i·x)", simp) apply (rule is_ub_thelub, simp) done lemma approx_strict [simp]: "approx i·(⊥::'a::bifinite) = ⊥" by (rule UU_I, rule approx_less) lemma approx_approx1: "i ≤ j ==> approx i·(approx j·x) = approx i·(x::'a::profinite)" apply (rule antisym_less) apply (rule monofun_cfun_arg [OF approx_less]) apply (rule sq_ord_eq_less_trans [OF approx_idem [symmetric]]) apply (rule monofun_cfun_arg) apply (rule monofun_cfun_fun) apply (erule chain_mono [OF chain_approx]) done lemma approx_approx2: "j ≤ i ==> approx i·(approx j·x) = approx j·(x::'a::profinite)" apply (rule antisym_less) apply (rule approx_less) apply (rule sq_ord_eq_less_trans [OF approx_idem [symmetric]]) apply (rule monofun_cfun_fun) apply (erule chain_mono [OF chain_approx]) done lemma approx_approx [simp]: "approx i·(approx j·x) = approx (min i j)·(x::'a::profinite)" apply (rule_tac x=i and y=j in linorder_le_cases) apply (simp add: approx_approx1 min_def) apply (simp add: approx_approx2 min_def) done lemma idem_fixes_eq_range: "∀x. f (f x) = f x ==> {x. f x = x} = {y. ∃x. y = f x}" by (auto simp add: eq_sym_conv) lemma finite_approx: "finite {y::'a::profinite. ∃x. y = approx n·x}" using finite_fixes_approx by (simp add: idem_fixes_eq_range) lemma finite_range_approx: "finite (range (λx::'a::profinite. approx n·x))" by (simp add: image_def finite_approx) lemma compact_approx [simp]: fixes x :: "'a::profinite" shows "compact (approx n·x)" proof (rule compactI2) fix Y::"nat => 'a" assume Y: "chain Y" have "finite_chain (λi. approx n·(Y i))" proof (rule finite_range_imp_finch) show "chain (λi. approx n·(Y i))" using Y by simp have "range (λi. approx n·(Y i)) ⊆ {x. approx n·x = x}" by clarsimp thus "finite (range (λi. approx n·(Y i)))" using finite_fixes_approx by (rule finite_subset) qed hence "∃j. (\<Squnion>i. approx n·(Y i)) = approx n·(Y j)" by (simp add: finite_chain_def maxinch_is_thelub Y) then obtain j where j: "(\<Squnion>i. approx n·(Y i)) = approx n·(Y j)" .. assume "approx n·x \<sqsubseteq> (\<Squnion>i. Y i)" hence "approx n·(approx n·x) \<sqsubseteq> approx n·(\<Squnion>i. Y i)" by (rule monofun_cfun_arg) hence "approx n·x \<sqsubseteq> (\<Squnion>i. approx n·(Y i))" by (simp add: contlub_cfun_arg Y) hence "approx n·x \<sqsubseteq> approx n·(Y j)" using j by simp hence "approx n·x \<sqsubseteq> Y j" using approx_less by (rule trans_less) thus "∃j. approx n·x \<sqsubseteq> Y j" .. qed lemma bifinite_compact_eq_approx: fixes x :: "'a::profinite" assumes x: "compact x" shows "∃i. approx i·x = x" proof - have chain: "chain (λi. approx i·x)" by simp have less: "x \<sqsubseteq> (\<Squnion>i. approx i·x)" by simp obtain i where i: "x \<sqsubseteq> approx i·x" using compactD2 [OF x chain less] .. with approx_less have "approx i·x = x" by (rule antisym_less) thus "∃i. approx i·x = x" .. qed lemma bifinite_compact_iff: "compact (x::'a::profinite) = (∃n. approx n·x = x)" apply (rule iffI) apply (erule bifinite_compact_eq_approx) apply (erule exE) apply (erule subst) apply (rule compact_approx) done lemma approx_induct: assumes adm: "adm P" and P: "!!n x. P (approx n·x)" shows "P (x::'a::profinite)" proof - have "P (\<Squnion>n. approx n·x)" by (rule admD [OF adm], simp, simp add: P) thus "P x" by simp qed lemma bifinite_less_ext: fixes x y :: "'a::profinite" shows "(!!i. approx i·x \<sqsubseteq> approx i·y) ==> x \<sqsubseteq> y" apply (subgoal_tac "(\<Squnion>i. approx i·x) \<sqsubseteq> (\<Squnion>i. approx i·y)", simp) apply (rule lub_mono, simp, simp, simp) done subsection {* Instance for continuous function space *} lemma finite_range_lemma: fixes h :: "'a::cpo -> 'b::cpo" fixes k :: "'c::cpo -> 'd::cpo" shows "[|finite {y. ∃x. y = h·x}; finite {y. ∃x. y = k·x}|] ==> finite {g. ∃f. g = (Λ x. k·(f·(h·x)))}" apply (rule_tac f="λg. {(h·x, y) |x y. y = g·x}" in finite_imageD) apply (rule_tac B="Pow ({y. ∃x. y = h·x} × {y. ∃x. y = k·x})" in finite_subset) apply (rule image_subsetI) apply (clarsimp, fast) apply simp apply (rule inj_onI) apply (clarsimp simp add: expand_set_eq) apply (rule ext_cfun, simp) apply (drule_tac x="h·x" in spec) apply (drule_tac x="k·(f·(h·x))" in spec) apply (drule iffD1, fast) apply clarsimp done instantiation "->" :: (profinite, profinite) profinite begin definition approx_cfun_def: "approx = (λn. Λ f x. approx n·(f·(approx n·x)))" instance apply (intro_classes, unfold approx_cfun_def) apply simp apply (simp add: lub_distribs eta_cfun) apply simp apply simp apply (rule finite_range_imp_finite_fixes) apply (intro finite_range_lemma finite_approx) done end instance "->" :: (profinite, bifinite) bifinite .. lemma approx_cfun: "approx n·f·x = approx n·(f·(approx n·x))" by (simp add: approx_cfun_def) end
lemma finite_range_imp_finite_fixes:
finite {x. ∃y. x = f y} ==> finite {x. f x = x}
lemma chain_approx:
chain approx
lemma lub_approx:
(LUB i. approx i) = (LAM x. x)
lemma approx_less:
approx i·x << x
lemma approx_strict:
approx i·UU = UU
lemma approx_approx1:
i ≤ j ==> approx i·(approx j·x) = approx i·x
lemma approx_approx2:
j ≤ i ==> approx i·(approx j·x) = approx j·x
lemma approx_approx:
approx i·(approx j·x) = approx (min i j)·x
lemma idem_fixes_eq_range:
∀x. f (f x) = f x ==> {x. f x = x} = {y. ∃x. y = f x}
lemma finite_approx:
finite {y. ∃x. y = approx n·x}
lemma finite_range_approx:
finite (range (Rep_CFun (approx n)))
lemma compact_approx:
compact (approx n·x)
lemma bifinite_compact_eq_approx:
compact x ==> ∃i. approx i·x = x
lemma bifinite_compact_iff:
compact x = (∃n. approx n·x = x)
lemma approx_induct:
[| adm P; !!n x. P (approx n·x) |] ==> P x
lemma bifinite_less_ext:
(!!i. approx i·x << approx i·y) ==> x << y
lemma finite_range_lemma:
[| finite {y. ∃x. y = h·x}; finite {y. ∃x. y = k·x} |]
==> finite {g. ∃f. g = (LAM x. k·(f·(h·x)))}
lemma approx_cfun:
approx n·f·x = approx n·(f·(approx n·x))