Theory Consistency
section βΉConsistencyβΊ
theory Consistency
imports
Soundness
begin
definition is_inconsistent_set :: "form set β bool" where
[iff]: "is_inconsistent_set π’ β· π’ β’ Fβoβ"
definition π¬β©0_is_inconsistent :: bool where
[iff]: "π¬β©0_is_inconsistent β· β’ Fβoβ"
definition is_wffo_consistent_with :: "form β form set β bool" where
[iff]: "is_wffo_consistent_with B π’ β· Β¬ is_inconsistent_set (π’ βͺ {B})"
subsection βΉExistence of a standard modelβΊ
text βΉWe construct a standard model in which βΉπ iβΊ is the set βΉ{0}βΊ:βΊ
primrec singleton_standard_domain_family (βΉπβ§SβΊ) where
"πβ§S i = 1"
| "πβ§S o = πΉ"
| "πβ§S (Ξ±βΞ²) = πβ§S Ξ± βΌ πβ§S Ξ²"
interpretation singleton_standard_frame: frame "πβ§S"
proof unfold_locales
{
fix Ξ±
have "πβ§S Ξ± β 0"
proof (induction Ξ±)
case (TFun Ξ² Ξ³)
from βΉπβ§S Ξ³ β 0βΊ obtain y where "y β elts (πβ§S Ξ³)"
by fastforce
then have "(βΞ»z β: πβ§S Ξ²β. y) β elts (πβ§S Ξ² βΌ πβ§S Ξ³)"
by (intro VPi_I)
then show ?case
by force
qed simp_all
}
then show "βΞ±. πβ§S Ξ± β 0"
by (intro allI)
qed simp_all
definition singleton_standard_constant_denotation_function (βΉπ₯β§SβΊ) where
[simp]: "π₯β§S k =
(
if
βΞ². is_Q_constant_of_type k Ξ²
then
let Ξ² = type_of_Q_constant k in qβΞ²ββπβ§Sβ
else
if
is_iota_constant k
then
βΞ»z β: πβ§S (iβo)β. 0
else
case k of (c, Ξ±) β SOME z. z β elts (πβ§S Ξ±)
)"
interpretation singleton_standard_premodel: premodel "πβ§S" "π₯β§S"
proof (unfold_locales)
show "βΞ±. π₯β§S (Q_constant_of_type Ξ±) = qβΞ±ββπβ§Sβ"
by simp
next
show "singleton_standard_frame.is_unique_member_selector (π₯β§S iota_constant)"
unfolding singleton_standard_frame.is_unique_member_selector_def proof
fix x
assume "x β elts (πβ§S i)"
then have "x = 0"
by simp
moreover have "(βΞ»z β: πβ§S (iβo)β. 0) β {0}βiββπβ§Sβ = 0"
using beta[OF singleton_standard_frame.one_element_function_is_domain_respecting]
unfolding singleton_standard_domain_family.simps(3) by blast
ultimately show "(π₯β§S iota_constant) β {x}βiββπβ§Sβ = x"
by fastforce
qed
next
show "βc Ξ±. Β¬ is_logical_constant (c, Ξ±) βΆ π₯β§S (c, Ξ±) β elts (πβ§S Ξ±)"
proof (intro allI impI)
fix c and Ξ±
assume "Β¬ is_logical_constant (c, Ξ±)"
then have "π₯β§S (c, Ξ±) = (SOME z. z β elts (πβ§S Ξ±))"
by auto
moreover have "βz. z β elts (πβ§S Ξ±)"
using eq0_iff and singleton_standard_frame.domain_nonemptiness by presburger
then have "(SOME z. z β elts (πβ§S Ξ±)) β elts (πβ§S Ξ±)"
using some_in_eq by auto
ultimately show "π₯β§S (c, Ξ±) β elts (πβ§S Ξ±)"
by auto
qed
qed
fun singleton_standard_wff_denotation_function (βΉπ±β§SβΊ) where
"π±β§S Ο (xβΞ±β) = Ο (x, Ξ±)"
| "π±β§S Ο (β¦cβ¦βΞ±β) = π₯β§S (c, Ξ±)"
| "π±β§S Ο (A Β· B) = (π±β§S Ο A) β (π±β§S Ο B)"
| "π±β§S Ο (Ξ»xβΞ±β. A) = (βΞ»z β: πβ§S Ξ±β. π±β§S (Ο((x, Ξ±) := z)) A)"
lemma singleton_standard_wff_denotation_function_closure:
assumes "frame.is_assignment πβ§S Ο"
and "A β wffsβΞ±β"
shows "π±β§S Ο A β elts (πβ§S Ξ±)"
using assms(2,1) proof (induction A arbitrary: Ο)
case (var_is_wff Ξ± x)
then show ?case
by simp
next
case (con_is_wff Ξ± c)
then show ?case
proof (cases "(c, Ξ±)" rule: constant_cases)
case non_logical
then show ?thesis
using singleton_standard_premodel.non_logical_constant_denotation
and singleton_standard_wff_denotation_function.simps(2) by presburger
next
case (Q_constant Ξ²)
then have "π±β§S Ο (β¦cβ¦βΞ±β) = qβΞ²ββπβ§Sβ"
by simp
moreover have "qβΞ²ββπβ§Sβ β elts (πβ§S (Ξ²βΞ²βo))"
using singleton_standard_domain_family.simps(3)
and singleton_standard_frame.identity_relation_is_domain_respecting by presburger
ultimately show ?thesis
using Q_constant by simp
next
case ΞΉ_constant
then have "π±β§S Ο (β¦cβ¦βΞ±β) = (βΞ»z β: πβ§S (iβo)β. 0)"
by simp
moreover have "(βΞ»z β: πβ§S (iβo)β. 0) β elts (πβ§S ((iβo)βi))"
by (simp add: VPi_I)
ultimately show ?thesis
using ΞΉ_constant by simp
qed
next
case (app_is_wff Ξ± Ξ² A B)
have "π±β§S Ο (A Β· B) = (π±β§S Ο A) β (π±β§S Ο B)"
using singleton_standard_wff_denotation_function.simps(3) .
moreover have "π±β§S Ο A β elts (πβ§S (Ξ±βΞ²))" and "π±β§S Ο B β elts (πβ§S Ξ±)"
using app_is_wff.IH and app_is_wff.prems by simp_all
ultimately show ?case
by (simp only: singleton_standard_frame.app_is_domain_respecting)
next
case (abs_is_wff Ξ² A Ξ± x)
have "π±β§S Ο (Ξ»xβΞ±β. A) = (βΞ»z β: πβ§S Ξ±β. π±β§S (Ο((x, Ξ±) := z)) A)"
using singleton_standard_wff_denotation_function.simps(4) .
moreover have "π±β§S (Ο((x, Ξ±) := z)) A β elts (πβ§S Ξ²)" if "z β elts (πβ§S Ξ±)" for z
using that and abs_is_wff.IH and abs_is_wff.prems by simp
ultimately show ?case
by (simp add: VPi_I)
qed
interpretation singleton_standard_model: standard_model "πβ§S" "π₯β§S" "π±β§S"
proof (unfold_locales)
show "singleton_standard_premodel.is_wff_denotation_function π±β§S"
by (simp add: singleton_standard_wff_denotation_function_closure)
next
show "βΞ± Ξ². πβ§S (Ξ±βΞ²) = πβ§S Ξ± βΌ πβ§S Ξ²"
using singleton_standard_domain_family.simps(3) by (intro allI)
qed
proposition standard_model_existence:
shows "ββ³. is_standard_model β³"
using singleton_standard_model.standard_model_axioms by auto
subsection βΉTheorem 5403 (Consistency Theorem)βΊ
proposition model_existence_implies_set_consistency:
assumes "is_hyps π’"
and "ββ³. is_general_model β³ β§ is_model_for β³ π’"
shows "Β¬ is_inconsistent_set π’"
proof (rule ccontr)
from assms(2) obtain π and π₯ and π± and β³
where "β³ = (π, π₯, π±)" and "is_model_for β³ π’" and "is_general_model β³" by fastforce
assume "Β¬ Β¬ is_inconsistent_set π’"
then have "π’ β’ Fβoβ"
by simp
with βΉis_general_model β³βΊ have "β³ β¨ Fβoβ"
using thm_5402(2)[OF assms(1) βΉis_model_for β³ π’βΊ] by simp
then have "π± Ο Fβoβ = βT" if "Ο β π" for Ο
using that and βΉβ³ = (π, π₯, π±)βΊ by force
moreover have "π± Ο Fβoβ = βF" if "Ο β π" for Ο
using βΉβ³ = (π, π₯, π±)βΊ and βΉis_general_model β³βΊ and that and general_model.prop_5401_d
by simp
ultimately have "βΟ. Ο β π"
by (auto simp add: inj_eq)
moreover have "βΟ. Ο β π"
proof -
let ?Ο = "Ξ»v. case v of (_, Ξ±) β SOME z. z β elts (π Ξ±)"
from βΉβ³ = (π, π₯, π±)βΊ and βΉis_general_model β³βΊ have "βΞ±. elts (π Ξ±) β {}"
using frame.domain_nonemptiness and premodel_def and general_model.axioms(1) by auto
with βΉβ³ = (π, π₯, π±)βΊ and βΉis_general_model β³βΊ have "?Ο β π"
using frame.is_assignment_def and premodel_def and general_model.axioms(1)
by (metis (mono_tags) case_prod_conv some_in_eq)
then show ?thesis
by (intro exI)
qed
ultimately show False ..
qed
proposition π¬β©0_is_consistent:
shows "Β¬ π¬β©0_is_inconsistent"
proof -
have "ββ³. is_general_model β³ β§ is_model_for β³ {}"
using standard_model_existence and standard_model.axioms(1) by blast
then show ?thesis
using model_existence_implies_set_consistency by simp
qed
lemmas thm_5403 = π¬β©0_is_consistent model_existence_implies_set_consistency
proposition principle_of_explosion:
assumes "is_hyps π’"
shows "is_inconsistent_set π’ β· (βA β (wffsβoβ). π’ β’ A)"
proof
assume "is_inconsistent_set π’"
show "βA β (wffsβoβ). π’ β’ A"
proof
fix A
assume "A β wffsβoβ"
from βΉis_inconsistent_set π’βΊ have "π’ β’ Fβoβ"
unfolding is_inconsistent_set_def .
then have "π’ β’ βπ΅βoβ. π΅βoβ"
unfolding false_is_forall .
with βΉA β wffsβoββΊ have "π’ β’ βS {(π΅, o) β£ A} (π΅βoβ)"
using "βI" by fastforce
then show "π’ β’ A"
by simp
qed
next
assume "βA β (wffsβoβ). π’ β’ A"
then have "π’ β’ Fβoβ"
using false_wff by (elim bspec)
then show "is_inconsistent_set π’"
unfolding is_inconsistent_set_def .
qed
end