Theory Calculator

theory Calculator
imports Stream_Processor
theory Calculator imports
  "../Stream_Processor/Stream_Processor"
begin

abbreviation put :: "'b ⇒ ('a, 'b) spν ⇒ ('a, 'b) spν"
where "put x c ≡ In (Put x c)"

type_synonym calc = "nat ⇒ (nat, nat) spν"

corec (friend) restart :: "calc ⇒ calc" where
  "restart f n = (if n > 0 then put n (restart f (n - 1)) else In (out (f 0)))"

lemma restart_simps: "⋀f. restart f n = (if n > 0 then put n (restart f (n - 1)) else f 0)"
by(subst restart.code) simp

lemma restart_cong: "f 0 = g 0 ⟹ restart f n = restart g n"
by(induction n)(simp_all add: restart.code)

corec f :: "calc" where
  "f n = put n (get (λv. if v ≠ 0 then f (2 * v + n) else restart f (v + n)))"

lemma f_sel [simp]: "out (f n) = Put n (get (λv. if v ≠ 0 then f (2 * v + n) else restart f (v + n)))"
by(simp add: f.code)

corec g :: "calc" where
  "g m = put (2 * m) (get (λv. if v = 0 then restart g (2 * m) else g (v + m)))"

lemma g_sel [simp]: "out (g m) = Put (2 * m) (get (λv. if v = 0 then restart g (2 * m) else g (v + m)))"
by(simp add: g.code)

lemma g_sim_f: "g m = f (2 * m)"
proof(coinduction arbitrary: m rule: restart.coinduct)
  case (Eq_nat_nat_spν m)
  have "restart g (m * 2) = restart (λ_. g 0) (m * 2)"
    and "restart f (m * 2) = restart (λ_. f 0) (m * 2)" by(rule restart_cong; simp)+
  then show ?case
    by(auto 4 3 intro!: spν.cong_get spν.cong_restart rel_funI exI simp add: mult.commute add_mult_distrib intro: spν.cong_base)
qed

export_code restart f g in Haskell module_name Calculator file code

end