Theory Resumption

theory Resumption
imports IO_Execution PF_Set
section {* The resumption monad *}

theory Resumption
imports
  IO_Execution
  PF_Set
  Misc
begin

context fixes dflt :: "'b" begin

definition terminal_default :: "('a, 'b) tllist ⇒ 'b"
where "terminal_default xs = (if tfinite xs then terminal xs else dflt)"

lemma terminal_default_TNil [simp, code]: "terminal_default (TNil b) = b"
by(simp add: terminal_default_def)

lemma terminal_default_TCons [simp, code]: "terminal_default (TCons x xs) = terminal_default xs"
by(simp add: terminal_default_def)

lemma terminal_default_ttl [simp]: "terminal_default (ttl xs) = terminal_default xs"
by(auto simp add: terminal_default_def ltl_llist_of_tllist[symmetric] simp del: ltl_llist_of_tllist)

end

context notes [[quick_and_dirty]] begin  (* TODO: remove once codatatype tactics work again as expected *)

codatatype (results: 'a, outputs: 'out, 'in) resumption
  = Done (result: "'a option")
  | Pause ("output": 'out) (resume: "'in ⇒ ('a, 'out, 'in) resumption")
where
  "resume (Done a) = (λinp. Done None)"
  
end

code_datatype Done Pause

primcorec bind_resumption :: 
  "('a, 'out, 'in) resumption
     ⇒ ('a ⇒ ('b, 'out, 'in) resumption) ⇒ ('b, 'out, 'in) resumption"
where
  "⟦ is_Done x; result x ≠ None ⟶ is_Done (f (the (result x))) ⟧ ⟹ is_Done (bind_resumption x f)"
| "result (bind_resumption x f) = result x ⤜ result ∘ f"
| "output (bind_resumption x f) = (if is_Done x then output (f (the (result x))) else output x)"
| "resume (bind_resumption x f) = (λinp. if is_Done x then resume (f (the (result x))) inp else bind_resumption (resume x inp) f)"

declare bind_resumption.sel [simp del]

adhoc_overloading bind bind_resumption

lemma is_Done_bind_resumption [simp]:
  "is_Done (x ⤜ f) ⟷ is_Done x ∧ (result x ≠ None ⟶ is_Done (f (the (result x))))"
by(simp add: bind_resumption_def)

lemma result_bind_resumption [simp]:
  "is_Done (x ⤜ f) ⟹ result (x ⤜ f) = result x ⤜ result ∘ f"
by(simp add: bind_resumption_def)

lemma output_bind_resumption [simp]:
  "¬ is_Done (x ⤜ f) ⟹ output (x ⤜ f) = (if is_Done x then output (f (the (result x))) else output x)"
by(simp add: bind_resumption_def)

lemma resume_bind_resumption [simp]:
  "¬ is_Done (x ⤜ f) ⟹
  resume (x ⤜ f) = 
  (if is_Done x then resume (f (the (result x))) 
   else (λinp. resume x inp ⤜ f))"
by(auto simp add: bind_resumption_def)

definition DONE :: "'a ⇒ ('a, 'out, 'in) resumption"
where "DONE = Done ∘ Some"

definition ABORT :: "('a, 'out, 'in) resumption"
where "ABORT = Done None"

lemma [simp]:
  shows is_Done_DONE: "is_Done (DONE a)"
  and is_Done_ABORT: "is_Done ABORT"
  and result_DONE: "result (DONE a) = Some a"
  and result_ABORT: "result ABORT = None"
  and DONE_inject: "DONE a = DONE b ⟷ a = b"
  and DONE_neq_ABORT: "DONE a ≠ ABORT"
  and ABORT_neq_DONE: "ABORT ≠ DONE a"
  and ABORT_eq_Done: "⋀a. ABORT = Done a ⟷ a = None"
  and Done_eq_ABORT: "⋀a. Done a = ABORT ⟷ a = None"
  and DONE_eq_Done: "⋀b. DONE a = Done b ⟷ b = Some a"
  and Done_eq_DONE: "⋀b. Done b = DONE a ⟷ b = Some a"
  and DONE_neq_Pause: "DONE a ≠ Pause out c"
  and Pause_neq_DONE: "Pause out c ≠ DONE a"
  and ABORT_neq_Pause: "ABORT ≠ Pause out c"
  and Pause_neq_ABORT: "Pause out c ≠ ABORT"
by(auto simp add: DONE_def ABORT_def)

lemma resume_ABORT [simp]:
  "resume (Done r) = (λinp. ABORT)"
by(simp add: ABORT_def)

declare resumption.sel(3)[simp del]

lemma results_DONE [simp]: "results (DONE x) = {x}"
by(simp add: DONE_def)

lemma Done_bind [code]:
  "Done a ⤜ f = (case a of None ⇒ Done None | Some a ⇒ f a)"
by(rule resumption.expand)(auto split: option.split)

lemma DONE_bind [simp]:
  "DONE a ⤜ f = f a"
by(simp add: DONE_def Done_bind)

lemma bind_resumption_Pause [simp, code]: fixes cont shows
  "Pause out cont ⤜ f
  = Pause out (λinp. cont inp ⤜ f)"
by(rule resumption.expand)(simp_all)

lemma bind_DONE [simp]:
  "x ⤜ DONE = x"
by(coinduction arbitrary: x)(auto simp add: split_beta o_def)

lemma bind_bind_resumption:
  fixes r :: "('a, 'in, 'out) resumption" 
  shows "(r ⤜ f) ⤜ g = do { x ← r; f x ⤜ g }"
apply(coinduction arbitrary: r rule: resumption.coinduct_strong)
apply(auto simp add: split_beta bind_eq_Some_conv)
apply(case_tac [!] "result r")
apply simp_all
done

lemmas resumption_monad = DONE_bind bind_DONE bind_bind_resumption

lemma ABORT_bind [simp]: "ABORT ⤜ f = ABORT"
by(simp add: ABORT_def Done_bind)

lemma bind_resumption_is_Done: "is_Done f ⟹ f ⤜ g = (if result f = None then ABORT else g (the (result f)))"
by(rule resumption.expand) auto

lemma bind_resumption_eq_Done_iff [simp]:
  "f ⤜ g = Done x ⟷ (∃y. f = DONE y ∧ g y = Done x) ∨ f = ABORT ∧ x = None"
by(cases f)(auto simp add: Done_bind split: option.split)

lemma bind_resumption_cong:
  assumes "x = y"
  and "⋀z. z ∈ results y ⟹ f z = g z"
  shows "x ⤜ f = y ⤜ g"
using assms(2) unfolding ‹x = y›
proof(coinduction arbitrary: y rule: resumption.coinduct_strong)
  case Eq_resumption thus ?case
    by(auto intro: resumption.set_sel simp add: is_Done_def rel_fun_def)
      (fastforce intro!: exI intro: resumption.set_sel(2) simp add: is_Done_def)
qed

primrec ensure :: "bool ⇒ (unit, 'out, 'in) resumption"
where
  "ensure True = DONE ()" 
| "ensure False = ABORT"

lemma is_Done_map_resumption [simp]:
  "is_Done (map_resumption f1 f2 r) ⟷ is_Done r"
by(cases r) simp_all

lemma result_map_resumption [simp]: 
  "is_Done r ⟹ result (map_resumption f1 f2 r) = map_option f1 (result r)"
by(clarsimp simp add: is_Done_def)

lemma output_map_resumption [simp]:
  "¬ is_Done r ⟹ output (map_resumption f1 f2 r) = f2 (output r)"
by(cases r) simp_all

lemma resume_map_resumption [simp]:
  "¬ is_Done r
  ⟹ resume (map_resumption f1 f2 r) = map_resumption f1 f2 ∘ resume r"
by(cases r) simp_all

(*
lemma "ctor_resumption f = ctor_resumption g ⟹ f = g"
apply(cases f)
apply(cases g)
apply simp
term ctor_resumption
apply(simp add: ctor_resumption_def)
term f

lemma results_induct:
  assumes major: "x ∈ results r"
  and Done: "⋀r. ⟦ is_Done r; ¬ Option.is_none (result r) ⟧ ⟹ P (the (result r)) r"
  and Pause: "⋀x r. ⟦ ¬ is_Done r; ⋀input. P x (resume r input) ⟧ ⟹ P x r"
  shows "P x r"
proof -
  have "∀x∈results r. P x r"
    apply(rule resumption.dtor_set1_induct)
     apply(case_tac z)
      apply(cut_tac r="Done (Some x)" in Done, simp, simp)
      apply(simp add: Done_def result_def resumption.dtor_ctor set1_pre_resumption_def Abs_resumption_pre_resumption_inverse sum_set_simps split: resumption.split_asm)
      apply(auto)[1]
thm Done[where r="Done (Some x)" for x, simplified, unfolded ]
      apply(rule Done[where r="Done (Some x)" for x, simplified])
      thm Abs_resumption_pre_resumption_inverse
    using assms
     apply(auto simp add: is_Done_def resume_def  set1_pre_resumption_def fsts_def snds_def case_tllist_def' collect_def sum_set_simps sum.set_map split: sum.splits)
 using major
apply(simp add: results_def)
*)


lemma rel_resumption_is_DoneD: "rel_resumption A B r1 r2 ⟹ is_Done r1 ⟷ is_Done r2"
by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust]) simp_all

lemma rel_resumption_resultD1:
  "⟦ rel_resumption A B r1 r2; is_Done r1 ⟧ ⟹ rel_option A (result r1) (result r2)"
by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust]) simp_all

