Theory Pcpo

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

section ‹Classes cpo and pcpo›

theory Pcpo
  imports Porder
begin

subsection ‹Complete partial orders›

text ‹The class cpo of chain complete partial orders›

class cpo = po +
  assumes cpo: "chain S  x. range S <<| x"
begin

text ‹in cpo's everthing equal to THE lub has lub properties for every chain›

lemma cpo_lubI: "chain S  range S <<| (i. S i)"
  by (fast dest: cpo elim: is_lub_lub)

lemma thelubE: "chain S; (i. S i) = l  range S <<| l"
  by (blast dest: cpo intro: is_lub_lub)

text ‹Properties of the lub›

lemma is_ub_thelub: "chain S  S x  (i. S i)"
  by (blast dest: cpo intro: is_lub_lub [THEN is_lub_rangeD1])

lemma is_lub_thelub: "chain S; range S <| x  (i. S i)  x"
  by (blast dest: cpo intro: is_lub_lub [THEN is_lubD2])

lemma lub_below_iff: "chain S  (i. S i)  x  (i. S i  x)"
  by (simp add: is_lub_below_iff [OF cpo_lubI] is_ub_def)

lemma lub_below: "chain S; i. S i  x  (i. S i)  x"
  by (simp add: lub_below_iff)

lemma below_lub: "chain S; x  S i  x  (i. S i)"
  by (erule below_trans, erule is_ub_thelub)

lemma lub_range_mono: "range X  range Y; chain Y; chain X  (i. X i)  (i. Y i)"
  apply (erule lub_below)
  apply (subgoal_tac "j. X i = Y j")
   apply clarsimp
   apply (erule is_ub_thelub)
  apply auto
  done

lemma lub_range_shift: "chain Y  (i. Y (i + j)) = (i. Y i)"
  apply (rule below_antisym)
   apply (rule lub_range_mono)
     apply fast
    apply assumption
   apply (erule chain_shift)
  apply (rule lub_below)
   apply assumption
  apply (rule_tac i="i" in below_lub)
   apply (erule chain_shift)
  apply (erule chain_mono)
  apply (rule le_add1)
  done

lemma maxinch_is_thelub: "chain Y  max_in_chain i Y = ((i. Y i) = Y i)"
  apply (rule iffI)
   apply (fast intro!: lub_eqI lub_finch1)
  apply (unfold max_in_chain_def)
  apply (safe intro!: below_antisym)
   apply (fast elim!: chain_mono)
  apply (drule sym)
  apply (force elim!: is_ub_thelub)
  done

text ‹the ⊑› relation between two chains is preserved by their lubs›

lemma lub_mono: "chain X; chain Y; i. X i  Y i  (i. X i)  (i. Y i)"
  by (fast elim: lub_below below_lub)

text ‹the = relation between two chains is preserved by their lubs›

lemma lub_eq: "(i. X i = Y i)  (i. X i) = (i. Y i)"
  by simp

lemma ch2ch_lub:
  assumes 1: "j. chain (λi. Y i j)"
  assumes 2: "i. chain (λj. Y i j)"
  shows "chain (λi. j. Y i j)"
  apply (rule chainI)
  apply (rule lub_mono [OF 2 2])
  apply (rule chainE [OF 1])
  done

lemma diag_lub:
  assumes 1: "j. chain (λi. Y i j)"
  assumes 2: "i. chain (λj. Y i j)"
  shows "(i. j. Y i j) = (i. Y i i)"
proof (rule below_antisym)
  have 3: "chain (λi. Y i i)"
    apply (rule chainI)
    apply (rule below_trans)
     apply (rule chainE [OF 1])
    apply (rule chainE [OF 2])
    done
  have 4: "chain (λi. j. Y i j)"
    by (rule ch2ch_lub [OF 1 2])
  show "(i. j. Y i j)  (i. Y i i)"
    apply (rule lub_below [OF 4])
    apply (rule lub_below [OF 2])
    apply (rule below_lub [OF 3])
    apply (rule below_trans)
     apply (rule chain_mono [OF 1 max.cobounded1])
    apply (rule chain_mono [OF 2 max.cobounded2])
    done
  show "(i. Y i i)  (i. j. Y i j)"
    apply (rule lub_mono [OF 3 4])
    apply (rule is_ub_thelub [OF 2])
    done
