Theory Refine_Monadic.RefineG_Domain
section ‹General Domain Theory›
theory RefineG_Domain
imports "../Refine_Misc"
begin
  subsection ‹General Order Theory Tools›
  lemma chain_f_apply: "Complete_Partial_Order.chain (fun_ord le) F
    ⟹ Complete_Partial_Order.chain le {y . ∃f∈F. y = f x}"
    unfolding Complete_Partial_Order.chain_def 
    by (auto simp: fun_ord_def)
  lemma ccpo_lift:
    assumes "class.ccpo lub le lt"
    shows "class.ccpo (fun_lub lub) (fun_ord le) (mk_less (fun_ord le))"
  proof -
    interpret ccpo: ccpo lub le lt by fact
    show ?thesis
      apply unfold_locales
      apply (simp_all only: mk_less_def fun_ord_def fun_lub_def)
      apply simp
      using ccpo.order_trans apply blast
      using ccpo.order.antisym apply blast
      using ccpo.ccpo_Sup_upper apply (blast intro: chain_f_apply)
      using ccpo.ccpo_Sup_least apply (blast intro: chain_f_apply)
      done
  qed
  
  
  lemma fun_lub_simps[simp]: 
    "fun_lub lub {} = (λx. lub {})"
    "fun_lub lub {f} = (λx. lub {f x})"
    unfolding fun_lub_def by auto
  subsection ‹Flat Ordering›
  lemma flat_ord_chain_cases: 
    assumes A: "Complete_Partial_Order.chain (flat_ord b) C"
    obtains "C={}" 
    | "C={b}" 
    | x where "x≠b" and "C={x}" 
    | x where "x≠b" and "C={b,x}"
  proof -
    have "∃m1 m2. C ⊆ {m1,m2} ∧ (m1 = b ∨ m2 = b)"
      apply (rule ccontr)
    proof clarsimp
      assume "∀m1 m2. C ⊆ {m1, m2} ⟶ m1≠b ∧ m2≠b"
      then obtain m1 m2 where "m1∈C" "m2∈C" 
        "m1≠m2" "m1≠b" "m2≠b"
        by blast
      with A show False
        unfolding Complete_Partial_Order.chain_def flat_ord_def
        by auto
    qed
    then obtain m where "C ⊆ {m,b}" by blast
    thus ?thesis 
      apply (cases "m=b") 
      using that
      apply blast+
      done
  qed
    
  lemma flat_lub_simps[simp]:
    "flat_lub b {} = b"
    "flat_lub b {x} = x"
    "flat_lub b (insert b X) = flat_lub b X"
    unfolding flat_lub_def
    by auto
  lemma flat_ord_simps[simp]:
    "flat_ord b b x" 
    by (simp add: flat_ord_def)