lemma rel_resumption_resultD2:
  "⟦ rel_resumption A B r1 r2; is_Done r2 ⟧ ⟹ rel_option A (result r1) (result r2)"
by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust]) simp_all

lemma rel_resumption_outputD1:
  "⟦ rel_resumption A B r1 r2; ¬ is_Done r1 ⟧ ⟹ B (output r1) (output r2)"
by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust]) simp_all

lemma rel_resumption_outputD2:
  "⟦ rel_resumption A B r1 r2; ¬ is_Done r2 ⟧ ⟹ B (output r1) (output r2)"
by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust]) simp_all

lemma rel_resumption_resumeD1:
  "⟦ rel_resumption A B r1 r2; ¬ is_Done r1 ⟧
  ⟹ rel_resumption A B (resume r1 inp) (resume r2 inp)"
by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust])(auto dest: rel_funD)

lemma rel_resumption_resumeD2:
  "⟦ rel_resumption A B r1 r2; ¬ is_Done r2 ⟧
  ⟹ rel_resumption A B (resume r1 inp) (resume r2 inp)"
by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust])(auto dest: rel_funD)

lemma rel_resumption_coinduct
  [consumes 1, case_names Done Pause,
   case_conclusion Done is_Done result,
   case_conclusion Pause "output" resume,
   coinduct pred: rel_resumption]:
  assumes X: "X r1 r2"
  and Done: "⋀r1 r2. X r1 r2 ⟹ (is_Done r1 ⟷ is_Done r2) ∧ (is_Done r1 ⟶ is_Done r2 ⟶ rel_option A (result r1) (result r2))"
  and Pause: "⋀r1 r2. ⟦ X r1 r2; ¬ is_Done r1; ¬ is_Done r2 ⟧ ⟹ B (output r1) (output r2) ∧ (∀inp. X (resume r1 inp) (resume r2 inp))" 
  shows "rel_resumption A B r1 r2"