qed

lemma ex_lub:
  assumes 1: "j. chain (λi. Y i j)"
  assumes 2: "i. chain (λj. Y i j)"
  shows "(i. j. Y i j) = (j. i. Y i j)"
  by (simp add: diag_lub 1 2)

end


subsection ‹Pointed cpos›

text ‹The class pcpo of pointed cpos›

class pcpo = cpo +
  assumes least: "x. y. x  y"
begin

definition bottom :: "'a"  ("")
  where "bottom = (THE x. y. x  y)"

lemma minimal [iff]: "  x"
  unfolding bottom_def
  apply (rule the1I2)
   apply (rule ex_ex1I)
    apply (rule least)
   apply (blast intro: below_antisym)
  apply simp
  done

end

text ‹Old "UU" syntax:›

syntax UU :: logic
translations "UU"  "CONST bottom"

text ‹Simproc to rewrite term = x to termx = .›
setup Reorient_Proc.add (fn Const_bottom _ => true | _ => false)
simproc_setup reorient_bottom (" = x") = K Reorient_Proc.proc

text ‹useful lemmas about term

lemma below_bottom_iff [simp]: "x    x = "
  by (simp add: po_eq_conv)

lemma eq_bottom_iff: "x =   x  "
  by simp

lemma bottomI: "x    x = "
  by (subst eq_bottom_iff)

lemma lub_eq_bottom_iff: "chain Y  (i. Y i) =   (i. Y i = )"
  by (simp only: eq_bottom_iff lub_below_iff)


subsection ‹Chain-finite and flat cpos›

text ‹further useful classes for HOLCF domains›

class chfin = po +
  assumes chfin: "chain Y  n. max_in_chain n Y"
begin

subclass cpo
  apply standard
  apply (frule chfin)
  apply (blast intro: lub_finch1)
  done

lemma chfin2finch: "chain Y  finite_chain Y"
  by (simp add: chfin finite_chain_def)

end

class flat = pcpo +
  assumes ax_flat: "x  y  x =   x = y"
begin

subclass chfin
proof
  fix Y
  assume *: "chain Y"
  show "n. max_in_chain n Y"
    apply (unfold max_in_chain_def)
    apply (cases "i. Y i = ")
     apply simp
    apply simp
    apply (erule exE)
    apply (rule_tac x="i" in exI)
    apply clarify
    using * apply (blast dest: chain_mono ax_flat)
    done
qed

lemma flat_below_iff: "x  y  x =   x = y"
  by (safe dest!: ax_flat)

lemma flat_eq: "a    a  b = (a = b)"
  by (safe dest!: ax_flat)

end

subsection ‹Discrete cpos›

class discrete_cpo = below +
  assumes discrete_cpo [simp]: "x  y  x = y"
begin

subclass po
  by standard simp_all

text ‹In a discrete cpo, every chain is constant›

lemma discrete_chain_const:
  assumes S: "chain S"
  shows "x. S = (λi. x)"
proof (intro exI ext)
  fix i :: nat
  from S le0 have "S 0  S i" by (rule chain_mono)
  then have "S 0 = S i" by simp
  then show "S i = S 0" by (rule sym)
qed

subclass chfin
proof
  fix S :: "nat  'a"
  assume S: "chain S"
  then have "x. S = (λi. x)"
    by (rule discrete_chain_const)
  then have "max_in_chain 0 S"
    by (auto simp: max_in_chain_def)
  then show "i. max_in_chain i S" ..
qed

end

end