section ‹Generative probabilistic values›
theory Generative_Probabilistic_Value imports
"~~/src/HOL/Probability/SPMF"
"../../AFP/MFMC_Countable/Rel_PMF_Characterisation"
PF_Set
Resumption
Generat
"../../BNF_Corec"
begin
interpretation rel_spmf_characterisation
by(unfold_locales)(rule rel_pmf_measureI)
context includes lifting_syntax begin
lemma rel_fun_curry:
"(A ===> B ===> C) f g ⟷ (rel_prod A B ===> C) (case_prod f) (case_prod g)"
by(auto simp add: rel_fun_def)
end
lemma tlength_eq_infinity_iff: "tlength xs = ∞ ⟷ ¬ tfinite xs"
including tllist.lifting by transfer(simp add: llength_eq_infty_conv_lfinite)
lemma transp_rel_fun: "⟦ is_equality Q; transp R ⟧ ⟹ transp (rel_fun Q R)"
by(rule transpI)(auto dest: transpD rel_funD simp add: is_equality_def)
lemma rel_fun_inf: "inf (rel_fun Q R) (rel_fun Q R') = rel_fun Q (inf R R')"
by(rule antisym)(auto elim: rel_fun_mono dest: rel_funD)
lemma eSuc_SUP: "A ≠ {} ⟹ eSuc (SUPREMUM A f) = (SUP x:A. eSuc (f x))"
by(subst eSuc_Sup)(simp_all)
lemma (in ccpo) Sup_image_mono:
assumes ccpo: "class.ccpo luba orda lessa"
and mono: "monotone orda op ≤ f"
and chain: "Complete_Partial_Order.chain orda A"
and "A ≠ {}"
shows "Sup (f ` A) ≤ (f (luba A))"
proof(rule ccpo_Sup_least)
from chain show "Complete_Partial_Order.chain op ≤ (f ` A)"
by(rule chain_imageI)(rule monotoneD[OF mono])
fix x
assume "x ∈ f ` A"
then obtain y where "x = f y" "y ∈ A" by blast
from ‹y ∈ A› have "orda y (luba A)" by(rule ccpo.ccpo_Sup_upper[OF ccpo chain])
hence "f y ≤ f (luba A)" by(rule monotoneD[OF mono])
thus "x ≤ f (luba A)" using ‹x = f y› by simp
qed
lemma (in ccpo) admissible_le_mono:
assumes "monotone op ≤ op ≤ f"
shows "ccpo.admissible Sup op ≤ (λx. x ≤ f x)"
proof(rule ccpo.admissibleI)
fix Y
assume chain: "Complete_Partial_Order.chain op ≤ Y"
and Y: "Y ≠ {}"
and le [rule_format]: "∀x∈Y. x ≤ f x"
have "⨆Y ≤ ⨆(f ` Y)" using chain
by(rule ccpo_Sup_least)(rule order_trans[OF le]; blast intro!: ccpo_Sup_upper chain_imageI[OF chain] intro: monotoneD[OF assms])
also have "… ≤ f (⨆Y)"
by(rule Sup_image_mono[OF _ assms chain Y, where lessa="op <"]) unfold_locales
finally show "⨆Y ≤ …" .
qed
lemma (in ccpo) fixp_induct_strong2:
assumes adm: "ccpo.admissible Sup op ≤ P"
and mono: "monotone op ≤ op ≤ f"
and bot: "P (⨆{})"
and step: "⋀x. ⟦ x ≤ ccpo_class.fixp f; x ≤ f x; P x ⟧ ⟹ P (f x)"
shows "P (ccpo_class.fixp f)"
proof(rule fixp_strong_induct[where P="λx. x ≤ f x ∧ P x", THEN conjunct2])
show "ccpo.admissible Sup op ≤ (λx. x ≤ f x ∧ P x)"
using admissible_le_mono adm by(rule admissible_conj)(rule mono)
next
show "⨆{} ≤ f (⨆{}) ∧ P (⨆{})"
by(auto simp add: bot chain_empty intro: ccpo_Sup_least)
next
fix x
assume "x ≤ ccpo_class.fixp f" "x ≤ f x ∧ P x"
thus "f x ≤ f (f x) ∧ P (f x)"
by(auto dest: monotoneD[OF mono] intro: step)
qed(rule mono)
context partial_function_definitions begin
lemma fixp_induct_strong2_uc:
fixes F :: "'c ⇒ 'c"
and U :: "'c ⇒ 'b ⇒ 'a"
and C :: "('b ⇒ 'a) ⇒ 'c"
and P :: "('b ⇒ 'a) ⇒ bool"
assumes mono: "⋀x. mono_body (λf. U (F (C f)) x)"
and eq: "f ≡ C (fixp_fun (λf. U (F (C f))))"
and inverse: "⋀f. U (C f) = f"
and adm: "ccpo.admissible lub_fun le_fun P"
and bot: "P (λ_. lub {})"
and step: "⋀f'. ⟦ le_fun (U f') (U f); le_fun (U f') (U (F f')); P (U f') ⟧ ⟹ P (U (F f'))"
shows "P (U f)"
unfolding eq inverse
apply (rule ccpo.fixp_induct_strong2[OF ccpo adm])
apply (insert mono, auto simp: monotone_def fun_ord_def bot fun_lub_def)[2]
apply (rule_tac f'5="C x" in step)
apply (simp_all add: inverse eq)
done
end
hide_const (open) Done
subsection {* Type definition *}
context notes [[quick_and_dirty]] begin
codatatype (results'_gpv: 'a, outs'_gpv: 'out, 'in) gpv
= GPV (the_gpv: "('a, 'out, 'in ⇒ ('a, 'out, 'in) gpv) generat spmf")
end
declare gpv.rel_eq [relator_eq]
text {* Reactive values are like generative, except that they take an input first. *}
type_synonym ('a, 'out, 'in) rpv = "'in ⇒ ('a, 'out, 'in) gpv"
print_translation -- ‹pretty printing for @{typ "('a, 'out, 'in) rpv"}› ‹
let
fun tr' [in1, Const (@{type_syntax gpv}, _) $ a $ out $ in2] =
if in1 = in2 then Syntax.const @{type_syntax rpv} $ a $ out $ in1
else raise Match;
in [(@{type_syntax "fun"}, K tr')]
end
›
typ "('a, 'out, 'in) rpv"
text ‹
Effectively, @{typ "('a, 'out, 'in) gpv"} and @{typ "('a, 'out, 'in) rpv"} are mutually recursive.
›
lemma eq_GPV_iff: "f = GPV g ⟷ the_gpv f = g"
by(cases f) auto
declare gpv.set[simp del]
lemma rel_gpv_def':
"rel_gpv A B gpv gpv' ⟷
(∃gpv''. (∀(x, y) ∈ results'_gpv gpv''. A x y) ∧ (∀(x, y) ∈ outs'_gpv gpv''. B x y) ∧
map_gpv fst fst gpv'' = gpv ∧ map_gpv snd snd gpv'' = gpv')"
unfolding rel_gpv_def by(auto simp add: BNF_Def.Grp_def)
definition results'_rpv :: "('a, 'out, 'in) rpv ⇒ 'a set"
where "results'_rpv rpv = range rpv ⤜ results'_gpv"
definition outs'_rpv :: "('a, 'out, 'in) rpv ⇒ 'out set"
where "outs'_rpv rpv = range rpv ⤜ outs'_gpv"
abbreviation rel_rpv
:: "('a ⇒ 'b ⇒ bool) ⇒ ('out ⇒ 'out' ⇒ bool)
⇒ ('in ⇒ ('a, 'out, 'in) gpv) ⇒ ('in ⇒ ('b, 'out', 'in) gpv) ⇒ bool"
where "rel_rpv A B ≡ rel_fun op = (rel_gpv A B)"
lemma in_results'_rpv [iff]: "x ∈ results'_rpv rpv ⟷ (∃input. x ∈ results'_gpv (rpv input))"
by(simp add: results'_rpv_def)
lemma in_outs_rpv [iff]: "out ∈ outs'_rpv rpv ⟷ (∃input. out ∈ outs'_gpv (rpv input))"
by(simp add: outs'_rpv_def)
lemma results'_GPV [simp]:
"results'_gpv (GPV r) =
(set_spmf r ⤜ generat_pures) ∪
((set_spmf r ⤜ generat_conts) ⤜ results'_rpv)"
by(auto simp add: gpv.set bind_UNION set_spmf_def)
lemma outs'_GPV [simp]:
"outs'_gpv (GPV r) =
(set_spmf r ⤜ generat_outs) ∪
((set_spmf r ⤜ generat_conts) ⤜ outs'_rpv)"
by(auto simp add: gpv.set bind_UNION set_spmf_def)
lemma outs'_gpv_unfold:
"outs'_gpv r =
(set_spmf (the_gpv r) ⤜ generat_outs) ∪
((set_spmf (the_gpv r) ⤜ generat_conts) ⤜ outs'_rpv)"
by(cases r) simp
lemma outs'_gpv_induct [consumes 1, case_names Out Cont, induct set: outs'_gpv]:
assumes x: "x ∈ outs'_gpv gpv"
and Out: "⋀generat gpv. ⟦ generat ∈ set_spmf (the_gpv gpv); x ∈ generat_outs generat ⟧ ⟹ P gpv"
and Cont: "⋀generat gpv c input.
⟦ generat ∈ set_spmf (the_gpv gpv); c ∈ generat_conts generat; x ∈ outs'_gpv (c input); P (c input) ⟧ ⟹ P gpv"
shows "P gpv"
using x
apply(induction y≡"x" gpv)
apply(rule Out, simp add: in_set_spmf, simp)
apply(erule imageE, rule Cont, simp add: in_set_spmf, simp, simp, simp)
.
lemma outs'_gpv_cases [consumes 1, case_names Out Cont, cases set: outs'_gpv]:
assumes "x ∈ outs'_gpv gpv"
obtains (Out) generat where "generat ∈ set_spmf (the_gpv gpv)" "x ∈ generat_outs generat"
| (Cont) generat c input where "generat ∈ set_spmf (the_gpv gpv)" "c ∈ generat_conts generat" "x ∈ outs'_gpv (c input)"
using assms by cases(auto simp add: in_set_spmf)
lemma outs'_gpvI [intro?]:
shows outs'_gpv_Out: "⟦ generat ∈ set_spmf (the_gpv gpv); x ∈ generat_outs generat ⟧ ⟹ x ∈ outs'_gpv gpv"
and outs'_gpv_Cont: "⟦ generat ∈ set_spmf (the_gpv gpv); c ∈ generat_conts generat; x ∈ outs'_gpv (c input) ⟧ ⟹ x ∈ outs'_gpv gpv"
by(auto intro: gpv.set_sel simp add: in_set_spmf)
lemma results'_gpv_induct [consumes 1, case_names Pure Cont, induct set: results'_gpv]:
assumes x: "x ∈ results'_gpv gpv"
and Pure: "⋀generat gpv. ⟦ generat ∈ set_spmf (the_gpv gpv); x ∈ generat_pures generat ⟧ ⟹ P gpv"
and Cont: "⋀generat gpv c input.
⟦ generat ∈ set_spmf (the_gpv gpv); c ∈ generat_conts generat; x ∈ results'_gpv (c input); P (c input) ⟧ ⟹ P gpv"
shows "P gpv"
using x
apply(induction y≡"x" gpv)
apply(rule Pure; simp add: in_set_spmf)
apply(erule imageE, rule Cont, simp add: in_set_spmf, simp, simp, simp)
.
lemma results'_gpv_cases [consumes 1, case_names Pure Cont, cases set: results'_gpv]:
assumes "x ∈ results'_gpv gpv"
obtains (Pure) generat where "generat ∈ set_spmf (the_gpv gpv)" "x ∈ generat_pures generat"
| (Cont) generat c input where "generat ∈ set_spmf (the_gpv gpv)" "c ∈ generat_conts generat" "x ∈ results'_gpv (c input)"
using assms by cases(auto simp add: in_set_spmf)
lemma results'_gpvI [intro?]:
shows results'_gpv_Pure: "⟦ generat ∈ set_spmf (the_gpv gpv); x ∈ generat_pures generat ⟧ ⟹ x ∈ results'_gpv gpv"
and results'_gpv_Cont: "⟦ generat ∈ set_spmf (the_gpv gpv); c ∈ generat_conts generat; x ∈ results'_gpv (c input) ⟧ ⟹ x ∈ results'_gpv gpv"
by(auto intro: gpv.set_sel simp add: in_set_spmf)
lemma left_unique_rel_gpv [transfer_rule]:
"⟦ left_unique A; left_unique B ⟧ ⟹ left_unique (rel_gpv A B)"
unfolding left_unique_alt_def gpv.rel_conversep[symmetric] gpv.rel_compp[symmetric]
by(subst gpv.rel_eq[symmetric])(rule gpv.rel_mono)
lemma right_unique_rel_gpv [transfer_rule]:
"⟦ right_unique A; right_unique B ⟧ ⟹ right_unique (rel_gpv A B)"
unfolding right_unique_alt_def gpv.rel_conversep[symmetric] gpv.rel_compp[symmetric]
by(subst gpv.rel_eq[symmetric])(rule gpv.rel_mono)
lemma bi_unique_rel_gpv [transfer_rule]:
"⟦ bi_unique A; bi_unique B ⟧ ⟹ bi_unique (rel_gpv A B)"
unfolding bi_unique_alt_def by(simp add: left_unique_rel_gpv right_unique_rel_gpv)
lemma left_total_rel_gpv [transfer_rule]:
"⟦ left_total A; left_total B ⟧ ⟹ left_total (rel_gpv A B)"
unfolding left_total_alt_def gpv.rel_conversep[symmetric] gpv.rel_compp[symmetric]
by(subst gpv.rel_eq[symmetric])(rule gpv.rel_mono)
lemma right_total_rel_gpv [transfer_rule]:
"⟦ right_total A; right_total B ⟧ ⟹ right_total (rel_gpv A B)"
unfolding right_total_alt_def gpv.rel_conversep[symmetric] gpv.rel_compp[symmetric]
by(subst gpv.rel_eq[symmetric])(rule gpv.rel_mono)
lemma bi_total_rel_gpv [transfer_rule]: "⟦ bi_total A; bi_total B ⟧ ⟹ bi_total (rel_gpv A B)"
unfolding bi_total_alt_def by(simp add: left_total_rel_gpv right_total_rel_gpv)
declare gpv.map_transfer[transfer_rule]
lemma if_distrib_map_gpv [if_distribs]:
"map_gpv f g (if b then gpv else gpv') = (if b then map_gpv f g gpv else map_gpv f g gpv')"
by simp
subsection {* Simple, derived operations *}
primcorec Done :: "'a ⇒ ('a, 'out, 'in) gpv"
where "the_gpv (Done a) = return_spmf (Pure a)"
primcorec Pause :: "'out ⇒ ('in ⇒ ('a, 'out, 'in) gpv) ⇒ ('a, 'out, 'in) gpv"
where "the_gpv (Pause out c) = return_spmf (IO out c)"
primcorec lift_spmf :: "'a spmf ⇒ ('a, 'out, 'in) gpv"
where "the_gpv (lift_spmf p) = map_spmf Pure p"
definition Fail :: "('a, 'out, 'in) gpv"
where "Fail = GPV (return_pmf None)"
definition React :: "('in ⇒ 'out × ('a, 'out, 'in) rpv) ⇒ ('a, 'out, 'in) rpv"
where "React f input = case_prod Pause (f input)"
definition rFail :: "('a, 'out, 'in) rpv"
where "rFail = (λ_. Fail)"
lemma Done_inject [simp]: "Done x = Done y ⟷ x = y"
by(simp add: Done.ctr)
lemma Pause_inject [simp]: "Pause out c = Pause out' c' ⟷ out = out' ∧ c = c'"
by(simp add: Pause.ctr)
lemma [simp]:
shows Done_neq_Pause: "Done x ≠ Pause out c"
and Pause_neq_Done: "Pause out c ≠ Done x"
by(simp_all add: Done.ctr Pause.ctr)
lemma outs'_gpv_Done [simp]: "outs'_gpv (Done x) = {}"
by(auto elim: outs'_gpv_cases)
lemma results'_gpv_Done [simp]: "results'_gpv (Done x) = {x}"
by(auto intro: results'_gpvI elim: results'_gpv_cases)
lemma outs'_gpv_Pause [simp]: "outs'_gpv (Pause out c) = insert out (⋃input. outs'_gpv (c input))"
by(auto 4 4 intro: outs'_gpvI elim: outs'_gpv_cases)
lemma results'_gpv_Pause [simp]: "results'_gpv (Pause out rpv) = results'_rpv rpv"
by(auto 4 4 intro: results'_gpvI elim: results'_gpv_cases)
lemma lift_spmf_return [simp]: "lift_spmf (return_spmf x) = Done x"
by(simp add: lift_spmf.ctr Done.ctr)
lemma lift_spmf_None [simp]: "lift_spmf (return_pmf None) = Fail"
by(rule gpv.expand)(simp add: Fail_def)
lemma the_gpv_lift_spmf [simp]: "the_gpv (lift_spmf r) = map_spmf Pure r"
by(simp)
lemma outs'_gpv_lift_spmf [simp]: "outs'_gpv (lift_spmf p) = {}"
by(auto 4 3 elim: outs'_gpv_cases)
lemma results'_gpv_lift_spmf [simp]: "results'_gpv (lift_spmf p) = set_spmf p"
by(auto 4 3 elim: results'_gpv_cases intro: results'_gpvI)
lemma lift_spmf_inject [simp]: "lift_spmf p = lift_spmf q ⟷ p = q"
by(auto simp add: lift_spmf.code dest!: pmf.inj_map_strong[rotated] option.inj_map_strong[rotated])
lemma map_lift_spmf: "map_gpv f g (lift_spmf p) = lift_spmf (map_spmf f p)"
by(rule gpv.expand)(simp add: gpv.map_sel spmf.map_comp o_def)
lemma [simp]:
shows Fail_neq_Pause: "Fail ≠ Pause out c"
and Pause_neq_Fail: "Pause out c ≠ Fail"
and Fail_neq_Done: "Fail ≠ Done x"
and Done_neq_Fail: "Done x ≠ Fail"
by(simp_all add: Fail_def Pause.ctr Done.ctr)
text {* Add @{typ unit} closure to circumvent SML value restriction *}
definition Fail' :: "unit ⇒ ('a, 'out, 'in) gpv"
where [code del]: "Fail' _ = Fail"
lemma Fail_code [code_unfold]: "Fail = Fail' ()"
by(simp add: Fail'_def)
lemma Fail'_code [code]:
"Fail' x = GPV (return_pmf None)"
by(simp add: Fail'_def Fail_def)
lemma Fail_sel [simp]:
"the_gpv Fail = return_pmf None"
by(simp add: Fail_def)
lemma Fail_eq_GPV_iff [simp]: "Fail = GPV f ⟷ f = return_pmf None"
by(auto simp add: Fail_def)
lemma outs'_gpv_Fail [simp]: "outs'_gpv Fail = {}"
by(auto elim: outs'_gpv_cases)
lemma results'_gpv_Fail [simp]: "results'_gpv Fail = {}"
by(auto elim: results'_gpv_cases)
lemma React_inject [iff]: "React f = React f' ⟷ f = f'"
by(auto simp add: React_def fun_eq_iff split_def intro: prod.expand)
lemma React_apply [simp]: "f input = (out, c) ⟹ React f input = Pause out c"
by(simp add: React_def)
lemma rFail_apply [simp]: "rFail input = Fail"
by(simp add: rFail_def)
lemma [simp]:
shows rFail_neq_React: "rFail ≠ React f"
and React_neq_rFail: "React f ≠ rFail"
by(simp_all add: React_def fun_eq_iff split_beta)
lemma rel_gpv_FailI [simp]: "rel_gpv A C Fail Fail"
by(subst gpv.rel_sel) simp
lemma rel_gpv_Done [iff]: "rel_gpv A C (Done x) (Done y) ⟷ A x y"
by(subst gpv.rel_sel) simp
lemma rel_gpv_Pause [iff]: "rel_gpv A C (Pause out c) (Pause out' c') ⟷ C out out' ∧ (∀x. rel_gpv A C (c x) (c' x))"
by(subst gpv.rel_sel)(simp add: rel_fun_def)
lemma rel_gpv_lift_spmf [iff]: "rel_gpv A C (lift_spmf p) (lift_spmf q) ⟷ rel_spmf A p q"
by(subst gpv.rel_sel)(simp add: spmf_rel_map)
context includes lifting_syntax begin
lemma Fail_parametric [transfer_rule]: "rel_gpv A C Fail Fail"
by(fact rel_gpv_FailI)
lemma Done_parametric [transfer_rule]: "(A ===> rel_gpv A C) Done Done"
by(rule rel_funI) simp
lemma Pause_parametric [transfer_rule]:
"(C ===> (op = ===> rel_gpv A C) ===> rel_gpv A C) Pause Pause"
by(simp add: rel_fun_def)
lemma lift_spmf_parametric [transfer_rule]:
"(rel_spmf A ===> rel_gpv A C) lift_spmf lift_spmf"
by(simp add: rel_fun_def)
end
lemma map_gpv_Done [simp]: "map_gpv f g (Done x) = Done (f x)"
by(simp add: Done.code)
lemma map_gpv_Pause [simp]: "map_gpv f g (Pause x c) = Pause (g x) (map_gpv f g ∘ c)"
by(simp add: Pause.code)
lemma map_gpv_Fail [simp]: "map_gpv f g Fail = Fail"
by(simp add: Fail_def)
subsection {* Monad structure *}
primcorec bind_gpv :: "('a, 'out, 'in) gpv ⇒ ('a ⇒ ('b, 'out, 'in) gpv) ⇒ ('b, 'out, 'in) gpv"
where
"the_gpv (bind_gpv r f) =
map_spmf (map_generat id id (op ∘ (case_sum id (λr. bind_gpv r f))))
(the_gpv r ⤜
(case_generat
(λx. map_spmf (map_generat id id (op ∘ Inl)) (the_gpv (f x)))
(λout c. return_spmf (IO out (λinput. Inr (c input))))))"
declare bind_gpv.sel [simp del]
adhoc_overloading Monad_Syntax.bind bind_gpv
lemma bind_gpv_unfold [code]:
"r ⤜ f = GPV (
do {
generat ← the_gpv r;
case generat of Pure x ⇒ the_gpv (f x)
| IO out c ⇒ return_spmf (IO out (λinput. c input ⤜ f))
})"
unfolding bind_gpv_def
apply(rule gpv.expand)
apply(simp add: map_spmf_bind_spmf)
apply(rule arg_cong[where f="bind_spmf (the_gpv r)"])
apply(auto split: generat.split simp add: map_spmf_bind_spmf fun_eq_iff spmf.map_comp o_def generat.map_comp id_def[symmetric] generat.map_id pmf.map_id option.map_id)
done
lemma bind_gpv_code_cong: "f = f' ⟹ bind_gpv f g = bind_gpv f' g" by simp
setup {* Code_Simp.map_ss (Simplifier.add_cong @{thm bind_gpv_code_cong}) *}
lemma bind_gpv_sel:
"the_gpv (r ⤜ f) =
do {
generat ← the_gpv r;
case generat of Pure x ⇒ the_gpv (f x)
| IO out c ⇒ return_spmf (IO out (λinput. bind_gpv (c input) f))
}"
by(subst bind_gpv_unfold) simp
lemma bind_gpv_sel' [simp]:
"the_gpv (r ⤜ f) =
do {
generat ← the_gpv r;
if is_Pure generat then the_gpv (f (result generat))
else return_spmf (IO (output generat) (λinput. bind_gpv (continuation generat input) f))
}"
unfolding bind_gpv_sel
by(rule arg_cong[where f="bind_spmf (the_gpv r)"])(simp add: fun_eq_iff split: generat.split)
lemma Done_bind_gpv [simp]: "Done a ⤜ f = f a"
by(rule gpv.expand)(simp)
lemma bind_gpv_Done [simp]: "f ⤜ Done = f"
proof(coinduction arbitrary: f rule: gpv.coinduct)
case (Eq_gpv f)
have *: "the_gpv f ⤜ (case_generat (λx. return_spmf (Pure x)) (λout c. return_spmf (IO out (λinput. Inr (c input))))) =
map_spmf (map_generat id id (op ∘ Inr)) (bind_spmf (the_gpv f) return_spmf)"
unfolding map_spmf_bind_spmf
by(rule arg_cong2[where f=bind_spmf])(auto simp add: fun_eq_iff split: generat.split)
show ?case
by(auto simp add: * bind_gpv.simps pmf.rel_map option.rel_map[abs_def] generat.rel_map[abs_def] simp del: bind_gpv_sel' intro!: rel_generatI rel_spmf_reflI)
qed
lemma if_distrib_bind_gpv2 [if_distribs]:
"bind_gpv gpv (λy. if b then f y else g y) = (if b then bind_gpv gpv f else bind_gpv gpv g)"
by simp
lemma lift_spmf_bind: "lift_spmf r ⤜ f = GPV (r ⤜ the_gpv ∘ f)"
by(coinduction arbitrary: r f rule: gpv.coinduct_strong)(auto simp add: bind_map_spmf o_def intro: rel_pmf_reflI rel_optionI rel_generatI)
lemma the_gpv_bind_gpv_lift_spmf [simp]:
"the_gpv (bind_gpv (lift_spmf p) f) = bind_spmf p (the_gpv ∘ f)"
by(simp add: bind_map_spmf o_def)
lemma lift_spmf_bind_spmf: "lift_spmf (p ⤜ f) = lift_spmf p ⤜ (λx. lift_spmf (f x))"
by(rule gpv.expand)(simp add: lift_spmf_bind o_def map_spmf_bind_spmf)
lemma GPV_bind:
"GPV f ⤜ g =
GPV (f ⤜ (λgenerat. case generat of Pure x ⇒ the_gpv (g x) | IO out c ⇒ return_spmf (IO out (λinput. c input ⤜ g))))"
by(subst bind_gpv_unfold) simp
lemma GPV_bind':
"GPV f ⤜ g = GPV (f ⤜ (λgenerat. if is_Pure generat then the_gpv (g (result generat)) else return_spmf (IO (output generat) (λinput. continuation generat input ⤜ g))))"
unfolding GPV_bind gpv.inject
by(rule arg_cong[where f="bind_spmf f"])(simp add: fun_eq_iff split: generat.split)
lemma bind_gpv_assoc:
fixes f :: "('a, 'out, 'in) gpv"
shows "(f ⤜ g) ⤜ h = f ⤜ (λx. g x ⤜ h)"
proof(coinduction arbitrary: f g h rule: gpv.coinduct_strong)
case (Eq_gpv f g h)
show ?case
apply(simp cong del: if_weak_cong)
apply(rule rel_spmf_bindI[where R="op ="])
apply(simp add: option.rel_eq pmf.rel_eq)
apply(fastforce intro: rel_pmf_return_pmfI rel_generatI rel_spmf_reflI)
done
qed
lemma map_gpv_bind_gpv: "map_gpv f g (bind_gpv gpv h) = bind_gpv (map_gpv id g gpv) (λx. map_gpv f g (h x))"
apply(coinduction arbitrary: gpv rule: gpv.coinduct_strong)
apply(simp add: bind_gpv.sel gpv.map_sel spmf_rel_map generat.rel_map o_def bind_map_spmf del: bind_gpv_sel')
apply(rule rel_spmf_bind_reflI)
apply(auto simp add: spmf_rel_map generat.rel_map split: generat.split intro!: rel_spmf_reflI generat.rel_refl rel_funI)
done
lemma map_gpv_id_bind_gpv: "map_gpv f id (bind_gpv gpv g) = bind_gpv gpv (map_gpv f id ∘ g)"
by(simp add: map_gpv_bind_gpv gpv.map_id o_def)
lemma map_gpv_conv_bind:
"map_gpv f (λx. x) x = bind_gpv x (λx. Done (f x))"
using map_gpv_bind_gpv[of f "λx. x" x Done] by(simp add: id_def[symmetric] gpv.map_id)
lemma bind_map_gpv: "bind_gpv (map_gpv f id gpv) g = bind_gpv gpv (g ∘ f)"
by(simp add: map_gpv_conv_bind id_def bind_gpv_assoc o_def)
lemma bind_gpv_Fail [simp]: "Fail ⤜ f = Fail"
by(subst bind_gpv_unfold)(simp add: Fail_def)
context includes lifting_syntax begin
lemma bind_gpv_parametric [transfer_rule]:
"(rel_gpv A C ===> (A ===> rel_gpv B C) ===> rel_gpv B C) bind_gpv bind_gpv"
unfolding bind_gpv_def by transfer_prover
end
lemma rel_gpv_bindI:
"⟦ rel_gpv A C gpv gpv'; ⋀x y. A x y ⟹ rel_gpv B C (f x) (g y) ⟧
⟹ rel_gpv B C (bind_gpv gpv f) (bind_gpv gpv' g)"
by(fact bind_gpv_parametric[THEN rel_funD, THEN rel_funD, OF _ rel_funI])
lemma bind_gpv_cong:
"⟦ gpv = gpv'; ⋀x. x ∈ results'_gpv gpv' ⟹ f x = g x ⟧ ⟹ bind_gpv gpv f = bind_gpv gpv' g"
apply(subst gpv.rel_eq[symmetric])
apply(rule rel_gpv_bindI[where A="eq_onp (λx. x ∈ results'_gpv gpv')"])
apply(subst (asm) gpv.rel_eq[symmetric])
apply(erule gpv.rel_mono_strong)
apply(simp add: eq_onp_def)
apply simp
apply(clarsimp simp add: gpv.rel_eq eq_onp_def)
done
definition bind_rpv :: "('a, 'in, 'out) rpv ⇒ ('a ⇒ ('b, 'in, 'out) gpv) ⇒ ('b, 'in, 'out) rpv"
where "bind_rpv rpv f = (λinput. bind_gpv (rpv input) f)"
lemma bind_rpv_apply [simp]: "bind_rpv rpv f input = bind_gpv (rpv input) f"
by(simp add: bind_rpv_def fun_eq_iff)
adhoc_overloading Monad_Syntax.bind bind_rpv
lemma bind_rpv_code_cong: "rpv = rpv' ⟹ bind_rpv rpv f = bind_rpv rpv' f" by simp
setup {* Code_Simp.map_ss (Simplifier.add_cong @{thm bind_rpv_code_cong}) *}
lemma bind_rpv_rDone [simp]: "bind_rpv rpv Done = rpv"
by(simp add: bind_rpv_def)
lemma bind_gpv_Pause [simp]: "bind_gpv (Pause out rpv) f = Pause out (bind_rpv rpv f)"
by(rule gpv.expand)(simp add: fun_eq_iff)
lemma bind_rpv_React [simp]: "bind_rpv (React f) g = React (apsnd (λrpv. bind_rpv rpv g) ∘ f)"
by(simp add: React_def split_beta fun_eq_iff)
lemma bind_rpv_assoc: "bind_rpv (bind_rpv rpv f) g = bind_rpv rpv ((λgpv. bind_gpv gpv g) ∘ f)"
by(simp add: fun_eq_iff bind_gpv_assoc o_def)
lemma bind_rpv_Done [simp]: "bind_rpv Done f = f"
by(simp add: bind_rpv_def)
subsection ‹ Embedding @{typ "'a spmf"} as a monad ›
context includes lifting_syntax begin
lemma neg_fun_distr3:
assumes 1: "left_unique R" "right_total R"
assumes 2: "right_unique S" "left_total S"
shows "(R OO R' ===> S OO S') ≤ ((R ===> S) OO (R' ===> S'))"
using functional_relation[OF 2] functional_converse_relation[OF 1]
unfolding rel_fun_def OO_def
apply clarify
apply (subst all_comm)
apply (subst all_conj_distrib[symmetric])
apply (intro choice)
by metis
end
locale spmf_to_gpv begin
text ‹
The lifting package cannot handle free term variables in the merging of transfer rules,
so for the embedding we define a specialised relator @{text "rel_gpv'"}
which acts only on the returned values.
›
definition rel_gpv' :: "('a ⇒ 'b ⇒ bool) ⇒ ('a, 'out, 'in) gpv ⇒ ('b, 'out, 'in) gpv ⇒ bool"
where "rel_gpv' A = rel_gpv A op ="
lemma rel_gpv'_eq [relator_eq]: "rel_gpv' op = = op ="
unfolding rel_gpv'_def gpv.rel_eq ..
lemma rel_gpv'_mono [relator_mono]: "A ≤ B ⟹ rel_gpv' A ≤ rel_gpv' B"
unfolding rel_gpv'_def by(rule gpv.rel_mono; simp)
lemma rel_gpv'_distr [relator_distr]: "rel_gpv' A OO rel_gpv' B = rel_gpv' (A OO B)"
unfolding rel_gpv'_def by (metis OO_eq gpv.rel_compp)
lemma left_unique_rel_gpv' [transfer_rule]: "left_unique A ⟹ left_unique (rel_gpv' A)"
unfolding rel_gpv'_def by(simp add: left_unique_rel_gpv left_unique_eq)
lemma right_unique_rel_gpv' [transfer_rule]: "right_unique A ⟹ right_unique (rel_gpv' A)"
unfolding rel_gpv'_def by(simp add: right_unique_rel_gpv right_unique_eq)
lemma bi_unique_rel_gpv' [transfer_rule]: "bi_unique A ⟹ bi_unique (rel_gpv' A)"
unfolding rel_gpv'_def by(simp add: bi_unique_rel_gpv bi_unique_eq)
lemma left_total_rel_gpv' [transfer_rule]: "left_total A ⟹ left_total (rel_gpv' A)"
unfolding rel_gpv'_def by(simp add: left_total_rel_gpv left_total_eq)
lemma right_total_rel_gpv' [transfer_rule]: "right_total A ⟹ right_total (rel_gpv' A)"
unfolding rel_gpv'_def by(simp add: right_total_rel_gpv right_total_eq)
lemma bi_total_rel_gpv' [transfer_rule]: "bi_total A ⟹ bi_total (rel_gpv' A)"
unfolding rel_gpv'_def by(simp add: bi_total_rel_gpv bi_total_eq)
text ‹
We cannot use @{text setup_lifting} because @{typ "('a, 'out, 'in) gpv"} contains
type variables which do not appear in @{typ "'a spmf"}.
›
definition cr_spmf_gpv :: "'a spmf ⇒ ('a, 'out, 'in) gpv ⇒ bool"
where "cr_spmf_gpv p gpv ⟷ gpv = lift_spmf p"
definition spmf_of_gpv :: "('a, 'out, 'in) gpv ⇒ 'a spmf"
where "spmf_of_gpv gpv = (THE p. gpv = lift_spmf p)"
lemma spmf_of_gpv_lift_spmf [simp]: "spmf_of_gpv (lift_spmf p) = p"
unfolding spmf_of_gpv_def by auto
lemma rel_spmf_setD2:
"⟦ rel_spmf A p q; y ∈ set_spmf q ⟧ ⟹ ∃x∈set_spmf p. A x y"
by(erule rel_spmfE) force
lemma rel_gpv_lift_spmf1: "rel_gpv A B (lift_spmf p) gpv ⟷ (∃q. gpv = lift_spmf q ∧ rel_spmf A p q)"
apply(subst gpv.rel_sel)
apply(simp add: spmf_rel_map rel_generat_Pure1)
apply safe
apply(rule exI[where x="map_spmf result (the_gpv gpv)"])
apply(clarsimp simp add: spmf_rel_map)
apply(rule conjI)
apply(rule gpv.expand)
apply(simp add: spmf.map_comp)
apply(subst map_spmf_cong[OF refl, where g=id])
apply(drule (1) rel_spmf_setD2)
apply clarsimp
apply simp
apply(erule rel_spmf_mono)
apply clarsimp
apply(clarsimp simp add: spmf_rel_map)
done
lemma rel_gpv_lift_spmf2: "rel_gpv A B gpv (lift_spmf q) ⟷ (∃p. gpv = lift_spmf p ∧ rel_spmf A p q)"
by(subst gpv.rel_flip[symmetric])(simp add: rel_gpv_lift_spmf1 pmf.rel_flip option.rel_conversep)
definition pcr_spmf_gpv :: "('a ⇒ 'b ⇒ bool) ⇒ 'a spmf ⇒ ('b, 'out, 'in) gpv ⇒ bool"
where "pcr_spmf_gpv A = cr_spmf_gpv OO rel_gpv A op ="
lemma pcr_cr_eq_spmf_gpv: "pcr_spmf_gpv op = = cr_spmf_gpv"
by(simp add: pcr_spmf_gpv_def gpv.rel_eq OO_eq)
lemma left_unique_cr_spmf_gpv: "left_unique cr_spmf_gpv"
by(rule left_uniqueI)(simp add: cr_spmf_gpv_def)
lemma left_unique_pcr_spmf_gpv [transfer_rule]:
"left_unique A ⟹ left_unique (pcr_spmf_gpv A)"
unfolding pcr_spmf_gpv_def by(intro left_unique_OO left_unique_cr_spmf_gpv left_unique_rel_gpv left_unique_eq)
lemma right_unique_cr_spmf_gpv: "right_unique cr_spmf_gpv"
by(rule right_uniqueI)(simp add: cr_spmf_gpv_def)
lemma right_unique_pcr_spmf_gpv [transfer_rule]:
"right_unique A ⟹ right_unique (pcr_spmf_gpv A)"
unfolding pcr_spmf_gpv_def by(intro right_unique_OO right_unique_cr_spmf_gpv right_unique_rel_gpv right_unique_eq)
lemma bi_unique_cr_spmf_gpv: "bi_unique cr_spmf_gpv"
by(simp add: bi_unique_alt_def left_unique_cr_spmf_gpv right_unique_cr_spmf_gpv)
lemma bi_unique_pcr_spmf_gpv [transfer_rule]: "bi_unique A ⟹ bi_unique (pcr_spmf_gpv A)"
by(simp add: bi_unique_alt_def left_unique_pcr_spmf_gpv right_unique_pcr_spmf_gpv)
lemma left_total_cr_spmf_gpv: "left_total cr_spmf_gpv"
by(rule left_totalI)(simp add: cr_spmf_gpv_def)
lemma left_total_pcr_spmf_gpv [transfer_rule]: "left_total A ==> left_total (pcr_spmf_gpv A)"
unfolding pcr_spmf_gpv_def by(intro left_total_OO left_total_cr_spmf_gpv left_total_rel_gpv left_total_eq)
unbundle lifting_syntax
lemma return_spmf_gpv_transfer':
"(op = ===> cr_spmf_gpv) return_spmf Done"
by(rule rel_funI)(simp add: cr_spmf_gpv_def)
lemma return_spmf_gpv_transfer [transfer_rule]:
"(A ===> pcr_spmf_gpv A) return_spmf Done"
unfolding pcr_spmf_gpv_def
apply(rewrite in "(⌑ ===> _) _ _" eq_OO[symmetric])
apply(rule pos_fun_distr[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp])
apply(rule relcomppI)
apply(rule return_spmf_gpv_transfer')
apply transfer_prover
done
lemma bind_spmf_gpv_transfer':
"(cr_spmf_gpv ===> (op = ===> cr_spmf_gpv) ===> cr_spmf_gpv) bind_spmf bind_gpv"
apply(clarsimp simp add: rel_fun_def cr_spmf_gpv_def)
apply(rule gpv.expand)
apply(simp add: bind_map_spmf map_spmf_bind_spmf o_def)
done
lemma bind_spmf_gpv_transfer [transfer_rule]:
"(pcr_spmf_gpv A ===> (A ===> pcr_spmf_gpv B) ===> pcr_spmf_gpv B) bind_spmf bind_gpv"
unfolding pcr_spmf_gpv_def
apply(rewrite in "(_ ===> (⌑ ===> _) ===> _) _ _" eq_OO[symmetric])
apply(rule fun_mono[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp])
apply(rule order.refl)
apply(rule fun_mono)
apply(rule neg_fun_distr3[OF left_unique_eq right_total_eq right_unique_cr_spmf_gpv left_total_cr_spmf_gpv])
apply(rule order.refl)
apply(rule fun_mono[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp])
apply(rule order.refl)
apply(rule pos_fun_distr)
apply(rule pos_fun_distr[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp])
apply(rule relcomppI)
apply(rule bind_spmf_gpv_transfer')
apply transfer_prover
done
lemma lift_spmf_gpv_transfer':
"(op = ===> cr_spmf_gpv) (λx. x) lift_spmf"
by(simp add: rel_fun_def cr_spmf_gpv_def)
lemma lift_spmf_gpv_transfer [transfer_rule]:
"(rel_spmf A ===> pcr_spmf_gpv A) (λx. x) lift_spmf"
unfolding pcr_spmf_gpv_def
apply(rewrite in "(⌑ ===> _) _ _" eq_OO[symmetric])
apply(rule pos_fun_distr[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp])
apply(rule relcomppI)
apply(rule lift_spmf_gpv_transfer')
apply transfer_prover
done
lemma fail_spmf_gpv_transfer': "cr_spmf_gpv (return_pmf None) Fail"
by(simp add: cr_spmf_gpv_def)
lemma fail_spmf_gpv_transfer [transfer_rule]: "pcr_spmf_gpv A (return_pmf None) Fail"
unfolding pcr_spmf_gpv_def
apply(rule relcomppI)
apply(rule fail_spmf_gpv_transfer')
apply transfer_prover
done
lemma map_spmf_gpv_transfer':
"(op = ===> R ===> cr_spmf_gpv ===> cr_spmf_gpv) (λf g. map_spmf f) map_gpv"
by(simp add: rel_fun_def cr_spmf_gpv_def map_lift_spmf)
lemma map_spmf_gpv_transfer [transfer_rule]:
"((A ===> B) ===> R ===> pcr_spmf_gpv A ===> pcr_spmf_gpv B) (λf g. map_spmf f) map_gpv"
unfolding pcr_spmf_gpv_def
apply(rewrite in "((⌑ ===> _) ===> _) _ _" eq_OO[symmetric])
apply(rewrite in "((_ ===> ⌑) ===> _) _ _" eq_OO[symmetric])
apply(rewrite in "(_ ===> ⌑ ===> _) _ _" OO_eq[symmetric])
apply(rule fun_mono[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp])
apply(rule neg_fun_distr3[OF left_unique_eq right_total_eq right_unique_eq left_total_eq])
apply(rule fun_mono[OF order.refl])
apply(rule pos_fun_distr)
apply(rule fun_mono[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp])
apply(rule order.refl)
apply(rule pos_fun_distr)
apply(rule pos_fun_distr[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp])
apply(rule relcomppI)
apply(unfold rel_fun_eq)
apply(rule map_spmf_gpv_transfer')
apply(unfold rel_fun_eq[symmetric])
apply transfer_prover
done
end
subsection {* Embedding resumptions *}
primcorec lift_resumption :: "('a, 'out, 'in) resumption ⇒ ('a, 'out, 'in) gpv"
where
"the_gpv (lift_resumption r) =
(case r of resumption.Done None ⇒ return_pmf None
| resumption.Done (Some x') => return_spmf (Pure x')
| resumption.Pause out c => map_spmf (map_generat id id (op ∘ lift_resumption)) (return_spmf (IO out c)))"
lemma the_gpv_lift_resumption:
"the_gpv (lift_resumption r) =
(if is_Done r then if Option.is_none (resumption.result r) then return_pmf None else return_spmf (Pure (the (resumption.result r)))
else return_spmf (IO (resumption.output r) (lift_resumption ∘ resume r)))"
by(simp split: option.split resumption.split)
declare lift_resumption.simps [simp del]
lemma lift_resumption_Done [code]:
"lift_resumption (resumption.Done x) = (case x of None ⇒ Fail | Some x' ⇒ Done x')"
by(rule gpv.expand)(simp add: the_gpv_lift_resumption split: option.split)
lemma lift_resumption_DONE [simp]:
"lift_resumption (DONE x) = Done x"
by(simp add: DONE_def lift_resumption_Done)
lemma lift_resumption_ABORT [simp]:
"lift_resumption ABORT = Fail"
by(simp add: ABORT_def lift_resumption_Done)
lemma lift_resumption_Pause [simp, code]:
"lift_resumption (resumption.Pause out c) = Pause out (lift_resumption ∘ c)"
by(rule gpv.expand)(simp add: the_gpv_lift_resumption)
subsection {* Assertions *}
definition assert_gpv :: "bool ⇒ (unit, 'out, 'in) gpv"
where "assert_gpv b = (if b then Done () else Fail)"
lemma assert_gpv_simps [simp]:
"assert_gpv True = Done ()"
"assert_gpv False = Fail"
by(simp_all add: assert_gpv_def)
lemma [simp]:
shows assert_gpv_eq_Done: "assert_gpv b = Done x ⟷ b"
and Done_eq_assert_gpv: "Done x = assert_gpv b ⟷ b"
and Pause_neq_assert_gpv: "Pause out rpv ≠ assert_gpv b"
and assert_gpv_neq_Pause: "assert_gpv b ≠ Pause out rpv"
and assert_gpv_eq_Fail: "assert_gpv b = Fail ⟷ ¬ b"
and Fail_eq_assert_gpv: "Fail = assert_gpv b ⟷ ¬ b"
by(simp_all add: assert_gpv_def)
lemma assert_gpv_inject [simp]: "assert_gpv b = assert_gpv b' ⟷ b = b'"
by(simp add: assert_gpv_def)
lemma assert_gpv_sel [simp]:
"the_gpv (assert_gpv b) = map_spmf Pure (assert_spmf b)"
by(simp add: assert_gpv_def)
lemma the_gpv_bind_assert [simp]:
"the_gpv (bind_gpv (assert_gpv b) f) =
bind_spmf (assert_spmf b) (the_gpv ∘ f)"
by(cases b) simp_all
primcorec try_gpv :: "('a, 'call, 'ret) gpv ⇒ ('a, 'call, 'ret) gpv ⇒ ('a, 'call, 'ret) gpv" ("TRY _ ELSE _" [0,60] 59)
where
"the_gpv (TRY gpv ELSE gpv') =
map_spmf (map_generat id id (λc input. case c input of Inl gpv ⇒ try_gpv gpv gpv' | Inr gpv' ⇒ gpv'))
(try_spmf (map_spmf (map_generat id id (map_fun id Inl)) (the_gpv gpv))
(map_spmf (map_generat id id (map_fun id Inr)) (the_gpv gpv')))"
lemma try_gpv_sel:
"the_gpv (TRY gpv ELSE gpv') =
TRY map_spmf (map_generat id id (λc input. TRY c input ELSE gpv')) (the_gpv gpv) ELSE the_gpv gpv'"
by(simp add: try_gpv_def map_try_spmf spmf.map_comp o_def generat.map_comp generat.map_ident id_def)
lemma try_gpv_Done [simp]: "TRY Done x ELSE gpv' = Done x"
by(rule gpv.expand)(simp)
lemma try_gpv_Fail [simp]: "TRY Fail ELSE gpv' = gpv'"
by(rule gpv.expand)(simp add: spmf.map_comp o_def generat.map_comp generat.map_ident)
lemma try_gpv_Pause [simp]: "TRY Pause out c ELSE gpv' = Pause out (λinput. TRY c input ELSE gpv')"
by(rule gpv.expand) simp
lemma try_gpv_Fail2 [simp]: "TRY gpv ELSE Fail = gpv"
by(coinduction arbitrary: gpv rule: gpv.coinduct_strong)
(auto simp add: spmf_rel_map generat.rel_map intro!: rel_spmf_reflI generat.rel_refl)
lemma lift_try_spmf: "lift_spmf (TRY p ELSE q) = TRY lift_spmf p ELSE lift_spmf q"
by(rule gpv.expand)(simp add: map_try_spmf spmf.map_comp o_def)
lemma try_assert_gpv: "TRY assert_gpv b ELSE gpv' = (if b then Done () else gpv')"
by(simp)
context includes lifting_syntax begin
lemma try_gpv_parametric [transfer_rule]:
"(rel_gpv A C ===> rel_gpv A C ===> rel_gpv A C) try_gpv try_gpv"
unfolding try_gpv_def by transfer_prover
end
lemma map_try_gpv: "map_gpv f g (TRY gpv ELSE gpv') = TRY map_gpv f g gpv ELSE map_gpv f g gpv'"
by(simp add: gpv.rel_map try_gpv_parametric[THEN rel_funD, THEN rel_funD] gpv.rel_refl gpv.rel_eq[symmetric])
lemma try_bind_assert_gpv:
"TRY (assert_gpv b ⤜ f) ELSE gpv = (if b then TRY (f ()) ELSE gpv else gpv)"
by(simp)
subsection {* Executing gpv: possibilistic trace semantics *}
type_synonym ('a, 'out, 'in) raw_trace = "(('out, 'in) event, 'a option) tllist"
context IO_execution_base begin
-- {* Work around the limitation that \isacommand{inductive\_set} does not support instantiated parameters *}
coinductive tracesp_gpv :: "('a, 'out, 'in) gpv ⇒ ('a, 'out, 'in) raw_trace ⇒ bool"
where
Fail: "None ∈ set_pmf (the_gpv r) ⟹ tracesp_gpv r (TNil None)"
| Pure: "Pure x ∈ set_spmf (the_gpv r) ⟹ tracesp_gpv r (TNil (Some x))"
| IO:
"⟦ IO out c ∈ set_spmf (the_gpv r); input ∈ wf_response out; tracesp_gpv (c input) tr ⟧
⟹ tracesp_gpv r (TCons (Event out input) tr)"
definition traces_gpv :: "('a, 'out, 'in) gpv ⇒ ('a, 'out, 'in) raw_trace set"
where "traces_gpv r ≡ {tr. tracesp_gpv r tr}"
lemma tracesp_gpv_traces_gpv_eq [pred_set_conv]: "tracesp_gpv r = (λtr. tr ∈ traces_gpv r)" by(simp add: traces_gpv_def)
context begin
local_setup {* Local_Theory.map_background_naming (Name_Space.mandatory_path "traces_gpv") *}
lemmas intros [intro?] = tracesp_gpv.intros[to_set]
and coinduct [consumes 1, case_names Fail Pure IO, coinduct set: traces_gpv] = tracesp_gpv.coinduct[to_set]
and cases [consumes 1, case_names Fail Pure IO, cases set: traces_gpv] = tracesp_gpv.cases[to_set]
and simps = tracesp_gpv.simps[to_set]
end
lemmas traces_gpv_intros [intro?] = tracesp_gpv.intros[to_set]
and traces_gpv_cases [consumes 1, case_names Fail Pure IO, cases set: traces_gpv] = tracesp_gpv.cases[to_set]
and traces_gpv_simps = tracesp_gpv.simps[to_set]
inductive_simps TNil_None_in_traces_gpv [to_set, simp]: "tracesp_gpv r (TNil None)"
inductive_simps TNil_Some_in_traces_gpv [to_set, simp]: "tracesp_gpv r (TNil (Some x))"
inductive_simps TCons_in_traces_gpv [to_set, simp]: "tracesp_gpv r (TCons (Event out input) tr)"
lemma traces_gpv_coinduct [consumes 1, case_names Fail Pure IO, case_conclusion IO wf_response ttl, coinduct set: traces_gpv]:
assumes "X r tr"
and Fail: "⋀r tr. ⟦ X r tr; tr = TNil None ⟧ ⟹ None ∈ set_pmf (the_gpv r)"
and Pure: "⋀r tr x. ⟦ X r tr; tr = TNil (Some x) ⟧ ⟹ Pure x ∈ set_spmf (the_gpv r)"
and IO: "⋀r out input tr tr'. ⟦ X r tr; tr = TCons (Event out input) tr' ⟧
⟹ input ∈ wf_response out ∧ (∃c. IO out c ∈ set_spmf (the_gpv r) ∧ (X (c input) tr' ∨ tr' ∈ traces_gpv (c input)))"
shows "tr ∈ traces_gpv r"
using ‹X r tr›
proof(coinduct rule: tracesp_gpv.coinduct[to_set])
case (tracesp_gpv r tr)
show ?case
proof(cases tr)
case (TNil xo)
show ?thesis
proof(cases xo)
case None
hence ?Fail using TNil tracesp_gpv by(simp add: Fail)
thus ?thesis ..
next
case (Some x)
hence ?Pure using TNil tracesp_gpv by(simp add: Pure)
thus ?thesis by simp
qed
next
case (TCons event tr')
obtain out input where "event = Event out input" by(cases event)
with IO[OF ‹X r tr›] TCons tracesp_gpv have ?IO by auto
thus ?thesis by simp
qed
qed
lemma traces_gpv_coinduct' [consumes 1, case_names Fail Pure IO, case_conclusion IO wf_response ttl]:
assumes "X r tr"
and Fail: "⋀r tr. ⟦ X r tr; is_TNil tr; Option.is_none (terminal tr) ⟧ ⟹ None ∈ set_pmf (the_gpv r)"
and Pure: "⋀r tr. ⟦ X r tr; is_TNil tr; ¬ Option.is_none (terminal tr) ⟧ ⟹ Pure (the (terminal tr)) ∈ set_spmf (the_gpv r)"
and IO: "⋀r tr. ⟦ X r tr; ¬ is_TNil tr ⟧
⟹ event_in (thd tr) ∈ wf_response (event_out (thd tr)) ∧ (∃c. IO (event_out (thd tr)) c ∈ set_spmf (the_gpv r) ∧ (X (c (event_in (thd tr))) (ttl tr) ∨ ttl tr ∈ traces_gpv (c (event_in (thd tr)))))"
shows "tr ∈ traces_gpv r"
using ‹X r tr›
apply(coinduct)
apply(drule Fail, simp, simp, simp)
apply(drule Pure, simp, simp, simp)
apply(drule IO, simp, simp)
done
lemma traces_gpv_cases':
assumes "tr ∈ traces_gpv r"
obtains (Fail) "is_TNil tr" "Option.is_none (terminal tr)" "None ∈ set_pmf (the_gpv r)"
| (Pure) "is_TNil tr" "¬ Option.is_none (terminal tr)" "Pure (the (terminal tr)) ∈ set_spmf (the_gpv r)"
| (IO) out c where "¬ is_TNil tr" "out = event_out (thd tr)" "IO out c ∈ set_spmf (the_gpv r)"
"event_in (thd tr) ∈ wf_response out" "ttl tr ∈ traces_gpv (c (event_in (thd tr)))"
using assms by cases auto
end
subsection {* Order for @{typ "('a, 'out, 'in) gpv"} *}
coinductive ord_gpv :: "('a, 'out, 'in) gpv ⇒ ('a, 'out, 'in) gpv ⇒ bool"
where
"ord_spmf (rel_generat op = op = (rel_fun op = ord_gpv)) f g ⟹ ord_gpv (GPV f) (GPV g)"
inductive_simps ord_gpv_simps [simp]:
"ord_gpv (GPV f) (GPV g)"
lemma ord_gpv_coinduct [consumes 1, case_names ord_gpv, coinduct pred: ord_gpv]:
assumes "X f g"
and step: "⋀f g. X f g ⟹ ord_spmf (rel_generat op = op = (rel_fun op = X)) (the_gpv f) (the_gpv g)"
shows "ord_gpv f g"
using `X f g`
by(coinduct)(auto dest: step simp add: eq_GPV_iff intro: ord_spmf_mono rel_generat_mono rel_fun_mono)
lemma ord_gpv_the_gpvD:
"ord_gpv f g ⟹ ord_spmf (rel_generat op = op = (rel_fun op = ord_gpv)) (the_gpv f) (the_gpv g)"
by(erule ord_gpv.cases) simp
lemma reflp_equality: "reflp op ="
by(simp add: reflp_def)
lemma ord_gpv_reflI [simp]: "ord_gpv f f"
by(coinduction arbitrary: f)(auto intro: ord_spmf_reflI simp add: rel_generat_same rel_fun_def)
lemma reflp_ord_gpv: "reflp ord_gpv"
by(rule reflpI)(rule ord_gpv_reflI)
lemma ord_gpv_trans:
assumes "ord_gpv f g" "ord_gpv g h"
shows "ord_gpv f h"
using assms
proof(coinduction arbitrary: f g h)
case (ord_gpv f g h)
have *: "ord_spmf (rel_generat op = op = (rel_fun op = (λf h. ∃g. ord_gpv f g ∧ ord_gpv g h))) (the_gpv f) (the_gpv h) =
ord_spmf (rel_generat (op = OO op =) (op = OO op =) (rel_fun op = (ord_gpv OO ord_gpv))) (the_gpv f) (the_gpv h)"
by(simp add: relcompp.simps[abs_def])
then show ?case using ord_gpv
by(auto elim!: ord_gpv.cases simp add: generat.rel_compp ord_spmf_compp fun.rel_compp)
qed
lemma ord_gpv_compp: "(ord_gpv OO ord_gpv) = ord_gpv"
by(auto simp add: fun_eq_iff intro: ord_gpv_trans)
lemma transp_ord_gpv [simp]: "transp ord_gpv"
by(blast intro: transpI ord_gpv_trans)
lemma ord_gpv_antisym:
"⟦ ord_gpv f g; ord_gpv g f ⟧ ⟹ f = g"
proof(coinduction arbitrary: f g)
case (Eq_gpv f g)
let ?R = "rel_generat op = op = (rel_fun op = ord_gpv)"
from ‹ord_gpv f g› have "ord_spmf ?R (the_gpv f) (the_gpv g)" by cases simp
moreover
from ‹ord_gpv g f› have "ord_spmf ?R (the_gpv g) (the_gpv f)" by cases simp
ultimately have "rel_spmf (inf ?R ?R¯¯) (the_gpv f) (the_gpv g)"
by(rule rel_spmf_inf)(auto 4 3 intro: transp_rel_generatI transp_ord_gpv reflp_ord_gpv reflp_equality reflp_fun1 is_equality_eq transp_rel_fun)
also have "inf ?R ?R¯¯ = rel_generat (inf op = op =) (inf op = op =) (rel_fun op = (inf ord_gpv ord_gpv¯¯))"
unfolding rel_generat_inf[symmetric] rel_fun_inf[symmetric]
by(simp add: generat.rel_conversep[symmetric] fun.rel_conversep)
finally show ?case by(simp add: inf_fun_def)
qed
lemma RFail_least [simp]: "ord_gpv Fail f"
by(coinduction arbitrary: f)(simp add: eq_GPV_iff)
subsection {* Bounds on interaction *}
context notes monotone_SUP[partial_function_mono] begin
declaration {* Partial_Function.init "lfp_strong" @{term lfp.fixp_fun} @{term lfp.mono_body}
@{thm lfp.fixp_rule_uc} @{thm lfp.fixp_strong_induct_uc} NONE *}
partial_function (lfp_strong) interaction_bound :: "('a, 'out, 'in) gpv ⇒ enat"
where
"interaction_bound gpv =
(SUP generat:set_spmf (the_gpv gpv). case generat of Pure _ ⇒ 0 | IO out c ⇒ eSuc (SUP input. interaction_bound (c input)))"
end
lemma interaction_bound_fixp_induct [case_names adm bottom step]:
"⟦ ccpo.admissible (fun_lub Sup) (fun_ord op ≤) P;
P (λ_. 0);
⋀interaction_bound'. ⟦ P interaction_bound'; ⋀gpv. interaction_bound' gpv ≤ interaction_bound gpv ⟧
⟹ P (λgpv. ⨆generat∈set_spmf (the_gpv gpv). case generat of Pure x ⇒ 0
| IO out c ⇒ eSuc (⨆input. interaction_bound' (c input))) ⟧
⟹ P interaction_bound"
by(erule interaction_bound.fixp_induct)(simp_all add: bot_enat_def fun_ord_def)
lemma interaction_bound_coinduct [consumes 1, case_names interaction_bound]:
assumes X: "X gpv n"
and *: "⋀gpv n out c input. ⟦ X gpv n; IO out c ∈ set_spmf (the_gpv gpv) ⟧
⟹ ∃n'. (X (c input) n' ∨ interaction_bound (c input) ≤ n') ∧ eSuc n' ≤ n"
shows "interaction_bound gpv ≤ n"
using X
proof(induction arbitrary: gpv n rule: interaction_bound_fixp_induct)
case adm show ?case by(intro cont_intro)
case bottom show ?case by simp
next
case (step interaction_bound')
{ fix out c
assume IO: "IO out c ∈ set_spmf (the_gpv gpv)"
from *[OF step.prems IO] obtain n' where n: "n = eSuc n'"
by(cases n rule: co.enat.exhaust) auto
moreover
{ fix input
have "∃n''. (X (c input) n'' ∨ interaction_bound (c input) ≤ n'') ∧ eSuc n'' ≤ n"
using step.prems IO ‹n = eSuc n'› by(auto 4 3 dest: *)
then have "interaction_bound' (c input) ≤ n'" using n
by(auto dest: step.IH intro: step.hyps[THEN order_trans] elim!: order_trans simp add: neq_zero_conv_eSuc) }
ultimately have "eSuc (⨆input. interaction_bound' (c input)) ≤ n"
by(auto intro: SUP_least) }
then show ?case by(auto intro!: SUP_least split: generat.split)
qed
lemma interaction_bound_IO:
"IO out c ∈ set_spmf (the_gpv gpv)
⟹ eSuc (interaction_bound (c input)) ≤ interaction_bound gpv"
by(rewrite in "_ ≤ ⌑" interaction_bound.simps)(auto intro!: SUP_upper2)
lemma interaction_bound_Done [simp]: "interaction_bound (Done x) = 0"
by(simp add: interaction_bound.simps)
lemma interaction_bound_Fail [simp]: "interaction_bound Fail = 0"
by(simp add: interaction_bound.simps bot_enat_def)
lemma interaction_bound_Pause [simp]: "interaction_bound (Pause out c) = eSuc (SUP input. interaction_bound (c input))"
by(simp add: interaction_bound.simps)
lemma interaction_bound_lift_spmf [simp]: "interaction_bound (lift_spmf p) = 0"
by(simp add: interaction_bound.simps SUP_constant bot_enat_def)
lemma interaction_bound_assert_gpv [simp]: "interaction_bound (assert_gpv b) = 0"
by(cases b) simp_all
lemma interaction_bound_bind:
defines "ib1 ≡ interaction_bound" and "ib2 ≡ interaction_bound"
shows "interaction_bound (p ⤜ f) ≤ ib1 p + (SUP x:results'_gpv p. ib2 (f x))"
proof(induction arbitrary: p rule: interaction_bound_fixp_induct)
case adm show ?case by simp
case bottom show ?case by simp
case (step interaction_bound') show ?case unfolding ib1_def ib2_def
proof(rule SUP_least)
fix generat'
assume "generat' ∈ set_spmf (the_gpv (p ⤜ f))"
then obtain generat where generat: "generat ∈ set_spmf (the_gpv p)"
and *: "case generat of Pure x ⇒ generat' ∈ set_spmf (the_gpv (f x))
| IO out c ⇒ generat' = IO out (λinput. c input ⤜ f)"
by(clarsimp simp add: set_bind_spmf bind_gpv.sel simp del: bind_gpv_sel')
(clarsimp split: generat.split_asm simp add: generat.map_comp o_def generat.map_id[unfolded id_def])
show "(case generat' of Pure x ⇒ 0 | IO out c ⇒ eSuc (⨆input. interaction_bound' (c input)))
≤ interaction_bound p + (⨆x∈results'_gpv p. interaction_bound (f x))"
(is "?lhs ≤ ?p + ?f")
proof(cases generat)
case (Pure x)
have "?lhs ≤ (case generat' of Pure x ⇒ 0 | IO out c ⇒ eSuc (⨆input. interaction_bound (c input)))"
by(cases generat')(auto intro: step.hyps SUP_mono)
also have "… ≤ (SUP generat':set_spmf (the_gpv (f x)). (case generat' of Pure x ⇒ 0 | IO out c ⇒ eSuc (⨆input. interaction_bound (c input))))"
using * Pure by(auto intro: SUP_upper)
also have "… ≤ 0 + ?f" using generat Pure
by(rewrite in "_ ≤ ⌑" interaction_bound.simps)(auto 4 3 intro: SUP_upper results'_gpv_Pure)
also have "… ≤ ?p + ?f" by simp
finally show ?thesis .
next
case (IO out c)
with * have "?lhs = eSuc (SUP input. interaction_bound' (c input ⤜ f))" by simp
also have "… ≤ eSuc (SUP input. interaction_bound (c input) + (SUP x:results'_gpv (c input). interaction_bound (f x)))"
by(auto intro: SUP_mono step.IH[unfolded ib1_def ib2_def])
also have "… ≤ (case IO out c of Pure (x :: 'a) ⇒ 0 | IO out c ⇒ eSuc (SUP input. interaction_bound (c input))) + (SUP input. SUP x:results'_gpv (c input). interaction_bound (f x))"
by(simp add: iadd_Suc)(metis (mono_tags, lifting) SUP_le_iff add_mono_thms_linordered_semiring(1) dual_order.refl)
also have "… ≤ ?p + ?f"
apply(rewrite in "_ ≤ ⌑" interaction_bound.simps)
apply(rule add_mono SUP_least SUP_upper generat[unfolded IO])+
apply(auto intro: results'_gpv_Cont[OF generat]simp add: IO)
done
finally show ?thesis .
qed
qed
qed
lemma interaction_bound_bind_lift_spmf [simp]:
"interaction_bound (lift_spmf p ⤜ f) = (SUP x:set_spmf p. interaction_bound (f x))"
by(subst (1 2) interaction_bound.simps)(simp add: set_bind_spmf bind_UNION SUP_UNION)
text ‹
There is no nice @{const interaction_bound} equation for @{const bind_gpv}, as it computes
an exact bound, but we only need an upper bound.
As @{typ enat} is hard to work with (and @{term ∞} does not constrain a gpv in any way),
we work with @{typ nat}.
›
inductive interaction_bounded_by :: "('a, 'out, 'in) gpv ⇒ nat ⇒ bool"
for gpv n where
interaction_bounded_by: "⟦ interaction_bound gpv ≤ enat n ⟧ ⟹ interaction_bounded_by gpv n"
lemmas interaction_bounded_byI = interaction_bounded_by
hide_fact (open) interaction_bounded_by
lemma interaction_bounded_by_mono:
"⟦ interaction_bounded_by gpv n; n ≤ m ⟧ ⟹ interaction_bounded_by gpv m"
unfolding interaction_bounded_by.simps by(erule order_trans) simp
lemma interaction_bounded_by_contD:
"⟦ interaction_bounded_by gpv n; IO out c ∈ set_spmf (the_gpv gpv) ⟧
⟹ n > 0 ∧ interaction_bounded_by (c input) (n - 1)"
unfolding interaction_bounded_by.simps
by(subst (asm) interaction_bound.simps)(auto simp add: SUP_le_iff eSuc_le_iff enat_eSuc_iff dest!: bspec)
lemma interaction_bounded_byI_epred:
assumes "⋀out c. IO out c ∈ set_spmf (the_gpv gpv) ⟹ n ≠ 0 ∧ (∀input. interaction_bounded_by (c input) (n - 1))"
shows "interaction_bounded_by gpv n"
unfolding interaction_bounded_by.simps
by(subst interaction_bound.simps)(auto intro!: SUP_least split: generat.split dest!: assms simp add: eSuc_le_iff enat_eSuc_iff gr0_conv_Suc neq_zero_conv_eSuc interaction_bounded_by.simps)
lemma interaction_bounded_by_IO:
"⟦ IO out c ∈ set_spmf (the_gpv gpv); interaction_bounded_by gpv n ⟧
⟹ n ≠ 0 ∧ interaction_bounded_by (c input) (n - 1)"
by(drule interaction_bound_IO[where input=input])(auto simp add: interaction_bounded_by.simps epred_conv_minus eSuc_le_iff enat_eSuc_iff)
named_theorems interaction_bound
lemma interaction_bounded_by_start:
"⟦ TERM gpv; interaction_bounded_by gpv m; interaction_bounded_by gpv m ⟹ thesis ⟧ ⟹ thesis"
by simp
lemma interaction_bounded_by_Done [iff]: "interaction_bounded_by (Done x) n"
by(simp add: interaction_bounded_by.simps)
lemma interaction_bounded_by_DoneI [interaction_bound]: "interaction_bounded_by (Done x) 0"
by simp
lemma interaction_bounded_by_Fail [iff]: "interaction_bounded_by Fail n"
by(simp add: interaction_bounded_by.simps)
lemma interaction_bounded_by_FailI [interaction_bound]: "interaction_bounded_by Fail 0"
by simp
lemma interaction_bounded_by_lift_spmf [iff]: "interaction_bounded_by (lift_spmf p) n"
by(simp add: interaction_bounded_by.simps)
lemma interaction_bounded_by_lift_spmfI [interaction_bound]: "interaction_bounded_by (lift_spmf p) 0"
by simp
lemma interaction_bounded_by_assert_gpv [simp]: "interaction_bounded_by (assert_gpv b) n"
by(cases b) simp_all
lemma interaction_bounded_by_assert_gpvI [interaction_bound]: "interaction_bounded_by (assert_gpv b) 0"
by simp
lemma interaction_bounded_by_Pause [simp]:
"interaction_bounded_by (Pause out c) n ⟷ 0 < n ∧ (∀input. interaction_bounded_by (c input) (n - 1))"
by(auto 4 3 simp add: interaction_bounded_by.simps eSuc_le_iff enat_eSuc_iff gr0_conv_Suc intro: SUP_least dest: order_trans[OF SUP_upper, rotated])
lemma interaction_bounded_by_bind_PauseI [interaction_bound]:
"(⋀input. interaction_bounded_by (c input ⤜ f) n)
⟹ interaction_bounded_by (Pause out c ⤜ f) (n + 1)"
by(simp add: plus_1_eSuc)
lemma interaction_bounded_by_bindI [interaction_bound]:
"⟦ interaction_bounded_by gpv n; ⋀x. x ∈ results'_gpv gpv ⟹ interaction_bounded_by (f x) m ⟧
⟹ interaction_bounded_by (gpv ⤜ f) (n + m)"
unfolding interaction_bounded_by.simps plus_enat_simps(1)[symmetric]
by(rule order_trans[OF interaction_bound_bind])(erule add_mono, simp add: SUP_least)
lemma interaction_bounded_by_bind_lift_spmf [iff]:
"interaction_bounded_by (lift_spmf p ⤜ f) n ⟷ (∀x∈set_spmf p. interaction_bounded_by (f x) n)"
by(simp add: interaction_bounded_by.simps SUP_le_iff)
lemma interaction_bounded_by_bind_lift_spmfI [interaction_bound]:
"(⋀x. x ∈ set_spmf p ⟹ interaction_bounded_by (f x) n)
⟹ interaction_bounded_by (lift_spmf p ⤜ f) n"
by(simp)
lemma interaction_bounded_by_bind_DoneI [interaction_bound]:
"interaction_bounded_by (f x) n ⟹ interaction_bounded_by (Done x ⤜ f) n"
by(simp)
lemma interaction_bounded_by_if [interaction_bound]:
"⟦ b ⟹ interaction_bounded_by gpv1 n; ¬ b ⟹ interaction_bounded_by gpv2 m ⟧
⟹ interaction_bounded_by (if b then gpv1 else gpv2) (max n m)"
by(auto 4 3 simp add: max_def not_le elim: interaction_bounded_by_mono)
subsection {* Typing *}
subsubsection {* Interface between gpvs and rpvs *}
type_synonym ('call, 'ret) ℐ = "'call set × ('call ⇒ 'ret set)"
fun outs_ℐ :: "('call, 'ret) ℐ ⇒ 'call set"
where "outs_ℐ (outs, c) = outs"
lemma outs_ℐ_conv: "outs_ℐ = fst"
by(simp add: fun_eq_iff)
fun responses_ℐ :: "('call, 'ret) ℐ ⇒ 'call ⇒ 'ret set"
where "responses_ℐ (outs, c) = c"
lemma responses_ℐ_conv: "responses_ℐ = snd"
by(simp add: fun_eq_iff)
abbreviation ℐ_top :: "('call, 'ret) ℐ"
where "ℐ_top ≡ (UNIV, λ_. UNIV)"
context begin
qualified inductive resultsp_gpv :: "('out, 'in) ℐ ⇒ 'a ⇒ ('a, 'out, 'in) gpv ⇒ bool"
for Γ x
where
Pure: "Pure x ∈ set_spmf (the_gpv gpv) ⟹ resultsp_gpv Γ x gpv"
| IO:
"⟦ IO out c ∈ set_spmf (the_gpv gpv); input ∈ responses_ℐ Γ out; resultsp_gpv Γ x (c input) ⟧
⟹ resultsp_gpv Γ x gpv"
definition results_gpv :: "('out, 'in) ℐ ⇒ ('a, 'out, 'in) gpv ⇒ 'a set"
where "results_gpv Γ gpv ≡ {x. resultsp_gpv Γ x gpv}"
lemma resultsp_gpv_results_gpv_eq [pred_set_conv]: "resultsp_gpv Γ x gpv ⟷ x ∈ results_gpv Γ gpv"
by(simp add: results_gpv_def)
context begin
local_setup {* Local_Theory.map_background_naming (Name_Space.mandatory_path "results_gpv") *}
lemmas intros [intro?] = resultsp_gpv.intros[to_set]
and Pure = Pure[to_set]
and IO = IO[to_set]
and induct [consumes 1, case_names Pure IO, induct set: results_gpv] = resultsp_gpv.induct[to_set]
and cases [consumes 1, case_names Pure IO, cases set: results_gpv] = resultsp_gpv.cases[to_set]
and simps = resultsp_gpv.simps[to_set]
end
inductive_simps results_gpv_GPV [to_set, simp]: "resultsp_gpv Γ x (GPV gpv)"
end
lemma results_gpv_Done [iff]: "results_gpv Γ (Done x) = {x}"
by(auto simp add: Done.ctr)
lemma results_gpv_Fail [iff]: "results_gpv Γ Fail = {}"
by(auto simp add: Fail_def)
lemma results_gpv_Pause [simp]:
"results_gpv Γ (Pause out c) = (⋃input∈responses_ℐ Γ out. results_gpv Γ (c input))"
by(auto simp add: Pause.ctr)
lemma results_gpv_lift_spmf [iff]: "results_gpv Γ (lift_spmf p) = set_spmf p"
by(auto simp add: lift_spmf.ctr)
lemma results_gpv_assert_gpv [simp]: "results_gpv Γ (assert_gpv b) = (if b then {()} else {})"
by auto
lemma results_gpv_bind_gpv [simp]:
"results_gpv Γ (gpv ⤜ f) = (⋃x∈results_gpv Γ gpv. results_gpv Γ (f x))"
(is "?lhs = ?rhs")
proof(intro set_eqI iffI)
fix x
assume "x ∈ ?lhs"
then show "x ∈ ?rhs"
proof(induction gpv'≡"gpv ⤜ f" arbitrary: gpv)
case Pure thus ?case
by(auto 4 3 simp add: set_bind_spmf split: if_split_asm intro: results_gpv.intros rev_bexI)
next
case (IO out c input)
from ‹IO out c ∈ _›
obtain generat where generat: "generat ∈ set_spmf (the_gpv gpv)"
and *: "IO out c ∈ set_spmf (if is_Pure generat then the_gpv (f (result generat))
else return_spmf (IO (output generat) (λinput. continuation generat input ⤜ f)))"
by(auto simp add: set_bind_spmf)
thus ?case
proof(cases generat)
case (Pure y)
with generat have "y ∈ results_gpv Γ gpv" by(auto intro: results_gpv.intros)
thus ?thesis using * Pure ‹input ∈ responses_ℐ Γ out› ‹x ∈ results_gpv Γ (c input)›
by(auto intro: results_gpv.IO)
next
case (IO out' c')
hence [simp]: "out' = out"
and c: "⋀input. c input = bind_gpv (c' input) f" using * by simp_all
from IO.hyps(4)[OF c] obtain y where y: "y ∈ results_gpv Γ (c' input)"
and "x ∈ results_gpv Γ (f y)" by blast
from y IO generat have "y ∈ results_gpv Γ gpv" using `input ∈ responses_ℐ Γ out`
by(auto intro: results_gpv.IO)
with ‹x ∈ results_gpv Γ (f y)› show ?thesis by blast
qed
qed
next
fix x
assume "x ∈ ?rhs"
then obtain y where y: "y ∈ results_gpv Γ gpv"
and x: "x ∈ results_gpv Γ (f y)" by blast
from y show "x ∈ ?lhs"
proof(induction)
case (Pure gpv)
with x show ?case
by cases(auto 4 4 intro: results_gpv.intros rev_bexI simp add: set_bind_spmf)
qed(auto 4 4 simp add: set_bind_spmf intro: rev_bexI results_gpv.IO)
qed
lemma results_gpv_ℐ_top: "results_gpv ℐ_top = results'_gpv"
proof(intro ext set_eqI iffI)
fix x gpv
assume "x ∈ results_gpv ℐ_top gpv"
then show "x ∈ results'_gpv gpv"
by induction(auto intro: results'_gpvI)
next
fix x gpv
assume "x ∈ results'_gpv gpv"
then show "x ∈ results_gpv ℐ_top gpv"
by induction(auto intro: results_gpv.intros elim!: generat.set_cases)
qed
lemma results'_bind_gpv [simp]:
"results'_gpv (bind_gpv gpv f) = (⋃x∈results'_gpv gpv. results'_gpv (f x))"
unfolding results_gpv_ℐ_top[symmetric] by simp
subsubsection {* Type judgements *}
coinductive WT_gpv :: "('out, 'in) ℐ ⇒ ('a, 'out, 'in) gpv ⇒ bool" ("((_)/ ⊢g (_) √)" [100, 0] 99)
for Γ
where
"(⋀out c. IO out c ∈ set_spmf gpv ⟹ out ∈ outs_ℐ Γ ∧ (∀input∈responses_ℐ Γ out. Γ ⊢g c input √))
⟹ Γ ⊢g GPV gpv √"
lemma WT_gpv_coinduct [consumes 1, case_names WT_gpv, case_conclusion WT_gpv out cont, coinduct pred: WT_gpv]:
assumes *: "X gpv"
and step: "⋀gpv out c.
⟦ X gpv; IO out c ∈ set_spmf (the_gpv gpv) ⟧
⟹ out ∈ outs_ℐ Γ ∧ (∀input ∈ responses_ℐ Γ out. X (c input) ∨ Γ ⊢g c input √)"
shows "Γ ⊢g gpv √"
using * by(coinduct)(auto dest: step simp add: eq_GPV_iff)
lemma WT_gpv_simps:
"Γ ⊢g GPV gpv √ ⟷
(∀out c. IO out c ∈ set_spmf gpv ⟶ out ∈ outs_ℐ Γ ∧ (∀input∈responses_ℐ Γ out. Γ ⊢g c input √))"
by(subst WT_gpv.simps) simp
lemma WT_gpvI:
"(⋀out c. IO out c ∈ set_spmf (the_gpv gpv) ⟹ out ∈ outs_ℐ Γ ∧ (∀input∈responses_ℐ Γ out. Γ ⊢g c input √))
⟹ Γ ⊢g gpv √"
by(cases gpv)(simp add: WT_gpv_simps)
lemma WT_gpvD:
assumes "Γ ⊢g gpv √"
shows WT_gpv_OutD: "IO out c ∈ set_spmf (the_gpv gpv) ⟹ out ∈ outs_ℐ Γ"
and WT_gpv_ContD: "⟦ IO out c ∈ set_spmf (the_gpv gpv); input ∈ responses_ℐ Γ out ⟧ ⟹ Γ ⊢g c input √"
using assms by(cases, fastforce)+
lemma WT_gpv_Done [iff]: "Γ ⊢g Done x √"
by(rule WT_gpvI) simp_all
lemma WT_gpv_Fail [iff]: "Γ ⊢g Fail √"
by(rule WT_gpvI) simp_all
lemma WT_gpv_PauseI:
"⟦ out ∈ outs_ℐ Γ; ⋀input. input ∈ responses_ℐ Γ out ⟹ Γ ⊢g c input √ ⟧
⟹ Γ ⊢g Pause out c √"
by(rule WT_gpvI) simp_all
lemma WT_gpv_Pause [iff]:
"Γ ⊢g Pause out c √ ⟷ out ∈ outs_ℐ Γ ∧ (∀input ∈ responses_ℐ Γ out. Γ ⊢g c input √)"
by(auto intro: WT_gpv_PauseI dest: WT_gpvD)
lemma WT_gpv_bindI:
"⟦ Γ ⊢g gpv √; ⋀x. x ∈ results_gpv Γ gpv ⟹ Γ ⊢g f x √ ⟧
⟹ Γ ⊢g gpv ⤜ f √"
proof(coinduction arbitrary: gpv)
case [rule_format]: (WT_gpv out c gpv)
from ‹IO out c ∈ _›
obtain generat where generat: "generat ∈ set_spmf (the_gpv gpv)"
and *: "IO out c ∈ set_spmf (if is_Pure generat then the_gpv (f (result generat))
else return_spmf (IO (output generat) (λinput. continuation generat input ⤜ f)))"
by(auto simp add: set_bind_spmf)
show ?case
proof(cases generat)
case (Pure y)
with generat have "y ∈ results_gpv Γ gpv" by(auto intro: results_gpv.Pure)
hence "Γ ⊢g f y √" by(rule WT_gpv)
with * Pure show ?thesis by(auto dest: WT_gpvD)
next
case (IO out' c')
hence [simp]: "out' = out"
and c: "⋀input. c input = bind_gpv (c' input) f" using * by simp_all
from generat IO have **: "IO out c' ∈ set_spmf (the_gpv gpv)" by simp
with ‹Γ ⊢g gpv √› have ?out by(auto dest: WT_gpvD)
moreover {
fix input
assume input: "input ∈ responses_ℐ Γ out"
with ‹Γ ⊢g gpv √› ** have "Γ ⊢g c' input √" by(rule WT_gpvD)
moreover {
fix y
assume "y ∈ results_gpv Γ (c' input)"
with ** input have "y ∈ results_gpv Γ gpv" by(rule results_gpv.IO)
hence "Γ ⊢g f y √" by(rule WT_gpv) }
moreover note calculation }
hence ?cont using c by blast
ultimately show ?thesis ..
qed
qed
lemma WT_gpv_bindD2:
assumes WT: "Γ ⊢g gpv ⤜ f √"
and x: "x ∈ results_gpv Γ gpv"
shows "Γ ⊢g f x √"
using x WT
proof induction
case (Pure gpv)
show ?case
proof(rule WT_gpvI)
fix out c
assume "IO out c ∈ set_spmf (the_gpv (f x))"
with Pure have "IO out c ∈ set_spmf (the_gpv (gpv ⤜ f))"
by(auto simp add: set_bind_spmf intro: rev_bexI)
with ‹Γ ⊢g gpv ⤜ f √› show "out ∈ outs_ℐ Γ ∧ (∀input∈responses_ℐ Γ out. Γ ⊢g c input √)"
by(auto dest: WT_gpvD)
qed
next
case (IO out c gpv input)
from ‹IO out c ∈ _›
have "IO out (λinput. bind_gpv (c input) f) ∈ set_spmf (the_gpv (gpv ⤜ f))"
by(auto simp add: set_bind_spmf intro: rev_bexI)
with IO.prems have "Γ ⊢g c input ⤜ f √" using ‹input ∈ _› by(rule WT_gpv_ContD)
thus ?case by(rule IO.IH)
qed
lemma WT_gpv_bindD1: "Γ ⊢g gpv ⤜ f √ ⟹ Γ ⊢g gpv √"
proof(coinduction arbitrary: gpv)
case (WT_gpv out c gpv)
from ‹IO out c ∈ _›
have "IO out (λinput. bind_gpv (c input) f) ∈ set_spmf (the_gpv (gpv ⤜ f))"
by(auto simp add: set_bind_spmf intro: rev_bexI)
with ‹Γ ⊢g gpv ⤜ f √› show ?case
by(auto simp del: bind_gpv_sel' dest: WT_gpvD)
qed
lemma WT_gpv_bind [simp]: "Γ ⊢g gpv ⤜ f √ ⟷ Γ ⊢g gpv √ ∧ (∀x∈results_gpv Γ gpv. Γ ⊢g f x √)"
by(blast intro: WT_gpv_bindI dest: WT_gpv_bindD1 WT_gpv_bindD2)
lemma WT_gpv_top [simp, intro!]: "ℐ_top ⊢g gpv √"
by(coinduction arbitrary: gpv)(auto)
lemma WT_gpv_lift_spmf [simp, intro!]: "ℐ ⊢g lift_spmf p √"
by(rule WT_gpvI) auto
lemma WT_gpv_coinduct_bind [consumes 1, case_names WT_gpv, case_conclusion WT_gpv out cont]:
assumes *: "X gpv"
and step: "⋀gpv out c. ⟦ X gpv; IO out c ∈ set_spmf (the_gpv gpv) ⟧
⟹ out ∈ outs_ℐ ℐ ∧ (∀input∈responses_ℐ ℐ out.
X (c input) ∨
ℐ ⊢g c input √ ∨
(∃(gpv' :: ('b, 'call, 'ret) gpv) f. c input = gpv' ⤜ f ∧ ℐ ⊢g gpv' √ ∧ (∀x∈results_gpv ℐ gpv'. X (f x))))"
shows "ℐ ⊢g gpv √"
proof -
fix x
def gpv' ≡ "Done x :: ('b, 'call, 'ret) gpv"
and f ≡ "(λ_. gpv) :: 'b ⇒ ('a, 'call, 'ret) gpv"
with * have "ℐ ⊢g gpv' √" and "⋀x. x ∈ results_gpv ℐ gpv' ⟹ X (f x)" by simp_all
then have "ℐ ⊢g gpv' ⤜ f √"
proof(coinduction arbitrary: gpv' f rule: WT_gpv_coinduct)
case [rule_format]: (WT_gpv out c gpv')
from ‹IO out c ∈ _›
obtain generat where generat: "generat ∈ set_spmf (the_gpv gpv')"
and *: "IO out c ∈ set_spmf (if is_Pure generat
then the_gpv (f (result generat))
else return_spmf (IO (output generat) (λinput. continuation generat input ⤜ f)))"
by(clarsimp simp add: set_bind_spmf)
show ?case
proof(cases generat)
case (Pure x)
from Pure * have IO: "IO out c ∈ set_spmf (the_gpv (f x))" by simp
from generat Pure have "x ∈ results_gpv ℐ gpv'" by (simp add: results_gpv.Pure)
then have "X (f x)" by(rule WT_gpv)
from step[OF this IO] show ?thesis by(auto 4 4 intro: exI[where x="Done _"])
next
case (IO out' c')
with * have [simp]: "out' = out"
and c: "c = (λinput. c' input ⤜ f)" by simp_all
from IO generat have IO: "IO out c' ∈ set_spmf (the_gpv gpv')" by simp
then have "⋀input. input ∈ responses_ℐ ℐ out ⟹ results_gpv ℐ (c' input) ⊆ results_gpv ℐ gpv'"
by(auto intro: results_gpv.IO)
with WT_gpvD[OF ‹ℐ ⊢g gpv' √› IO] show ?thesis unfolding c using WT_gpv(2) by blast
qed
qed
then show ?thesis unfolding gpv'_def f_def by simp
qed
subsection {* Sub-gpvs *}
context begin
qualified inductive sub_gpvsp :: "('out, 'in) ℐ ⇒ ('a, 'out, 'in) gpv ⇒ ('a, 'out, 'in) gpv ⇒ bool"
for ℐ x
where
base:
"⟦ IO out c ∈ set_spmf (the_gpv gpv); input ∈ responses_ℐ ℐ out; x = c input ⟧
⟹ sub_gpvsp ℐ x gpv"
| cont:
"⟦ IO out c ∈ set_spmf (the_gpv gpv); input ∈ responses_ℐ ℐ out; sub_gpvsp ℐ x (c input) ⟧
⟹ sub_gpvsp ℐ x gpv"
qualified lemma sub_gpvsp_base:
"⟦ IO out c ∈ set_spmf (the_gpv gpv); input ∈ responses_ℐ ℐ out ⟧
⟹ sub_gpvsp ℐ (c input) gpv"
by(rule base) simp_all
definition sub_gpvs :: "('out, 'in) ℐ ⇒ ('a, 'out, 'in) gpv ⇒ ('a, 'out, 'in) gpv set"
where "sub_gpvs ℐ gpv ≡ {x. sub_gpvsp ℐ x gpv}"
lemma sub_gpvsp_sub_gpvs_eq [pred_set_conv]: "sub_gpvsp ℐ x gpv ⟷ x ∈ sub_gpvs ℐ gpv"
by(simp add: sub_gpvs_def)
context begin
local_setup {* Local_Theory.map_background_naming (Name_Space.mandatory_path "sub_gpvs") *}
lemmas intros [intro?] = sub_gpvsp.intros[to_set]
and base = sub_gpvsp_base[to_set]
and cont = cont[to_set]
and induct [consumes 1, case_names Pure IO, induct set: sub_gpvs] = sub_gpvsp.induct[to_set]
and cases [consumes 1, case_names Pure IO, cases set: sub_gpvs] = sub_gpvsp.cases[to_set]
and simps = sub_gpvsp.simps[to_set]
end
end
lemma WT_sub_gpvsD:
assumes "ℐ ⊢g gpv √" and "gpv' ∈ sub_gpvs ℐ gpv"
shows "ℐ ⊢g gpv' √"
using assms(2,1) by(induction)(auto dest: WT_gpvD)
lemma WT_sub_gpvsI:
"⟦ ⋀out c. IO out c ∈ set_spmf (the_gpv gpv) ⟹ out ∈ outs_ℐ Γ;
⋀gpv'. gpv' ∈ sub_gpvs Γ gpv ⟹ Γ ⊢g gpv' √ ⟧
⟹ Γ ⊢g gpv √"
by(rule WT_gpvI)(auto intro: sub_gpvs.base)
subsection {* Losslessness *}
text {* A gpv is lossless iff we are guaranteed to get a result after a finite number of interactions
that respect the interface *}
inductive lossless_gpv :: "('out, 'in) ℐ ⇒ ('a, 'out, 'in) gpv ⇒ bool"
for ℐ
where
lossless_gpv:
"⟦ lossless_spmf p;
⋀out c input. ⟦ IO out c ∈ set_spmf p; input ∈ responses_ℐ ℐ out ⟧
⟹ lossless_gpv ℐ (c input) ⟧
⟹ lossless_gpv ℐ (GPV p)"
hide_fact lossless_gpv
lemma lossless_gpvI [intro?]:
"⟦ lossless_spmf (the_gpv gpv);
⋀out c input. ⟦ IO out c ∈ set_spmf (the_gpv gpv); input ∈ responses_ℐ ℐ out ⟧
⟹ lossless_gpv ℐ (c input) ⟧
⟹ lossless_gpv ℐ gpv"
by(cases gpv)(auto intro: lossless_gpv.intros)
lemma lossless_gpvD:
assumes "lossless_gpv ℐ gpv"
shows lossless_gpv_lossless_spmfD: "lossless_spmf (the_gpv gpv)"
and lossless_gpv_continuationD:
"⟦ IO out c ∈ set_spmf (the_gpv gpv); input ∈ responses_ℐ ℐ out ⟧ ⟹ lossless_gpv ℐ (c input)"
using assms by(auto elim: lossless_gpv.cases)
lemma lossless_gpv_Done [iff]: "lossless_gpv ℐ (Done x)"
by(rule lossless_gpvI) auto
lemma lossless_gpv_Fail [iff]: "¬ lossless_gpv ℐ Fail"
by(auto dest: lossless_gpvD)
lemma lossless_gpv_Pause [simp]:
"lossless_gpv ℐ (Pause out c) ⟷ (∀input ∈ responses_ℐ ℐ out. lossless_gpv ℐ (c input))"
by(auto dest: lossless_gpvD intro: lossless_gpvI)
lemma lossless_gpv_assert_gpv [iff]: "lossless_gpv ℐ (assert_gpv b) ⟷ b"
by(cases b) simp_all
lemma lossless_gpv_lift_spmf [iff]: "lossless_gpv ℐ (lift_spmf p) ⟷ lossless_spmf p"
by(auto dest: lossless_gpvD intro: lossless_gpvI)
lemma lossless_bind_gpvI:
"⟦ lossless_gpv ℐ gpv; ⋀x. x ∈ results_gpv ℐ gpv ⟹ lossless_gpv ℐ (f x) ⟧
⟹ lossless_gpv ℐ (gpv ⤜ f)"
proof(induction rule: lossless_gpv.induct)
case (lossless_gpv p)
{ fix x
assume "Pure x ∈ set_spmf p"
hence "x ∈ results_gpv ℐ (GPV p)" by simp
hence "lossless_gpv ℐ (f x)" by(rule lossless_gpv.prems) }
with ‹lossless_spmf p› show ?case unfolding GPV_bind
apply(intro lossless_gpv.intros)
apply(fastforce dest: lossless_gpvD split: generat.split)
apply(clarsimp simp add: set_bind_spmf; split generat.split_asm)
apply(auto dest: lossless_gpvD intro!: lossless_gpv)
done
qed
lemma lossless_bind_gpvD1: "lossless_gpv ℐ (gpv ⤜ f) ⟹ lossless_gpv ℐ gpv"
proof(induction gpv'≡"gpv ⤜ f" arbitrary: gpv rule: lossless_gpv.induct)
case (lossless_gpv p)
obtain p' where gpv: "gpv = GPV p'" by(cases gpv)
from lossless_gpv.hyps gpv have "lossless_spmf p'" by(simp add: GPV_bind)
then show ?case unfolding gpv
proof(rule lossless_gpv.intros)
fix out c input
assume "IO out c ∈ set_spmf p'" "input ∈ responses_ℐ ℐ out"
hence "IO out (λinput. c input ⤜ f) ∈ set_spmf p" using lossless_gpv.hyps gpv
by(auto simp add: GPV_bind set_bind_spmf intro: rev_bexI)
thus "lossless_gpv ℐ (c input)" using ‹input ∈ _› by(rule lossless_gpv.hyps) simp
qed
qed
lemma lossless_bind_gpvD2:
assumes "lossless_gpv Γ (gpv ⤜ f)"
and "x ∈ results_gpv Γ gpv"
shows "lossless_gpv Γ (f x)"
using assms(2,1)
proof(induction)
case (Pure gpv)
thus ?case
by -(rule lossless_gpvI, auto 4 4 dest: lossless_gpvD simp add: set_bind_spmf intro: rev_bexI)
qed(auto 4 4 dest: lossless_gpvD simp add: set_bind_spmf intro: rev_bexI)
lemma lossless_bind_gpv [simp]:
"lossless_gpv Γ (gpv ⤜ f) ⟷ lossless_gpv Γ gpv ∧ (∀x∈results_gpv Γ gpv. lossless_gpv Γ (f x))"
by(blast intro: lossless_bind_gpvI dest: lossless_bind_gpvD1 lossless_bind_gpvD2)
lemma lossless_WT_gpv_induct [consumes 2, case_names lossless_gpv]:
assumes lossless: "lossless_gpv ℐ gpv"
and WT: "ℐ ⊢g gpv √"
and step: "⋀p. ⟦
lossless_spmf p;
⋀out c. IO out c ∈ set_spmf p ⟹ out ∈ outs_ℐ ℐ;
⋀out c input. ⟦IO out c ∈ set_spmf p; out ∈ outs_ℐ ℐ ⟹ input ∈ responses_ℐ ℐ out⟧ ⟹ lossless_gpv ℐ (c input);
⋀out c input. ⟦IO out c ∈ set_spmf p; out ∈ outs_ℐ ℐ ⟹ input ∈ responses_ℐ ℐ out⟧ ⟹ ℐ ⊢g c input √;
⋀out c input. ⟦IO out c ∈ set_spmf p; out ∈ outs_ℐ ℐ ⟹ input ∈ responses_ℐ ℐ out⟧ ⟹ P (c input)⟧
⟹ P (GPV p)"
shows "P gpv"
using lossless WT
apply(induction)
apply(erule step)
apply(auto elim: WT_gpvD simp add: WT_gpv_simps)
done
lemma lossless_gpv_induct_strong [consumes 1, case_names lossless_gpv]:
assumes gpv: "lossless_gpv ℐ gpv"
and step:
"⋀p. ⟦ lossless_spmf p;
⋀gpv. gpv ∈ sub_gpvs ℐ (GPV p) ⟹ lossless_gpv ℐ gpv;
⋀gpv. gpv ∈ sub_gpvs ℐ (GPV p) ⟹ P gpv ⟧
⟹ P (GPV p)"
shows "P gpv"
proof -
def gpv' ≡ gpv
then have "gpv' ∈ insert gpv (sub_gpvs ℐ gpv)" by simp
with gpv have "lossless_gpv ℐ gpv' ∧ P gpv'"
proof(induction arbitrary: gpv')
case (lossless_gpv p)
from ‹gpv' ∈ insert (GPV p) _› show ?case
proof(rule insertE)
assume "gpv' = GPV p"
moreover have "lossless_gpv ℐ (GPV p)"
by(auto 4 3 intro: lossless_gpvI lossless_gpv.hyps)
moreover have "P (GPV p)" using lossless_gpv.hyps(1)
by(rule step)(fastforce elim: sub_gpvs.cases lossless_gpv.IH[THEN conjunct1] lossless_gpv.IH[THEN conjunct2])+
ultimately show ?case by simp
qed(fastforce elim: sub_gpvs.cases lossless_gpv.IH[THEN conjunct1] lossless_gpv.IH[THEN conjunct2])
qed
thus ?thesis by(simp add: gpv'_def)
qed
lemma lossless_sub_gpvsI:
assumes spmf: "lossless_spmf (the_gpv gpv)"
and sub: "⋀gpv'. gpv' ∈ sub_gpvs ℐ gpv ⟹ lossless_gpv ℐ gpv'"
shows "lossless_gpv ℐ gpv"
using spmf by(rule lossless_gpvI)(rule sub[OF sub_gpvs.base])
lemma lossless_sub_gpvsD:
assumes "lossless_gpv ℐ gpv" "gpv' ∈ sub_gpvs ℐ gpv"
shows "lossless_gpv ℐ gpv'"
using assms(2,1) by(induction)(auto dest: lossless_gpvD)
lemma lossless_WT_gpv_induct_strong [consumes 2, case_names lossless_gpv]:
assumes lossless: "lossless_gpv ℐ gpv"
and WT: "ℐ ⊢g gpv √"
and step: "⋀p. ⟦ lossless_spmf p;
⋀out c. IO out c ∈ set_spmf p ⟹ out ∈ outs_ℐ ℐ;
⋀gpv. gpv ∈ sub_gpvs ℐ (GPV p) ⟹ lossless_gpv ℐ gpv;
⋀gpv. gpv ∈ sub_gpvs ℐ (GPV p) ⟹ ℐ ⊢g gpv √;
⋀gpv. gpv ∈ sub_gpvs ℐ (GPV p) ⟹ P gpv ⟧
⟹ P (GPV p)"
shows "P gpv"
using lossless WT
apply(induction rule: lossless_gpv_induct_strong)
apply(erule step)
apply(auto elim: WT_gpvD dest: WT_sub_gpvsD)
done
lemma try_gpv_lossless [simp]:
fixes gpv :: "('a, 'call, 'ret) gpv"
assumes "lossless_gpv ℐ_top gpv" -- ‹TODO: generalise to arbitrary typings ?›
shows "TRY gpv ELSE gpv' = gpv"
using assms
apply induction
apply(rule gpv.expand)
apply(simp add: spmf.map_comp generat.map_comp o_def)
apply(rule map_spmf_idI)
subgoal for p generat by(cases generat) simp_all
done
lemma try_gpv_bind_lossless:
assumes "lossless_gpv ℐ_top gpv" -- ‹TODO: generalise to arbitrary typings?›
shows "TRY bind_gpv gpv f ELSE gpv' = bind_gpv gpv (λx. TRY f x ELSE gpv')"
using assms
apply induction
apply(rule gpv.expand)
apply(simp add: map_try_spmf spmf.map_comp o_def generat.map_comp map_spmf_bind_spmf bind_gpv.sel try_gpv_sel del: bind_gpv_sel' try_gpv.sel)
apply(subst try_spmf_bind_spmf_lossless, simp)
apply(rule bind_spmf_cong, simp)
apply(clarsimp simp add: spmf.map_comp o_def map_try_spmf generat.map_comp generat.map_ident[abs_def] split: generat.split)
done
subsection {* Inlining *}
lemma gpv_coinduct_bind [consumes 1, case_names Eq_gpv]:
fixes gpv gpv' :: "('a, 'call, 'ret) gpv"
assumes *: "R gpv gpv'"
and step: "⋀gpv gpv'. R gpv gpv'
⟹ rel_spmf (rel_generat op = op = (rel_fun op = (λgpv gpv'. R gpv gpv' ∨ gpv = gpv' ∨
(∃gpv2 :: ('b, 'call, 'ret) gpv. ∃gpv2' :: ('c, 'call, 'ret) gpv. ∃f f'. gpv = bind_gpv gpv2 f ∧ gpv' = bind_gpv gpv2' f' ∧
rel_gpv (λx y. R (f x) (f' y)) op = gpv2 gpv2'))))
(the_gpv gpv) (the_gpv gpv')"
shows "gpv = gpv'"
proof -
fix x y
def gpv1 ≡ "Done x :: ('b, 'call, 'ret) gpv"
and f ≡ "(λ_. gpv) :: 'b ⇒ ('a, 'call, 'ret) gpv"
and gpv1' ≡ "Done y :: ('c, 'call, 'ret) gpv"
and f' ≡ "(λ_. gpv') :: 'c ⇒ ('a, 'call, 'ret) gpv"
from * have "rel_gpv (λx y. R (f x) (f' y)) op = gpv1 gpv1'"
by(simp add: gpv1_def gpv1'_def f_def f'_def)
then have "gpv1 ⤜ f = gpv1' ⤜ f'"
proof(coinduction arbitrary: gpv1 gpv1' f f' rule: gpv.coinduct_strong)
case (Eq_gpv gpv1 gpv1' f f')
from Eq_gpv[simplified gpv.rel_sel] show ?case unfolding bind_gpv.sel spmf_rel_map
apply(rule rel_spmf_bindI)
subgoal for generat generat'
apply(cases generat generat' rule: generat.exhaust[case_product generat.exhaust]; clarsimp simp add: o_def spmf_rel_map generat.rel_map)
subgoal premises Pure for x y
using step[OF ‹R (f x) (f' y)›] apply -
apply(assumption | rule rel_spmf_mono rel_generat_mono rel_fun_mono refl)+
apply(fastforce intro: exI[where x="Done _"])+
done
subgoal by(fastforce simp add: rel_fun_def)
done
done
qed
thus ?thesis by(simp add: gpv1_def gpv1'_def f_def f'_def)
qed
text {*
Inlining one gpv into another. This may throw out arbitrarily many
interactions between the two gpvs if the inlined one does not call its callee.
So we define it as the coiteration of a least-fixpoint search operator.
*}
context
fixes callee :: "'s ⇒ 'call ⇒ ('ret × 's, 'call', 'ret') gpv"
notes [[function_internals]]
begin
partial_function (spmf) inline1
:: "('a, 'call, 'ret) gpv ⇒ 's
⇒ ('a × 's + 'call' × ('ret × 's, 'call', 'ret') rpv × ('a, 'call, 'ret) rpv) spmf"
where
"inline1 gpv s =
the_gpv gpv ⤜
case_generat (λx. return_spmf (Inl (x, s)))
(λout rpv. the_gpv (callee s out) ⤜
case_generat (λ(x, y). inline1 (rpv x) y)
(λout rpv'. return_spmf (Inr (out, rpv', rpv))))"
lemma inline1_unfold:
"inline1 gpv s =
the_gpv gpv ⤜
case_generat (λx. return_spmf (Inl (x, s)))
(λout rpv. the_gpv (callee s out) ⤜
case_generat (λ(x, y). inline1 (rpv x) y)
(λout rpv'. return_spmf (Inr (out, rpv', rpv))))"
by(fact inline1.simps)
lemma inline1_fixp_induct [case_names adm bottom step]:
assumes "ccpo.admissible (fun_lub lub_spmf) (fun_ord (ord_spmf op =)) (λinline1'. P (λgpv s. inline1' (gpv, s)))"
and "P (λ_ _. return_pmf None)"
and "⋀inline1'. P inline1' ⟹ P (λgpv s. the_gpv gpv ⤜ case_generat (λx. return_spmf (Inl (x, s))) (λout rpv. the_gpv (callee s out) ⤜ case_generat (λ(x, y). inline1' (rpv x) y) (λout rpv'. return_spmf (Inr (out, rpv', rpv)))))"
shows "P inline1"
using assms by(rule inline1.fixp_induct[unfolded curry_conv[abs_def]])
lemma inline1_fixp_induct_strong [case_names adm bottom step]:
assumes "ccpo.admissible (fun_lub lub_spmf) (fun_ord (ord_spmf op =)) (λinline1'. P (λgpv s. inline1' (gpv, s)))"
and "P (λ_ _. return_pmf None)"
and "⋀inline1'. ⟦ ⋀gpv s. ord_spmf op = (inline1' gpv s) (inline1 gpv s); P inline1' ⟧
⟹ P (λgpv s. the_gpv gpv ⤜ case_generat (λx. return_spmf (Inl (x, s))) (λout rpv. the_gpv (callee s out) ⤜ case_generat (λ(x, y). inline1' (rpv x) y) (λout rpv'. return_spmf (Inr (out, rpv', rpv)))))"
shows "P inline1"
using assms by(rule spmf.fixp_strong_induct_uc[where P="λf. P (curry f)" and U=case_prod and C=curry, OF inline1.mono inline1_def, simplified curry_case_prod, simplified curry_conv[abs_def] fun_ord_def split_paired_All prod.case case_prod_eta, OF refl]) blast+
lemma inline1_fixp_induct_strong2 [case_names adm bottom step]:
assumes "ccpo.admissible (fun_lub lub_spmf) (fun_ord (ord_spmf op =)) (λinline1'. P (λgpv s. inline1' (gpv, s)))"
and "P (λ_ _. return_pmf None)"
and "⋀inline1'.
⟦ ⋀gpv s. ord_spmf op = (inline1' gpv s) (inline1 gpv s);
⋀gpv s. ord_spmf op = (inline1' gpv s) (the_gpv gpv ⤜ case_generat (λx. return_spmf (Inl (x, s))) (λout rpv. the_gpv (callee s out) ⤜ case_generat (λ(x, y). inline1' (rpv x) y) (λout rpv'. return_spmf (Inr (out, rpv', rpv)))));
P inline1' ⟧
⟹ P (λgpv s. the_gpv gpv ⤜ case_generat (λx. return_spmf (Inl (x, s))) (λout rpv. the_gpv (callee s out) ⤜ case_generat (λ(x, y). inline1' (rpv x) y) (λout rpv'. return_spmf (Inr (out, rpv', rpv)))))"
shows "P inline1"
using assms
by(rule spmf.fixp_induct_strong2_uc[where P="λf. P (curry f)" and U=case_prod and C=curry, OF inline1.mono inline1_def, simplified curry_case_prod, simplified curry_conv[abs_def] fun_ord_def split_paired_All prod.case case_prod_eta, OF refl]) blast+
lemma inline1_Fail [simp]: "inline1 Fail s = return_pmf None"
by(rewrite inline1.simps) simp
lemma inline1_Done [simp]: "inline1 (Done x) s = return_spmf (Inl (x, s))"
by(rewrite inline1.simps) simp
lemma inline1_lift_spmf [simp]: "inline1 (lift_spmf p) s = map_spmf (λx. Inl (x, s)) p"
by(rewrite inline1.simps)(simp add: bind_map_spmf o_def map_spmf_conv_bind_spmf)
lemma inline1_Pause:
"inline1 (Pause out c) s =
the_gpv (callee s out) ⤜ (λreact. case react of Pure (x, s') ⇒ inline1 (c x) s' | IO out' c' ⇒ return_spmf (Inr (out', c', c)))"
by(rewrite inline1.simps) simp
lemma inline1_bind_gpv:
fixes gpv f s
defines [simp]: "inline11 ≡ inline1" and [simp]: "inline12 ≡ inline1" and [simp]: "inline13 ≡ inline1"
shows "inline11 (bind_gpv gpv f) s = bind_spmf (inline12 gpv s)
(λres. case res of Inl (x, s') ⇒ inline13 (f x) s' | Inr (out, rpv', rpv) ⇒ return_spmf (Inr (out, rpv', bind_rpv rpv f)))"
(is "?lhs = ?rhs")
proof(rule spmf.leq_antisym)
note [intro!] = ord_spmf_bind_reflI and [split] = generat.split
show "ord_spmf op = ?lhs ?rhs" unfolding inline11_def
proof(induction arbitrary: gpv s f rule: inline1_fixp_induct)
case adm show ?case by simp
case bottom show ?case by simp
case (step inline1')
show ?case unfolding inline12_def
apply(rewrite inline1.simps; clarsimp simp add: bind_rpv_def)
apply(rule conjI; clarsimp)
subgoal premises Pure for x
apply(rewrite inline1.simps; clarsimp)
subgoal for out c ret s' using step.IH[of "Done x" "λ_. c ret" s'] by simp
done
subgoal for out c ret s' using step.IH[of "c ret" f s'] by(simp cong del: sum.case_cong_weak)
done
qed
show "ord_spmf op = ?rhs ?lhs" unfolding inline12_def
proof(induction arbitrary: gpv s rule: inline1_fixp_induct)
case adm show ?case by simp
case bottom show ?case by simp
case (step inline1')
show ?case unfolding inline11_def
apply(rewrite inline1.simps; clarsimp simp add: bind_rpv_def)
apply(rule conjI; clarsimp)
subgoal by(rewrite inline1.simps; simp)
subgoal for out c ret s' using step.IH[of "c ret" s'] by(simp cong del: sum.case_cong_weak)
done
qed
qed
end
declare [[bnf_internals]]
context notes [[quick_and_dirty]] begin
codatatype ('a, 'call, 'ret, 'x) gpv' = GPV' (the_gpv': "('a, 'call, 'ret ⇒ ('a, 'call, 'ret, 'x) gpv') generat spmf")
end
primcorec gpv_of_gpv' :: "('a, 'call, 'ret, 'x) gpv' ⇒ ('a, 'call, 'ret) gpv"
where "the_gpv (gpv_of_gpv' gpv') = map_spmf (map_generat id id (op ∘ gpv_of_gpv')) (the_gpv' gpv')"
primcorec gpv'_of_gpv :: "('a, 'call, 'ret) gpv ⇒ ('a, 'call, 'ret, 'x) gpv'"
where "the_gpv' (gpv'_of_gpv gpv') = map_spmf (map_generat id id (op ∘ gpv'_of_gpv)) (the_gpv gpv')"
lemma gpv_of_gpv'_inverse [simp]: "gpv'_of_gpv (gpv_of_gpv' gpv') = gpv'"
apply(coinduction arbitrary: gpv')
apply(auto simp add: spmf_rel_map generat.rel_map rel_fun_def intro!: rel_spmf_reflI generat.rel_refl)
done
lemma gpv'_of_gpv_inverse [simp]: "gpv_of_gpv' (gpv'_of_gpv gpv) = gpv"
apply(coinduction arbitrary: gpv)
apply(auto simp add: spmf_rel_map generat.rel_map rel_fun_def intro!: rel_spmf_reflI generat.rel_refl)
done
lemma td_gpv': "type_definition gpv_of_gpv' gpv'_of_gpv UNIV"
by unfold_locales simp_all
setup_lifting td_gpv'
lift_definition bind_gpv' :: "('a, 'call, 'ret) gpv ⇒ ('a ⇒ ('b, 'call, 'ret, 'a) gpv') ⇒ ('b, 'call, 'ret, 'a) gpv'" is bind_gpv .
thm bind_gpv'.rep_eq
lemma gpv_of_gpv'_GPV' [simp]: "gpv_of_gpv' (GPV' p) = GPV (map_spmf (map_generat id id (op ∘ gpv_of_gpv')) p)"
by(rule gpv.expand) simp
lemma gpv'_of_gpv_GPV [simp]: "gpv'_of_gpv (GPV p) = GPV' (map_spmf (map_generat id id (op ∘ gpv'_of_gpv)) p)"
by(rule gpv'.expand) simp
abbreviation cast_gpv' :: "('a, 'call, 'ret, 'b) gpv' ⇒ ('a, 'call, 'ret, 'c) gpv'"
where "cast_gpv' gpv' ≡ gpv'_of_gpv (gpv_of_gpv' gpv')"
context includes lifting_syntax begin
lemma GPV'_transfer [transfer_rule]:
"(rel_spmf (rel_generat op = op = (op = ===> cr_gpv')) ===> cr_gpv') GPV GPV'"
by(simp add: cr_gpv'_def gpv.rel_eq[symmetric] rel_fun_def spmf_rel_map generat.rel_map)
lemma the_gpv'_transfer [transfer_rule]:
"(cr_gpv' ===> rel_spmf (rel_generat op = op = (op = ===> cr_gpv'))) the_gpv the_gpv'"
by(auto simp add: cr_gpv'_def spmf_rel_map generat.rel_map rel_fun_def intro!: rel_spmf_reflI generat.rel_refl)
end
lemma bind_gpv'_simps:
"bind_gpv' r f =
GPV' (map_spmf (map_generat id id (op ∘ (case_sum id (λr. bind_gpv' r f))))
(bind_spmf (the_gpv r)
(case_generat
(λx. map_spmf (map_generat id id (op ∘ Inl)) (the_gpv' (f x)))
(λout c. return_spmf (IO out (λinput. Inr (c input)))))))"
by transfer(rule bind_gpv.ctr)
lemma bind_gpv'_sel:
"the_gpv' (bind_gpv' r f) =
map_spmf (map_generat id id (λrpv x. case_sum id (λr. bind_gpv' r f) (rpv x)))
(bind_spmf (the_gpv r)
(case_generat
(λx. map_spmf (map_generat id id (λr x. Inl (r x))) (the_gpv' (f x)))
(λout c. return_spmf (IO out (λinput. Inr (c input))))))"
by(subst bind_gpv'_simps)(simp add: o_def[abs_def])
lemma gpv_of_gpv'_bind_gpv' [simp]: "gpv_of_gpv' (bind_gpv' gpv f) = bind_gpv gpv (gpv_of_gpv' ∘ f)"
by transfer(simp add: o_def)
lemma gpv'_of_gpv_bind_gpv [simp]: "gpv'_of_gpv (bind_gpv gpv f) = bind_gpv' gpv (gpv'_of_gpv ∘ f)"
by transfer (simp add: o_def)
friend_of_corec bind_gpv'
where "bind_gpv' r f =
GPV' (map_spmf (map_generat id id (op ∘ (case_sum id (λr. bind_gpv' r f))))
(bind_spmf (the_gpv r)
(case_generat (λx. map_spmf (map_generat id id (op ∘ Inl)) (the_gpv' (f x)))
(λout c. return_spmf (IO out (λinput. Inr (c input)))))))"
subgoal by(rule bind_gpv'_simps)
apply(fold relator_eq)
apply transfer_prover+
done
lift_definition inline1'
:: "('s ⇒ 'call ⇒ ('ret × 's, 'call', 'ret') gpv) ⇒ ('a, 'call, 'ret, 'ret × 's) gpv' ⇒ 's
⇒ ('a × 's + 'call' × ('ret × 's, 'call', 'ret') rpv × ('ret ⇒ ('a, 'call, 'ret, 'ret × 's) gpv')) spmf"
is inline1 .
lemma inline1_conv_inline1':
"inline1 oracle gpv' s = map_spmf (map_sum id (map_prod id (map_prod id (map_fun id gpv_of_gpv')))) (inline1' oracle (gpv'_of_gpv gpv') s)"
by transfer(simp add: id_def[symmetric] map_fun_id prod.map_id0 sum.map_id0)
corec inline' :: "('s ⇒ 'call ⇒ ('ret × 's, 'call', 'ret') gpv) ⇒ ('a, 'call, 'ret, 'ret × 's) gpv' ⇒ 's ⇒ ('a × 's, 'call', 'ret', 'ret × 's) gpv'"
where
"inline' oracle c s = GPV' (
map_spmf (λresult. case result of Inl xs ⇒ Pure xs
| Inr (out, oracle', rpv) ⇒ IO out (λinput. bind_gpv' (oracle' input) (λ(x, s'). inline' oracle (rpv x) s'))) (inline1' oracle c s))"
lemma inline'_sel:
"the_gpv' (inline' oracle c s) =
map_spmf (λresult. case result of Inl xs ⇒ Pure xs
| Inr (out, oracle', rpv) ⇒ IO out (λinput. bind_gpv' (oracle' input) (λ(x, s'). inline' oracle (rpv x) s'))) (inline1' oracle c s)"
by(subst inline'.code) simp
definition inline :: "('s ⇒ 'call ⇒ ('ret × 's, 'call', 'ret') gpv) ⇒ ('a, 'call, 'ret) gpv ⇒ 's ⇒ ('a × 's, 'call', 'ret') gpv"
where "inline oracle c s = gpv_of_gpv' (inline' oracle (gpv'_of_gpv c) s)"
lemma inline_sel:
"the_gpv (inline oracle c s) =
map_spmf (λresult. case result of Inl xs ⇒ Pure xs
| Inr (out, oracle', rpv) ⇒ IO out (λinput. bind_gpv (oracle' input) (λ(x, s'). inline oracle (rpv x) s'))) (inline1 oracle c s)"
by(auto simp add: inline_def inline'_sel spmf.map_comp o_def inline1_conv_inline1' split_def intro!: arg_cong2[where f=map_spmf] ext split: sum.split)
lemma inline_Fail [simp]: "inline oracle Fail s = Fail"
by(rule gpv.expand)(simp add: inline_sel)
lemma inline_Done [simp]: "inline oracle (Done x) s = Done (x, s)"
by(rule gpv.expand)(simp add: inline_sel)
lemma inline1'_bind_gpv:
"inline1' callee (gpv'_of_gpv (bind_gpv gpv f)) s =
bind_spmf (inline1 callee gpv s) (λres. case res of Inl (x, s') ⇒ inline1' callee (gpv'_of_gpv (f x)) s' | Inr (out, rpv', rpv) ⇒ return_spmf (Inr (out, rpv', λinput. gpv'_of_gpv (bind_gpv (rpv input) f))))"
by transfer(simp add: inline1_bind_gpv bind_rpv_def cong del: sum.case_cong)
lemma inline'_bind_gpv [simp]:
"inline' oracle (gpv'_of_gpv (bind_gpv gpv f)) s = cast_gpv' (bind_gpv' (inline oracle gpv s) (λ(x, s'). cast_gpv' (inline' oracle (gpv'_of_gpv (f x)) s')))"
by(coinduction arbitrary: gpv s rule: inline'.coinduct)
(auto 4 4 simp add: bind_gpv_assoc inline'_sel inline_sel inline1_conv_inline1' bind_gpv'_sel spmf_rel_map spmf.map_comp o_def generat.map_comp map_spmf_bind_spmf inline1'_bind_gpv bind_map_spmf intro!: rel_spmf_bind_reflI rel_spmf_reflI rel_funI gpv'.cong_bind_gpv'[OF refl] gpv'.cong_refl[OF refl] split: sum.split intro: gpv'.cong_base)
lemma inline_bind_gpv [simp]:
"inline oracle (bind_gpv gpv f) s = bind_gpv (inline oracle gpv s) (λ(x, s'). inline oracle (f x) s')"
unfolding inline_def inline'_bind_gpv by(simp add: o_def split_def)
lemma set_inline1_lift_spmf1: "set_spmf (inline1 (λs x. lift_spmf (p s x)) gpv s) ⊆ range Inl"
apply(induction arbitrary: gpv s rule: inline1_fixp_induct)
subgoal by(rule cont_intro ccpo_class.admissible_leI)+
apply(auto simp add: set_bind_spmf o_def bind_UNION split: generat.split_asm)+
done
lemma in_set_inline1_lift_spmf1: "y ∈ set_spmf (inline1 (λs x. lift_spmf (p s x)) gpv s) ⟹ ∃r s'. y = Inl (r, s')"
by(drule set_inline1_lift_spmf1[THEN subsetD]) auto
lemma inline_lift_spmf1:
fixes p defines "callee ≡ λs c. lift_spmf (p s c)"
shows "inline callee gpv s = lift_spmf (map_spmf projl (inline1 callee gpv s))"
by(rule gpv.expand)(auto simp add: inline_sel spmf.map_comp callee_def intro!: map_spmf_cong[OF refl] dest: in_set_inline1_lift_spmf1)
context includes lifting_syntax begin
lemma inline'_transfer [transfer_rule]:
"(op = ===> cr_gpv' ===> op = ===> cr_gpv') inline inline'"
by(simp add: rel_fun_def cr_gpv'_def inline_def)
lemma inline1_parametric [transfer_rule]:
"((S ===> C ===> rel_gpv (rel_prod op = S) C') ===> rel_gpv A C ===> S
===> rel_spmf (rel_sum (rel_prod A S) (rel_prod C' (rel_prod (rel_rpv (rel_prod op = S) C') (rel_rpv A C)))))
inline1 inline1"
(is "(_ ===> ?R) _ _")
proof(rule rel_funI)
show "?R (inline1 callee) (inline1 callee')"
if [transfer_rule]: "(S ===> C ===> rel_gpv (rel_prod op = S) C') callee callee'"
for callee callee'
unfolding inline1_def
by(unfold rel_fun_curry case_prod_curry)(rule fixp_spmf_parametric[OF inline1.mono inline1.mono]; transfer_prover)
qed
lemma rel_gpv'_transfer [transfer_rule]: "(op = ===> op = ===> R ===> cr_gpv' ===> cr_gpv' ===> op =) (λA B _. rel_gpv A B) rel_gpv'"
apply(intro rel_funI; clarsimp simp add: cr_gpv'_def; rule iffI)
subgoal for A C r r' gpv gpv'
apply(coinduction arbitrary: gpv gpv')
subgoal for gpv gpv' by(cases gpv gpv' rule: gpv'.exhaust[case_product gpv'.exhaust])(auto simp add: spmf_rel_map generat.rel_map rel_fun_def elim!: rel_spmf_mono rel_generat_mono)
done
subgoal for A C r r' gpv gpv'
by(coinduction arbitrary: gpv gpv')(auto 4 4 elim!: gpv'.rel_cases simp add: spmf_rel_map generat.rel_map rel_fun_def elim!: rel_spmf_mono rel_generat_mono)
done
lemma gpv_of_gpv'_parametric [transfer_rule]: "(rel_gpv' A C R ===> rel_gpv A C) gpv_of_gpv' gpv_of_gpv'"
by(rule rel_funI) transfer
lemma gpv'_of_gpv_parametric [transfer_rule]: "(rel_gpv A C ===> rel_gpv' A C R) gpv'_of_gpv gpv'_of_gpv"
by(rule rel_funI) transfer
lifting_update gpv'.lifting
lifting_forget gpv'.lifting
lemma gpv'_of_gpv_map_gpv:
"gpv'_of_gpv (map_gpv f g gpv) = map_gpv' f g h (gpv'_of_gpv gpv)"
unfolding gpv'.rel_eq[symmetric] gpv'.rel_map
by(rule gpv'_of_gpv_parametric[THEN rel_funD])(simp add: gpv.rel_map gpv.rel_refl)
lemma gpv_of_gpv'_map_gpv':
"gpv_of_gpv' (map_gpv' f g h gpv') = map_gpv f g (gpv_of_gpv' gpv')"
unfolding gpv.rel_eq[symmetric] gpv.rel_map
by(rule gpv_of_gpv'_parametric[where R="λ_ _. True", THEN rel_funD])(simp add: gpv'.rel_map gpv'.rel_refl)
lemma inline1'_parametric [transfer_rule]:
"((S ===> C ===> rel_gpv (rel_prod op = S) C') ===> rel_gpv' A C R ===> S
===> rel_spmf (rel_sum (rel_prod A S) (rel_prod C' (rel_prod (rel_rpv (rel_prod op = S) C') (op = ===> rel_gpv' A C R')))))
inline1' inline1'"
unfolding inline1'_def map_fun_def o_def by transfer_prover
declare gpv'.rel_eq [relator_eq]
declare gpv'.dtor_unfold_transfer[transfer_rule]
lemma bind_gpv'_parametric [transfer_rule]:
"(rel_gpv A C ===> (A ===> rel_gpv' B C R) ===> rel_gpv' B C R') (λgpv f. cast_gpv' (bind_gpv' gpv f)) bind_gpv'"
including gpv'.lifting unfolding rel_fun_def
by transfer(rule bind_gpv_parametric[unfolded rel_fun_def])
lemma map_gpv'_id_bind_gpv':
"map_gpv' f id h (bind_gpv' gpv g) = cast_gpv' (bind_gpv' gpv (map_gpv' f id id ∘ g))"
apply(rule sym)
apply(unfold gpv'.rel_eq[symmetric] gpv'.rel_map)
apply(rule bind_gpv'_parametric[where A="op =" and R="op =", THEN rel_funD, THEN rel_funD])
apply(simp_all add: gpv'.rel_map gpv'.rel_refl rel_fun_def gpv.rel_refl)
done
end
context
fixes callee1 :: "'s1 ⇒ 'c1 ⇒ ('r1 × 's1, 'c, 'r) gpv"
and callee2 :: "'s2 ⇒ 'c2 ⇒ ('r2 × 's2, 'c1, 'r1) gpv"
begin
partial_function (spmf) inline2 :: "('a, 'c2, 'r2) gpv ⇒ 's2 ⇒ 's1
⇒ ('a × ('s2 × 's1) + 'c × ('r1 × 's1, 'c, 'r) rpv × ('r2 × 's2, 'c1, 'r1) rpv × ('a, 'c2, 'r2) rpv) spmf"
where
"inline2 gpv s2 s1 =
bind_spmf (the_gpv gpv)
(case_generat (λx. return_spmf (Inl (x, s2, s1)))
(λout rpv. bind_spmf (inline1 callee1 (callee2 s2 out) s1)
(case_sum (λ((r2, s2), s1). inline2 (rpv r2) s2 s1)
(λ(x, rpv'', rpv'). return_spmf (Inr (x, rpv'', rpv', rpv))))))"
lemma inline2_fixp_induct [case_names adm bottom step]:
assumes "ccpo.admissible (fun_lub lub_spmf) (fun_ord (ord_spmf op =)) (λinline2. P (λgpv s2 s1. inline2 ((gpv, s2), s1)))"
and "P (λ_ _ _. return_pmf None)"
and "⋀inline2'. P inline2' ⟹
P (λgpv s2 s1. bind_spmf (the_gpv gpv) (λgenerat. case generat of
Pure x ⇒ return_spmf (Inl (x, s2, s1))
| IO out rpv ⇒ bind_spmf (inline1 callee1 (callee2 s2 out) s1) (λlr. case lr of
Inl ((r2, s2), c) ⇒ inline2' (rpv r2) s2 c
| Inr (x, rpv'', rpv') ⇒ return_spmf (Inr (x, rpv'', rpv', rpv)))))"
shows "P inline2"
using assms unfolding split_def by(rule inline2.fixp_induct[unfolded curry_conv[abs_def] split_def])
lemma inline1_inline_conv_inline2:
fixes gpv' :: "('r2 × 's2, 'c1, 'r1) gpv"
shows "inline1 callee1 (inline callee2 gpv s2) s1 =
map_spmf (map_sum (λ(x, (s2, s1)). ((x, s2), s1))
(λ(x, rpv'', rpv', rpv). (x, rpv'', λr1. rpv' r1 ⤜ (λ(r2, s2). inline callee2 (rpv r2) s2))))
(inline2 gpv s2 s1)"
(is "?lhs = ?rhs")
proof(rule spmf.leq_antisym)
def inline1_1 ≡ "inline1 :: ('s1 ⇒ 'c1 ⇒ ('r1 × 's1, 'c, 'r) gpv) ⇒ ('r2 × 's2, 'c1, 'r1) gpv ⇒ 's1 ⇒ _"
have "ord_spmf op = ?lhs ?rhs"
-- ‹ We need in the inductive step that the approximation behaves well with @{const bind_gpv}
because the corecursion goes through @{const bind_gpv}. So we have to thread it through the induction
and do one half of the proof from @{thm [source] inline1_bind_gpv} again. We cannot inline
@{thm [source] inline1_bind_gpv} in this proof here because the types are too specific.
›
and "ord_spmf op = (inline1 callee1 (gpv' ⤜ f) s1')
(do {
res ← inline1_1 callee1 gpv' s1';
case res of Inl (x, s') ⇒ inline1 callee1 (f x) s'
| Inr (out, rpv', rpv) ⇒ return_spmf (Inr (out, rpv', rpv ⤜ f))
})" for gpv' and f :: "_ ⇒ ('a × 's2, 'c1, 'r1) gpv" and s1'
proof(induction arbitrary: gpv s2 s1 gpv' f s1' rule: inline1_fixp_induct_strong2)
case adm thus ?case
apply(rule cont_intro)+
subgoal for a b c d by(cases d; clarsimp)
done
case (step inline1')
note step_IH = step.IH[unfolded inline1_1_def] and step_hyps = step.hyps[unfolded inline1_1_def]
{ case 1
have inline1: "ord_spmf op =
(inline1 callee2 gpv s2 ⤜ (λlr. case lr of Inl as2 ⇒ return_spmf (Inl (as2, s1))
| Inr (out1, rpv', rpv) ⇒ the_gpv (callee1 s1 out1) ⤜ (λgenerat. case generat of
Pure (r1, s1) ⇒ inline1' (bind_gpv (rpv' r1) (λ(r2, s2). inline callee2 (rpv r2) s2)) s1
| IO out rpv'' ⇒ return_spmf (Inr (out, rpv'', λr1. bind_gpv (rpv' r1) (λ(r2, s2). inline callee2 (rpv r2) s2)) ))))
(the_gpv gpv ⤜ (λgenerat. case generat of Pure x ⇒ return_spmf (Inl ((x, s2), s1))
| IO out2 rpv ⇒ inline1_1 callee1 (callee2 s2 out2) s1 ⤜ (λlr. case lr of
Inl ((r2, s2), s1) ⇒
map_spmf (map_sum (λ(x, s2, s1). ((x, s2), s1)) (λ(x, rpv'', rpv', rpv). (x, rpv'', λr1. bind_gpv (rpv' r1) (λ(r2, s2). inline callee2 (rpv r2) s2))))
(inline2 (rpv r2) s2 s1)
| Inr (out, rpv'', rpv') ⇒
return_spmf (Inr (out, rpv'', λr1. bind_gpv (rpv' r1) (λ(r2, s2). inline callee2 (rpv r2) s2))))))"
proof(induction arbitrary: gpv s2 s1 rule: inline1_fixp_induct)
case step2: (step inline1'')
note step2_IH = step2.IH[unfolded inline1_1_def]
show ?case unfolding inline1_1_def
apply(rewrite in "ord_spmf _ _ ⌑" inline1.simps)
apply(clarsimp intro!: ord_spmf_bind_reflI split: generat.split)
apply(rule conjI)
subgoal by(rewrite in "ord_spmf _ _ ⌑" inline2.simps)(clarsimp simp add: map_spmf_bind_spmf o_def split: generat.split sum.split intro!: ord_spmf_bind_reflI spmf.leq_trans[OF step2_IH])
subgoal by(clarsimp intro!: ord_spmf_bind_reflI step_IH[THEN spmf.leq_trans] split: generat.split sum.split simp add: bind_rpv_def)
done
qed simp_all
show ?case
apply(rewrite in "ord_spmf _ ⌑ _" inline_sel)
apply(rewrite in "ord_spmf _ _ ⌑" inline2.simps)
apply(clarsimp simp add: map_spmf_bind_spmf bind_map_spmf o_def intro!: ord_spmf_bind_reflI split: generat.split)
apply(rule spmf.leq_trans[OF spmf.leq_trans, OF _ inline1])
apply(auto intro!: ord_spmf_bind_reflI split: sum.split generat.split simp add: inline1_1_def map_spmf_bind_spmf)
done }
{ case 2
show ?case unfolding inline1_1_def
by(rewrite inline1.simps)(auto simp del: bind_gpv_sel' simp add: bind_gpv.sel map_spmf_bind_spmf bind_map_spmf o_def bind_rpv_def intro!: ord_spmf_bind_reflI step_IH(2)[THEN spmf.leq_trans] step_hyps(2) split: generat.split sum.split) }
qed simp_all
thus "ord_spmf op = ?lhs ?rhs" by -
show "ord_spmf op = ?rhs ?lhs"
proof(induction arbitrary: gpv s2 s1 rule: inline2_fixp_induct)
case adm show ?case by simp
case bottom show ?case by simp
case (step inline2')
show ?case
apply(rewrite in "ord_spmf _ _ ⌑" inline1.simps)
apply(rewrite inline_sel)
apply(rewrite in "ord_spmf _ ⌑ _" inline1.simps)
apply(rewrite in "ord_spmf _ _ ⌑" inline1.simps)
apply(clarsimp simp add: map_spmf_bind_spmf bind_map_spmf intro!: ord_spmf_bind_reflI split: generat.split)
apply(rule conjI)
subgoal
apply clarsimp
apply(rule step.IH[THEN spmf.leq_trans])
apply(rewrite in "ord_spmf _ ⌑ _" inline1.simps)
apply(rewrite inline_sel)
apply(simp add: bind_map_spmf)
done
subgoal by(clarsimp intro!: ord_spmf_bind_reflI split: generat.split sum.split simp add: o_def inline1_bind_gpv bind_rpv_def step.IH)
done
qed
qed
lemma inline1_inline_conv_inline2':
"inline1 (λ(s2, s1) c2. map_gpv (λ((r, s2), s1). (r, s2, s1)) id (inline callee1 (callee2 s2 c2) s1)) gpv (s2, s1) =
map_spmf (map_sum id (λ(x, rpv'', rpv', rpv). (x, λr. bind_gpv (rpv'' r)
(λ(r1, s1). map_gpv (λ((r2, s2), s1). (r2, s2, s1)) id (inline callee1 (rpv' r1) s1)), rpv)))
(inline2 gpv s2 s1)"
(is "?lhs = ?rhs")
proof(rule spmf.leq_antisym)
show "ord_spmf op = ?lhs ?rhs"
proof(induction arbitrary: gpv s2 s1 rule: inline1_fixp_induct)
case (step inline1') show ?case
by(rewrite inline2.simps)(auto simp add: map_spmf_bind_spmf o_def inline_sel gpv.map_sel bind_map_spmf id_def[symmetric] gpv.map_id map_gpv_bind_gpv split_def intro!: ord_spmf_bind_reflI step.IH[THEN spmf.leq_trans] split: generat.split sum.split)
qed simp_all
show "ord_spmf op = ?rhs ?lhs"
proof(induction arbitrary: gpv s2 s1 rule: inline2_fixp_induct)
case (step inline2')
show ?case
apply(rewrite in "ord_spmf _ _ ⌑" inline1.simps)
apply(clarsimp simp add: map_spmf_bind_spmf bind_rpv_def o_def gpv.map_sel bind_map_spmf inline_sel map_gpv_bind_gpv id_def[symmetric] gpv.map_id split_def split: generat.split sum.split intro!: ord_spmf_bind_reflI)
apply(rule spmf.leq_trans[OF spmf.leq_trans, OF _ step.IH])
apply(auto simp add: split_def id_def[symmetric] intro!: ord_spmf_reflI)
done
qed simp_all
qed
context includes gpv'.lifting begin
lift_definition inline2' :: "('a, 'c2, 'r2) gpv ⇒ 's2 ⇒ 's1
⇒ ('a × ('s2 × 's1) + 'c × ('r1 × 's1, 'c, 'r) rpv × ('r2 × 's2, 'c1, 'r1) rpv × ('r2 ⇒ ('a, 'c2, 'r2, 'foo) gpv')) spmf"
is "inline2" .
lemma inline1'_inline_conv_inline2:
"inline1' callee1 (gpv'_of_gpv (inline callee2 gpv s2)) s1 =
map_spmf (map_sum (λ(x, (s2, s1)). ((x, s2), s1))
(λ(x, rpv'', rpv', rpv). (x, rpv'', λr1. gpv'_of_gpv (bind_gpv (rpv' r1) (λ(r2, s2). inline callee2 (gpv_of_gpv' (rpv r2)) s2)))))
(inline2' gpv s2 s1)"
by transfer(rule inline1_inline_conv_inline2)
lemma inline1'_inline_conv_inline2':
"inline1' (λ(s2, s1) c2. map_gpv (λ((r, s2), s1). (r, s2, s1)) id (inline callee1 (callee2 s2 c2) s1)) (gpv'_of_gpv gpv) (s2, s1) =
map_spmf (map_sum id (λ(x, rpv'', rpv', rpv). (x, λr. bind_gpv (rpv'' r)
(λ(r1, s1). map_gpv (λ((r2, s2), s1). (r2, s2, s1)) id (inline callee1 (rpv' r1) s1)), rpv)))
(inline2' gpv s2 s1)"
by transfer(rule inline1_inline_conv_inline2')
lemma bind_gpv'_assoc:
"bind_gpv' (bind_gpv gpv f) g = cast_gpv' (bind_gpv' gpv (λx. cast_gpv' (bind_gpv' (f x) g)))"
by transfer(rule bind_gpv_assoc)
end
lemma inline'_assoc:
"cast_gpv' (inline' callee1 (gpv'_of_gpv (inline callee2 gpv s2)) s1) =
map_gpv' (λ(r, s2, s1). ((r, s2), s1)) id id (inline' (λ(s2, s1) c2. map_gpv (λ((r, s2), s1). (r, s2, s1)) id (inline callee1 (callee2 s2 c2) s1)) (gpv'_of_gpv gpv) (s2, s1))"
apply(coinduction arbitrary: s2 s1 gpv rule: gpv'.coinduct_upto)
apply simp
apply(subst inline'_sel)
apply(subst gpv'.map_sel)
apply(subst inline'_sel)
apply(subst inline1'_inline_conv_inline2)
apply(subst inline1'_inline_conv_inline2')
apply(unfold spmf.map_comp o_def case_sum_map_sum spmf_rel_map generat.rel_map)
apply(rule rel_spmf_reflI)
apply(split sum.split)
apply(clarsimp simp add: split_beta bind_gpv'_assoc gpv_of_gpv'_map_gpv' o_def rel_fun_def intro!: gpv'.cong_bind_gpv'[OF refl])
apply(subst bind_gpv_assoc[symmetric])
apply(subst bind_gpv_assoc[symmetric])
apply(simp add: split_beta)
apply(subst map_gpv'_id_bind_gpv')
apply(simp add: o_def)
-- ‹We have to reassociate the return tuple such that type types for @{term bind_gpv} are right›
apply(subgoal_tac "bind_gpv (aa xa) (λx. inline callee1 (ab (fst x)) (snd x)) =
map_gpv (λ(r2, s2, s1). ((r2, s2), s1)) id (map_gpv (λ((r2, s2), s1). (r2, s2, s1)) id (bind_gpv (aa xa) (λx. inline callee1 (ab (fst x)) (snd x))))")
apply(erule forw_subst[where P="λx. gpv'.v1.congclp _ (gpv'_of_gpv (bind_gpv x _)) _"])
apply(subst bind_map_gpv)
apply(subst map_gpv_id_bind_gpv)
apply(subst gpv'_of_gpv_bind_gpv)
apply(rule gpv'.cong_bind_gpv')
apply(simp add: o_def)
apply(clarsimp simp add: rel_fun_def)
apply(rule gpv'.cong_base)
apply(rule exI conjI)+
apply(rule refl)
apply simp
apply(simp add: gpv.map_comp o_def split_def id_def[symmetric] gpv.map_id)
done
lemma inline_assoc:
"inline callee1 (inline callee2 gpv s2) s1 =
map_gpv (λ(r, s2, s1). ((r, s2), s1)) id (inline (λ(s2, s1) c2. map_gpv (λ((r, s2), s1). (r, s2, s1)) id (inline callee1 (callee2 s2 c2) s1)) gpv (s2, s1))"
apply(subst (2 4) inline_def)
apply(subst gpv_of_gpv'_map_gpv'[symmetric])
apply(subst inline'_assoc[symmetric])
apply simp
done
end
lemma set_inline2_lift_spmf1: "set_spmf (inline2 (λs x. lift_spmf (p s x)) callee gpv s s') ⊆ range Inl"
apply(induction arbitrary: gpv s s' rule: inline2_fixp_induct)
subgoal by(rule cont_intro ccpo_class.admissible_leI)+
apply(auto simp add: set_bind_spmf o_def bind_UNION split: generat.split_asm sum.split_asm dest!: in_set_inline1_lift_spmf1)
apply blast
done
lemma in_set_inline2_lift_spmf1: "y ∈ set_spmf (inline2 (λs x. lift_spmf (p s x)) callee gpv s s') ⟹ ∃r s s'. y = Inl (r, s, s')"
by(drule set_inline2_lift_spmf1[THEN subsetD]) auto
lemma interaction_bound_inline1':
assumes "(⋀s x. interaction_bound (callee s x) ≤ q)"
shows "interaction_bound gpv ≤ p
⟹ set_spmf (inline1 callee gpv s) ⊆ {Inr (out, c', rpv) | out c' rpv.
(∀input. eSuc (interaction_bound (c' input)) ≤ q) ∧
(∀x. eSuc (interaction_bound (rpv x)) ≤ p)}
∪ range Inl"
proof(induction arbitrary: gpv s rule: inline1_fixp_induct)
{ case adm show ?case by(rule cont_intro ccpo_class.admissible_leI)+ }
{ case bottom show ?case by simp }
case (step inline1')
have *: "⋀out c input. IO out c ∈ set_spmf (the_gpv gpv) ⟹ eSuc (interaction_bound (c input)) ≤ p"
by(erule interaction_bound_IO[THEN order_trans])(rule step.prems)
have **: "⋀out c s x input. IO out c ∈ set_spmf (the_gpv (callee s x)) ⟹ eSuc (interaction_bound (c input)) ≤ q"
using interaction_bound_IO assms by(rule order_trans)
show ?case
by(auto 6 4 simp add: set_bind_spmf bind_UNION del: subsetI intro!: UN_least step.IH assms * ** order_trans[OF ile_eSuc *] split: generat.split)
qed
lemma interaction_bound_inline1:
"⟦ Inr (out, c', rpv) ∈ set_spmf (inline1 callee gpv s);
interaction_bound gpv ≤ p; ⋀s x. interaction_bound (callee s x) ≤ q ⟧
⟹ eSuc (interaction_bound (c' input)) ≤ q ∧ eSuc (interaction_bound (rpv x)) ≤ p"
using interaction_bound_inline1'[where callee=callee and q=q and gpv=gpv and p=p and s=s] by auto
lemma interaction_bounded_by_inline1:
"⟦ Inr (out, c', rpv) ∈ set_spmf (inline1 callee gpv s);
interaction_bounded_by gpv p; ⋀s x. interaction_bounded_by (callee s x) q ⟧
⟹ q ≠ 0 ∧ interaction_bounded_by (c' input) (q - 1) ∧ p ≠ 0 ∧ interaction_bounded_by (rpv x) (p - 1)"
unfolding interaction_bounded_by.simps
apply(drule (1) interaction_bound_inline1[where input=input and x=x], assumption)
apply(cases p q rule: nat.exhaust[case_product nat.exhaust])
apply(simp_all add: zero_enat_def[symmetric] eSuc_enat[symmetric])
done
lemma interaction_bounded_by_inline [interaction_bound]:
fixes callee :: "'s ⇒ 'call ⇒ ('ret × 's, 'call', 'ret') gpv"
and gpv :: "('a, 'call, 'ret) gpv"
assumes p: "interaction_bounded_by gpv p"
and q: "⋀s x. interaction_bounded_by (callee s x) q"
shows "interaction_bounded_by (inline callee gpv s) (p * q)"
proof -
fix x
def gpv' ≡ "Done (x, s) :: ('ret × 's, 'call', 'ret') gpv"
and r ≡ "0 :: nat" and rpv ≡ "(λ_. gpv) :: ('a, 'call, 'ret) rpv"
with p have "interaction_bounded_by gpv' r" "⋀y. interaction_bounded_by (rpv y) p"
by simp_all
hence "interaction_bound (bind_gpv gpv' (λ(y, s). inline callee (rpv y) s)) ≤ enat (r + p * q)"
proof(coinduction arbitrary: gpv' rpv p r rule: interaction_bound_coinduct)
case [rule_format]: (interaction_bound out c input gpv' rpv p r)
from ‹IO out c ∈ set_spmf (the_gpv (gpv' ⤜ (λ(y, s). inline callee (rpv y) s)))›
obtain generat where generat: "generat ∈ set_spmf (the_gpv gpv')"
and *: "IO out c ∈ set_spmf (if is_Pure generat then the_gpv (case result generat of (y, s) ⇒ inline callee (rpv y) s)
else return_spmf (IO (output generat) (λinput. continuation generat input ⤜ (λ(y, s). inline callee (rpv y) s))))"
by(clarsimp simp add: set_bind_spmf)
show ?case
proof(cases generat)
case (Pure ys)
obtain y s where ys [simp]: "ys = (y, s)" by(cases ys)
from Pure * have "IO out c ∈ set_spmf (the_gpv (inline callee (rpv y) s))" by simp
then obtain c' rpv' where c'rpv: "Inr (out, c', rpv') ∈ set_spmf (inline1 callee (rpv y) s)"
and c: "c = (λinput. c' input ⤜ (λ(x, s). inline callee (rpv' x) s))"
by(clarsimp simp add: inline_sel split: sum.split_asm)
from interaction_bounded_by_inline1[OF c'rpv interaction_bound(2) q]
have "q ≠ 0" and c': "interaction_bounded_by (c' input) (q - 1)"
and "p ≠ 0" and rpv': "⋀x. interaction_bounded_by (rpv' x) (p - 1)" by blast+
let ?n = "(q - 1) + (p - 1) * q"
have "eSuc ?n ≤ r + p * q" using ‹q ≠ 0› ‹p ≠ 0›
by(clarsimp simp add: eSuc_enat gr0_conv_Suc)
then show ?thesis unfolding c using c' rpv' q by blast
next
case (IO out' c')
with * have c: "c = (λinput. c' input ⤜ (λ(y, s). inline callee (rpv y) s))"
and [simp]: "out' = out" by simp_all
from ‹interaction_bounded_by gpv' r› generat IO
have "r > 0" and c': "interaction_bounded_by (c' input) (r - 1)"
by(auto dest: interaction_bounded_by_IO)
let ?n = "(r - 1) + p * q"
have "eSuc ?n ≤ r + p * q" using ‹r > 0› by(clarsimp simp add: gr0_conv_Suc eSuc_enat)
then show ?thesis unfolding c using c' interaction_bound(2) by blast
qed
qed
then show ?thesis unfolding gpv'_def rpv_def r_def interaction_bounded_by.simps by simp
qed
context
fixes ℐ :: "('call, 'ret) ℐ"
and ℐ' :: "('call', 'ret') ℐ"
and callee :: "'s ⇒ 'call ⇒ ('ret × 's, 'call', 'ret') gpv"
assumes results: "⋀s x. x ∈ outs_ℐ ℐ ⟹ results_gpv ℐ' (callee s x) ⊆ responses_ℐ ℐ x × UNIV"
begin
lemma inline1_in_sub_gpvs_callee:
assumes "Inr (out, callee', rpv') ∈ set_spmf (inline1 callee gpv s)"
and WT: "ℐ ⊢g gpv √"
shows "∃call∈outs_ℐ ℐ. ∃s. ∀x ∈ responses_ℐ ℐ' out. callee' x ∈ sub_gpvs ℐ' (callee s call)"
proof -
from WT
have "set_spmf (inline1 callee gpv s) ⊆ {Inr (out, callee', rpv') | out callee' rpv'.
∃call∈outs_ℐ ℐ. ∃s. ∀x ∈ responses_ℐ ℐ' out. callee' x ∈ sub_gpvs ℐ' (callee s call)} ∪ range Inl"
(is "?concl (inline1 callee) gpv s")
proof(induction arbitrary: gpv s rule: inline1_fixp_induct)
case adm show ?case by(intro cont_intro ccpo_class.admissible_leI)
case bottom show ?case by simp
case (step inline1')
{ fix out c
assume IO: "IO out c ∈ set_spmf (the_gpv gpv)"
from step.prems IO have out: "out ∈ outs_ℐ ℐ" by(rule WT_gpvD)
{ fix x s'
assume Pure: "Pure (x, s') ∈ set_spmf (the_gpv (callee s out))"
then have "(x, s') ∈ results_gpv ℐ' (callee s out)" by(rule results_gpv.Pure)
with out have "x ∈ responses_ℐ ℐ out" by(auto dest: results)
with step.prems IO have "ℐ ⊢g c x √" by(rule WT_gpvD)
hence "?concl inline1' (c x) s'" by(rule step.IH)
} moreover {
fix out' c'
assume "IO out' c' ∈ set_spmf (the_gpv (callee s out))"
hence "∀x∈responses_ℐ ℐ' out'. c' x ∈ sub_gpvs ℐ' (callee s out)"
by(auto intro: sub_gpvs.base)
then have "∃call∈outs_ℐ ℐ. ∃s. ∀x∈responses_ℐ ℐ' out'. c' x ∈ sub_gpvs ℐ' (callee s call)"
using out by blast
} moreover note calculation }
then show ?case using step.prems
by(auto del: subsetI simp add: set_bind_spmf bind_UNION intro!: UN_least split: generat.split)
qed
thus ?thesis using assms by fastforce
qed
lemma inline1_in_sub_gpvs:
assumes "Inr (out, callee', rpv') ∈ set_spmf (inline1 callee gpv s)"
and "(x, s') ∈ results_gpv ℐ' (callee' input)"
and "input ∈ responses_ℐ ℐ' out"
and "ℐ ⊢g gpv √"
shows "rpv' x ∈ sub_gpvs ℐ gpv"
proof -
from ‹ℐ ⊢g gpv √›
have "set_spmf (inline1 callee gpv s) ⊆ {Inr (out, callee', rpv') | out callee' rpv'.
∀input ∈ responses_ℐ ℐ' out. ∀(x, s')∈results_gpv ℐ' (callee' input). rpv' x ∈ sub_gpvs ℐ gpv}
∪ range Inl" (is "?concl (inline1 callee) gpv s" is "_ ⊆ ?rhs gpv s")
proof(induction arbitrary: gpv s rule: inline1_fixp_induct)
case adm show ?case by(intro cont_intro ccpo_class.admissible_leI)
case bottom show ?case by simp
next
case (step inline1')
{ fix out c
assume IO: "IO out c ∈ set_spmf (the_gpv gpv)"
from step.prems IO have out: "out ∈ outs_ℐ ℐ" by(rule WT_gpvD)
{ fix x s'
assume Pure: "Pure (x, s') ∈ set_spmf (the_gpv (callee s out))"
then have "(x, s') ∈ results_gpv ℐ' (callee s out)" by(rule results_gpv.Pure)
with out have "x ∈ responses_ℐ ℐ out" by(auto dest: results)
with step.prems IO have "ℐ ⊢g c x √" by(rule WT_gpvD)
hence "?concl inline1' (c x) s'" by(rule step.IH)
also have "… ⊆ ?rhs gpv s'" using IO Pure
by(fastforce intro: sub_gpvs.cont dest: WT_gpv_OutD[OF step.prems] results[THEN subsetD, OF _ results_gpv.Pure])
finally have "set_spmf (inline1' (c x) s') ⊆ …" .
} moreover {
fix out' c' input x s'
assume "IO out' c' ∈ set_spmf (the_gpv (callee s out))"
and "input ∈ responses_ℐ ℐ' out'" and "(x, s') ∈ results_gpv ℐ' (c' input)"
then have "c x ∈ sub_gpvs ℐ gpv" using IO
by(auto intro!: sub_gpvs.base dest: WT_gpv_OutD[OF step.prems] results[THEN subsetD, OF _ results_gpv.IO])
} moreover note calculation }
then show ?case
by(auto simp add: set_bind_spmf bind_UNION intro!: UN_least split: generat.split del: subsetI)
qed
with assms show ?thesis by fastforce
qed
context
assumes WT: "⋀x s. x ∈ outs_ℐ ℐ ⟹ ℐ' ⊢g callee s x √"
begin
lemma WT_gpv_inline1:
assumes "Inr (out, rpv, rpv') ∈ set_spmf (inline1 callee gpv s)"
and "ℐ ⊢g gpv √"
shows "out ∈ outs_ℐ ℐ'" (is "?thesis1")
and "input ∈ responses_ℐ ℐ' out ⟹ ℐ' ⊢g rpv input √" (is "PROP ?thesis2")
and "⟦ input ∈ responses_ℐ ℐ' out; (x, s') ∈ results_gpv ℐ' (rpv input) ⟧ ⟹ ℐ ⊢g rpv' x √" (is "PROP ?thesis3")
proof -
from ‹ℐ ⊢g gpv √›
have "set_spmf (inline1 callee gpv s) ⊆ {Inr (out, rpv, rpv') | out rpv rpv'. out ∈ outs_ℐ ℐ'} ∪ range Inl"
proof(induction arbitrary: gpv s rule: inline1_fixp_induct)
{ case adm show ?case by(intro cont_intro ccpo_class.admissible_leI) }
{ case bottom show ?case by simp }
case (step inline1')
{ fix out c
assume IO: "IO out c ∈ set_spmf (the_gpv gpv)"
from step.prems IO have out: "out ∈ outs_ℐ ℐ" by(rule WT_gpvD)
{ fix x s'
assume Pure: "Pure (x, s') ∈ set_spmf (the_gpv (callee s out))"
then have "(x, s') ∈ results_gpv ℐ' (callee s out)" by(rule results_gpv.Pure)
with out have "x ∈ responses_ℐ ℐ out" by(auto dest: results)
with step.prems IO have "ℐ ⊢g c x √" by(rule WT_gpvD)
} moreover {
fix out' c'
from out have "ℐ' ⊢g callee s out √" by(rule WT)
moreover assume "IO out' c' ∈ set_spmf (the_gpv (callee s out))"
ultimately have "out' ∈ outs_ℐ ℐ'" by(rule WT_gpvD)
} moreover note calculation }
then show ?case
by(auto del: subsetI simp add: set_bind_spmf bind_UNION intro!: UN_least split: generat.split intro!: step.IH[THEN order_trans])
qed
then show ?thesis1 using assms by auto
assume "input ∈ responses_ℐ ℐ' out"
with inline1_in_sub_gpvs_callee[OF ‹Inr _ ∈ _›] ‹ℐ ⊢g gpv √›
obtain out' s where "out' ∈ outs_ℐ ℐ"
and *: "rpv input ∈ sub_gpvs ℐ' (callee s out')" by auto
from ‹out' ∈ _› have "ℐ' ⊢g callee s out' √" by(rule WT)
then show "ℐ' ⊢g rpv input √" using * by(rule WT_sub_gpvsD)
assume "(x, s') ∈ results_gpv ℐ' (rpv input)"
with ‹Inr _ ∈ _› have "rpv' x ∈ sub_gpvs ℐ gpv"
using ‹input ∈ _› ‹ℐ ⊢g gpv √› by(rule inline1_in_sub_gpvs)
with ‹ℐ ⊢g gpv √› show "ℐ ⊢g rpv' x √" by(rule WT_sub_gpvsD)
qed
lemma WT_gpv_inline:
assumes "ℐ ⊢g gpv √"
shows "ℐ' ⊢g inline callee gpv s √"
using assms
proof(coinduction arbitrary: gpv s rule: WT_gpv_coinduct_bind)
case (WT_gpv out c gpv)
from ‹IO out c ∈ _› obtain callee' rpv'
where Inr: "Inr (out, callee', rpv') ∈ set_spmf (inline1 callee gpv s)"
and c: "c = (λinput. callee' input ⤜ (λ(x, s). inline callee (rpv' x) s))"
by(clarsimp simp add: inline_sel split: sum.split_asm)
from Inr ‹ℐ ⊢g gpv √› have ?out by(rule WT_gpv_inline1)
moreover have "?cont TYPE('ret × 's)" (is "∀input∈_. _ ∨ _ ∨ ?case' input")
proof(rule ballI disjI2)+
fix input
assume "input ∈ responses_ℐ ℐ' out"
with Inr ‹ℐ ⊢g gpv √ ›have "ℐ' ⊢g callee' input √"
and "⋀x s'. (x, s') ∈ results_gpv ℐ' (callee' input) ⟹ ℐ ⊢g rpv' x √"
by(blast intro: WT_gpv_inline1)+
then show "?case' input" by(subst c)(auto 4 4)
qed
ultimately show "?case TYPE('ret × 's)" ..
qed
end
context
fixes gpv :: "('a, 'call, 'ret) gpv"
assumes gpv: "lossless_gpv ℐ gpv" "ℐ ⊢g gpv √"
begin
lemma lossless_spmf_inline1:
assumes lossless: "⋀s x. x ∈ outs_ℐ ℐ ⟹ lossless_spmf (the_gpv (callee s x))"
shows "lossless_spmf (inline1 callee gpv s)"
using gpv
proof(induction arbitrary: s rule: lossless_WT_gpv_induct)
case (lossless_gpv p)
show ?case using ‹lossless_spmf p›
apply(subst inline1_unfold)
apply(auto split: generat.split intro: lossless lossless_gpv.hyps dest: results[THEN subsetD, rotated, OF results_gpv.Pure] intro: lossless_gpv.IH)
done
qed
lemma lossless_gpv_inline1:
assumes *: "Inr (out, rpv, rpv') ∈ set_spmf (inline1 callee gpv s)"
and **: "input ∈ responses_ℐ ℐ' out"
and lossless: "⋀s x. x ∈ outs_ℐ ℐ ⟹ lossless_gpv ℐ' (callee s x)"
shows "lossless_gpv ℐ' (rpv input)"
proof -
from inline1_in_sub_gpvs_callee[OF * gpv(2)] **
obtain out' s where "out' ∈ outs_ℐ ℐ" and ***: "rpv input ∈ sub_gpvs ℐ' (callee s out')" by blast
from ‹out' ∈ _› have "lossless_gpv ℐ' (callee s out')" by(rule lossless)
thus ?thesis using *** by(rule lossless_sub_gpvsD)
qed
lemma lossless_results_inline1:
assumes "Inr (out, rpv, rpv') ∈ set_spmf (inline1 callee gpv s)"
and "(x, s') ∈ results_gpv ℐ' (rpv input)"
and "input ∈ responses_ℐ ℐ' out"
shows "lossless_gpv ℐ (rpv' x)"
proof -
from assms gpv(2) have "rpv' x ∈ sub_gpvs ℐ gpv" by(rule inline1_in_sub_gpvs)
with gpv(1) show "lossless_gpv ℐ (rpv' x)" by(rule lossless_sub_gpvsD)
qed
end
lemmas lossless_inline1[rotated 2] = lossless_spmf_inline1 lossless_gpv_inline1 lossless_results_inline1
lemma lossless_inline[rotated]:
fixes gpv :: "('a, 'call, 'ret) gpv"
assumes gpv: "lossless_gpv ℐ gpv" "ℐ ⊢g gpv √"
and lossless: "⋀s x. x ∈ outs_ℐ ℐ ⟹ lossless_gpv ℐ' (callee s x)"
shows "lossless_gpv ℐ' (inline callee gpv s)"
using gpv
proof(induction arbitrary: s rule: lossless_WT_gpv_induct_strong)
case (lossless_gpv p)
have lp: "lossless_gpv ℐ (GPV p)" by(rule lossless_sub_gpvsI)(auto intro: lossless_gpv.hyps)
moreover have wp: "ℐ ⊢g GPV p √" by(rule WT_sub_gpvsI)(auto intro: lossless_gpv.hyps)
ultimately have "lossless_spmf (the_gpv (inline callee (GPV p) s))"
by(auto simp add: inline_sel intro: lossless_spmf_inline1 lossless_gpv_lossless_spmfD[OF lossless])
moreover {
fix out c input
assume IO: "IO out c ∈ set_spmf (the_gpv (inline callee (GPV p) s))"
and "input ∈ responses_ℐ ℐ' out"
from IO obtain callee' rpv
where Inr: "Inr (out, callee', rpv) ∈ set_spmf (inline1 callee (GPV p) s)"
and c: "c = (λinput. callee' input ⤜ (λ(x, y). inline callee (rpv x) y))"
by(clarsimp simp add: inline_sel split: sum.split_asm)
from Inr ‹input ∈ _› lossless lp wp have "lossless_gpv ℐ' (callee' input)" by(rule lossless_inline1)
moreover {
fix x s'
assume "(x, s') ∈ results_gpv ℐ' (callee' input)"
with Inr have "rpv x ∈ sub_gpvs ℐ (GPV p)" using ‹input ∈ _› wp by(rule inline1_in_sub_gpvs)
hence "lossless_gpv ℐ' (inline callee (rpv x) s')" by(rule lossless_gpv.IH)
} ultimately have "lossless_gpv ℐ' (c input)" unfolding c by clarsimp
} ultimately show ?case by(rule lossless_gpvI)
qed
end
definition id_oracle :: "'s ⇒ 'call ⇒ ('ret × 's, 'call, 'ret) gpv"
where "id_oracle s x = Pause x (λx. Done (x, s))"
lemma inline1_id_oracle:
"inline1 id_oracle gpv s =
map_spmf (λgenerat. case generat of Pure x ⇒ Inl (x, s) | IO out c ⇒ Inr (out, λx. Done (x, s), c)) (the_gpv gpv)"
by(subst inline1.simps)(auto simp add: id_oracle_def map_spmf_conv_bind_spmf intro!: bind_spmf_cong split: generat.split)
lemma inline_id_oracle [simp]: "inline id_oracle gpv s = map_gpv (λx. (x, s)) id gpv"
by(coinduction arbitrary: gpv s)(auto 4 3 simp add: inline_sel inline1_id_oracle spmf_rel_map gpv.map_sel o_def generat.rel_map intro!: rel_spmf_reflI rel_funI split: generat.split)
subsection {* Running GPVs *}
type_synonym ('call, 'ret, 's) callee = "'s ⇒ 'call ⇒ ('ret × 's) spmf"
context fixes callee :: "('call, 'ret, 's) callee" notes [[function_internals]] begin
partial_function (spmf) exec_gpv :: "('a, 'call, 'ret) gpv ⇒ 's ⇒ ('a × 's) spmf"
where
"exec_gpv c s =
the_gpv c ⤜
case_generat (λx. return_spmf (x, s))
(λout c. callee s out ⤜ (λ(x, y). exec_gpv (c x) y))"
abbreviation run_gpv :: "('a, 'call, 'ret) gpv ⇒ 's ⇒ 'a spmf"
where "run_gpv gpv s ≡ map_spmf fst (exec_gpv gpv s)"
lemma exec_gpv_fixp_induct [case_names adm bottom step]:
assumes "ccpo.admissible (fun_lub lub_spmf) (fun_ord (ord_spmf op =)) (λf. P (λc s. f (c, s)))"
and "P (λ_ _. return_pmf None)"
and "⋀exec_gpv. P exec_gpv ⟹
P (λc s. the_gpv c ⤜ case_generat (λx. return_spmf (x, s)) (λout c. callee s out ⤜ (λ(x, y). exec_gpv (c x) y)))"
shows "P exec_gpv"
using assms(1)
by(rule exec_gpv.fixp_induct[unfolded curry_conv[abs_def]])(simp_all add: assms(2-))
lemma exec_gpv_fixp_induct_strong [case_names adm bottom step]:
assumes "ccpo.admissible (fun_lub lub_spmf) (fun_ord (ord_spmf op =)) (λf. P (λc s. f (c, s)))"
and "P (λ_ _. return_pmf None)"
and "⋀exec_gpv'. ⟦ ⋀c s. ord_spmf op = (exec_gpv' c s) (exec_gpv c s); P exec_gpv' ⟧
⟹ P (λc s. the_gpv c ⤜ case_generat (λx. return_spmf (x, s)) (λout c. callee s out ⤜ (λ(x, y). exec_gpv' (c x) y)))"
shows "P exec_gpv"
using assms
by(rule spmf.fixp_strong_induct_uc[where P="λf. P (curry f)" and U=case_prod and C=curry, OF exec_gpv.mono exec_gpv_def, simplified curry_case_prod, simplified curry_conv[abs_def] fun_ord_def split_paired_All prod.case case_prod_eta, OF refl]) blast
lemma exec_gpv_fixp_induct_strong2 [case_names adm bottom step]:
assumes "ccpo.admissible (fun_lub lub_spmf) (fun_ord (ord_spmf op =)) (λf. P (λc s. f (c, s)))"
and "P (λ_ _. return_pmf None)"
and "⋀exec_gpv'.
⟦ ⋀c s. ord_spmf op = (exec_gpv' c s) (exec_gpv c s);
⋀c s. ord_spmf op = (exec_gpv' c s) (the_gpv c ⤜ case_generat (λx. return_spmf (x, s)) (λout c. callee s out ⤜ (λ(x, y). exec_gpv' (c x) y)));
P exec_gpv' ⟧
⟹ P (λc s. the_gpv c ⤜ case_generat (λx. return_spmf (x, s)) (λout c. callee s out ⤜ (λ(x, y). exec_gpv' (c x) y)))"
shows "P exec_gpv"
using assms
by(rule spmf.fixp_induct_strong2_uc[where P="λf. P (curry f)" and U=case_prod and C=curry, OF exec_gpv.mono exec_gpv_def, simplified curry_case_prod, simplified curry_conv[abs_def] fun_ord_def split_paired_All prod.case case_prod_eta, OF refl]) blast+
end
lemma exec_gpv_conv_inline1:
"exec_gpv callee gpv s = map_spmf projl (inline1 (λs c. lift_spmf (callee s c) :: (_, unit, unit) gpv) gpv s)"
by(induction arbitrary: gpv s rule: parallel_fixp_induct_2_2[OF partial_function_definitions_spmf partial_function_definitions_spmf exec_gpv.mono inline1.mono exec_gpv_def inline1_def, unfolded lub_spmf_empty, case_names adm bottom step])
(auto simp add: map_spmf_bind_spmf o_def spmf.map_comp bind_map_spmf split_def intro!: bind_spmf_cong[OF refl] split: generat.split)
lemma exec_gpv_simps:
"exec_gpv callee gpv s =
the_gpv gpv ⤜
case_generat (λx. return_spmf (x, s))
(λout rpv. callee s out ⤜ (λ(x, y). exec_gpv callee (rpv x) y))"
by(fact exec_gpv.simps)
lemma exec_gpv_lift_spmf [simp]:
"exec_gpv callee (lift_spmf p) s = bind_spmf p (λx. return_spmf (x, s))"
by(simp add: exec_gpv_conv_inline1 spmf.map_comp o_def map_spmf_conv_bind_spmf)
lemma exec_gpv_Done [simp]: "exec_gpv callee (Done x) s = return_spmf (x, s)"
by(simp add: exec_gpv_conv_inline1)
lemma exec_gpv_Fail [simp]: "exec_gpv callee Fail s = return_pmf None"
by(simp add: exec_gpv_conv_inline1)
lemma if_distrib_exec_gpv [if_distribs]:
"exec_gpv callee (if b then x else y) s = (if b then exec_gpv callee x s else exec_gpv callee y s)"
by simp
lemmas exec_gpv_fixp_parallel_induct [case_names adm bottom step] =
parallel_fixp_induct_2_2[OF partial_function_definitions_spmf partial_function_definitions_spmf exec_gpv.mono exec_gpv.mono exec_gpv_def exec_gpv_def, unfolded lub_spmf_empty]
context includes lifting_syntax begin
lemma exec_gpv_parametric [transfer_rule]:
fixes A :: "'a ⇒ 'b ⇒ bool"
and CALL :: "'call ⇒ 'call' ⇒ bool"
and S :: "'s ⇒ 's' ⇒ bool"
shows
"((S ===> CALL ===> rel_spmf (rel_prod (op = :: 'ret ⇒ _) S)) ===> rel_gpv A CALL ===> S ===> rel_spmf (rel_prod A S))
exec_gpv exec_gpv"
apply(rule rel_funI)+
apply(unfold spmf_rel_map exec_gpv_conv_inline1)
apply(rule rel_spmf_mono_strong)
apply(erule inline1_parametric[THEN rel_funD, THEN rel_funD, THEN rel_funD, rotated])
prefer 3
apply(drule in_set_inline1_lift_spmf1)+
apply fastforce
subgoal by simp
subgoal premises [transfer_rule] by transfer_prover
done
end
lemma exec_gpv_bind: "exec_gpv callee (c ⤜ f) s = exec_gpv callee c s ⤜ (λ(x, s') ⇒ exec_gpv callee (f x) s')"
by(auto simp add: exec_gpv_conv_inline1 inline1_bind_gpv map_spmf_bind_spmf o_def bind_map_spmf intro!: bind_spmf_cong[OF refl] dest: in_set_inline1_lift_spmf1)
lemma exec_gpv_map_gpv_id:
"exec_gpv oracle (map_gpv f id gpv) σ = map_spmf (apfst f) (exec_gpv oracle gpv σ)"
proof(rule sym)
def gpv' ≡ "map_gpv f id gpv"
have [transfer_rule]: "rel_gpv (λx y. y = f x) op = gpv gpv'"
unfolding gpv'_def by(simp add: gpv.rel_map gpv.rel_refl)
have "rel_spmf (rel_prod (λx y. y = f x) op =) (exec_gpv oracle gpv σ) (exec_gpv oracle gpv' σ)"
by transfer_prover
thus "map_spmf (apfst f) (exec_gpv oracle gpv σ) = exec_gpv oracle (map_gpv f id gpv) σ"
unfolding spmf_rel_eq[symmetric] gpv'_def spmf_rel_map by(rule rel_spmf_mono) clarsimp
qed
lemma exec_gpv_Pause [simp]:
"exec_gpv callee (Pause out f) s = callee s out ⤜ (λ(x, s'). exec_gpv callee (f x) s')"
by(simp add: inline1_Pause map_spmf_bind_spmf bind_map_spmf o_def exec_gpv_conv_inline1 split_def)
lemma exec_gpv_inline:
fixes callee :: "'s ⇒ 'c ⇒ ('r × 's) spmf"
and gpv :: "'s' ⇒ 'c' ⇒ ('r' × 's', 'c, 'r) gpv"
shows "exec_gpv callee (inline gpv c' s') s =
map_spmf (λ(x, s', s). ((x, s'), s)) (exec_gpv (λ(s', s) y. map_spmf (λ((x, s'), s). (x, s', s)) (exec_gpv callee (gpv s' y) s)) c' (s', s))"
(is "?lhs = ?rhs")
proof -
have "?lhs = map_spmf projl (map_spmf (map_sum (λ(x, s2, y). ((x, s2), y))
(λ(x, rpv'' :: ('r × 's, unit, unit) rpv, rpv', rpv). (x, rpv'', λr1. bind_gpv (rpv' r1) (λ(r2, y). inline gpv (rpv r2) y))))
(inline2 (λs c. lift_spmf (callee s c)) gpv c' s' s))"
unfolding exec_gpv_conv_inline1 by(simp add: inline1_inline_conv_inline2)
also have "… = map_spmf (λ(x, s', s). ((x, s'), s)) (map_spmf projl (map_spmf (map_sum id
(λ(x, rpv'' :: ('r × 's, unit, unit) rpv, rpv', rpv). (x, λr. bind_gpv (rpv'' r) (λ(r1, s1). map_gpv (λ((r2, s2), s1). (r2, s2, s1)) id (inline (λs c. lift_spmf (callee s c)) (rpv' r1) s1)), rpv)))
(inline2 (λs c. lift_spmf (callee s c)) gpv c' s' s)))"
unfolding spmf.map_comp by(rule map_spmf_cong[OF refl])(auto dest!: in_set_inline2_lift_spmf1)
also have "… = ?rhs" unfolding exec_gpv_conv_inline1
by(subst inline1_inline_conv_inline2'[symmetric])(simp add: spmf.map_comp split_def inline_lift_spmf1 map_lift_spmf)
finally show ?thesis .
qed
locale callee_invariant_base =
fixes callee :: "'s ⇒ 'a ⇒ ('b × 's) spmf"
and I :: "'s ⇒ bool"
locale callee_invariant = callee_invariant_base
+
assumes callee_invariant: "⋀s x y s'. ⟦ (y, s') ∈ set_spmf (callee s x); I s ⟧ ⟹ I s'"
begin
lemma exec_gpv_invariant':
"I s ⟹ set_spmf (exec_gpv callee gpv s) ⊆ {(x, s'). I s'}"
proof(induction arbitrary: gpv s rule: exec_gpv_fixp_induct)
case adm show ?case by(intro cont_intro ccpo_class.admissible_leI)
case bottom show ?case by simp
case step thus ?case
by(auto simp add: set_bind_spmf bind_UNION intro!: UN_least del: subsetI split: generat.split dest!: callee_invariant)
qed
lemma exec_gpv_invariant:
"⟦ (x, s') ∈ set_spmf (exec_gpv callee gpv s); I s ⟧ ⟹ I s'"
by(drule exec_gpv_invariant') blast
lemma interaction_bounded_by_exec_gpv_count':
fixes count
assumes bound: "interaction_bounded_by gpv n"
and count: "⋀s x y s'. ⟦ (y, s') ∈ set_spmf (callee s x); I s ⟧ ⟹ count s' ≤ Suc (count s)"
and I: "I s"
shows "set_spmf (exec_gpv callee gpv s) ⊆ {(x, s'). count s' ≤ n + count s}"
using bound I
proof(induction arbitrary: gpv s n rule: exec_gpv_fixp_induct)
case adm show ?case by(intro cont_intro ccpo_class.admissible_leI)
case bottom show ?case by simp
case (step exec_gpv')
{ fix out c input s'
assume out: "IO out c ∈ set_spmf (the_gpv gpv)"
and input: "(input, s') ∈ set_spmf (callee s out)"
from step.prems out have "n > 0"
and bound': "interaction_bounded_by (c input) (n - 1)"
by(auto dest: interaction_bounded_by_contD)
from input ‹I s› have "I s'" by(rule callee_invariant)
with bound' have "set_spmf (exec_gpv' (c input) s') ⊆ {(x, s''). count s'' ≤ n - 1 + count s'}"
by(rule step.IH)
also have "… ⊆ {(x, s''). count s'' ≤ n + count s}" using ‹n > 0› count[OF input ‹I s›]
by(cases n rule: nat.exhaust)(auto)
also note calculation }
then show ?case
by(clarsimp simp add: set_bind_spmf bind_UNION del: subsetI intro!: UN_least split: generat.split)
qed
lemma interaction_bounded_by_exec_gpv_count:
fixes count
assumes bound: "interaction_bounded_by gpv n"
and xs': "(x, s') ∈ set_spmf (exec_gpv callee gpv s)"
and count: "⋀s x y s'. ⟦ (y, s') ∈ set_spmf (callee s x); I s ⟧ ⟹ count s' ≤ Suc (count s)"
and I: "I s"
shows "count s' ≤ n + count s"
using bound count I by(rule interaction_bounded_by_exec_gpv_count'[THEN subsetD, OF _ _ _ xs', unfolded mem_Collect_eq prod.case])
lemma interaction_bounded_by_exec_gpv_bad':
fixes bad
assumes bound: "interaction_bounded_by gpv n"
and I: "I s"
and not_bad: "¬ bad s"
and prob_bad: "⋀s x. ⟦ I s; ¬ bad s ⟧ ⟹ spmf (map_spmf (bad ∘ snd) (callee s x)) True ≤ k"
shows "ennreal (spmf (map_spmf (bad ∘ snd) (exec_gpv callee gpv s)) True) ≤ ennreal k * n"
proof -
from prob_bad[OF I not_bad] have k_nonneg: "0 ≤ k" by (meson order_trans pmf_nonneg)
show ?thesis using bound I not_bad
proof(induction arbitrary: n s gpv rule: exec_gpv_fixp_induct)
case adm show ?case by(intro cont_intro ccpo_class.admissible_leI)
case bottom show ?case using k_nonneg by(cases n)simp_all
case (step exec_gpv')
let ?M = "restrict_space (measure_spmf (the_gpv gpv)) {x. ¬ is_Pure x}"
let ?O = "λgenerat. measure_spmf (callee s (output generat))"
have "ennreal (spmf (map_spmf (bad ∘ snd) (the_gpv gpv ⤜ case_generat (λx. return_spmf (x, s)) (λout c. callee s out ⤜ (λ(x, y). exec_gpv' (c x) y)))) True)
= ∫⇧+ generat. ennreal (spmf (map_spmf (bad ∘ snd) (case generat of Pure x ⇒ return_spmf (x, s)
| IO out c ⇒ callee s out ⤜ (λ(x, y). exec_gpv' (c x) y))) True)
∂measure_spmf (the_gpv gpv)"
by(simp add: map_spmf_bind_spmf ennreal_spmf_bind)
also have "… = ∫⇧+ generat. ennreal (spmf (bind_spmf (callee s (output generat)) (λ(x, s'). map_spmf (bad ∘ snd) (exec_gpv' (continuation generat x) s'))) True) ∂?M"
by(auto simp add: nn_integral_restrict_space ‹¬ bad s› map_spmf_bind_spmf split_def o_def intro: nn_integral_cong split: generat.split)
also have "… = ∫⇧+ generat. ∫⇧+ (x, s'). ennreal (spmf (map_spmf (bad ∘ snd) (exec_gpv' (continuation generat x) s')) True) ∂?O generat ∂?M"
by(simp add: ennreal_spmf_bind split_def)
also have "… ≤ ∫⇧+ generat. ∫⇧+ (x, s'). (if bad s' then 1 else ennreal k * (n - 1)) ∂?O generat ∂?M"
proof(clarsimp simp add: AE_restrict_space_iff not_is_Pure_conv intro!: nn_integral_mono_AE split del: if_split)
fix out c x s'
assume IO: "IO out c ∈ set_spmf (the_gpv gpv)"
and xs': "(x, s') ∈ set_spmf (callee s out)"
from xs' ‹I s› have "I s'" by(rule callee_invariant)
show "ennreal (spmf (map_spmf (bad ∘ snd) (exec_gpv' (c x) s')) True) ≤ (if bad s' then 1 else ennreal k * (n - 1))"
proof(cases "bad s'")
case False
from step.prems(1) IO have "interaction_bounded_by (c x) (n - 1)"
by(blast dest: interaction_bounded_by_contD)
then show ?thesis using ‹I s'› False step.IH[of "c x" "n - 1" s'] by(simp add: o_def)
qed(simp add: pmf_le_1)
qed
also have "… = ∫⇧+ generat. ∫⇧+ b. indicator {True} b + ennreal (k * (n - 1)) * indicator {False} b ∂measure_spmf (map_spmf (bad ∘ snd) (callee s (output generat))) ∂?M"
(is "_ = ∫⇧+ generat. ∫⇧+ _. _ ∂?O' generat ∂_")
by(auto intro!: nn_integral_cong simp add: ennreal_mult[symmetric] ennreal_of_nat_eq_real_of_nat k_nonneg)
also have "… = ∫⇧+ generat. (∫⇧+ b. indicator {True} b ∂?O' generat) + ennreal k * (n - 1) * ∫⇧+ b. indicator {False} b ∂?O' generat ∂?M"
by(subst nn_integral_add)(simp_all add: k_nonneg nn_integral_cmult o_def ennreal_of_nat_eq_real_of_nat ennreal_mult)
also have "… = ∫⇧+ generat. spmf (map_spmf (bad ∘ snd) (callee s (output generat))) True + k * (n - 1) * spmf (map_spmf (bad ∘ snd) (callee s (output generat))) False ∂?M"
by(simp del: nn_integral_map_spmf add: emeasure_spmf_single ennreal_mult k_nonneg ennreal_of_nat_eq_real_of_nat)
also have "… ≤ ∫⇧+ generat. ennreal k * n ∂?M"
proof(rule nn_integral_mono_AE; clarsimp simp add: AE_restrict_space_iff not_is_Pure_conv ennreal_of_nat_eq_real_of_nat ennreal_mult[symmetric] k_nonneg ennreal_plus[symmetric] simp del: One_nat_def ennreal_plus)
fix out c
assume "IO out c ∈ set_spmf (the_gpv gpv)"
hence "n > 0" using step.prems(1) by(blast dest: interaction_bounded_by_contD)
have "spmf (map_spmf (bad ∘ snd) (callee s out)) True ≤ k" (is "?o True ≤ _")
using ‹I s› ‹¬ bad s› by(rule prob_bad)
hence "?o True ≤ k" by(simp del: o_apply)
hence "?o True + k * (n - 1) * ?o False ≤ k + k * (n - 1) * 1"
by(rule add_mono)(rule mult_left_mono, simp_all add: pmf_le_1 k_nonneg)
also have "… ≤ k * n" using `n > 0`
by(cases n)(auto simp add: zero_enat_def gr0_conv_Suc eSuc_enat[symmetric] field_simps)
finally show "?o True + k * (n - 1) * ?o False ≤ k * n" .
qed
also have "… ≤ k * n * 1"
by(simp add: k_nonneg emeasure_restrict_space measure_spmf.emeasure_eq_measure space_restrict_space measure_spmf.subprob_measure_le_1 mult_left_mono ennreal_of_nat_eq_real_of_nat ennreal_mult del: mult_1_right)
finally show ?case by(simp del: o_apply add: ennreal_of_nat_eq_real_of_nat ennreal_mult k_nonneg)
qed
qed
end
interpretation oi_True: callee_invariant callee "λ_. True" for callee
by unfold_locales (simp_all)
inductive WT_callee :: "('call, 'ret) ℐ ⇒ ('call ⇒ ('ret × 's) spmf) ⇒ bool" ("(_) ⊢c/ (_) √" [100, 0] 99)
for ℐ callee
where
WT_callee:
"⟦ ⋀call ret s. ⟦ call ∈ outs_ℐ ℐ; (ret, s) ∈ set_spmf (callee call) ⟧ ⟹ ret ∈ responses_ℐ ℐ call ⟧
⟹ ℐ ⊢c callee √"
lemmas WT_calleeI = WT_callee
hide_fact WT_callee
lemma WT_calleeD: "⟦ ℐ ⊢c callee √; (ret, s) ∈ set_spmf (callee out); out ∈ outs_ℐ ℐ ⟧ ⟹ ret ∈ responses_ℐ ℐ out"
by(rule WT_callee.cases)
lemma WT_callee_top [intro!, simp]: "ℐ_top ⊢c callee √"
by(rule WT_calleeI) simp
lemma (in callee_invariant) lossless_exec_gpv:
assumes gpv: "lossless_gpv ℐ gpv"
and callee: "⋀s out. ⟦ out ∈ outs_ℐ ℐ; I s ⟧ ⟹ lossless_spmf (callee s out)"
and WT_gpv: "ℐ ⊢g gpv √"
and WT_callee: "⋀s. I s ⟹ ℐ ⊢c callee s √"
and I: "I s"
shows "lossless_spmf (exec_gpv callee gpv s)"
using gpv WT_gpv I
proof(induction arbitrary: s rule: lossless_WT_gpv_induct)
case (lossless_gpv gpv)
show ?case using lossless_gpv.hyps lossless_gpv.prems
by(subst exec_gpv.simps)(fastforce split: generat.split simp add: callee intro!: lossless_gpv.IH intro: WT_calleeD[OF WT_callee] elim!: callee_invariant)
qed
lemma lossless_exec_gpv:
"⟦ lossless_gpv ℐ gpv; ⋀s out. out ∈ outs_ℐ ℐ ⟹ lossless_spmf (callee s out);
ℐ ⊢g gpv √; ⋀s. ℐ ⊢c callee s √ ⟧
⟹ lossless_spmf (exec_gpv callee gpv s)"
by(rule oi_True.lossless_exec_gpv; simp)
lemma (in callee_invariant) in_set_spmf_exec_gpv_into_results_gpv:
assumes *: "(x, s') ∈ set_spmf (exec_gpv callee gpv s)"
and WT_gpv : "ℐ ⊢g gpv √"
and WT_callee: "⋀s. I s ⟹ ℐ ⊢c callee s √"
and I: "I s"
shows "x ∈ results_gpv ℐ gpv"
proof -
have "set_spmf (exec_gpv callee gpv s) ⊆ results_gpv ℐ gpv × UNIV"
using WT_gpv I
proof(induction arbitrary: gpv s rule: exec_gpv_fixp_induct)
{ case adm show ?case by(intro cont_intro ccpo_class.admissible_leI) }
{ case bottom show ?case by simp }
case (step exec_gpv')
{ fix out c ret s'
assume IO: "IO out c ∈ set_spmf (the_gpv gpv)"
and ret: "(ret, s') ∈ set_spmf (callee s out)"
from step.prems(1) IO have "out ∈ outs_ℐ ℐ" by(rule WT_gpvD)
with WT_callee[OF ‹I s›] ret have "ret ∈ responses_ℐ ℐ out" by(rule WT_calleeD)
with step.prems(1) IO have "ℐ ⊢g c ret √" by(rule WT_gpvD)
moreover from ret ‹I s› have "I s'" by(rule callee_invariant)
ultimately have "set_spmf (exec_gpv' (c ret) s') ⊆ results_gpv ℐ (c ret) × UNIV"
by(rule step.IH)
also have "… ⊆ results_gpv ℐ gpv × UNIV" using IO ‹ret ∈ _›
by(auto intro: results_gpv.IO)
finally have "set_spmf (exec_gpv' (c ret) s') ⊆ results_gpv ℐ gpv × UNIV" . }
then show ?case using step.prems
by(auto simp add: set_bind_spmf bind_UNION intro!: UN_least del: subsetI split: generat.split intro: results_gpv.Pure)
qed
thus "x ∈ results_gpv ℐ gpv" using * by blast+
qed
end