Theory Calculi_And_Annotations
theory Calculi_And_Annotations
imports Disjunctive_Consequence_Relations
begin
section ‹Calculi and Derivations›
context
begin
no_notation Extended.extended.Pinf (‹∞›)
typedef 'a infinite_llist = ‹{ l :: 'a llist. llength l = ∞ }›
morphisms to_llist Abs_infinite_llist
using llength_inf_llist
by blast
setup_lifting type_definition_infinite_llist
lift_definition llmap :: ‹('a ⇒ 'b) ⇒ 'a infinite_llist ⇒ 'b infinite_llist› is lmap
by auto
lift_definition llnth :: ‹'a infinite_llist ⇒ nat ⇒ 'a› is lnth .
lift_definition Liminf_infinite_llist :: ‹'a set infinite_llist ⇒ 'a set› is Liminf_llist .
lift_definition Limsup_infinite_llist :: ‹'a set infinite_llist ⇒ 'a set› is Limsup_llist .
lift_definition Sup_infinite_llist :: ‹'a set infinite_llist ⇒ 'a set› is Sup_llist .
lift_definition llhd :: ‹'a infinite_llist ⇒ 'a› is lhd .
lemma llength_of_to_llist_is_infinite: ‹llength (to_llist L) = ∞›
using to_llist
by auto
end
locale sound_inference_system =
inference_system Inf + sound_cons: consequence_relation bot entails_sound
for
Inf :: "'f inference set" and
bot :: "'f" and
entails_sound :: "'f set ⇒ 'f set ⇒ bool" (infix "⊨s" 50)
+ assumes
sound: "ι ∈ Inf ⟹ set (prems_of ι) ⊨s {concl_of ι}"
no_notation IArray.sub (infixl "!!" 100)
definition is_derivation :: "('f set ⇒ 'f set ⇒ bool) ⇒ ('f set infinite_llist) ⇒ bool" where
"is_derivation R Ns ≡ ∀i. R (llnth Ns i) (llnth Ns (Suc i))"
definition terminates :: "'f set infinite_llist ⇒ bool" where
"terminates Ns ≡ ∃i. ∀j>i. llnth Ns j = llnth Ns i"
abbreviation ‹lim_inf ≡ Liminf_infinite_llist›
abbreviation limit :: "'f set infinite_llist ⇒ 'f set" where "limit Ns ≡ lim_inf Ns"
abbreviation lim_sup :: ‹'f set infinite_llist ⇒ 'f set› where
‹lim_sup Ns ≡ Limsup_infinite_llist Ns›
locale calculus = inference_system Inf + consequence_relation bot entails
for
bot :: "'f" and
Inf :: ‹'f inference set› and
entails :: "'f set ⇒ 'f set ⇒ bool" (infix "⊨" 50)
+ fixes
Red⇩I :: "'f set ⇒ 'f inference set" and
Red⇩F :: "'f set ⇒ 'f set"
assumes
Red_I_to_Inf: "Red⇩I N ⊆ Inf" and
Red_F_Bot: "N ⊨ {bot} ⟹ N - Red⇩F N ⊨ {bot}" and
Red_F_of_subset: "N ⊆ N' ⟹ Red⇩F N ⊆ Red⇩F N'" and
Red_I_of_subset: "N ⊆ N' ⟹ Red⇩I N ⊆ Red⇩I N'" and
Red_F_of_Red_F_subset: "N' ⊆ Red⇩F N ⟹ Red⇩F N ⊆ Red⇩F (N - N')" and
Red_I_of_Red_F_subset: "N' ⊆ Red⇩F N ⟹ Red⇩I N ⊆ Red⇩I (N - N')" and
Red_I_of_Inf_to_N: "ι ∈ Inf ⟹ concl_of ι ∈ N ⟹ ι ∈ Red⇩I N"
begin
definition saturated :: "'f set ⇒ bool" where
"saturated N ⟷ Inf_from N ⊆ Red⇩I N"
definition Red⇩I_strict :: "'f set ⇒ 'f inference set" where
"Red⇩I_strict N = {ι. ι ∈ Red⇩I N ∨ (ι ∈ Inf ∧ bot ∈ N)}"
definition Red⇩F_strict :: "'f set ⇒ 'f set" where
"Red⇩F_strict N = {C. C ∈ Red⇩F N ∨ (bot ∈ N ∧ C ≠ bot)}"
lemma strict_calc_if_nobot:
"∀N. bot ∉ Red⇩F N ⟹ calculus bot Inf entails Red⇩I_strict Red⇩F_strict"
proof
fix N
show ‹Red⇩I_strict N ⊆ Inf› unfolding Red⇩I_strict_def using Red_I_to_Inf by blast
next
fix N
assume
bot_notin: "∀N. bot ∉ Red⇩F N" and
entails_bot: ‹N ⊨ {bot}›
show ‹N - Red⇩F_strict N ⊨ {bot}›
proof (cases "bot ∈ N")
assume bot_in: "bot ∈ N"
have ‹bot ∉ Red⇩F N› using bot_notin by blast
then have ‹bot ∉ Red⇩F_strict N› unfolding Red⇩F_strict_def by blast
then have ‹Red⇩F_strict N = UNIV - {bot}›
unfolding Red⇩F_strict_def using bot_in by blast
then have ‹N - Red⇩F_strict N = {bot}› using bot_in by blast
then show ‹N - Red⇩F_strict N ⊨ {bot}› using entails_reflexive[of bot] by simp
next
assume ‹bot ∉ N›
then have ‹Red⇩F_strict N = Red⇩F N› unfolding Red⇩F_strict_def by blast
then show ‹N - Red⇩F_strict N ⊨ {bot}› using Red_F_Bot[OF entails_bot] by simp
qed
next
fix N N' :: "'f set"
assume ‹N ⊆ N'›
then show ‹Red⇩F_strict N ⊆ Red⇩F_strict N'›
unfolding Red⇩F_strict_def using Red_F_of_subset by blast
next
fix N N' :: "'f set"
assume ‹N ⊆ N'›
then show ‹Red⇩I_strict N ⊆ Red⇩I_strict N'›
unfolding Red⇩I_strict_def using Red_I_of_subset by blast
next
fix N' N
assume
bot_notin: "∀N. bot ∉ Red⇩F N" and
subs_red: "N' ⊆ Red⇩F_strict N"
have ‹bot ∉ Red⇩F_strict N›
using bot_notin unfolding Red⇩F_strict_def by blast
then have nbot_in: ‹bot ∉ N'› using subs_red by blast
show ‹Red⇩F_strict N ⊆ Red⇩F_strict (N - N')›
proof (cases "bot ∈ N")
case True
then have bot_in: "bot ∈ N - N'" using nbot_in by blast
then show ?thesis unfolding Red⇩F_strict_def using bot_notin by force
next
case False
then have eq_red: "Red⇩F_strict N = Red⇩F N" unfolding Red⇩F_strict_def by simp
then have "N' ⊆ Red⇩F N" using subs_red by simp
then have "Red⇩F N ⊆ Red⇩F (N - N')" using Red_F_of_Red_F_subset by simp
then show ?thesis using eq_red Red⇩F_strict_def by blast
qed
next
fix N' N
assume
"∀N. bot ∉ Red⇩F N" and
subs_red: "N' ⊆ Red⇩F_strict N"
then have bot_notin: "bot ∉ N'" unfolding Red⇩F_strict_def by blast
then show "Red⇩I_strict N ⊆ Red⇩I_strict (N - N')"
proof (cases "bot ∈ N")
case True
then show ?thesis
unfolding Red⇩I_strict_def using bot_notin Red_I_to_Inf by fastforce
next
case False
then show ?thesis
using bot_notin Red_I_to_Inf subs_red Red_I_of_Red_F_subset
unfolding Red⇩I_strict_def Red⇩F_strict_def by simp
qed
next
fix ι N
assume "ι ∈ Inf"
then show "concl_of ι ∈ N ⟹ ι ∈ Red⇩I_strict N"
unfolding Red⇩I_strict_def using Red_I_of_Inf_to_N Red_I_to_Inf by simp
qed
definition weakly_fair :: "'f set infinite_llist ⇒ bool" where
‹weakly_fair Ns ≡ Inf_from (Liminf_infinite_llist Ns) ⊆ Sup_infinite_llist (llmap Red⇩I Ns)›
abbreviation fair :: "'f set infinite_llist ⇒ bool" where "fair N ≡ weakly_fair N"
definition derive :: "'f set ⇒ 'f set ⇒ bool" (infix "⊳" 50) where
"M ⊳ N ≡ (M - N ⊆ Red⇩F N)"
lemma derive_refl: "M ⊳ M" unfolding derive_def by simp
lemma deriv_red_in: ‹M ⊳ N ⟹ Red⇩F M ⊆ N ∪ Red⇩F N›
proof -
fix M N
assume deriv: ‹M ⊳ N›
then have ‹M ⊆ N ∪ Red⇩F N›
unfolding derive_def by blast
then have red_m_in: ‹Red⇩F M ⊆ Red⇩F (N ∪ Red⇩F N)›
using Red_F_of_subset by blast
have ‹Red⇩F (N ∪ Red⇩F N) ⊆ Red⇩F (N ∪ Red⇩F N - (Red⇩F N - N))›
using Red_F_of_Red_F_subset[of "Red⇩F N - N" "N ∪ Red⇩F N"]
Red_F_of_subset[of "N" "N ∪ Red⇩F N"] by fast
then have ‹Red⇩F (N ∪ Red⇩F N) ⊆ Red⇩F N›
by (metis Diff_subset_conv Red_F_of_subset Un_Diff_cancel lfp.leq_trans subset_refl sup.commute)
then show ‹Red⇩F M ⊆ N ∪ Red⇩F N› using red_m_in by blast
qed
lemma derive_trans: "M ⊳ N ⟹ N ⊳ N' ⟹ M ⊳ N'"
using deriv_red_in by (smt Diff_subset_conv derive_def subset_trans sup.absorb_iff2)
end
locale sound_calculus = sound_inference_system Inf bot entails_sound +
consequence_relation bot entails
for
bot :: "'f" and
Inf :: ‹'f inference set› and
entails :: "'f set ⇒ 'f set ⇒ bool" (infix "⊨" 50) and
entails_sound :: "'f set ⇒ 'f set ⇒ bool" (infix "⊨s" 50)
+ fixes
Red⇩I :: "'f set ⇒ 'f inference set" and
Red⇩F :: "'f set ⇒ 'f set"
assumes
Red_I_to_Inf: "Red⇩I N ⊆ Inf" and
Red_F_Bot: "N ⊨ {bot} ⟹ N - Red⇩F N ⊨ {bot}" and
Red_F_of_subset: "N ⊆ N' ⟹ Red⇩F N ⊆ Red⇩F N'" and
Red_I_of_subset: "N ⊆ N' ⟹ Red⇩I N ⊆ Red⇩I N'" and
Red_F_of_Red_F_subset: "N' ⊆ Red⇩F N ⟹ Red⇩F N ⊆ Red⇩F (N - N')" and
Red_I_of_Red_F_subset: "N' ⊆ Red⇩F N ⟹ Red⇩I N ⊆ Red⇩I (N - N')" and
Red_I_of_Inf_to_N: "ι ∈ Inf ⟹ concl_of ι ∈ N ⟹ ι ∈ Red⇩I N"
begin
sublocale calculus bot Inf entails
by (simp add: calculus.intro calculus_axioms.intro Red_F_Bot
Red_F_of_Red_F_subset Red_F_of_subset Red_I_of_Inf_to_N Red_I_of_Red_F_subset Red_I_of_subset
Red_I_to_Inf consequence_relation_axioms)
end
locale statically_complete_calculus = calculus +
assumes statically_complete: "saturated N ⟹ N ⊨ {bot} ⟹ bot ∈ N"
begin
lemma inf_from_subs: "M ⊆ N ⟹ Inf_from M ⊆ Inf_from N"
unfolding Inf_from_def by blast
lemma nobot_in_Red: ‹bot ∉ Red⇩F N›
proof -
have ‹UNIV ⊨ {bot}›
using entails_reflexive[of bot] entails_subsets[of "{bot}" UNIV "{bot}" "{bot}"] by fast
then have non_red_entails_bot: ‹UNIV - (Red⇩F UNIV) ⊨ {bot}› using Red_F_Bot[of UNIV] by simp
have ‹Inf_from UNIV ⊆ Red⇩I UNIV›
unfolding Inf_from_def using Red_I_of_Inf_to_N[of _ UNIV] by blast
then have sat_non_red: ‹saturated (UNIV - Red⇩F UNIV)›
unfolding saturated_def Inf_from_def using Red_I_of_Red_F_subset[of "Red⇩F UNIV" UNIV] by blast
have ‹bot ∉ Red⇩F UNIV›
using statically_complete[OF sat_non_red non_red_entails_bot] by fast
then show ?thesis using Red_F_of_subset[of _ UNIV] by auto
qed
interpretation strict_calculus:
statically_complete_calculus bot Inf entails Red⇩I_strict Red⇩F_strict
proof -
interpret strict_calc: calculus bot Inf entails Red⇩I_strict Red⇩F_strict
using strict_calc_if_nobot nobot_in_Red by blast
have ‹saturated N ⟹ strict_calc.saturated N›
unfolding saturated_def strict_calc.saturated_def Red⇩I_strict_def by blast
have ‹strict_calc.saturated N ⟹ N ⊨ {bot} ⟹ bot ∈ N› for N
proof -
assume
strict_sat: "strict_calc.saturated N" and
entails_bot: "N ⊨ {bot}"
have ‹bot ∉ N ⟹ Red⇩I_strict N = Red⇩I N› unfolding Red⇩I_strict_def by simp
then have ‹bot ∉ N ⟹ saturated N›
unfolding saturated_def using strict_sat by (simp add: strict_calc.saturated_def)
then have ‹bot ∉ N ⟹ bot ∈ N›
using statically_complete[OF _ entails_bot] by simp
then show ‹bot ∈ N› by auto
qed
then show ‹statically_complete_calculus bot Inf entails Red⇩I_strict Red⇩F_strict›
unfolding statically_complete_calculus_def statically_complete_calculus_axioms_def
using strict_calc.calculus_axioms by blast
qed
end
locale dynamically_complete_calculus = calculus +
assumes dynamically_complete:
‹is_derivation (⊳) Ns ⟹ fair Ns ⟹ llhd Ns ⊨ {bot} ⟹ ∃i. bot ∈ llnth Ns i›
section ‹Annotated Formulas and Consequence Relations›