using X
apply(rule resumption.rel_coinduct)
apply(unfold rel_fun_def)
apply(rule conjI)
 apply(erule Done[THEN conjunct1])
apply(rule conjI)
 apply(erule Done[THEN conjunct2])
apply(rule impI)+
apply(drule (2) Pause)
apply blast
done

subsection {* setup for @{text "partial_function"} *}

context includes lifting_syntax begin

coinductive resumption_ord :: "('a, 'out, 'in) resumption ⇒ ('a, 'out, 'in) resumption ⇒ bool"
where
  Done_Done: "flat_ord None a a' ⟹ resumption_ord (Done a) (Done a')"
| Done_Pause: "resumption_ord ABORT (Pause out c)"
| Pause_Pause: "(op = ===> resumption_ord) c c' ⟹ resumption_ord (Pause out c) (Pause out c')"

inductive_simps resumption_ord_simps [simp]:
  "resumption_ord (Pause out c) r"
  "resumption_ord r (Done a)"

lemma resumption_ord_is_DoneD:
  "⟦ resumption_ord r r'; is_Done r' ⟧ ⟹ is_Done r"
by(cases r')(auto simp add: fun_ord_def)

lemma resumption_ord_resultD:
  "⟦ resumption_ord r r'; is_Done r' ⟧ ⟹ flat_ord None (result r) (result r')"
by(cases r')(auto simp add: flat_ord_def)

lemma resumption_ord_outputD:
  "⟦ resumption_ord r r'; ¬ is_Done r ⟧ ⟹ output r = output r'"
by(cases r) auto

lemma resumption_ord_resumeD:
  "⟦ resumption_ord r r'; ¬ is_Done r ⟧ ⟹ (op = ===> resumption_ord) (resume r) (resume r')"
by(cases r) auto

lemma resumption_ord_abort:
  "⟦ resumption_ord r r'; is_Done r; ¬ is_Done r' ⟧ ⟹ result r = None"
by(auto elim: resumption_ord.cases)

lemma resumption_ord_coinduct [consumes 1, case_names Done Abort Pause, case_conclusion Pause "output" resume, coinduct pred: resumption_ord]:
  assumes "X r r'"
  and Done: "⋀r r'. ⟦ X r r'; is_Done r' ⟧ ⟹ is_Done r ∧ flat_ord None (result r) (result r')"
  and Abort: "⋀r r'. ⟦ X r r'; ¬ is_Done r'; is_Done r ⟧ ⟹ result r = None"
  and Pause: "⋀r r'. ⟦ X r r'; ¬ is_Done r; ¬ is_Done r' ⟧ 
  ⟹ output r = output r' ∧ (op = ===> (λr r'. X r r' ∨ resumption_ord r r')) (resume r) (resume r')"
  shows "resumption_ord r r'"
using `X r r'`
proof coinduct
  case (resumption_ord r r')
  thus ?case
    by(cases r r' rule: resumption.exhaust[case_product resumption.exhaust])(auto dest: Done Pause Abort)
qed

end

lemma resumption_ort_ABORT [intro!, simp]: "resumption_ord ABORT r"
by(cases r)(simp_all add: flat_ord_def resumption_ord.Done_Pause)

lemma resumption_ord_refl: "resumption_ord r r"
by(coinduction arbitrary: r)(auto simp add: flat_ord_def)

lemma resumption_ord_antisym:
  "⟦ resumption_ord r r'; resumption_ord r' r ⟧
  ⟹ r = r'"
proof(coinduction arbitrary: r r' rule: resumption.coinduct_strong)
  case (Eq_resumption r r')
  thus ?case
    by cases(auto simp add: flat_ord_def rel_fun_def)
qed

lemma resumption_ord_trans:
  "⟦ resumption_ord r r'; resumption_ord r' r'' ⟧
  ⟹ resumption_ord r r''"
proof(coinduction arbitrary: r r' r'')
  case (Done r r' r'')
  thus ?case by(auto 4 4 elim: resumption_ord.cases simp add: flat_ord_def)
next
  case (Abort r r' r'')
  thus ?case by(auto 4 4 elim: resumption_ord.cases simp add: flat_ord_def)
next
  case (Pause r r' r'')
  hence "resumption_ord r r'" "resumption_ord r' r''" by simp_all
  thus ?case using `¬ is_Done r` `¬ is_Done r''`
    by(cases)(auto simp add: rel_fun_def)
qed

primcorec resumption_lub :: "('a, 'out, 'in) resumption set ⇒ ('a, 'out, 'in) resumption"
where
  "∀r ∈ R. is_Done r ⟹ is_Done (resumption_lub R)"
| "result (resumption_lub R) = flat_lub None (result ` R)"
| "output (resumption_lub R) = (THE out. out ∈ output ` (R ∩ {r. ¬ is_Done r}))"
| "resume (resumption_lub R) = (λinp. resumption_lub ((λc. c inp) ` resume ` (R ∩ {r. ¬ is_Done r})))"

lemma is_Done_resumption_lub [simp]:
  "is_Done (resumption_lub R) ⟷ (∀r ∈ R. is_Done r)"
by(simp add: resumption_lub_def)

lemma result_resumption_lub [simp]:
  "∀r ∈ R. is_Done r ⟹ result (resumption_lub R) = flat_lub None (result ` R)"
by(simp add: resumption_lub_def)

lemma output_resumption_lub [simp]:
  "∃r∈R. ¬ is_Done r ⟹ output (resumption_lub R) = (THE out. out ∈ output ` (R ∩ {r. ¬ is_Done r}))"
by(simp add: resumption_lub_def)

lemma resume_resumption_lub [simp]:
  "∃r∈R. ¬ is_Done r
  ⟹ resume (resumption_lub R) inp = 
     resumption_lub ((λc. c inp) ` resume ` (R ∩ {r. ¬ is_Done r}))"
by(simp add: resumption_lub_def)

context
  fixes R state inp R'
  defines R'_def: "R' ≡ (λc. c inp) ` resume ` (R ∩ {r. ¬ is_Done r})"
  assumes chain: "Complete_Partial_Order.chain resumption_ord R"
begin

lemma resumption_ord_chain_resume: "Complete_Partial_Order.chain resumption_ord R'"
proof(rule chainI)
  fix r' r''
  assume "r' ∈ R'"
    and "r'' ∈ R'"
  then obtain 𝗋' 𝗋'' 
    where r': "r' = resume 𝗋' inp" "𝗋' ∈ R" "¬ is_Done 𝗋'"
    and r'': "r'' = resume 𝗋'' inp" "𝗋'' ∈ R" "¬ is_Done 𝗋''"
    by(auto simp add: R'_def)
  from chain `𝗋' ∈ R` `𝗋'' ∈ R`
  have "resumption_ord 𝗋' 𝗋'' ∨ resumption_ord 𝗋'' 𝗋'"
    by(auto elim: chainE)
  with r' r''
  have "resumption_ord (resume 𝗋' inp) (resume 𝗋'' inp) ∨
        resumption_ord (resume 𝗋'' inp) (resume 𝗋' inp)"
    by(auto elim: resumption_ord.cases simp add: rel_fun_def)
  with r' r''
  show "resumption_ord r' r'' ∨ resumption_ord r'' r'" by auto
qed

end

lemma resumption_partial_function_definition:
  "partial_function_definitions resumption_ord resumption_lub"
proof
  fix r
  show "resumption_ord r r" by(rule resumption_ord_refl)
next
  fix r r' r''
  assume "resumption_ord r r'" "resumption_ord r' r''"
  thus "resumption_ord r r''" by(rule resumption_ord_trans)
next
  fix r r'
  assume "resumption_ord r r'" "resumption_ord r' r"
  thus "r = r'" by(rule resumption_ord_antisym)
next
  fix R and r :: "('a, 'b, 'c) resumption"
  assume "Complete_Partial_Order.chain resumption_ord R" "r ∈ R"
  thus "resumption_ord r (resumption_lub R)"
  proof(coinduction arbitrary: r R)
    case (Done r R)
    note chain = `Complete_Partial_Order.chain resumption_ord R`
      and r = `r ∈ R`
    from `is_Done (resumption_lub R)` have A: "∀r ∈ R. is_Done r" by simp
    with r obtain a' where "r = Done a'" by(cases r) auto
    { fix r'
      assume "a' ≠ None"
      hence "(THE x. x ∈ result ` R ∧ x ≠ None) = a'"
        using r A `r = Done a'`
        by(auto 4 3 intro!: the_equality intro: rev_image_eqI elim: chainE[OF chain] simp add: flat_ord_def is_Done_def) 
    }
    with A r `r = Done a'` show ?case
      by(cases a')(auto simp add: flat_ord_def flat_lub_def)
  next
    case (Abort r R)
    hence chain: "Complete_Partial_Order.chain resumption_ord R" and "r ∈ R" by simp_all
    from `r ∈ R` `¬ is_Done (resumption_lub R)` `is_Done r`
    show ?case by(auto elim: chainE[OF chain] dest: resumption_ord_abort resumption_ord_is_DoneD)
  next
    case (Pause r R)
    hence chain: "Complete_Partial_Order.chain resumption_ord R"
      and r: "r ∈ R" by simp_all
    have ?resume 
      using r `¬ is_Done r` resumption_ord_chain_resume[OF chain]
      by(auto simp add: rel_fun_def bexI)
    moreover
    from r `¬ is_Done r` have "output (resumption_lub R) = output r"
      by(auto 4 4 simp add: bexI intro!: the_equality elim: chainE[OF chain] dest: resumption_ord_outputD)
    ultimately show ?case by simp
  qed
next
  fix R and r :: "('a, 'b, 'c) resumption"
  assume "Complete_Partial_Order.chain resumption_ord R" "⋀r'. r' ∈ R ⟹ resumption_ord r' r"
  thus "resumption_ord (resumption_lub R) r"
  proof(coinduction arbitrary: R r)
    case (Done R r)
    hence chain: "Complete_Partial_Order.chain resumption_ord R"
      and ub: "∀r'∈R. resumption_ord r' r" by simp_all
    from `is_Done r` ub have is_Done: "∀r' ∈ R. is_Done r'"
      and ub': "⋀r'. r' ∈ result ` R ⟹ flat_ord None r' (result r)"
      by(auto dest: resumption_ord_is_DoneD resumption_ord_resultD)
    from is_Done have chain': "Complete_Partial_Order.chain (flat_ord None) (result ` R)"
      by(auto 5 2 intro!: chainI elim: chainE[OF chain] dest: resumption_ord_resultD)
    hence "flat_ord None (flat_lub None (result ` R)) (result r)"
      by(rule partial_function_definitions.lub_least[OF flat_interpretation])(rule ub')
    thus ?case using is_Done by simp
  next
    case (Abort R r)
    hence chain: "Complete_Partial_Order.chain resumption_ord R"
      and ub: "∀r'∈R. resumption_ord r' r" by simp_all
    from `¬ is_Done r` `is_Done (resumption_lub R)` ub
    show ?case by(auto simp add: flat_lub_def dest: resumption_ord_abort)
  next
    case (Pause R r)
    hence chain: "Complete_Partial_Order.chain resumption_ord R"
      and ub: "⋀r'. r'∈R ⟹ resumption_ord r' r" by simp_all
    from `¬ is_Done (resumption_lub R)` have exR: "∃r ∈ R. ¬ is_Done r" by simp
    then obtain r' where r': "r' ∈ R" "¬ is_Done r'" by auto
    with ub[of r'] have "output r = output r'" by(auto dest: resumption_ord_outputD)
    also have [symmetric]: "output (resumption_lub R) = output r'" using exR r'
      by(auto 4 4 elim: chainE[OF chain] dest: resumption_ord_outputD)
    finally have ?output ..
    moreover 
    { fix inp r''
      assume "r'' ∈ R" "¬ is_Done r''"
      with ub[of r'']
      have "resumption_ord (resume r'' inp) (resume r inp)"
        by(auto dest!: resumption_ord_resumeD simp add: rel_fun_def) }
    with exR resumption_ord_chain_resume[OF chain] r'
    have ?resume by(auto simp add: rel_fun_def)
    ultimately show ?case ..
  qed
qed

interpretation resumption:
  partial_function_definitions resumption_ord resumption_lub
by (rule resumption_partial_function_definition)

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

abbreviation "mono_resumption ≡ monotone (fun_ord resumption_ord) resumption_ord"

lemma mono_resumption_resume:
  assumes "mono_resumption B"
  shows "mono_resumption (λf. resume (B f) inp)"
proof
  fix f g :: "'a ⇒ ('b, 'c, 'd) resumption"
  assume fg: "fun_ord resumption_ord f g"
  hence "resumption_ord (B f) (B g)" by(rule monotoneD[OF assms])
  with resumption_ord_resumeD[OF this]
  show "resumption_ord (resume (B f) inp) (resume (B g) inp)"
    by(cases "is_Done (B f)")(auto simp add: rel_fun_def is_Done_def)
qed

lemma bind_resumption_mono [partial_function_mono]:
  assumes mf: "mono_resumption B"
  and mg: "⋀y. mono_resumption (C y)"
  shows "mono_resumption (λf. do { y ← B f; C y f })"
proof(rule monotoneI)
  fix f g :: "'a ⇒ ('b, 'c, 'd) resumption"
  assume *: "fun_ord resumption_ord f g"
  def f'  "B f" and g'  "B g"
    and h  "λx. C x f" and k  "λx. C x g"
  with mf[THEN monotoneD, OF *] mg[THEN monotoneD, OF *]
  have "resumption_ord f' g'" "⋀x. resumption_ord (h x) (k x)" by auto
  thus "resumption_ord (f' ⤜ h) (g' ⤜ k)"
  proof(coinduction arbitrary: f' g' h k)
    case (Done f' g' h k)
    hence le: "resumption_ord f' g'"
      and mg: "⋀y. resumption_ord (h y) (k y)" by simp_all
    from `is_Done (g' ⤜ k)`
    have done_Bg: "is_Done g'" 
      and "result g' ≠ None ⟹ is_Done (k (the (result g')))" by simp_all
    moreover
    have "is_Done f'" using le done_Bg by(rule resumption_ord_is_DoneD)
    moreover
    from le done_Bg have "flat_ord None (result f') (result g')"
      by(rule resumption_ord_resultD)
    hence "result f' ≠ None ⟹ result g' = result f'"
      by(auto simp add: flat_ord_def)
    moreover
    have "resumption_ord (h (the (result f'))) (k (the (result f')))" by(rule mg)
    ultimately show ?case
      by(subst (1 2) result_bind_resumption)(auto dest: resumption_ord_is_DoneD resumption_ord_resultD simp add: flat_ord_def bind_eq_None_conv)
  next
    case (Abort f' g' h k)
    hence "resumption_ord (h (the (result f'))) (k (the (result f')))" by simp
    thus ?case using Abort
      by(cases "is_Done g'")(auto 4 4 simp add: bind_eq_None_conv flat_ord_def dest: resumption_ord_abort resumption_ord_resultD resumption_ord_is_DoneD)
  next
    case (Pause f' g' h k)
    hence ?output
      by(auto 4 4 dest: resumption_ord_outputD resumption_ord_is_DoneD resumption_ord_resultD resumption_ord_abort simp add: flat_ord_def)
    moreover have ?resume
    proof(cases "is_Done f'")
      case False
      with Pause show ?thesis
        by(auto simp add: rel_fun_def dest: resumption_ord_is_DoneD intro: resumption_ord_resumeD[THEN rel_funD] intro!: exI)
    next
      case True
      hence "is_Done g'" using Pause by(auto dest: resumption_ord_abort)
      thus ?thesis using True Pause resumption_ord_resultD[OF `resumption_ord f' g'`]
        by(auto intro!: rel_funI simp add: bind_resumption_is_Done flat_ord_def intro: resumption_ord_resumeD[THEN rel_funD] exI[where x=f'] exI[where x=g'])
    qed
    ultimately show ?case ..
  qed
qed

subsection {* Execution resumption *}

context IO_execution_base begin

context
  fixes interp_raw :: "'out ⇒ real_world ⇒ ('in × real_world) option"
  defines interp_raw_def: "interp_raw ≡ map_fun id Rep_IO interp"
begin

partial_function (option) interp_resumption_raw :: "('a, 'out, 'in) resumption ⇒ real_world ⇒ ('a × real_world) option"
where
  "interp_resumption_raw r state = 
   (if is_Done r then case result r of None ⇒ None | Some x ⇒ Some (x, state)
    else interp_raw (output r) state ⤜ (λ(input, state). interp_resumption_raw (resume r input) state))"

context includes IO.lifting lifting_syntax begin

lift_definition interp_resumption :: "('a, 'out, 'in) resumption ⇒ 'a IO" is interp_resumption_raw .

lemma interp_resumption_Done [simp]:
  "interp_resumption (Done x) = (case x of None ⇒ fail_IO | Some x' ⇒ return_IO x')"
including undefined_transfer
by transfer (simp add: interp_resumption_raw.simps fun_eq_iff split: option.split)

lemma interp_resumption_DONE [simp]:
  "interp_resumption (DONE x) = return_IO x"
by(simp add: DONE_def)

lemma interp_resumption_ABORT:
  "interp_resumption ABORT = fail_IO"
by(simp add: ABORT_def)

lemma interp_resumption_Pause [simp]:
  "interp_resumption (Pause out c) = interp out ⤜ interp_resumption ∘ c"
proof -
  have [transfer_rule]: "(op = ===> pcr_IO op =) interp_raw interp"
    unfolding interp_raw_def IO.pcr_cr_eq by(auto simp add: cr_IO_def)
  show ?thesis
    by transfer(simp add: interp_resumption_raw.simps fun_eq_iff o_def cong: prod.case_cong)
qed

end

end

end

subsection {* Code generator setup *}
text {*
  Implement resumption concurrency as IO monad in Haskell and as side effects in SML/OCaml/Scala.
*}

-- {* use type variable @{typ 'inp} instead of the usual @{typ 'in} because the code generator forgets to rename type variables for @{text code_abort} constants and OCaml does not like the type variable @{typ 'inp} *}
definition resumption_ABORT :: "unit ⇒ ('a, 'out, 'inp) resumption"
  where [simp, code del]: "resumption_ABORT _ = ABORT"
lemma ABORT_code [code, code_unfold]: "ABORT = resumption_ABORT ()" by simp
declare [[code abort: resumption_ABORT]]

subsection {* Setup for lifting and transfer *}

declare resumption.rel_eq [id_simps, relator_eq]
declare resumption.rel_mono [relator_mono]

lemma rel_resumption_OO [relator_distr]:
  "rel_resumption A B OO rel_resumption C D = rel_resumption (A OO C) (B OO D)" 
  (is "?lhs = ?rhs")
proof(intro ext iffI)
  fix r1 r2
  assume "?lhs r1 r2"
  thus "?rhs r1 r2"
    by(coinduct r1 r2 rule: rel_resumption_coinduct)(auto dest: rel_resumption_is_DoneD rel_resumption_resultD1 rel_resumption_outputD1 rel_resumption_outputD2 rel_resumption_resumeD1 rel_resumption_resumeD2 simp add: option.rel_compp)
next
  def r12 
    "corec_resumption (λ(r1, r2). is_Done r1) 
       (λ(r1, r2). do { a1 ← result r1; a2 ← result r2; Some (SOME ac. A a1 ac ∧ C ac a2) })
       (λ(r1, r2). SOME bd. B (output r1) bd ∧ D bd (output r2))
       (λ(r1, r2). λinp. Inr (resume r1 inp, resume r2 inp))
    :: ('a, 'b, 'c) resumption × ('d, 'e, 'c) resumption ⇒ ('f, 'g, 'c) resumption"

  fix r1 r2
  assume "?rhs r1 r2"
  hence "rel_resumption A B r1 (r12 (r1, r2))"
  proof(coinduction arbitrary: r1 r2 rule: rel_resumption_coinduct)
    case (Done r1 r2)
    have ?is_Done by(simp add: r12_def)
    moreover have ?result using Done
      by(cases "result r1")(fastforce simp add: r12_def option.rel_compp option_rel_Some1 dest!: rel_resumption_resultD1 intro: someI2)+
    ultimately show ?case ..
  next
    case (Pause r1 r2)
    hence ?output by(auto dest!: rel_resumption_outputD1 simp add: r12_def intro: someI2)
    moreover have ?resume (is "All ?C")
    proof
      fix inp
      from rel_resumption_resumeD1[OF `rel_resumption (A OO C) (B OO D) r1 r2`, of inp] `¬ is_Done r1`
      show "?C inp" by(auto simp add: r12_def rel_fun_def)
    qed
    ultimately show ?case ..
  qed
  moreover
  from `?rhs r1 r2`
  have "rel_resumption C D (r12 (r1, r2)) r2"
  proof(coinduction arbitrary: r1 r2 rule: rel_resumption_coinduct)
    case (Done r1 r2)
    hence ?is_Done by(auto simp add: r12_def dest: rel_resumption_is_DoneD)
    moreover have ?result using Done
      by(cases "result r2")(fastforce simp add: r12_def option.rel_compp option_rel_Some2 dest!: rel_resumption_resultD2 intro: someI2)+
    ultimately show ?case ..
  next
    case (Pause r1 r2)
    hence ?output by(fastforce dest!: rel_resumption_outputD2 simp add: r12_def intro: someI2)
    moreover have ?resume (is "All ?C")
    proof
      fix inp
      from Pause have "¬ is_Done r1" by(simp add: r12_def)
      with rel_resumption_resumeD1[OF `rel_resumption (A OO C) (B OO D) r1 r2`, of inp]
      show "?C inp" by(auto simp add: r12_def rel_fun_def)
    qed
    ultimately show ?case ..
  qed
  ultimately show "?lhs r1 r2" ..
qed

lemma left_total_rel_resumption [transfer_rule]:
  "⟦ left_total R1; left_total R2 ⟧ ⟹ left_total (rel_resumption R1 R2)"
  unfolding left_total_def
proof
  fix r :: "('a, 'c, 'e) resumption"
  assume R1: "∀x. ∃y. R1 x y" and R2: "∀x. ∃y. R2 x y"
  have "rel_resumption R1 R2 r (map_resumption (λx. SOME y. R1 x y) (λx. SOME y. R2 x y) r)"
  proof(coinduction arbitrary: r rule: rel_resumption_coinduct)
    case (Done r)
    have ?is_Done by simp
    moreover have ?result using R1
      by(cases "result r")(auto intro: someI)
    ultimately show ?case ..
  next
    case (Pause r)
    thus ?case using R2 by(auto intro: someI)
  qed
  thus "∃r'. rel_resumption R1 R2 r r'" ..
qed

lemma left_unique_rel_resumption [transfer_rule]:
  assumes R1: "left_unique R1"
  and R2: "left_unique R2"
  shows "left_unique (rel_resumption R1 R2)"
  unfolding left_unique_def
proof(intro strip)
  fix r1 r2 r3
  assume "rel_resumption R1 R2 r1 r3" "rel_resumption R1 R2 r2 r3"
  thus "r1 = r2"
  proof(coinduction arbitrary: r1 r2 r3)
    case (Eq_resumption r1 r2 r3)
    hence ?is_Done by(auto dest: rel_resumption_is_DoneD)
    moreover have ?Done using Eq_resumption
      option.left_unique_rel[OF R1, unfolded left_unique_def]
      by(auto dest!: rel_resumption_resultD1)
    moreover have ?Pause using Eq_resumption R2[unfolded left_unique_def]
      by(auto 4 3 dest: rel_resumption_outputD1 rel_resumption_resumeD1 simp add: rel_fun_def)
    ultimately show ?case by simp
  qed
qed

lemma right_total_rel_resumption [transfer_rule]:
  "⟦ right_total R1; right_total R2 ⟧ ⟹ right_total (rel_resumption R1 R2)"
  unfolding right_total_def
proof
  fix r :: "('b, 'd, 'e) resumption"
  assume R1: "∀y. ∃x. R1 x y" and R2: "∀y. ∃x. R2 x y"
  have "rel_resumption R1 R2 (map_resumption (λy. SOME x. R1 x y) (λy. SOME x. R2 x y) r) r"
  proof(coinduction arbitrary: r rule: rel_resumption_coinduct)
    case (Done r)
    have ?is_Done by simp
    moreover have ?result using R1
      by(cases "result r")(auto intro: someI)
    ultimately show ?case ..
  next
    case (Pause r)
    thus ?case using R2
      by(auto intro: someI)
  qed
  thus "∃r'. rel_resumption R1 R2 r' r" ..
qed

lemma right_unique_rel_resumption [transfer_rule]:
  assumes R1: "right_unique R1"
  and R2: "right_unique R2"
  shows "right_unique (rel_resumption R1 R2)"
  unfolding right_unique_def
proof(intro strip)
  fix r1 r2 r3
  assume "rel_resumption R1 R2 r1 r2" "rel_resumption R1 R2 r1 r3"
  thus "r2 = r3"
  proof(coinduction arbitrary: r1 r2 r3)
    case (Eq_resumption r1 r2 r3)
    hence ?is_Done by(auto dest: rel_resumption_is_DoneD)
    moreover have ?Done using Eq_resumption option.right_unique_rel[OF R1, unfolded right_unique_def]
      by(auto dest!: rel_resumption_resultD2)
    moreover have ?Pause using Eq_resumption R2[unfolded right_unique_def]
      by(auto 4 3 dest: rel_resumption_outputD2 rel_resumption_resumeD2 simp add: rel_fun_def)
    ultimately show ?case by simp
  qed
qed

lemma bi_total_rel_resumption [transfer_rule]:
  "⟦ bi_total A; bi_total B ⟧ ⟹ bi_total (rel_resumption A B)"
unfolding bi_total_alt_def
by(blast intro: left_total_rel_resumption right_total_rel_resumption)

lemma bi_unique_rel_resumption [transfer_rule]:
  "⟦ bi_unique A; bi_unique B ⟧ ⟹ bi_unique (rel_resumption A B)"
unfolding bi_unique_alt_def
by(blast intro: left_unique_rel_resumption right_unique_rel_resumption)

lemma Quotient_resumption [quot_map]:
  assumes Q1: "Quotient R1 Abs1 Rep1 T1"
  and Q2: "Quotient R2 Abs2 Rep2 T2"
  shows "Quotient (rel_resumption R1 R2) (map_resumption Abs1 Abs2) (map_resumption Rep1 Rep2) (rel_resumption T1 T2)"
unfolding Quotient_alt_def
proof(intro conjI allI impI)
  have 1: "⋀x y. rel_option T1 x y ⟹ map_option Abs1 x = y" 
    using option.Quotient[OF Q1] by(simp add: Quotient_alt_def)
  have 2: "⋀x y. T2 x y ⟹ Abs2 x = y"
    using Q2 by(simp add: Quotient_alt_def)

  fix r1 r2
  assume "rel_resumption T1 T2 r1 r2"
  thus "map_resumption Abs1 Abs2 r1 = r2"
    by(coinduction arbitrary: r1 r2)(auto 4 3 dest: rel_resumption_is_DoneD rel_resumption_resultD1 1 rel_resumption_outputD1 2 rel_resumption_resumeD1)
next
  from option.Quotient[OF Q1] have 1: "⋀y. rel_option T1 (map_option Rep1 y) y"
    by(simp add: Quotient_alt_def)
  from Q2 have 2: "⋀x. T2 (Rep2 x) x" by(simp add: Quotient_alt_def)

  fix r
  show "rel_resumption T1 T2 (map_resumption Rep1 Rep2 r) r"
    by(coinduction arbitrary: r rule: rel_resumption_coinduct)(auto simp add: 1 2)
next
  have 1: "⋀x y. rel_option R1 x y ⟷ rel_option T1 x (map_option Abs1 x) ∧ rel_option T1 y (map_option Abs1 y) ∧ map_option Abs1 x = map_option Abs1 y"
    using option.Quotient[OF Q1] by(simp add: Quotient_alt_def)
  have 2: "⋀x y. R2 x y ⟷ T2 x (Abs2 x) ∧ T2 y (Abs2 y) ∧ Abs2 x = Abs2 y"
    using Q2 by(simp add: Quotient_alt_def)

  fix r1 r2
  show "rel_resumption R1 R2 r1 r2 ⟷
        rel_resumption T1 T2 r1 (map_resumption Abs1 Abs2 r1) ∧
        rel_resumption T1 T2 r2 (map_resumption Abs1 Abs2 r2) ∧
        map_resumption Abs1 Abs2 r1 = map_resumption Abs1 Abs2 r2" (is "?lhs ⟷ ?rhs")
  proof
    assume ?rhs
    thus ?lhs
    proof(coinduction arbitrary: r1 r2 rule: rel_resumption_coinduct)
      case (Done r1 r2)
      hence ?is_Done by(metis rel_resumption_is_DoneD)
      moreover have ?result using Done
        by(metis rel_resumption_resultD1 result_map_resumption 1)
      ultimately show ?case ..
    next
      case (Pause r1 r2)
      hence ?output by(metis rel_resumption_outputD1 output_map_resumption 2)
      moreover
      from `¬ is_Done r1` obtain out1 c1 where r1: "r1 = Pause out1 c1" by(cases r1) auto
      from `¬ is_Done r2` obtain out2 c2 where r2: "r2 = Pause out2 c2" by(cases r2) auto
      have ?resume (is "All ?C")
      proof
        fix inp
        from Pause r1 r2 
        have c1: "rel_resumption T1 T2 (c1 inp) (map_resumption Abs1 Abs2 (c1 inp))"
          and c2: "rel_resumption T1 T2 (c2 inp) (map_resumption Abs1 Abs2 (c2 inp))"
          and eq: "map_resumption Abs1 Abs2 ∘ c1 = map_resumption Abs1 Abs2 ∘ c2"
          by(simp_all add: rel_fun_def)
        thus "?C inp" by(simp add: fun_eq_iff r1 r2)
      qed
      ultimately show ?case ..
    qed
  next
    assume ?lhs
    hence "rel_resumption T1 T2 r1 (map_resumption Abs1 Abs2 r1)"
    proof(coinduction arbitrary: r1 r2 rule: rel_resumption_coinduct)
      case Done
      thus ?case by(auto dest!: rel_resumption_resultD1 simp add: 1)
    next
      case (Pause r1 r2)
      hence ?output by(auto 4 3 dest: rel_resumption_outputD1 simp add: 2)
      moreover have ?resume using Pause
        by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust])(auto 4 3 simp add: rel_fun_def)
      ultimately show ?case ..
    qed
    moreover from `?lhs` have "rel_resumption T1 T2 r2 (map_resumption Abs1 Abs2 r2)"
    proof(coinduction arbitrary: r1 r2 rule: rel_resumption_coinduct)
      case Done
      thus ?case by(auto dest!: rel_resumption_resultD2 simp add: 1)
    next
      case (Pause r1 r2)
      hence ?output by(auto dest: rel_resumption_outputD2 simp add: 2)
      moreover have ?resume using Pause
        by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust])(auto 4 3 simp add: rel_fun_def)
      ultimately show ?case ..
    qed
    moreover from `?lhs` have "map_resumption Abs1 Abs2 r1 = map_resumption Abs1 Abs2 r2"
    proof(coinduction arbitrary: r1 r2)
      case (Eq_resumption r1 r2)
      hence ?is_Done ?Done by(auto dest: rel_resumption_is_DoneD rel_resumption_resultD1 simp add: 1)
      moreover
      have ?Pause using Eq_resumption
        by(cases r1 r2 rule: resumption.exhaust[case_product resumption.exhaust])(auto dest: rel_resumption_outputD1 simp add: 2 rel_fun_def)
      ultimately show ?case by simp
    qed
    ultimately show ?rhs by blast
  qed
qed

end