Theory Refine_Pfun

Up to index of Isabelle/HOL/Collections/Refine_Monadic

theory Refine_Pfun
imports Refine_Basic Refine_Det
header {* \isaheader{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 "op ≤" "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[OF nrec.ccpo])
apply (unfold fun_lub_def)
apply (intro allI impI)
apply (rule Sup_least)
apply auto
done

(*
lemma fixp_induct_nrec:
fixes F :: "'c => 'c" and
U :: "'c => 'b => 'a nres" and
C :: "('b => 'a nres) => 'c" and
P :: "'b => 'a => bool"
assumes mono: "!!x. nrec_mono (λf. U (F (C f)) x)"
assumes eq: "f ≡ C (nrec_ffix (λf. U (F (C f))))"
assumes inverse2: "!!f. U (C f) = f"
assumes step: "!!f x. (!!x. U f x ≤ SPEC (P x)) ==> U (F f) x ≤ SPEC (P x)"
assumes defined: "RETURN y ≤ U f x"
shows "U f x ≤ SPEC (P x)"
using step defined
nrec.fixp_induct_uc[of U F C, OF mono eq inverse2 nrec_admissible]
by blast

lemma fixp_induct_nrec':
fixes F :: "'c => 'c" and
U :: "'c => 'b => 'a nres" and
C :: "('b => 'a nres) => 'c" and
P :: "'b => 'a => bool"
assumes mono: "!!x. nrec_mono (λf. U (F (C f)) x)"
assumes eq: "f ≡ C (nrec_ffix (λf. U (F (C f))))"
assumes inverse2: "!!f. U (C f) = f"
assumes step: "!!f x0. (!!x0. U f x0 ≤ SPEC (P x0))
==> U (F f) x0 ≤ SPEC (P x0)"
assumes defined: "RETURN y ≤ U f x"
shows "P x y"
proof -
note defined
also have "∀x0. U f x0 ≤ SPEC (P x0)"
apply (rule nrec.fixp_induct_uc[of U F C, OF mono eq inverse2
nrec_admissible])
using step by blast
hence "U f x ≤ SPEC (P x)" by simp
finally show "P x y" by auto
qed
*)

(* TODO/FIXME: Induction rule seems not to work here !
Function package expects induction rule where conclusion is a binary
predicate as free variable.
*)


declaration {* Partial_Function.init "nrec" @{term nrec.fixp_fun}
@{term nrec.mono_body} @{thm nrec.fixp_rule_uc}
(*SOME @{thm fixp_induct_nrec'}*) NONE *}


lemma bind_mono_pfun[partial_function_mono]:
"[| nrec.mono_body B; !!y. nrec.mono_body (λf. C y f) |] ==>
nrec.mono_body (λ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 "op ≤" "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 (op ≤::'b dres => _ => _) = op ≤"
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[OF drec.ccpo])
apply (unfold fun_lub_def)
apply clarsimp
apply (drule_tac x=x in point_chainI)
apply (erule dres_Sup_chain_cases)
apply (simp only:)
apply simp
apply (simp only:)
apply auto []
apply (simp only:)
apply force
done
qed

declaration {* Partial_Function.init "drec" @{term drec.fixp_fun}
@{term drec.mono_body} @{thm drec.fixp_rule_uc} NONE *}


lemma drec_bind_mono_pfun[partial_function_mono]:
"[| drec.mono_body B; !!y. drec.mono_body (λf. C y f) |] ==>
drec.mono_body (λf. dbind (B f) (λy. C y f))"

apply rule
apply (rule dbind_mono)
apply (blast dest: monotoneD)+
done

end