Theory Refine_Monadic.Refine_Pfun
section ‹Partial Function Package Setup›
theory Refine_Pfun
imports Refine_Basic Refine_Det
begin
text ‹
  In this theory, we set up the partial function package to be used 
  with our refinement framework.
›
subsection ‹Nondeterministic Result Monad›
interpretation nrec:
  partial_function_definitions "(≤)" "Sup::'a nres set ⇒ 'a nres"
  by unfold_locales (auto simp add: Sup_upper Sup_least)
lemma nrec_admissible: "nrec.admissible (λ(f::'a ⇒ 'b nres).
  (∀x0. f x0 ≤ SPEC (P x0)))"
  apply (rule ccpo.admissibleI)
  apply (unfold fun_lub_def)
  apply (intro allI impI)
  apply (rule Sup_least)
  apply auto
  done
    
declaration ‹Partial_Function.init "nrec" @{term nrec.fixp_fun}
  @{term nrec.mono_body} @{thm nrec.fixp_rule_uc} @{thm nrec.fixp_induct_uc}
   (NONE)›
lemma bind_mono_pfun[partial_function_mono]:
  fixes C :: "'a ⇒ ('b ⇒ 'c nres) ⇒ ('d nres)"
  shows
  "⟦ monotone (fun_ord (≤)) (≤) B; 
    ⋀y. monotone (fun_ord (≤)) (≤) (λf. C y f) ⟧ ⟹ 
     monotone (fun_ord (≤)) (≤) (λf. bind (B f) (λy. C y f))"
  apply rule
  apply (rule Refine_Basic.bind_mono)
  apply (blast dest: monotoneD)+
  done
subsection ‹Deterministic Result Monad›
interpretation drec:
  partial_function_definitions "(≤)" "Sup::'a dres set ⇒ 'a dres"
  by unfold_locales (auto simp add: Sup_upper Sup_least)
lemma drec_admissible: "drec.admissible (λ(f::'a ⇒ 'b dres).
  (∀x. P x ⟶ 
    (f x ≠ dFAIL ∧ 
    (∀r. f x = dRETURN r ⟶ Q x r))))"
proof -
  have [simp]: "fun_ord ((≤) ::'b dres ⇒ _ ⇒ _) = (≤)"
    apply (intro ext)
    unfolding fun_ord_def le_fun_def
    by (rule refl)
  have [simp]: "⋀A x. {y. ∃f∈A. y = f x} = (λf. f x)`A" by auto
  show ?thesis
    apply (rule ccpo.admissibleI)
    apply (unfold fun_lub_def)
    apply clarsimp
    apply (drule_tac x=x in point_chainI)
    apply (erule dres_Sup_chain_cases)
    apply auto
    apply (metis (poly_guards_query) SUP_bot_conv(1))
    apply (metis (poly_guards_query) SUP_bot_conv(1))
    apply metis
    done
qed
declaration ‹Partial_Function.init "drec" @{term drec.fixp_fun}
  @{term drec.mono_body} @{thm drec.fixp_rule_uc} @{thm drec.fixp_induct_uc} 
  NONE›
lemma drec_bind_mono_pfun[partial_function_mono]:
  fixes C :: "'a ⇒ ('b ⇒ 'c dres) ⇒ ('d dres)"
  shows
  "⟦ monotone (fun_ord (≤)) (≤) B; 
    ⋀y. monotone (fun_ord (≤)) (≤) (λf. C y f) ⟧ ⟹ 
     monotone (fun_ord (≤)) (≤) (λf. dbind (B f) (λy. C y f))"
  apply rule
  apply (rule dbind_mono)
  apply (blast dest: monotoneD)+
  done
end