Theory IO_Execution

theory IO_Execution
imports IO_Monad TLList
subsection {* Semantics of the IO monad *}

theory IO_Execution imports
  IO_Monad
  "../../AFP/Coinductive/TLList"
begin

hide_const (open) cont

lemma lnth_ltl_ldropn_conv_ldropn:
  "enat n < llength xs ⟹ LCons (lnth xs n) (ltl (ldropn n xs)) = ldropn n xs"
by(simp add: ltl_ldropn ldrop_eSuc_ltl[symmetric] ldropn_Suc_conv_ldropn)

lemma lprefix_lappend2_iff:
  "llength xs = llength us
  ⟹ lprefix (lappend xs ys) (lappend us vs) ⟷ xs = us ∧ (lfinite us ⟶ lprefix ys vs)"
by(metis lappend_eq_lappend_conv lappend_lprefixE lprefix_lappend_same)

lemma tlength_eq_0_iff:
  "tlength xs = 0 ⟷ is_TNil xs"
by(cases xs) simp_all

lemma tdropn_Suc': "tdropn (Suc n) xs = tdropn n (ttl xs)"
by(simp add: tdropn_Suc split: tllist.split)

context
  includes tllist.lifting lifting_syntax
begin

lemma tllist_set_rel_transfer [transfer_rule]:
  "(rel_set (pcr_tllist op = op =) ===> rel_set (pcr_tllist op = op =) ===> op =)
   (rel_set (λ(xs, a) (ys, b). xs = ys ∧ (lfinite ys ⟶ a = b))) op ="
unfolding tllist.pcr_cr_eq
by(rule Quotient_set_rel_eq)(rule Quotient_tllist)

lemma epred_tlength: "epred (tlength xs) = tlength (ttl xs)"
by(transfer)(simp add: epred_llength) 

lift_definition tcast :: "('a, 'b) tllist ⇒ ('a, 'c) tllist"
is "λ(xs, a). (xs, undefined)" by clarsimp

lemma tcast_TNil: "tcast (TNil x) = TNil undefined"
by transfer simp

lemma tcast_TCons [simp]: "tcast (TCons x xs) = TCons x (tcast xs)"
by transfer clarsimp

lemma terminal_tcast: "terminal (tcast xs) = undefined"
by transfer clarsimp

lemma is_TNil_tcast [simp]: "is_TNil (tcast xs) ⟷ is_TNil xs"
by transfer clarsimp

lemma thd_tcast [simp]: "thd (tcast xs) = thd xs"
by transfer clarsimp

lemma ttl_tcast [simp]: "ttl (tcast xs) = tcast (ttl xs)"
by transfer clarsimp

lemma tset_tcast [simp]: "tset (tcast xs) = tset xs"
by transfer clarsimp

lemma tcast_lappendt [simp]: "tcast (lappendt xs ys) = lappendt xs (tcast ys)"
by transfer auto

lemma tcast_tappend [simp]: "tcast (tappend xs f) = tappend xs (tcast ∘ f)"
by transfer(auto simp add: split_beta)

lemma llist_of_tllist_tcast [simp]: "llist_of_tllist (tcast xs) = llist_of_tllist xs"
by transfer clarsimp

lemma tllist_all2_tcast_iff:
  "tllist_all2 A B (tcast xs) (tcast ys) ⟷ 
  llist_all2 A (llist_of_tllist xs) (llist_of_tllist ys) ∧ (tfinite xs ⟶ tfinite ys ⟶ B undefined undefined)"
by transfer(auto dest: llist_all2_lfiniteD)

lemma tllist_all2_tcastI:
  "⟦ llist_all2 A (llist_of_tllist xs) (llist_of_tllist ys); 
     ⟦ tfinite xs; tfinite ys ⟧ ⟹ B undefined undefined ⟧
  ⟹ tllist_all2 A B (tcast xs) (tcast ys)"
by(simp add: tllist_all2_tcast_iff)



