header {* Incremental coinduction *}
theory Incremental
imports
Main
begin
locale Incremental =
fixes Retr :: "('a :: complete_lattice) => 'a"
and Cl :: "('a :: complete_lattice) => 'a"
assumes mono_Retr: "mono Retr"
and mono_Cl: "mono Cl"
and idem_Cl: "Cl (Cl r) = Cl r"
and ext_Cl: "r ≤ Cl r"
and Cl_Retr: "Cl r = r ==> Cl (inf r (Retr r)) = inf r (Retr r)"
begin
lemma leq_Cl:
assumes "Cl s = s" and "r ≤ s"
shows "Cl r ≤ s"
by (metis assms monoE mono_Cl)
lemma Cl_respectful_Retr:
assumes rs: "r ≤ s" and rrs: "r ≤ Retr s"
shows "Cl r ≤ Retr (Cl s)"
proof-
let ?ss = "Cl s"
have "Cl ?ss = ?ss" using idem_Cl .
from Cl_Retr[OF this] have 1: "Cl (inf ?ss (Retr ?ss)) = inf ?ss (Retr ?ss)" .
have "r ≤ ?ss" using rs ext_Cl by (metis order_trans)
moreover have "r ≤ Retr ?ss" using rrs ext_Cl mono_Retr by (metis mono_def order_trans)
ultimately have "r ≤ inf ?ss (Retr ?ss)" by simp
from leq_Cl[OF 1 this] show ?thesis by simp
qed
lemma
assumes respectful: "!! r s. [|r ≤ s; r ≤ Retr s|] ==> Cl r ≤ Retr (Cl s)"
and r: "Cl r = r"
shows "Cl (inf r (Retr r)) = inf r (Retr r)"
proof-
have 1: "Cl (inf r (Retr r)) ≤ r" by (metis inf_le1 monoE mono_Cl r)
have "Cl (inf r (Retr r)) ≤ Retr (Cl r)" by (rule respectful) auto
hence "Cl (inf r (Retr r)) ≤ Retr r" unfolding r .
with 1 show ?thesis by (metis antisym ext_Cl inf.bounded_iff)
qed
lemma Cl_sound_Retr:
assumes "r ≤ Retr (Cl r)"
shows "r ≤ gfp Retr"
proof-
have "Cl r ≤ Retr (Cl (Cl r))" apply(rule Cl_respectful_Retr)
using ext_Cl assms by auto
hence "Cl r ≤ Retr (Cl r)" unfolding idem_Cl .
hence "Cl r ≤ gfp Retr" by (metis gfp_upperbound)
thus ?thesis using ext_Cl by (metis order.trans)
qed
lemma mono_Retr_Cl: "mono (Retr o Cl)"
using mono_Retr mono_Cl unfolding mono_def by (metis comp_apply)
lemma gfp_Retr_Cl: "gfp (Retr o Cl) = gfp Retr"
proof(rule order_class.antisym)
show "gfp (Retr o Cl) ≤ gfp Retr"
apply(rule Cl_sound_Retr) using gfp_lemma2[OF mono_Retr_Cl] by simp
next
show "gfp Retr ≤ gfp (Retr o Cl)"
apply(rule gfp_mono) using ext_Cl mono_Retr unfolding mono_def by auto
qed
definition G :: "'a => 'a" where "G r ≡ gfp (λ s. Retr (Cl (sup r s)))"
lemma mono_pre_G: "mono (λ s. Retr (Cl (sup r s)))"
using mono_Retr mono_Cl unfolding mono_def by (metis dual_order.refl sup_mono)
lemma G_unfold: "G r = Retr (Cl (sup r (G r)))"
using gfp_unfold[OF mono_pre_G] unfolding G_def .
lemma G_mono[simp, intro]: "r ≤ s ==> G r ≤ G s"
unfolding G_def apply(rule gfp_mono) by (metis monoE mono_Cl mono_Retr order_refl sup.mono)
lemma mono_G: "mono G"
using G_mono unfolding mono_def by auto
lemma G_initialize: "gfp Retr = G bot"
unfolding G_def gfp_Retr_Cl[symmetric] o_def by auto
lemma G_accumulate: "r ≤ G s <-> r ≤ G (sup s r)"
proof
assume "r ≤ G s" thus "r ≤ G (sup s r)"
by (metis G_mono le_sup_iff sup.orderI sup_absorb2 sup_idem)
next
assume r: "r ≤ G (sup s r)"
have 0: "sup s (G (sup s r)) = sup (sup s r) (G (sup s r))"
using r by (metis sup_absorb2 sup_commute sup_left_commute)
have "G (sup s r) ≤ G s"
unfolding G_def apply(rule gfp_upperbound) unfolding G_def[symmetric] 0
unfolding G_unfold[symmetric] by simp
thus "r ≤ G s" using r by (metis order.trans)
qed
lemma G_compose:
assumes 11: "s1 ≤ G r1" and 22: "s2 ≤ G r2" and 12: "r1 ≤ sup r s2" and 21: "r2 ≤ sup r s1"
shows "sup s1 s2 ≤ G r"
proof-
let ?K = "G (sup r (sup s1 s2))"
have "s1 ≤ G (sup r s2)" by (metis "11" "12" G_mono order.trans)
also have "... ≤ ?K" by (metis G_mono le_sup_iff sup.orderI sup_idem)
finally have s1: "s1 ≤ ?K" .
have "s2 ≤ G (sup r s1)" by (metis "22" "21" G_mono order.trans)
also have "... ≤ ?K" by (metis G_mono le_sup_iff sup.orderI sup_idem)
finally have s2: "s2 ≤ ?K" .
from s1 s2 have "sup s1 s2 ≤ ?K" by (metis sup_least)
thus ?thesis unfolding G_accumulate[symmetric] .
qed
definition ded (infix "\<turnstile>" 40) where "r \<turnstile> s ≡ s ≤ Cl (sup r (G r))"
lemma gfp_le_G: "gfp Retr ≤ G r"
by (metis G_initialize bot.extremum mono_G mono_def)
theorem Ax:
assumes "s ≤ Cl (sup r (gfp Retr))"
shows "r \<turnstile> s"
using assms gfp_le_G unfolding ded_def
by (metis (hide_lams, mono_tags) eq_iff monoD mono_Cl order.trans sup.mono)
theorem Split:
assumes "!! s. s ∈ S ==> r \<turnstile> s"
shows "r \<turnstile> Sup S"
using assms unfolding ded_def by (metis Sup_least)
theorem Coind:
assumes "s ≤ Retr p" and "sup r s \<turnstile> p"
shows "r \<turnstile> s"
proof-
let ?rs = "sup r s"
have "s ≤ Retr (Cl (sup ?rs (G ?rs)))" using assms unfolding ded_def
by (metis (full_types) monoD mono_Retr order_trans)
hence "s ≤ G r" apply(subst G_accumulate) unfolding G_unfold[symmetric] .
thus ?thesis unfolding ded_def
by (metis ext_Cl order.trans sup.bounded_iff)
qed
theorem soundness_ded:
assumes "gfp Retr \<turnstile> s"
shows "s ≤ gfp Retr"
using assms unfolding ded_def
by (metis Cl_respectful_Retr Cl_sound_Retr G_unfold dual_order.antisym dual_order.refl
ext_Cl gfp_le_G le_supI2 sup_least)
theorem Ax': assumes "s ≤ Cl r" shows "r \<turnstile> s"
apply(rule Ax) using assms by (metis (full_types) le_sup_iff monoD mono_Cl sup_absorb2 sup_ge1)
end
end