Theory Bifinite

Up to index of Isabelle/HOLCF

theory Bifinite
imports Deflation

(*  Title:      HOLCF/Bifinite.thy
    Author:     Brian Huffman
*)

header {* Bifinite domains and approximation *}

theory Bifinite
imports Deflation
begin

subsection {* Omega-profinite and bifinite domains *}

class profinite =
  fixes approx :: "nat => 'a -> 'a"
  assumes chain_approx [simp]: "chain approx"
  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 approx_less: "approx i·x \<sqsubseteq> x"
proof -
  have "chain (λi. approx i·x)" by simp
  hence "approx i·x \<sqsubseteq> (\<Squnion>i. approx i·x)" by (rule is_ub_thelub)
  thus "approx i·x \<sqsubseteq> x" by simp
qed

lemma finite_deflation_approx: "finite_deflation (approx i)"
proof
  fix x :: 'a
  show "approx i·(approx i·x) = approx i·x"
    by (rule approx_idem)
  show "approx i·x \<sqsubseteq> x"
    by (rule approx_less)
  show "finite {x. approx i·x = x}"
    by (rule finite_fixes_approx)
qed

interpretation approx: finite_deflation "approx i"
by (rule finite_deflation_approx)

lemma (in deflation) deflation: "deflation d" ..

lemma deflation_approx: "deflation (approx i)"
by (rule approx.deflation)

lemma lub_approx [simp]: "(\<Squnion>i. approx i) = (Λ x. x)"
by (rule ext_cfun, simp add: contlub_cfun_fun)

lemma approx_strict [simp]: "approx i·⊥ = ⊥"
by (rule UU_I, rule approx_less)

lemma approx_approx1:
  "i ≤ j ==> approx i·(approx j·x) = approx i·x"
apply (rule deflation_less_comp1 [OF deflation_approx deflation_approx])
apply (erule chain_mono [OF chain_approx])
done

lemma approx_approx2:
  "j ≤ i ==> approx i·(approx j·x) = approx j·x"
apply (rule deflation_less_comp2 [OF deflation_approx deflation_approx])
apply (erule chain_mono [OF chain_approx])
done

lemma approx_approx [simp]:
  "approx i·(approx j·x) = approx (min i j)·x"
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 finite_image_approx: "finite ((λx. approx n·x) ` A)"
by (rule approx.finite_image)

lemma finite_range_approx: "finite (range (λx. approx i·x))"
by (rule approx.finite_range)

lemma compact_approx [simp]: "compact (approx n·x)"
by (rule approx.compact)

lemma profinite_compact_eq_approx: "compact x ==> ∃i. approx i·x = x"
by (rule admD2, simp_all)

lemma profinite_compact_iff: "compact x <-> (∃n. approx n·x = x)"
 apply (rule iffI)
  apply (erule profinite_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"
proof -
  have "P (\<Squnion>n. approx n·x)"
    by (rule admD [OF adm], simp, simp add: P)
  thus "P x" by simp
qed

lemma profinite_less_ext: "(!!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_cfun_lemma:
  assumes a: "finite (range (λx. a·x))"
  assumes b: "finite (range (λy. b·y))"
  shows "finite (range (λf. Λ x. b·(f·(a·x))))"  (is "finite (range ?h)")
proof (rule finite_imageD)
  let ?f = "λg. range (λx. (a·x, g·x))"
  show "finite (?f ` range ?h)"
  proof (rule finite_subset)
    let ?B = "Pow (range (λx. a·x) × range (λy. b·y))"
    show "?f ` range ?h ⊆ ?B"
      by clarsimp
    show "finite ?B"
      by (simp add: a b)
  qed
  show "inj_on ?f (range ?h)"
  proof (rule inj_onI, rule ext_cfun, clarsimp)
    fix x f g
    assume "range (λx. (a·x, b·(f·(a·x)))) = range (λx. (a·x, b·(g·(a·x))))"
    hence "range (λx. (a·x, b·(f·(a·x)))) ⊆ range (λx. (a·x, b·(g·(a·x))))"
      by (rule equalityD1)
    hence "(a·x, b·(f·(a·x))) ∈ range (λx. (a·x, b·(g·(a·x))))"
      by (simp add: subset_eq)
    then obtain y where "(a·x, b·(f·(a·x))) = (a·y, b·(g·(a·y)))"
      by (rule rangeE)
    thus "b·(f·(a·x)) = b·(g·(a·x))"
      by clarsimp
  qed
qed

instantiation "->" :: (profinite, profinite) profinite
begin

definition
  approx_cfun_def:
    "approx = (λn. Λ f x. approx n·(f·(approx n·x)))"

instance proof
  show "chain (approx :: nat => ('a -> 'b) -> ('a -> 'b))"
    unfolding approx_cfun_def by simp
next
  fix x :: "'a -> 'b"
  show "(\<Squnion>i. approx i·x) = x"
    unfolding approx_cfun_def
    by (simp add: lub_distribs eta_cfun)
next
  fix i :: nat and x :: "'a -> 'b"
  show "approx i·(approx i·x) = approx i·x"
    unfolding approx_cfun_def by simp
next
  fix i :: nat
  show "finite {x::'a -> 'b. approx i·x = x}"
    apply (rule finite_range_imp_finite_fixes)
    apply (simp add: approx_cfun_def)
    apply (intro finite_range_cfun_lemma finite_range_approx)
    done
qed

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