lift_definition timage :: "('a ⇒ ('b, 'c) tllist set) ⇒ ('b, 'a) tllist ⇒ ('b, 'c) tllist set"
 is "λf (xs, a). if lfinite xs then apfst (lappend xs) ` f a else {(xs, undefined)}"
by(fastforce intro!: rel_setI dest: rel_setD1 rel_setD2 intro: rev_bexI)

lemma timage_TCons [simp]:
  "timage f (TCons x xs) = TCons x ` timage f xs"
by transfer(fastforce intro: rel_setI rev_bexI)

lemma timage_TNil [simp]:  "timage f (TNil x) = f x"
by transfer(fastforce intro: rel_setI rev_bexI)

lemma timage_is_TNil: "is_TNil xs ⟹ timage f xs = f (terminal xs)"
by(cases xs) simp_all

lemma timage_tinfite [simp]:
  "¬ tfinite xs ⟹ timage f xs = {tcast xs}"
by transfer(auto intro: rel_setI)

lemma timage_tfinite:
  "tfinite xs ⟹ timage f xs = lappendt (llist_of_tllist xs) ` f (terminal xs)"
by transfer(auto intro: rel_setI elim!: rev_bexI)

lemma timage_cases:
  assumes "ys ∈ timage f xs"
  obtains (tfinite) zs where "tfinite xs" "ys = lappendt (llist_of_tllist xs) zs" "zs ∈ f (terminal xs)"
  | (tinfinite) "¬ tfinite xs" "ys = tcast xs"
using assms by(cases "tfinite xs")(auto simp add: timage_tfinite)

lemma timage_tfiniteE:
  assumes "ys ∈ timage f xs"
  and "tfinite xs"
  obtains zs where "ys = lappendt (llist_of_tllist xs) zs" "zs ∈ f (terminal xs)"
using assms by(auto simp add: timage_tfinite)

lemma timage_tfiniteI:
  "⟦ tfinite xs; zs ∈ f (terminal xs) ⟧ ⟹ lappendt (llist_of_tllist xs) zs ∈ timage f xs"
by(simp add: timage_tfinite)

lemma timage_tinfiniteI:
  "¬ tfinite xs ⟹ tcast xs ∈ timage f xs"
by simp

lemma timage_lappendt [simp]:
  "lfinite xs ⟹ timage f (lappendt xs ys) = lappendt xs ` timage f ys"
by transfer(fastforce intro: rel_setI elim!: rev_bexI simp add: lappend_assoc)

lemma tllist_all2_lappendtI:
  "⟦ llist_all2 A xs xs'; ⟦ lfinite xs; lfinite xs' ⟧ ⟹ tllist_all2 A B ys ys' ⟧
  ⟹ tllist_all2 A B (lappendt xs ys) (lappendt xs' ys')"
by transfer(auto 4 3 dest: llist_all2_lfiniteD intro: llist_all2_lappendI)

lemma timage_parametric [transfer_rule]:
  "((A ===> rel_set (tllist_all2 B C)) ===> tllist_all2 B A ===> rel_set (tllist_all2 B C)) timage timage"
by(rule rel_funI)+(blast intro!: rel_setI dest: tllist_all2_tfiniteD tllist_all2_tfinite1_terminalD tllist_all2_tfinite2_terminalD rel_funD rel_setD1 rel_setD2 intro: rev_bexI timage_tfiniteI tllist_all2_lappendtI llist_all2_tllist_of_llistI tllist_all2_tcastI tllist_all2D_llist_all2_llist_of_tllist timage_tinfiniteI elim!: timage_cases)

lifting_update tllist.lifting
lifting_forget tllist.lifting

end


datatype ('a, 'b) event = Event (event_out: 'a) (event_in: 'b)

datatype_compat event

type_synonym ('a, 'out, 'in) trace = "(('out, 'in) event, 'a option) tllist"

locale IO_execution_base =
  fixes wf_response :: "'out ⇒ 'in set"
  and interp :: "'out ⇒ 'in IO"

definition unit_wf_response :: "unit ⇒ unit set"
where "unit_wf_response _ = {()}"

definition unit_interp :: "unit ⇒ unit IO"
where "unit_interp = return_IO"

end