diff --git a/thys/Epistemic_Logic/Epistemic_Logic.thy b/thys/Epistemic_Logic/Epistemic_Logic.thy --- a/thys/Epistemic_Logic/Epistemic_Logic.thy +++ b/thys/Epistemic_Logic/Epistemic_Logic.thy @@ -1,1605 +1,1411 @@ (* File: Epistemic_Logic.thy Author: Asta Halkjær From This work is a formalization of epistemic logic with countably many agents. It includes proofs of soundness and completeness for the axiom system K. The completeness proof is based on the textbook "Reasoning About Knowledge" by Fagin, Halpern, Moses and Vardi (MIT Press 1995). The extensions of system K (T, KB, K4, S4, S5) and their completeness proofs are based on the textbook "Modal Logic" by Blackburn, de Rijke and Venema (Cambridge University Press 2001). *) theory Epistemic_Logic imports "HOL-Library.Countable" begin section \Syntax\ type_synonym id = string datatype 'i fm = FF ("\<^bold>\") | Pro id | Dis \'i fm\ \'i fm\ (infixr "\<^bold>\" 30) | Con \'i fm\ \'i fm\ (infixr "\<^bold>\" 35) | Imp \'i fm\ \'i fm\ (infixr "\<^bold>\" 25) | K 'i \'i fm\ abbreviation TT ("\<^bold>\") where \TT \ \<^bold>\ \<^bold>\ \<^bold>\\ abbreviation Neg ("\<^bold>\ _" [40] 40) where \Neg p \ p \<^bold>\ \<^bold>\\ abbreviation \L i p \ \<^bold>\ K i (\<^bold>\ p)\ section \Semantics\ -datatype ('i, 'w) kripke = Kripke (\: \'w \ id \ bool\) (\: \'i \ 'w \ 'w set\) +datatype ('i, 'w) kripke = + Kripke (\: \'w set\) (\: \'w \ id \ bool\) (\: \'i \ 'w \ 'w set\) primrec semantics :: \('i, 'w) kripke \ 'w \ 'i fm \ bool\ ("_, _ \ _" [50, 50] 50) where \(M, w \ \<^bold>\) = False\ | \(M, w \ Pro x) = \ M w x\ | \(M, w \ (p \<^bold>\ q)) = ((M, w \ p) \ (M, w \ q))\ | \(M, w \ (p \<^bold>\ q)) = ((M, w \ p) \ (M, w \ q))\ | \(M, w \ (p \<^bold>\ q)) = ((M, w \ p) \ (M, w \ q))\ -| \(M, w \ K i p) = (\v \ \ M i w. M, v \ p)\ +| \(M, w \ K i p) = (\v \ \ M \ \ M i w. M, v \ p)\ section \S5 Axioms\ definition reflexive :: \('i, 'w) kripke \ bool\ where - \reflexive M \ \i w. w \ \ M i w\ + \reflexive M \ \i. \w \ \ M. w \ \ M i w\ definition symmetric :: \('i, 'w) kripke \ bool\ where - \symmetric M \ \i v w. v \ \ M i w \ w \ \ M i v\ + \symmetric M \ \i. \v \ \ M. \w \ \ M. v \ \ M i w \ w \ \ M i v\ definition transitive :: \('i, 'w) kripke \ bool\ where - \transitive M \ \i v w u. w \ \ M i v \ u \ \ M i w \ u \ \ M i v\ + \transitive M \ \i. \u \ \ M. \v \ \ M. \w \ \ M. + w \ \ M i v \ u \ \ M i w \ u \ \ M i v\ abbreviation equivalence :: \('i, 'w) kripke \ bool\ where \equivalence M \ reflexive M \ symmetric M \ transitive M\ lemma Imp_intro [intro]: \(M, w \ p \ M, w \ q) \ M, w \ (p \<^bold>\ q)\ by simp theorem distribution: \M, w \ (K i p \<^bold>\ K i (p \<^bold>\ q) \<^bold>\ K i q)\ proof assume \M, w \ (K i p \<^bold>\ K i (p \<^bold>\ q))\ then have \M, w \ K i p\ \M, w \ K i (p \<^bold>\ q)\ by simp_all - then have \\v \ \ M i w. M, v \ p\ \\v \ \ M i w. M, v \ (p \<^bold>\ q)\ + then have \\v \ \ M \ \ M i w. M, v \ p\ \\v \ \ M \ \ M i w. M, v \ (p \<^bold>\ q)\ by simp_all - then have \\v \ \ M i w. M, v \ q\ + then have \\v \ \ M \ \ M i w. M, v \ q\ by simp then show \M, w \ K i q\ by simp qed theorem generalization: assumes valid: \\(M :: ('i, 'w) kripke) w. M, w \ p\ shows \(M :: ('i, 'w) kripke), w \ K i p\ proof - have \\w' \ \ M i w. M, w' \ p\ using valid by blast then show \M, w \ K i p\ by simp qed theorem truth: - assumes \reflexive M\ + assumes \reflexive M\ \w \ \ M\ shows \M, w \ (K i p \<^bold>\ p)\ proof assume \M, w \ K i p\ - then have \\v \ \ M i w. M, v \ p\ + then have \\v \ \ M \ \ M i w. M, v \ p\ by simp moreover have \w \ \ M i w\ - using \reflexive M\ unfolding reflexive_def by blast + using \reflexive M\ \w \ \ M\ unfolding reflexive_def by blast ultimately show \M, w \ p\ - by simp + using \w \ \ M\ by simp qed theorem pos_introspection: - assumes \transitive M\ + assumes \transitive M\ \w \ \ M\ shows \M, w \ (K i p \<^bold>\ K i (K i p))\ proof assume \M, w \ K i p\ - then have \\v \ \ M i w. M, v \ p\ + then have \\v \ \ M \ \ M i w. M, v \ p\ by simp - then have \\v \ \ M i w. \u \ \ M i v. M, u \ p\ - using \transitive M\ unfolding transitive_def by blast - then have \\v \ \ M i w. M, v \ K i p\ + then have \\v \ \ M \ \ M i w. \u \ \ M \ \ M i v. M, u \ p\ + using \transitive M\ \w \ \ M\ unfolding transitive_def by blast + then have \\v \ \ M \ \ M i w. M, v \ K i p\ by simp then show \M, w \ K i (K i p)\ by simp qed theorem neg_introspection: - assumes \symmetric M\ \transitive M\ + assumes \symmetric M\ \transitive M\ \w \ \ M\ shows \M, w \ (\<^bold>\ K i p \<^bold>\ K i (\<^bold>\ K i p))\ proof assume \M, w \ \<^bold>\ (K i p)\ - then obtain u where \u \ \ M i w\ \\ (M, u \ p)\ + then obtain u where \u \ \ M i w\ \\ (M, u \ p)\ \u \ \ M\ by auto - moreover have \\v \ \ M i w. u \ \ M i v\ - using \u \ \ M i w\ \symmetric M\ \transitive M\ + moreover have \\v \ \ M \ \ M i w. u \ \ M \ \ M i v\ + using \u \ \ M i w\ \symmetric M\ \transitive M\ \u \ \ M\ \w \ \ M\ unfolding symmetric_def transitive_def by blast - ultimately have \\v \ \ M i w. M, v \ \<^bold>\ K i p\ + ultimately have \\v \ \ M \ \ M i w. M, v \ \<^bold>\ K i p\ by auto then show \M, w \ K i (\<^bold>\ K i p)\ by simp qed -lemma ex_equiv_model: \\M. reflexive M \ symmetric M \ transitive M\ -proof (intro exI conjI) - let ?M = \(Kripke undefined (\_ w. {w}))\ - show \reflexive ?M\ - unfolding reflexive_def by simp - show \symmetric ?M\ - unfolding symmetric_def by auto - show \transitive ?M\ - unfolding transitive_def by simp -qed - section \Normal Modal Logic\ primrec eval :: \(id \ bool) \ ('i fm \ bool) \ 'i fm \ bool\ where \eval _ _ \<^bold>\ = False\ | \eval g _ (Pro x) = g x\ | \eval g h (p \<^bold>\ q) = (eval g h p \ eval g h q)\ | \eval g h (p \<^bold>\ q) = (eval g h p \ eval g h q)\ | \eval g h (p \<^bold>\ q) = (eval g h p \ eval g h q)\ | \eval _ h (K i p) = h (K i p)\ abbreviation \tautology p \ \g h. eval g h p\ inductive AK :: \('i fm \ bool) \ 'i fm \ bool\ ("_ \ _" [50, 50] 50) for A :: \'i fm \ bool\ where A1: \tautology p \ A \ p\ | A2: \A \ (K i p \<^bold>\ K i (p \<^bold>\ q) \<^bold>\ K i q)\ | Ax: \A p \ A \ p\ | R1: \A \ p \ A \ (p \<^bold>\ q) \ A \ q\ | R2: \A \ p \ A \ K i p\ section \Soundness\ -lemma eval_semantics: \eval (pi w) (\q. Kripke pi r, w \ q) p = (Kripke pi r, w \ p)\ +lemma eval_semantics: \eval (pi w) (\q. Kripke W pi r, w \ q) p = (Kripke W pi r, w \ p)\ by (induct p) simp_all lemma tautology: assumes \tautology p\ shows \M, w \ p\ proof - - from assms have \eval (g w) (\q. Kripke g r, w \ q) p\ for g r + from assms have \eval (g w) (\q. Kripke W g r, w \ q) p\ for W g r by simp - then have \Kripke g r, w \ p\ for g r + then have \Kripke W g r, w \ p\ for W g r using eval_semantics by fast then show \M, w \ p\ by (metis kripke.collapse) qed theorem soundness: fixes M :: \('i, 'w) kripke\ - assumes \\(M :: ('i, 'w) kripke) w p. A p \ P M \ M, w \ p\ - shows \A \ p \ P M \ M, w \ p\ - by (induct p arbitrary: w rule: AK.induct) (simp_all add: assms tautology) + assumes \\(M :: ('i, 'w) kripke) w p. A p \ P M \ w \ \ M \ M, w \ p\ + shows \A \ p \ P M \ w \ \ M \ M, w \ p\ + by (induct p arbitrary: w rule: AK.induct) (auto simp: assms tautology) section \Derived rules\ lemma K_A2': \A \ (K i (p \<^bold>\ q) \<^bold>\ K i p \<^bold>\ K i q)\ proof - have \A \ (K i p \<^bold>\ K i (p \<^bold>\ q) \<^bold>\ K i q)\ using A2 by fast moreover have \A \ ((P \<^bold>\ Q \<^bold>\ R) \<^bold>\ (Q \<^bold>\ P \<^bold>\ R))\ for P Q R by (simp add: A1) ultimately show ?thesis using R1 by fast qed lemma K_map: assumes \A \ (p \<^bold>\ q)\ shows \A \ (K i p \<^bold>\ K i q)\ proof - note \A \ (p \<^bold>\ q)\ then have \A \ K i (p \<^bold>\ q)\ using R2 by fast moreover have \A \ (K i (p \<^bold>\ q) \<^bold>\ K i p \<^bold>\ K i q)\ using K_A2' by fast ultimately show ?thesis using R1 by fast qed lemma K_LK: \A \ (L i (\<^bold>\ p) \<^bold>\ \<^bold>\ K i p)\ proof - have \A \ (p \<^bold>\ \<^bold>\ \<^bold>\ p)\ by (simp add: A1) moreover have \A \ ((P \<^bold>\ Q) \<^bold>\ (\<^bold>\ Q \<^bold>\ \<^bold>\ P))\ for P Q using A1 by force ultimately show ?thesis using K_map R1 by fast qed primrec imply :: \'i fm list \ 'i fm \ 'i fm\ where \imply [] q = q\ | \imply (p # ps) q = (p \<^bold>\ imply ps q)\ lemma K_imply_head: \A \ imply (p # ps) p\ proof - have \tautology (imply (p # ps) p)\ by (induct ps) simp_all then show ?thesis using A1 by blast qed lemma K_imply_Cons: assumes \A \ imply ps q\ shows \A \ imply (p # ps) q\ proof - have \A \ (imply ps q \<^bold>\ imply (p # ps) q)\ by (simp add: A1) with R1 assms show ?thesis . qed lemma K_right_mp: assumes \A \ imply ps p\ \A \ imply ps (p \<^bold>\ q)\ shows \A \ imply ps q\ proof - have \tautology (imply ps p \<^bold>\ imply ps (p \<^bold>\ q) \<^bold>\ imply ps q)\ by (induct ps) simp_all with A1 have \A \ (imply ps p \<^bold>\ imply ps (p \<^bold>\ q) \<^bold>\ imply ps q)\ . then show ?thesis using assms R1 by blast qed lemma tautology_imply_superset: assumes \set ps \ set qs\ shows \tautology (imply ps r \<^bold>\ imply qs r)\ proof (rule ccontr) assume \\ tautology (imply ps r \<^bold>\ imply qs r)\ then obtain g h where \\ eval g h (imply ps r \<^bold>\ imply qs r)\ by blast then have \eval g h (imply ps r)\ \\ eval g h (imply qs r)\ by simp_all then consider (np) \\p \ set ps. \ eval g h p\ | (r) \\p \ set ps. eval g h p\ \eval g h r\ by (induct ps) auto then show False proof cases case np then have \\p \ set qs. \ eval g h p\ using \set ps \ set qs\ by blast then have \eval g h (imply qs r)\ by (induct qs) simp_all then show ?thesis using \\ eval g h (imply qs r)\ by blast next case r then have \eval g h (imply qs r)\ by (induct qs) simp_all then show ?thesis using \\ eval g h (imply qs r)\ by blast qed qed lemma K_imply_weaken: assumes \A \ imply ps q\ \set ps \ set ps'\ shows \A \ imply ps' q\ proof - have \tautology (imply ps q \<^bold>\ imply ps' q)\ using \set ps \ set ps'\ tautology_imply_superset by blast then have \A \ (imply ps q \<^bold>\ imply ps' q)\ using A1 by blast then show ?thesis using \A \ imply ps q\ R1 by blast qed lemma imply_append: \imply (ps @ ps') q = imply ps (imply ps' q)\ by (induct ps) simp_all lemma K_ImpI: assumes \A \ imply (p # G) q\ shows \A \ imply G (p \<^bold>\ q)\ proof - have \set (p # G) \ set (G @ [p])\ by simp then have \A \ imply (G @ [p]) q\ using assms K_imply_weaken by blast then have \A \ imply G (imply [p] q)\ using imply_append by metis then show ?thesis by simp qed lemma K_Boole: assumes \A \ imply ((\<^bold>\ p) # G) \<^bold>\\ shows \A \ imply G p\ proof - have \A \ imply G (\<^bold>\ \<^bold>\ p)\ using assms K_ImpI by blast moreover have \tautology (imply G (\<^bold>\ \<^bold>\ p) \<^bold>\ imply G p)\ by (induct G) simp_all then have \A \ (imply G (\<^bold>\ \<^bold>\ p) \<^bold>\ imply G p)\ using A1 by blast ultimately show ?thesis using R1 by blast qed lemma K_DisE: assumes \A \ imply (p # G) r\ \A \ imply (q # G) r\ \A \ imply G (p \<^bold>\ q)\ shows \A \ imply G r\ proof - have \tautology (imply (p # G) r \<^bold>\ imply (q # G) r \<^bold>\ imply G (p \<^bold>\ q) \<^bold>\ imply G r)\ by (induct G) auto then have \A \ (imply (p # G) r \<^bold>\ imply (q # G) r \<^bold>\ imply G (p \<^bold>\ q) \<^bold>\ imply G r)\ using A1 by blast then show ?thesis using assms R1 by blast qed lemma K_mp: \A \ imply (p # (p \<^bold>\ q) # G) q\ by (meson K_imply_head K_imply_weaken K_right_mp set_subset_Cons) lemma K_swap: assumes \A \ imply (p # q # G) r\ shows \A \ imply (q # p # G) r\ using assms K_ImpI by (metis imply.simps(1-2)) lemma K_DisL: assumes \A \ imply (p # ps) q\ \A \ imply (p' # ps) q\ shows \A \ imply ((p \<^bold>\ p') # ps) q\ proof - have \A \ imply (p # (p \<^bold>\ p') # ps) q\ \A \ imply (p' # (p \<^bold>\ p') # ps) q\ using assms K_swap K_imply_Cons by blast+ moreover have \A \ imply ((p \<^bold>\ p') # ps) (p \<^bold>\ p')\ using K_imply_head by blast ultimately show ?thesis using K_DisE by blast qed lemma K_distrib_K_imp: assumes \A \ K i (imply G q)\ shows \A \ imply (map (K i) G) (K i q)\ proof - have \A \ (K i (imply G q) \<^bold>\ imply (map (K i) G) (K i q))\ proof (induct G) case Nil then show ?case by (simp add: A1) next case (Cons a G) have \A \ (K i a \<^bold>\ K i (imply (a # G) q) \<^bold>\ K i (imply G q))\ by (simp add: A2) moreover have \A \ ((K i a \<^bold>\ K i (imply (a # G) q) \<^bold>\ K i (imply G q)) \<^bold>\ (K i (imply G q) \<^bold>\ imply (map (K i) G) (K i q)) \<^bold>\ (K i a \<^bold>\ K i (imply (a # G) q) \<^bold>\ imply (map (K i) G) (K i q)))\ by (simp add: A1) ultimately have \A \ (K i a \<^bold>\ K i (imply (a # G) q) \<^bold>\ imply (map (K i) G) (K i q))\ using Cons R1 by blast moreover have \A \ ((K i a \<^bold>\ K i (imply (a # G) q) \<^bold>\ imply (map (K i) G) (K i q)) \<^bold>\ (K i (imply (a # G) q) \<^bold>\ K i a \<^bold>\ imply (map (K i) G) (K i q)))\ by (simp add: A1) ultimately have \A \ (K i (imply (a # G) q) \<^bold>\ K i a \<^bold>\ imply (map (K i) G) (K i q))\ using R1 by blast then show ?case by simp qed then show ?thesis using assms R1 by blast qed section \Completeness\ subsection \Consistent sets\ definition consistent :: \('i fm \ bool) \ 'i fm set \ bool\ where \consistent A S \ \S'. set S' \ S \ A \ imply S' \<^bold>\\ lemma inconsistent_subset: assumes \consistent A V\ \\ consistent A ({p} \ V)\ obtains V' where \set V' \ V\ \A \ imply (p # V') \<^bold>\\ proof - obtain V' where V': \set V' \ ({p} \ V)\ \p \ set V'\ \A \ imply V' \<^bold>\\ using assms unfolding consistent_def by blast then have *: \A \ imply (p # V') \<^bold>\\ using K_imply_Cons by blast let ?S = \removeAll p V'\ have \set (p # V') \ set (p # ?S)\ by auto then have \A \ imply (p # ?S) \<^bold>\\ using * K_imply_weaken by blast moreover have \set ?S \ V\ using V'(1) by (metis Diff_subset_conv set_removeAll) ultimately show ?thesis using that by blast qed lemma consistent_deriv: assumes \consistent A V\ \A \ p\ shows \consistent A ({p} \ V)\ using assms by (metis R1 consistent_def imply.simps(2) inconsistent_subset) lemma consistent_consequent: assumes \consistent A V\ \p \ V\ \A \ (p \<^bold>\ q)\ shows \consistent A ({q} \ V)\ proof - have \\V'. set V' \ V \ \ A \ imply (p # V') \<^bold>\\ using \consistent A V\ \p \ V\ unfolding consistent_def by (metis insert_subset list.simps(15)) then have \\V'. set V' \ V \ \ A \ imply (q # V') \<^bold>\\ using \A \ (p \<^bold>\ q)\ K_imply_head K_right_mp by (metis imply.simps(1-2)) then show ?thesis using \consistent A V\ inconsistent_subset by metis qed lemma consistent_consequent': assumes \consistent A V\ \p \ V\ \tautology (p \<^bold>\ q)\ shows \consistent A ({q} \ V)\ using assms consistent_consequent A1 by blast lemma consistent_disjuncts: assumes \consistent A V\ \(p \<^bold>\ q) \ V\ shows \consistent A ({p} \ V) \ consistent A ({q} \ V)\ proof (rule ccontr) assume \\ ?thesis\ then have \\ consistent A ({p} \ V)\ \\ consistent A ({q} \ V)\ by blast+ then obtain S' T' where S': \set S' \ V\ \A \ imply (p # S') \<^bold>\\ and T': \set T' \ V\ \A \ imply (q # T') \<^bold>\\ using \consistent A V\ inconsistent_subset by metis from S' have p: \A \ imply (p # S' @ T') \<^bold>\\ by (metis K_imply_weaken Un_upper1 append_Cons set_append) moreover from T' have q: \A \ imply (q # S' @ T') \<^bold>\\ by (metis K_imply_head K_right_mp R1 imply.simps(2) imply_append) ultimately have \A \ imply ((p \<^bold>\ q) # S' @ T') \<^bold>\\ using K_DisL by blast then have \A \ imply (S' @ T') \<^bold>\\ using S'(1) T'(1) p q \consistent A V\ \(p \<^bold>\ q) \ V\ unfolding consistent_def by (metis Un_subset_iff insert_subset list.simps(15) set_append) moreover have \set (S' @ T') \ V\ by (simp add: S'(1) T'(1)) ultimately show False using \consistent A V\ unfolding consistent_def by blast qed lemma exists_finite_inconsistent: assumes \\ consistent A ({\<^bold>\ p} \ V)\ obtains W where \{\<^bold>\ p} \ W \ {\<^bold>\ p} \ V\ \(\<^bold>\ p) \ W\ \finite W\ \\ consistent A ({\<^bold>\ p} \ W)\ proof - obtain W' where W': \set W' \ {\<^bold>\ p} \ V\ \A \ imply W' \<^bold>\\ using assms unfolding consistent_def by blast let ?S = \removeAll (\<^bold>\ p) W'\ have \\ consistent A ({\<^bold>\ p} \ set ?S)\ unfolding consistent_def using W'(2) by auto moreover have \finite (set ?S)\ by blast moreover have \{\<^bold>\ p} \ set ?S \ {\<^bold>\ p} \ V\ using W'(1) by auto moreover have \(\<^bold>\ p) \ set ?S\ by simp ultimately show ?thesis by (meson that) qed lemma inconsistent_imply: assumes \\ consistent A ({\<^bold>\ p} \ set G)\ shows \A \ imply G p\ using assms K_Boole K_imply_weaken unfolding consistent_def by (metis insert_is_Un list.simps(15)) -lemma consistent_axioms: - assumes \\(M :: ('i :: countable, 'w) kripke) w p. A p \ P M \ M, w \ p\ \\M. P M\ - shows \consistent A {}\ -proof (rule ccontr) - assume \\ ?thesis\ - then have \A \ \<^bold>\\ - unfolding consistent_def by simp - then have \M, w \ \<^bold>\\ for M :: \('i, 'w) kripke\ and w - using soundness assms by (metis semantics.simps(1)) - then show False - by simp -qed - subsection \Maximal consistent sets\ definition maximal :: \('i fm \ bool) \ 'i fm set \ bool\ where \maximal A S \ \p. p \ S \ \ consistent A ({p} \ S)\ theorem deriv_in_maximal: assumes \consistent A V\ \maximal A V\ \A \ p\ shows \p \ V\ using assms R1 inconsistent_subset unfolding consistent_def maximal_def by (metis imply.simps(2)) theorem exactly_one_in_maximal: assumes \consistent A V\ \maximal A V\ shows \p \ V \ (\<^bold>\ p) \ V\ proof assume \p \ V\ then show \(\<^bold>\ p) \ V\ using assms K_mp unfolding consistent_def maximal_def by (metis empty_subsetI insert_subset list.set(1) list.simps(15)) next assume \(\<^bold>\ p) \ V\ have \A \ (p \<^bold>\ \<^bold>\ p)\ by (simp add: A1) then have \(p \<^bold>\ \<^bold>\ p) \ V\ using assms deriv_in_maximal by blast then have \consistent A ({p} \ V) \ consistent A ({\<^bold>\ p} \ V)\ using assms consistent_disjuncts by blast then show \p \ V\ using \maximal A V\ \(\<^bold>\ p) \ V\ unfolding maximal_def by blast qed theorem consequent_in_maximal: assumes \consistent A V\ \maximal A V\ \p \ V\ \(p \<^bold>\ q) \ V\ shows \q \ V\ proof - have \\V'. set V' \ V \ \ A \ imply (p # (p \<^bold>\ q) # V') \<^bold>\\ using \consistent A V\ \p \ V\ \(p \<^bold>\ q) \ V\ unfolding consistent_def by (metis insert_subset list.simps(15)) then have \\V'. set V' \ V \ \ A \ imply (q # V') \<^bold>\\ by (meson K_mp K_ImpI K_imply_weaken K_right_mp set_subset_Cons) then have \consistent A ({q} \ V)\ using \consistent A V\ inconsistent_subset by metis then show ?thesis using \maximal A V\ unfolding maximal_def by fast qed theorem ax_in_maximal: assumes \consistent A V\ \maximal A V\ \A p\ shows \p \ V\ using assms deriv_in_maximal Ax by blast theorem mcs_properties: - assumes \consistent A V\ \maximal A V\ + assumes \consistent A V\ and \maximal A V\ shows \A \ p \ p \ V\ and \p \ V \ (\<^bold>\ p) \ V\ and \p \ V \ (p \<^bold>\ q) \ V \ q \ V\ using assms deriv_in_maximal exactly_one_in_maximal consequent_in_maximal by blast+ subsection \Lindenbaum extension\ instantiation fm :: (countable) countable begin instance by countable_datatype end primrec extend :: \('i fm \ bool) \ 'i fm set \ (nat \ 'i fm) \ nat \ 'i fm set\ where \extend A S f 0 = S\ | \extend A S f (Suc n) = (if consistent A ({f n} \ extend A S f n) then {f n} \ extend A S f n else extend A S f n)\ definition Extend :: \('i fm \ bool) \ 'i fm set \ (nat \ 'i fm) \ 'i fm set\ where \Extend A S f \ \n. extend A S f n\ lemma Extend_subset: \S \ Extend A S f\ unfolding Extend_def using Union_upper extend.simps(1) range_eqI by metis lemma extend_bound: \(\n \ m. extend A S f n) = extend A S f m\ by (induct m) (simp_all add: atMost_Suc) lemma consistent_extend: \consistent A S \ consistent A (extend A S f n)\ by (induct n) simp_all lemma UN_finite_bound: assumes \finite A\ \A \ (\n. f n)\ shows \\m :: nat. A \ (\n \ m. f n)\ using assms proof (induct rule: finite_induct) case (insert x A) then obtain m where \A \ (\n \ m. f n)\ by fast then have \A \ (\n \ (m + k). f n)\ for k by fastforce moreover obtain m' where \x \ f m'\ using insert(4) by blast ultimately have \{x} \ A \ (\n \ m + m'. f n)\ by auto then show ?case by blast qed simp lemma consistent_Extend: assumes \consistent A S\ shows \consistent A (Extend A S f)\ unfolding Extend_def proof (rule ccontr) assume \\ consistent A (\n. extend A S f n)\ then obtain S' where \A \ imply S' \<^bold>\\ \set S' \ (\n. extend A S f n)\ unfolding consistent_def by blast then obtain m where \set S' \ (\n \ m. extend A S f n)\ using UN_finite_bound by (metis List.finite_set) then have \set S' \ extend A S f m\ using extend_bound by blast moreover have \consistent A (extend A S f m)\ using assms consistent_extend by blast ultimately show False unfolding consistent_def using \A \ imply S' \<^bold>\\ by blast qed lemma maximal_Extend: assumes \surj f\ shows \maximal A (Extend A S f)\ proof (rule ccontr) assume \\ maximal A (Extend A S f)\ then obtain p where \p \ Extend A S f\ \consistent A ({p} \ Extend A S f)\ unfolding maximal_def using assms consistent_Extend by blast obtain k where n: \f k = p\ using \surj f\ unfolding surj_def by metis then have \p \ extend A S f (Suc k)\ using \p \ Extend A S f\ unfolding Extend_def by blast then have \\ consistent A ({p} \ extend A S f k)\ using n by fastforce moreover have \{p} \ extend A S f k \ {p} \ Extend A S f\ unfolding Extend_def by blast ultimately have \\ consistent A ({p} \ Extend A S f)\ unfolding consistent_def by fastforce then show False using \consistent A ({p} \ Extend A S f)\ by blast qed lemma maximal_extension: fixes V :: \('i :: countable) fm set\ assumes \consistent A V\ obtains W where \V \ W\ \consistent A W\ \maximal A W\ proof - let ?W = \Extend A V from_nat\ have \V \ ?W\ using Extend_subset by blast moreover have \consistent A ?W\ using assms consistent_Extend by blast moreover have \maximal A ?W\ using assms maximal_Extend surj_from_nat by blast ultimately show ?thesis using that by blast qed -lemma ex_mcs: - assumes \\(M :: ('i :: countable, 'w) kripke) w p. A p \ P M \ M, w \ p\ \\M. P M\ - shows \\x. x \ {V. consistent A V \ maximal A V}\ -proof - let ?V = \Extend A {\<^bold>\} from_nat\ - have \consistent A {}\ - using assms consistent_axioms by blast - moreover have \A \ \<^bold>\\ - by (simp add: A1) - ultimately have \consistent A {\<^bold>\}\ - using consistent_deriv by auto - then have \consistent A ?V\ - using consistent_Extend by blast - moreover have \maximal A ?V\ - using maximal_Extend surj_from_nat by blast - ultimately show \?V \ {V. consistent A V \ maximal A V}\ - by blast -qed - subsection \Canonical model\ abbreviation pi :: \'i fm set \ id \ bool\ where \pi V x \ Pro x \ V\ abbreviation known :: \'i fm set \ 'i \ 'i fm set\ where \known V i \ {p. K i p \ V}\ abbreviation reach :: \('i fm \ bool) \ 'i \ 'i fm set \ 'i fm set set\ where - \reach A i V \ {W. known V i \ W \ consistent A W \ maximal A W}\ + \reach A i V \ {W. known V i \ W}\ + +abbreviation mcss :: \('i fm \ bool) \ 'i fm set set\ where + \mcss A \ {W. consistent A W \ maximal A W}\ lemma truth_lemma: fixes A and p :: \('i :: countable) fm\ - defines \M \ Kripke pi (reach A)\ - assumes \consistent A V\ \maximal A V\ + defines \M \ Kripke (mcss A) pi (reach A)\ + assumes \consistent A V\ and \maximal A V\ shows \(p \ V \ M, V \ p) \ ((\<^bold>\ p) \ V \ M, V \ \<^bold>\ p)\ using assms unfolding M_def proof (induct p arbitrary: V) case FF then show ?case proof (intro conjI impI iffI) assume \\<^bold>\ \ V\ then have False using \consistent A V\ K_imply_head unfolding consistent_def by (metis bot.extremum insert_subset list.set(1) list.simps(15)) - then show \Kripke pi (reach A), V \ \<^bold>\\ .. + then show \Kripke (mcss A) pi (reach A), V \ \<^bold>\\ .. next - assume \Kripke pi (reach A), V \ \<^bold>\ \<^bold>\\ + assume \Kripke (mcss A) pi (reach A), V \ \<^bold>\ \<^bold>\\ then show \(\<^bold>\ \<^bold>\) \ V\ using \consistent A V\ \maximal A V\ unfolding maximal_def by (meson K_Boole inconsistent_subset consistent_def) qed simp_all next case (Pro x) then show ?case proof (intro conjI impI iffI) assume \(\<^bold>\ Pro x) \ V\ - then show \Kripke pi (reach A), V \ \<^bold>\ Pro x\ + then show \Kripke (mcss A) pi (reach A), V \ \<^bold>\ Pro x\ using \consistent A V\ \maximal A V\ exactly_one_in_maximal by auto next - assume \Kripke pi (reach A), V \ \<^bold>\ Pro x\ + assume \Kripke (mcss A) pi (reach A), V \ \<^bold>\ Pro x\ then show \(\<^bold>\ Pro x) \ V\ using \consistent A V\ \maximal A V\ exactly_one_in_maximal by auto qed (simp_all add: \maximal A V\ maximal_def) next case (Dis p q) - have \(p \<^bold>\ q) \ V \ Kripke pi (reach A), V \ (p \<^bold>\ q)\ + have \(p \<^bold>\ q) \ V \ Kripke (mcss A) pi (reach A), V \ (p \<^bold>\ q)\ proof assume \(p \<^bold>\ q) \ V\ then have \consistent A ({p} \ V) \ consistent A ({q} \ V)\ using \consistent A V\ consistent_disjuncts by blast then have \p \ V \ q \ V\ using \maximal A V\ unfolding maximal_def by fast - then show \Kripke pi (reach A), V \ (p \<^bold>\ q)\ + then show \Kripke (mcss A) pi (reach A), V \ (p \<^bold>\ q)\ using Dis by simp qed - moreover have \(\<^bold>\ (p \<^bold>\ q)) \ V \ Kripke pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ + moreover have \(\<^bold>\ (p \<^bold>\ q)) \ V \ Kripke (mcss A) pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ proof assume \(\<^bold>\ (p \<^bold>\ q)) \ V\ then have \consistent A ({\<^bold>\ q} \ V)\ \consistent A ({\<^bold>\ p} \ V)\ using \consistent A V\ consistent_consequent' by fastforce+ then have \(\<^bold>\ p) \ V\ \(\<^bold>\ q) \ V\ using \maximal A V\ unfolding maximal_def by fast+ - then show \Kripke pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ + then show \Kripke (mcss A) pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ using Dis by simp qed ultimately show ?case using exactly_one_in_maximal Dis by auto next case (Con p q) - have \(p \<^bold>\ q) \ V \ Kripke pi (reach A), V \ (p \<^bold>\ q)\ + have \(p \<^bold>\ q) \ V \ Kripke (mcss A) pi (reach A), V \ (p \<^bold>\ q)\ proof assume \(p \<^bold>\ q) \ V\ then have \consistent A ({p} \ V)\ \consistent A ({q} \ V)\ using \consistent A V\ consistent_consequent' by fastforce+ then have \p \ V\ \q \ V\ using \maximal A V\ unfolding maximal_def by fast+ - then show \Kripke pi (reach A), V \ (p \<^bold>\ q)\ + then show \Kripke (mcss A) pi (reach A), V \ (p \<^bold>\ q)\ using Con by simp qed - moreover have \(\<^bold>\ (p \<^bold>\ q)) \ V \ Kripke pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ + moreover have \(\<^bold>\ (p \<^bold>\ q)) \ V \ Kripke (mcss A) pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ proof assume \(\<^bold>\ (p \<^bold>\ q)) \ V\ then have \consistent A ({\<^bold>\ p \<^bold>\ \<^bold>\ q} \ V)\ using \consistent A V\ consistent_consequent' by fastforce then have \consistent A ({\<^bold>\ p} \ V) \ consistent A ({\<^bold>\ q} \ V)\ using \consistent A V\ \maximal A V\ consistent_disjuncts unfolding maximal_def by blast then have \(\<^bold>\ p) \ V \ (\<^bold>\ q) \ V\ using \maximal A V\ unfolding maximal_def by fast - then show \Kripke pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ + then show \Kripke (mcss A) pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ using Con by simp qed ultimately show ?case using exactly_one_in_maximal Con by auto next case (Imp p q) - have \(p \<^bold>\ q) \ V \ Kripke pi (reach A), V \ (p \<^bold>\ q)\ + have \(p \<^bold>\ q) \ V \ Kripke (mcss A) pi (reach A), V \ (p \<^bold>\ q)\ proof assume \(p \<^bold>\ q) \ V\ then have \consistent A ({\<^bold>\ p \<^bold>\ q} \ V)\ using \consistent A V\ consistent_consequent' by fastforce then have \consistent A ({\<^bold>\ p} \ V) \ consistent A ({q} \ V)\ using \consistent A V\ \maximal A V\ consistent_disjuncts unfolding maximal_def by blast then have \(\<^bold>\ p) \ V \ q \ V\ using \maximal A V\ unfolding maximal_def by fast - then show \Kripke pi (reach A), V \ (p \<^bold>\ q)\ + then show \Kripke (mcss A) pi (reach A), V \ (p \<^bold>\ q)\ using Imp by simp qed - moreover have \(\<^bold>\ (p \<^bold>\ q)) \ V \ Kripke pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ + moreover have \(\<^bold>\ (p \<^bold>\ q)) \ V \ Kripke (mcss A) pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ proof assume \(\<^bold>\ (p \<^bold>\ q)) \ V\ then have \consistent A ({p} \ V)\ \consistent A ({\<^bold>\ q} \ V)\ using \consistent A V\ consistent_consequent' by fastforce+ then have \p \ V\ \(\<^bold>\ q) \ V\ using \maximal A V\ unfolding maximal_def by fast+ - then show \Kripke pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ + then show \Kripke (mcss A) pi (reach A), V \ \<^bold>\ (p \<^bold>\ q)\ using Imp by simp qed ultimately show ?case using exactly_one_in_maximal Imp \consistent A V\ by auto next case (K i p) - then have \K i p \ V \ Kripke pi (reach A), V \ K i p\ + then have \K i p \ V \ Kripke (mcss A) pi (reach A), V \ K i p\ by auto - moreover have \(Kripke pi (reach A), V \ K i p) \ K i p \ V\ + moreover have \(Kripke (mcss A) pi (reach A), V \ K i p) \ K i p \ V\ proof (intro allI impI) - assume \Kripke pi (reach A), V \ K i p\ + assume \Kripke (mcss A) pi (reach A), V \ K i p\ have \\ consistent A ({\<^bold>\ p} \ known V i)\ proof assume \consistent A ({\<^bold>\ p} \ known V i)\ then obtain W where W: \{\<^bold>\ p} \ known V i \ W\ \consistent A W\ \maximal A W\ using \consistent A V\ maximal_extension by blast - then have \Kripke pi (reach A), W \ \<^bold>\ p\ + then have \Kripke (mcss A) pi (reach A), W \ \<^bold>\ p\ using K \consistent A V\ by blast - moreover have \W \ reach A i V\ - using W by simp - ultimately have \Kripke pi (reach A), V \ \<^bold>\ K i p\ + moreover have \W \ reach A i V\ \W \ mcss A\ + using W by simp_all + ultimately have \Kripke (mcss A) pi (reach A), V \ \<^bold>\ K i p\ by auto then show False - using \Kripke pi (reach A), V \ K i p\ by auto + using \Kripke (mcss A) pi (reach A), V \ K i p\ by auto qed then obtain W where W: \{\<^bold>\ p} \ W \ {\<^bold>\ p} \ known V i\ \(\<^bold>\ p) \ W\ \finite W\ \\ consistent A ({\<^bold>\ p} \ W)\ using exists_finite_inconsistent by metis obtain L where L: \set L = W\ using \finite W\ finite_list by blast then have \A \ imply L p\ using W(4) inconsistent_imply by blast then have \A \ K i (imply L p)\ using R2 by fast then have \A \ imply (map (K i) L) (K i p)\ using K_distrib_K_imp by fast then have \imply (map (K i) L) (K i p) \ V\ using deriv_in_maximal K.prems(1, 2) by blast then show \K i p \ V\ using L W(1-2) proof (induct L arbitrary: W) case (Cons a L) then have \K i a \ V\ by auto then have \imply (map (K i) L) (K i p) \ V\ using Cons(2) \consistent A V\ \maximal A V\ consequent_in_maximal by auto then show ?case using Cons by auto qed simp qed - moreover have \(Kripke pi (reach A), V \ \<^bold>\ K i p) \ (\<^bold>\ K i p) \ V\ + moreover have \(Kripke (mcss A) pi (reach A), V \ \<^bold>\ K i p) \ (\<^bold>\ K i p) \ V\ using \consistent A V\ \maximal A V\ exactly_one_in_maximal calculation(1) by (metis (no_types, lifting) semantics.simps(1, 5)) - moreover have \(\<^bold>\ K i p) \ V \ Kripke pi (reach A), V \ \<^bold>\ K i p\ + moreover have \(\<^bold>\ K i p) \ V \ Kripke (mcss A) pi (reach A), V \ \<^bold>\ K i p\ using \consistent A V\ \maximal A V\ calculation(2) exactly_one_in_maximal by auto ultimately show ?case by blast qed lemma canonical_model: - assumes \consistent A S\ \p \ S\ - defines \V \ Extend A S from_nat\ and \M \ Kripke pi (reach A)\ - shows \M, V \ p\ \consistent A V\ \maximal A V\ + assumes \consistent A S\ and \p \ S\ + defines \V \ Extend A S from_nat\ and \M \ Kripke (mcss A) pi (reach A)\ + shows \M, V \ p\ and \consistent A V\ and \maximal A V\ proof - have \consistent A V\ using \consistent A S\ unfolding V_def using consistent_Extend by blast have \maximal A V\ unfolding V_def using maximal_Extend surj_from_nat by blast { fix x assume \x \ S\ then have \x \ V\ unfolding V_def using Extend_subset by blast then have \M, V \ x\ unfolding M_def using truth_lemma \consistent A V\ \maximal A V\ by blast } then show \M, V \ p\ using \p \ S\ by blast+ show \consistent A V\ \maximal A V\ by fact+ qed subsection \Completeness\ lemma imply_completeness: - assumes valid: \\(M :: ('i, ('i :: countable) fm set) kripke) w. - list_all (\q. M, w \ q) G \ M, w \ p\ - shows \A \ imply G p\ -proof (rule K_Boole, rule ccontr) - assume \\ A \ imply ((\<^bold>\ p) # G) \<^bold>\\ + assumes valid: \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. + (\q \ G. M, w \ q) \ M, w \ p\ + shows \\qs. set qs \ G \ (A \ imply qs p)\ +proof (rule ccontr) + assume \\qs. set qs \ G \ A \ imply qs p\ + then have *: \\qs. set qs \ G \ \ A \ imply ((\<^bold>\ p) # qs) \<^bold>\\ + using K_Boole by blast - let ?S = \set ((\<^bold>\ p) # G)\ + let ?S = \{\<^bold>\ p} \ G\ let ?V = \Extend A ?S from_nat\ - let ?M = \Kripke pi (reach A)\ + let ?M = \Kripke (mcss A) pi (reach A)\ have \consistent A ?S\ - unfolding consistent_def using \\ A \ imply ((\<^bold>\ p) # G) \<^bold>\\ K_imply_weaken by blast - then have \?M, ?V \ (\<^bold>\ p)\ \list_all (\p. ?M, ?V \ p) G\ - unfolding list_all_def using canonical_model by fastforce+ - then have \?M, ?V \ p\ - using valid by blast + using * by (metis K_imply_Cons consistent_def inconsistent_subset) + then have \?M, ?V \ (\<^bold>\ p)\ \\q \ G. ?M, ?V \ q\ + using canonical_model by fastforce+ + moreover have \?V \ mcss A\ + using \consistent A ?S\ consistent_Extend maximal_Extend surj_from_nat by blast + ultimately have \?M, ?V \ p\ + using valid by simp then show False using \?M, ?V \ (\<^bold>\ p)\ by simp qed theorem completeness: - assumes \\(M :: ('i :: countable, 'i fm set) kripke) w. M, w \ p\ + assumes \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. M, w \ p\ shows \A \ p\ - using assms imply_completeness[where G=\[]\] by auto + using assms imply_completeness[where G=\{}\] by auto section \System K\ abbreviation SystemK :: \'i fm \ bool\ ("\\<^sub>K _" [50] 50) where \\\<^sub>K p \ (\_. False) \ p\ -lemma soundness\<^sub>K: \\\<^sub>K p \ M, w \ p\ +lemma soundness\<^sub>K: \\\<^sub>K p \ w \ \ M \ M, w \ p\ using soundness by metis -abbreviation \valid\<^sub>K p \ \(M :: (nat, nat fm set) kripke) w. M, w \ p\ +abbreviation \valid\<^sub>K p \ \(M :: (nat, nat fm set) kripke). \w \ \ M. M, w \ p\ theorem main\<^sub>K: \valid\<^sub>K p \ \\<^sub>K p\ proof assume \valid\<^sub>K p\ with completeness show \\\<^sub>K p\ by blast next assume \\\<^sub>K p\ with soundness\<^sub>K show \valid\<^sub>K p\ - by (intro allI) + by fast qed -corollary \valid\<^sub>K p \ M, w \ p\ -proof - assume \valid\<^sub>K p\ - then have \\\<^sub>K p\ - unfolding main\<^sub>K . - with soundness\<^sub>K show \M, w \ p\ . +corollary + assumes \valid\<^sub>K p\ and \w \ \ M\ + shows \M, w \ p\ +proof - + have \\\<^sub>K p\ + using assms(1) unfolding main\<^sub>K . + with soundness\<^sub>K assms(2) show \M, w \ p\ by fast qed section \System T\ text \Also known as System M\ inductive AxT :: \'i fm \ bool\ where \AxT (K i p \<^bold>\ p)\ abbreviation SystemT :: \'i fm \ bool\ ("\\<^sub>T _" [50] 50) where \\\<^sub>T p \ AxT \ p\ -lemma soundness_AxT: \AxT p \ reflexive M \ M, w \ p\ +lemma soundness_AxT: \AxT p \ reflexive M \ w \ \ M \ M, w \ p\ by (induct p rule: AxT.induct) (meson truth) -lemma soundness\<^sub>T: \\\<^sub>T p \ reflexive M \ M, w \ p\ +lemma soundness\<^sub>T: \\\<^sub>T p \ reflexive M \ w \ \ M \ M, w \ p\ using soundness soundness_AxT . -typedef 'i mcs\<^sub>T = \{V :: ('i :: countable) fm set. consistent AxT V \ maximal AxT V}\ - using ex_mcs[where A=AxT and P=reflexive] soundness_AxT ex_equiv_model by fast - -abbreviation \pi\<^sub>T \ pi o Rep_mcs\<^sub>T\ -abbreviation \reach\<^sub>T i V \ Abs_mcs\<^sub>T ` (reach AxT i (Rep_mcs\<^sub>T V))\ - -lemma mcs\<^sub>T_equiv: - assumes \consistent AxT V\ \maximal AxT V\ - shows \(Kripke pi (reach AxT), V \ p) = (Kripke pi\<^sub>T reach\<^sub>T, Abs_mcs\<^sub>T V \ p)\ - using assms by (induct p arbitrary: V) (simp_all add: Abs_mcs\<^sub>T_inverse) - lemma AxT_reflexive: - assumes \\p. AxT p \ A p\ \consistent A V\ \maximal A V\ + assumes \\p. AxT p \ A p\ and \consistent A V\ and \maximal A V\ shows \V \ reach A i V\ proof - have \(K i p \<^bold>\ p) \ V\ for p using assms ax_in_maximal AxT.intros by metis then have \p \ V\ if \K i p \ V\ for p using that assms consequent_in_maximal by blast then show ?thesis using assms by blast qed -lemma mcs\<^sub>T_reflexive: \reflexive (Kripke pi\<^sub>T reach\<^sub>T)\ +lemma mcs\<^sub>T_reflexive: + assumes \\p. AxT p \ A p\ + shows \reflexive (Kripke (mcss A) pi (reach A))\ unfolding reflexive_def -proof (intro allI) - fix i :: \'i :: countable\ and V :: \'i mcs\<^sub>T\ - let ?V = \Rep_mcs\<^sub>T V\ - have \consistent AxT ?V\ \maximal AxT ?V\ - using Rep_mcs\<^sub>T by blast+ - then have \?V \ reach AxT i ?V\ - using AxT_reflexive by fast - then have \V \ reach\<^sub>T i V\ - using Rep_mcs\<^sub>T_inverse by (metis (no_types, lifting) image_iff) - then show \V \ \ (Kripke pi\<^sub>T reach\<^sub>T) i V\ +proof safe + fix i V + assume \V \ \ (Kripke (mcss A) pi (reach A))\ + then have \consistent A V\ \maximal A V\ + by simp_all + with AxT_reflexive assms have \V \ reach A i V\ . + then show \V \ \ (Kripke (mcss A) pi (reach A)) i V\ by simp qed lemma imply_completeness_T: - assumes valid: \\(M :: ('i, ('i :: countable) mcs\<^sub>T) kripke) w. - reflexive M \ list_all (\q. M, w \ q) G \ M, w \ p\ - shows \AxT \ imply G p\ -proof (rule K_Boole, rule ccontr) - assume \\ AxT \ imply ((\<^bold>\ p) # G) \<^bold>\\ + assumes valid: \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. + reflexive M \ (\q \ G. M, w \ q) \ M, w \ p\ + shows \\qs. set qs \ G \ (AxT \ imply qs p)\ +proof (rule ccontr) + assume \\qs. set qs \ G \ AxT \ imply qs p\ + then have *: \\qs. set qs \ G \ \ AxT \ imply ((\<^bold>\ p) # qs) \<^bold>\\ + using K_Boole by blast - let ?S = \set ((\<^bold>\ p) # G)\ + let ?S = \{\<^bold>\ p} \ G\ let ?V = \Extend AxT ?S from_nat\ - let ?M = \Kripke pi (reach AxT)\ - let ?V\<^sub>T = \Abs_mcs\<^sub>T ?V\ - let ?M\<^sub>T = \Kripke pi\<^sub>T reach\<^sub>T\ + let ?M = \Kripke (mcss AxT) pi (reach AxT)\ have \consistent AxT ?S\ - unfolding consistent_def using \\ AxT \ imply ((\<^bold>\ p) # G) \<^bold>\\ K_imply_weaken by blast - then have \?M, ?V \ (\<^bold>\ p)\ \list_all (\p. ?M, ?V \ p) G\ \consistent AxT ?V\ \maximal AxT ?V\ + using * by (metis K_imply_Cons consistent_def inconsistent_subset) + then have \?M, ?V \ (\<^bold>\ p)\ \\q \ G. ?M, ?V \ q\ \consistent AxT ?V\ \maximal AxT ?V\ using canonical_model unfolding list_all_def by fastforce+ - then have \?M\<^sub>T, ?V\<^sub>T \ (\<^bold>\ p)\ \list_all (\p. ?M\<^sub>T, ?V\<^sub>T \ p) G\ - using \consistent AxT ?V\ \maximal AxT ?V\ mcs\<^sub>T_equiv unfolding list_all_def by blast+ - moreover have \reflexive ?M\<^sub>T\ + moreover have \reflexive ?M\ using mcs\<^sub>T_reflexive by fast - ultimately have \?M\<^sub>T, ?V\<^sub>T \ p\ - using valid by blast + ultimately have \?M, ?V \ p\ + using valid by auto then show False - using \?M\<^sub>T, ?V\<^sub>T \ (\<^bold>\ p)\ by simp + using \?M, ?V \ (\<^bold>\ p)\ by simp qed lemma completeness\<^sub>T: - assumes \\(M :: ('i, ('i :: countable) mcs\<^sub>T) kripke) w. reflexive M \ M, w \ p\ + assumes \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. reflexive M \ M, w \ p\ shows \\\<^sub>T p\ - using assms imply_completeness_T[where G=\[]\] by auto + using assms imply_completeness_T[where G=\{}\] by auto -abbreviation \valid\<^sub>T p \ \(M :: (nat, nat mcs\<^sub>T) kripke) w. reflexive M \ M, w \ p\ +abbreviation \valid\<^sub>T p \ \(M :: (nat, nat fm set) kripke). \w \ \ M. reflexive M \ M, w \ p\ theorem main\<^sub>T: \valid\<^sub>T p \ \\<^sub>T p\ using soundness\<^sub>T completeness\<^sub>T by fast corollary - assumes \reflexive M\ + assumes \reflexive M\ \w \ \ M\ shows \valid\<^sub>T p \ M, w \ p\ using assms soundness\<^sub>T completeness\<^sub>T by fast section \System KB\ inductive AxB :: \'i fm \ bool\ where \AxB (p \<^bold>\ K i (L i p))\ abbreviation SystemKB :: \'i fm \ bool\ ("\\<^sub>K\<^sub>B _" [50] 50) where \\\<^sub>K\<^sub>B p \ AxB \ p\ -lemma soundness_AxB: \AxB p \ symmetric M \ M, w \ p\ +lemma soundness_AxB: \AxB p \ symmetric M \ w \ \ M \ M, w \ p\ unfolding symmetric_def by (induct p rule: AxB.induct) auto -lemma soundness\<^sub>K\<^sub>B: \\\<^sub>K\<^sub>B p \ symmetric M \ M, w \ p\ +lemma soundness\<^sub>K\<^sub>B: \\\<^sub>K\<^sub>B p \ symmetric M \ w \ \ M \ M, w \ p\ using soundness soundness_AxB . -typedef 'i mcs\<^sub>K\<^sub>B = \{V :: ('i :: countable) fm set. consistent AxB V \ maximal AxB V}\ - using ex_mcs[where A=AxB and P=symmetric] soundness_AxB ex_equiv_model by fast - -abbreviation \pi\<^sub>K\<^sub>B \ pi o Rep_mcs\<^sub>K\<^sub>B\ -abbreviation \reach\<^sub>K\<^sub>B i V \ Abs_mcs\<^sub>K\<^sub>B ` (reach AxB i (Rep_mcs\<^sub>K\<^sub>B V))\ - -lemma mcs\<^sub>K\<^sub>B_equiv: - assumes \consistent AxB V\ \maximal AxB V\ - shows \(Kripke pi (reach AxB), V \ p) = (Kripke pi\<^sub>K\<^sub>B reach\<^sub>K\<^sub>B, Abs_mcs\<^sub>K\<^sub>B V \ p)\ - using assms by (induct p arbitrary: V) (simp_all add: Abs_mcs\<^sub>K\<^sub>B_inverse) - lemma AxB_symmetric': - assumes \\p. AxB p \ A p\ \consistent A V\ \maximal A V\ \consistent A W\ \maximal A W\ + assumes \\p. AxB p \ A p\ \consistent A V\ \maximal A V\ \consistent A W\ \maximal A W\ and \W \ reach A i V\ shows \V \ reach A i W\ proof - have \\p. K i p \ W \ p \ V\ proof (intro allI impI, rule ccontr) fix p assume \K i p \ W\ \p \ V\ then have \(\<^bold>\ p) \ V\ using assms(2-3) exactly_one_in_maximal by fast then have \K i (L i (\<^bold>\ p)) \ V\ using assms(1-3) ax_in_maximal AxB.intros consequent_in_maximal by fast then have \L i (\<^bold>\ p) \ W\ using \W \ reach A i V\ by fast then have \(\<^bold>\ K i p) \ W\ using assms(4-5) by (meson K_LK consistent_consequent maximal_def) then show False using \K i p \ W\ assms(4-5) exactly_one_in_maximal by fast qed then have \known W i \ V\ by blast then show ?thesis using assms(2-3) by simp qed lemma AxB_symmetric: - assumes \\p. AxB p \ A p\ \consistent A V\ \maximal A V\ \consistent A W\ \maximal A W\ + assumes \\p. AxB p \ A p\ \consistent A V\ \maximal A V\ \consistent A W\ \maximal A W\ shows \W \ reach A i V \ V \ reach A i W\ using assms AxB_symmetric'[where V=V and W=W] AxB_symmetric'[where V=W and W=V] by (intro iffI) blast+ -lemma mcs\<^sub>K\<^sub>B_symmetric: \symmetric (Kripke pi\<^sub>K\<^sub>B reach\<^sub>K\<^sub>B)\ +lemma mcs\<^sub>K\<^sub>B_symmetric: + assumes \\p. AxB p \ A p\ + shows \symmetric (Kripke (mcss A) pi (reach A))\ unfolding symmetric_def -proof (intro allI) - fix i :: \'i :: countable\ and V W :: \'i mcs\<^sub>K\<^sub>B\ - let ?V = \Rep_mcs\<^sub>K\<^sub>B V\ and ?W = \Rep_mcs\<^sub>K\<^sub>B W\ - - have \consistent AxB ?V\ \maximal AxB ?V\ - \consistent AxB ?W\ \maximal AxB ?W\ - using Rep_mcs\<^sub>K\<^sub>B by blast+ - then have \?V \ reach AxB i ?W \ ?W \ reach AxB i ?V\ - using AxB_symmetric' by fast - then have \V \ reach\<^sub>K\<^sub>B i W \ W \ reach\<^sub>K\<^sub>B i V\ - by (intro iffI) - (metis (no_types, lifting) Abs_mcs\<^sub>K\<^sub>B_inverse Rep_mcs\<^sub>K\<^sub>B_inverse image_iff mem_Collect_eq)+ - then show \(W \ \ (Kripke pi\<^sub>K\<^sub>B reach\<^sub>K\<^sub>B) i V) = (V \ \ (Kripke pi\<^sub>K\<^sub>B reach\<^sub>K\<^sub>B) i W)\ +proof (intro allI ballI) + fix i V W + assume \V \ \ (Kripke (mcss A) pi (reach A))\ \W \ \ (Kripke (mcss A) pi (reach A))\ + then have \consistent A V\ \maximal A V\ \consistent A W\ \maximal A W\ + by simp_all + with AxB_symmetric assms have \W \ reach A i V \ V \ reach A i W\ . + then show + \(W \ \ (Kripke (mcss A) pi (reach A)) i V) = (V \ \ (Kripke (mcss A) pi (reach A)) i W)\ by simp qed lemma imply_completeness_KB: - assumes valid: \\(M :: ('i :: countable, 'i mcs\<^sub>K\<^sub>B) kripke) w. - symmetric M \ list_all (\q. M, w \ q) G \ M, w \ p\ - shows \AxB \ imply G p\ -proof (rule K_Boole, rule ccontr) - assume \\ AxB \ imply ((\<^bold>\ p) # G) \<^bold>\\ + assumes valid: \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. + symmetric M \ (\q \ G. M, w \ q) \ M, w \ p\ + shows \\qs. set qs \ G \ (AxB \ imply qs p)\ +proof (rule ccontr) + assume \\qs. set qs \ G \ AxB \ imply qs p\ + then have *: \\qs. set qs \ G \ \ AxB \ imply ((\<^bold>\ p) # qs) \<^bold>\\ + using K_Boole by blast - let ?S = \set ((\<^bold>\ p) # G)\ + let ?S = \{\<^bold>\ p} \ G\ let ?V = \Extend AxB ?S from_nat\ - let ?M = \Kripke pi (reach AxB)\ - let ?V\<^sub>K\<^sub>B = \Abs_mcs\<^sub>K\<^sub>B ?V\ - let ?M\<^sub>K\<^sub>B = \Kripke pi\<^sub>K\<^sub>B reach\<^sub>K\<^sub>B\ + let ?M = \Kripke (mcss AxB) pi (reach AxB)\ have \consistent AxB ?S\ - unfolding consistent_def using \\ AxB \ imply ((\<^bold>\ p) # G) \<^bold>\\ K_imply_weaken by blast - then have \?M, ?V \ (\<^bold>\ p)\ \list_all (\p. ?M, ?V \ p) G\ \consistent AxB ?V\ \maximal AxB ?V\ + using * by (metis K_imply_Cons consistent_def inconsistent_subset) + then have \?M, ?V \ (\<^bold>\ p)\ \\q \ G. ?M, ?V \ q\ \consistent AxB ?V\ \maximal AxB ?V\ using canonical_model unfolding list_all_def by fastforce+ - then have \?M\<^sub>K\<^sub>B, ?V\<^sub>K\<^sub>B \ (\<^bold>\ p)\ \list_all (\p. ?M\<^sub>K\<^sub>B, ?V\<^sub>K\<^sub>B \ p) G\ - using \consistent AxB ?V\ \maximal AxB ?V\ mcs\<^sub>K\<^sub>B_equiv unfolding list_all_def by blast+ - moreover have \symmetric ?M\<^sub>K\<^sub>B\ + moreover have \symmetric ?M\ using mcs\<^sub>K\<^sub>B_symmetric by fast - ultimately have \?M\<^sub>K\<^sub>B, ?V\<^sub>K\<^sub>B \ p\ - using valid by blast + ultimately have \?M, ?V \ p\ + using valid by auto then show False - using \?M\<^sub>K\<^sub>B, ?V\<^sub>K\<^sub>B \ (\<^bold>\ p)\ by simp + using \?M, ?V \ (\<^bold>\ p)\ by simp qed lemma completeness\<^sub>K\<^sub>B: - assumes \\(M :: ('i, ('i :: countable) mcs\<^sub>K\<^sub>B) kripke) w. symmetric M \ M, w \ p\ + assumes \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. symmetric M \ M, w \ p\ shows \\\<^sub>K\<^sub>B p\ - using assms imply_completeness_KB[where G=\[]\] by auto + using assms imply_completeness_KB[where G=\{}\] by auto -abbreviation \valid\<^sub>K\<^sub>B p \ \(M :: (nat, nat mcs\<^sub>K\<^sub>B) kripke) w. symmetric M \ M, w \ p\ +abbreviation \valid\<^sub>K\<^sub>B p \ \(M :: (nat, nat fm set) kripke). \w \ \ M. symmetric M \ M, w \ p\ theorem main\<^sub>K\<^sub>B: \valid\<^sub>K\<^sub>B p \ \\<^sub>K\<^sub>B p\ using soundness\<^sub>K\<^sub>B completeness\<^sub>K\<^sub>B by fast corollary - assumes \symmetric M\ + assumes \symmetric M\ \w \ \ M\ shows \valid\<^sub>K\<^sub>B p \ M, w \ p\ using assms soundness\<^sub>K\<^sub>B completeness\<^sub>K\<^sub>B by fast section \System K4\ inductive Ax4 :: \'i fm \ bool\ where \Ax4 (K i p \<^bold>\ K i (K i p))\ abbreviation SystemK4 :: \'i fm \ bool\ ("\\<^sub>K\<^sub>4 _" [50] 50) where \\\<^sub>K\<^sub>4 p \ Ax4 \ p\ -lemma soundness_Ax4: \Ax4 p \ transitive M \ M, w \ p\ +lemma soundness_Ax4: \Ax4 p \ transitive M \ w \ \ M \ M, w \ p\ by (induct p rule: Ax4.induct) (meson pos_introspection) -lemma soundness\<^sub>K\<^sub>4: \\\<^sub>K\<^sub>4 p \ transitive M \ M, w \ p\ +lemma soundness\<^sub>K\<^sub>4: \\\<^sub>K\<^sub>4 p \ transitive M \ w \ \ M \ M, w \ p\ using soundness soundness_Ax4 . -typedef 'i mcs\<^sub>K\<^sub>4 = \{V :: ('i :: countable) fm set. consistent Ax4 V \ maximal Ax4 V}\ - using ex_mcs[where A=Ax4 and P=transitive] soundness_Ax4 ex_equiv_model by fast - -abbreviation \pi\<^sub>K\<^sub>4 \ pi o Rep_mcs\<^sub>K\<^sub>4\ -abbreviation \reach\<^sub>K\<^sub>4 i V \ Abs_mcs\<^sub>K\<^sub>4 ` (reach Ax4 i (Rep_mcs\<^sub>K\<^sub>4 V))\ - -lemma mcs\<^sub>K\<^sub>4_equiv: - assumes \consistent Ax4 V\ \maximal Ax4 V\ - shows \(Kripke pi (reach Ax4), V \ p) = (Kripke pi\<^sub>K\<^sub>4 reach\<^sub>K\<^sub>4, Abs_mcs\<^sub>K\<^sub>4 V \ p)\ - using assms by (induct p arbitrary: V) (simp_all add: Abs_mcs\<^sub>K\<^sub>4_inverse) - lemma Ax4_transitive: - assumes \\p. Ax4 p \ A p\ \consistent A V\ \maximal A V\ + assumes \\p. Ax4 p \ A p\ \consistent A V\ \maximal A V\ and \W \ reach A i V\ \U \ reach A i W\ shows \U \ reach A i V\ proof - have \(K i p \<^bold>\ K i (K i p)) \ V\ for p using assms(1-3) ax_in_maximal Ax4.intros by metis then have \K i (K i p) \ V\ if \K i p \ V\ for p using that assms(2-3) consequent_in_maximal by blast then show ?thesis using assms(4-5) by blast qed -lemma mcs\<^sub>K\<^sub>4_transitive: \transitive (Kripke pi\<^sub>K\<^sub>4 reach\<^sub>K\<^sub>4)\ +lemma mcs\<^sub>K\<^sub>4_transitive: + assumes \\p. Ax4 p \ A p\ + shows \transitive (Kripke (mcss A) pi (reach A))\ unfolding transitive_def proof safe - fix i :: \'i :: countable\ and V W U :: \'i mcs\<^sub>K\<^sub>4\ - let ?V = \Rep_mcs\<^sub>K\<^sub>4 V\ and ?W = \Rep_mcs\<^sub>K\<^sub>4 W\ and ?U = \Rep_mcs\<^sub>K\<^sub>4 U\ - - assume \W \ \ (Kripke pi\<^sub>K\<^sub>4 reach\<^sub>K\<^sub>4) i V\ \U \ \ (Kripke pi\<^sub>K\<^sub>4 reach\<^sub>K\<^sub>4) i W\ - then have \?W \ reach Ax4 i ?V\ \?U \ reach Ax4 i ?W\ - using Abs_mcs\<^sub>K\<^sub>4_inverse by auto - moreover have \consistent Ax4 ?V\ \maximal Ax4 ?V\ - using Rep_mcs\<^sub>K\<^sub>4 by blast+ - ultimately have \?U \ reach Ax4 i ?V\ - using Ax4_transitive[where V=\?V\] by blast - then have \U \ reach\<^sub>K\<^sub>4 i V\ - using Rep_mcs\<^sub>K\<^sub>4_inverse by (metis (no_types, lifting) image_iff) - then show \U \ \ (Kripke pi\<^sub>K\<^sub>4 reach\<^sub>K\<^sub>4) i V\ + fix i U V W + assume \V \ \ (Kripke (mcss A) pi (reach A))\ + then have \consistent A V\ \maximal A V\ + by simp_all + moreover assume + \W \ \ (Kripke (mcss A) pi (reach A)) i V\ + \U \ \ (Kripke (mcss A) pi (reach A)) i W\ + ultimately have \U \ reach A i V\ + using Ax4_transitive[where V=V and W=W and U=U] assms by simp + then show \U \ \ (Kripke (mcss A) pi (reach A)) i V\ by simp qed lemma imply_completeness_K4: - assumes valid: \\(M :: ('i :: countable, 'i mcs\<^sub>K\<^sub>4) kripke) w. - transitive M \ list_all (\q. M, w \ q) G \ M, w \ p\ - shows \Ax4 \ imply G p\ -proof (rule K_Boole, rule ccontr) - assume \\ Ax4 \ imply ((\<^bold>\ p) # G) \<^bold>\\ + assumes valid: \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. + transitive M \ (\q \ G. M, w \ q) \ M, w \ p\ + shows \\qs. set qs \ G \ (Ax4 \ imply qs p)\ +proof (rule ccontr) + assume \\qs. set qs \ G \ Ax4 \ imply qs p\ + then have *: \\qs. set qs \ G \ \ Ax4 \ imply ((\<^bold>\ p) # qs) \<^bold>\\ + using K_Boole by blast - let ?S = \set ((\<^bold>\ p) # G)\ + let ?S = \{\<^bold>\ p} \ G\ let ?V = \Extend Ax4 ?S from_nat\ - let ?M = \Kripke pi (reach Ax4)\ - let ?V\<^sub>K\<^sub>4 = \Abs_mcs\<^sub>K\<^sub>4 ?V\ - let ?M\<^sub>K\<^sub>4 = \Kripke pi\<^sub>K\<^sub>4 reach\<^sub>K\<^sub>4\ + let ?M = \Kripke (mcss Ax4) pi (reach Ax4)\ have \consistent Ax4 ?S\ - unfolding consistent_def using \\ Ax4 \ imply ((\<^bold>\ p) # G) \<^bold>\\ K_imply_weaken by blast - then have \?M, ?V \ (\<^bold>\ p)\ \list_all (\p. ?M, ?V \ p) G\ \consistent Ax4 ?V\ \maximal Ax4 ?V\ + using * by (metis K_imply_Cons consistent_def inconsistent_subset) + then have \?M, ?V \ (\<^bold>\ p)\ \\q \ G. ?M, ?V \ q\ \consistent Ax4 ?V\ \maximal Ax4 ?V\ using canonical_model unfolding list_all_def by fastforce+ - then have \?M\<^sub>K\<^sub>4, ?V\<^sub>K\<^sub>4 \ (\<^bold>\ p)\ \list_all (\p. ?M\<^sub>K\<^sub>4, ?V\<^sub>K\<^sub>4 \ p) G\ - using \consistent Ax4 ?V\ \maximal Ax4 ?V\ mcs\<^sub>K\<^sub>4_equiv unfolding list_all_def by blast+ - moreover have \transitive ?M\<^sub>K\<^sub>4\ + moreover have \transitive ?M\ using mcs\<^sub>K\<^sub>4_transitive by fast - ultimately have \?M\<^sub>K\<^sub>4, ?V\<^sub>K\<^sub>4 \ p\ - using valid by blast + ultimately have \?M, ?V \ p\ + using valid by auto then show False - using \?M\<^sub>K\<^sub>4, ?V\<^sub>K\<^sub>4 \ (\<^bold>\ p)\ by simp + using \?M, ?V \ (\<^bold>\ p)\ by simp qed lemma completeness\<^sub>K\<^sub>4: - assumes \\(M :: ('i, ('i :: countable) mcs\<^sub>K\<^sub>4) kripke) w. transitive M \ M, w \ p\ + assumes \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. transitive M \ M, w \ p\ shows \\\<^sub>K\<^sub>4 p\ - using assms imply_completeness_K4[where G=\[]\] by auto + using assms imply_completeness_K4[where G=\{}\] by auto -abbreviation \valid\<^sub>K\<^sub>4 p \ \(M :: (nat, nat mcs\<^sub>K\<^sub>4) kripke) w. transitive M \ M, w \ p\ +abbreviation \valid\<^sub>K\<^sub>4 p \ \(M :: (nat, nat fm set) kripke). \w \ \ M. transitive M \ M, w \ p\ theorem main\<^sub>K\<^sub>4: \valid\<^sub>K\<^sub>4 p \ \\<^sub>K\<^sub>4 p\ using soundness\<^sub>K\<^sub>4 completeness\<^sub>K\<^sub>4 by fast corollary - assumes \transitive M\ + assumes \transitive M\ \w \ \ M\ shows \valid\<^sub>K\<^sub>4 p \ M, w \ p\ using assms soundness\<^sub>K\<^sub>4 completeness\<^sub>K\<^sub>4 by fast section \System S4\ abbreviation Or :: \('a \ bool) \ ('a \ bool) \ 'a \ bool\ (infixl \\\ 65) where \A \ A' \ \x. A x \ A' x\ abbreviation SystemS4 :: \'i fm \ bool\ ("\\<^sub>S\<^sub>4 _" [50] 50) where \\\<^sub>S\<^sub>4 p \ AxT \ Ax4 \ p\ -lemma soundness_AxT4: \(AxT \ Ax4) p \ reflexive M \ transitive M \ M, w \ p\ +lemma soundness_AxT4: \(AxT \ Ax4) p \ reflexive M \ transitive M \ w \ \ M \ M, w \ p\ using soundness_AxT soundness_Ax4 by fast -lemma soundness\<^sub>S\<^sub>4: \\\<^sub>S\<^sub>4 p \ reflexive M \ transitive M \ M, w \ p\ +lemma soundness\<^sub>S\<^sub>4: \\\<^sub>S\<^sub>4 p \ reflexive M \ transitive M \ w \ \ M \ M, w \ p\ using soundness soundness_AxT4 . -typedef 'i mcs\<^sub>S\<^sub>4 = - \{V :: ('i :: countable) fm set. consistent (AxT \ Ax4) V \ maximal (AxT \ Ax4) V}\ - using ex_mcs[where A=\AxT \ Ax4\ and P=\\M. reflexive M \ transitive M\] - soundness_AxT4 ex_equiv_model by fast - -abbreviation \pi\<^sub>S\<^sub>4 \ pi o Rep_mcs\<^sub>S\<^sub>4\ -abbreviation \reach\<^sub>S\<^sub>4 i V \ Abs_mcs\<^sub>S\<^sub>4 ` (reach (AxT \ Ax4) i (Rep_mcs\<^sub>S\<^sub>4 V))\ - -lemma mcs\<^sub>S\<^sub>4_equiv: - assumes \consistent (AxT \ Ax4) V\ \maximal (AxT \ Ax4) V\ - shows \(Kripke pi (reach (AxT \ Ax4)), V \ p) = (Kripke pi\<^sub>S\<^sub>4 reach\<^sub>S\<^sub>4, Abs_mcs\<^sub>S\<^sub>4 V \ p)\ - using assms by (induct p arbitrary: V) (simp_all add: Abs_mcs\<^sub>S\<^sub>4_inverse) - -lemma mcs\<^sub>S\<^sub>4_reflexive: \reflexive (Kripke pi\<^sub>S\<^sub>4 reach\<^sub>S\<^sub>4)\ - unfolding reflexive_def -proof (intro allI) - fix i :: \'i :: countable\ and V :: \'i mcs\<^sub>S\<^sub>4\ - let ?V = \Rep_mcs\<^sub>S\<^sub>4 V\ - have \consistent (AxT \ Ax4) ?V\ \maximal (AxT \ Ax4) ?V\ - using Rep_mcs\<^sub>S\<^sub>4 by blast+ - then have \?V \ reach (AxT \ Ax4) i ?V\ - using AxT_reflexive[where V=\?V\] by presburger - then have \V \ reach\<^sub>S\<^sub>4 i V\ - using Rep_mcs\<^sub>S\<^sub>4_inverse by (metis (no_types, lifting) image_iff) - then show \V \ \ (Kripke pi\<^sub>S\<^sub>4 reach\<^sub>S\<^sub>4) i V\ - by simp -qed +lemma imply_completeness_S4: + assumes valid: \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. + reflexive M \ transitive M \ (\q \ G. M, w \ q) \ M, w \ p\ + shows \\qs. set qs \ G \ (AxT \ Ax4 \ imply qs p)\ +proof (rule ccontr) + assume \\qs. set qs \ G \ AxT \ Ax4 \ imply qs p\ + then have *: \\qs. set qs \ G \ \ AxT \ Ax4 \ imply ((\<^bold>\ p) # qs) \<^bold>\\ + using K_Boole by blast -lemma mcs\<^sub>S\<^sub>4_transitive: \transitive (Kripke pi\<^sub>S\<^sub>4 reach\<^sub>S\<^sub>4)\ - unfolding transitive_def -proof safe - fix i :: \'i :: countable\ and V W U :: \'i mcs\<^sub>S\<^sub>4\ - let ?V = \Rep_mcs\<^sub>S\<^sub>4 V\ and ?W = \Rep_mcs\<^sub>S\<^sub>4 W\ and ?U = \Rep_mcs\<^sub>S\<^sub>4 U\ - - assume \W \ \ (Kripke pi\<^sub>S\<^sub>4 reach\<^sub>S\<^sub>4) i V\ \U \ \ (Kripke pi\<^sub>S\<^sub>4 reach\<^sub>S\<^sub>4) i W\ - then have \?W \ reach (AxT \ Ax4) i ?V\ \?U \ reach (AxT \ Ax4) i ?W\ - using Abs_mcs\<^sub>S\<^sub>4_inverse by auto - moreover have \consistent (AxT \ Ax4) ?V\ \maximal (AxT \ Ax4) ?V\ - using Rep_mcs\<^sub>S\<^sub>4 by blast+ - ultimately have \?U \ reach (AxT \ Ax4) i ?V\ - using Ax4_transitive[where V=\?V\] by presburger - then have \U \ reach\<^sub>S\<^sub>4 i V\ - using Rep_mcs\<^sub>S\<^sub>4_inverse by (metis (no_types, lifting) image_iff) - then show \U \ \ (Kripke pi\<^sub>S\<^sub>4 reach\<^sub>S\<^sub>4) i V\ - by simp -qed - -lemma imply_completeness_S4: - assumes valid: \\(M :: ('i :: countable, 'i mcs\<^sub>S\<^sub>4) kripke) w. - reflexive M \ transitive M \ list_all (\q. M, w \ q) G \ M, w \ p\ - shows \AxT \ Ax4 \ imply G p\ -proof (rule K_Boole, rule ccontr) - assume \\ AxT \ Ax4 \ imply ((\<^bold>\ p) # G) \<^bold>\\ - - let ?S = \set ((\<^bold>\ p) # G)\ + let ?S = \{\<^bold>\ p} \ G\ let ?V = \Extend (AxT \ Ax4) ?S from_nat\ - let ?M = \Kripke pi (reach (AxT \ Ax4))\ - let ?V\<^sub>S\<^sub>4 = \Abs_mcs\<^sub>S\<^sub>4 ?V\ - let ?M\<^sub>S\<^sub>4 = \Kripke pi\<^sub>S\<^sub>4 reach\<^sub>S\<^sub>4\ + let ?M = \Kripke (mcss (AxT \ Ax4)) pi (reach (AxT \ Ax4))\ have \consistent (AxT \ Ax4) ?S\ - unfolding consistent_def using \\ AxT \ Ax4 \ imply ((\<^bold>\ p) # G) \<^bold>\\ K_imply_weaken by blast + using * by (metis (no_types, lifting) K_imply_Cons consistent_def inconsistent_subset) then have - \?M, ?V \ (\<^bold>\ p)\ \list_all (\p. ?M, ?V \ p) G\ + \?M, ?V \ (\<^bold>\ p)\ \\q \ G. ?M, ?V \ q\ \consistent (AxT \ Ax4) ?V\ \maximal (AxT \ Ax4) ?V\ using canonical_model unfolding list_all_def by fastforce+ - then have \?M\<^sub>S\<^sub>4, ?V\<^sub>S\<^sub>4 \ (\<^bold>\ p)\ \list_all (\p. ?M\<^sub>S\<^sub>4, ?V\<^sub>S\<^sub>4 \ p) G\ - using \consistent (AxT \ Ax4) ?V\ \maximal (AxT \ Ax4) ?V\ mcs\<^sub>S\<^sub>4_equiv - unfolding list_all_def by blast+ - moreover have \reflexive ?M\<^sub>S\<^sub>4\ \transitive ?M\<^sub>S\<^sub>4\ - using mcs\<^sub>S\<^sub>4_reflexive mcs\<^sub>S\<^sub>4_transitive by fast+ - ultimately have \?M\<^sub>S\<^sub>4, ?V\<^sub>S\<^sub>4 \ p\ - using valid by blast + moreover have \reflexive ?M\ \transitive ?M\ + by (simp_all add: mcs\<^sub>T_reflexive mcs\<^sub>K\<^sub>4_transitive) + ultimately have \?M, ?V \ p\ + using valid by auto then show False - using \?M\<^sub>S\<^sub>4, ?V\<^sub>S\<^sub>4 \ (\<^bold>\ p)\ by simp + using \?M, ?V \ (\<^bold>\ p)\ by simp qed lemma completeness\<^sub>S\<^sub>4: - assumes \\(M :: ('i, ('i :: countable) mcs\<^sub>S\<^sub>4) kripke) w. + assumes \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. reflexive M \ transitive M \ M, w \ p\ shows \\\<^sub>S\<^sub>4 p\ - using assms imply_completeness_S4[where G=\[]\] by auto + using assms imply_completeness_S4[where G=\{}\] by auto -abbreviation \valid\<^sub>S\<^sub>4 p \ \(M :: (nat, nat mcs\<^sub>S\<^sub>4) kripke) w. +abbreviation \valid\<^sub>S\<^sub>4 p \ \(M :: (nat, nat fm set) kripke). \w \ \ M. reflexive M \ transitive M \ M, w \ p\ theorem main\<^sub>S\<^sub>4: \valid\<^sub>S\<^sub>4 p \ \\<^sub>S\<^sub>4 p\ using soundness\<^sub>S\<^sub>4 completeness\<^sub>S\<^sub>4 by fast corollary - assumes \reflexive M\ \transitive M\ + assumes \reflexive M\ \transitive M\ \w \ \ M\ shows \valid\<^sub>S\<^sub>4 p \ M, w \ p\ using assms soundness\<^sub>S\<^sub>4 completeness\<^sub>S\<^sub>4 by fast section \System S5\ abbreviation SystemS5 :: \'i fm \ bool\ ("\\<^sub>S\<^sub>5 _" [50] 50) where \\\<^sub>S\<^sub>5 p \ AxT \ AxB \ Ax4 \ p\ abbreviation AxTB4 :: \'i fm \ bool\ where \AxTB4 \ AxT \ AxB \ Ax4\ -lemma soundness_AxTB4: \AxTB4 p \ equivalence M \ M, w \ p\ +lemma soundness_AxTB4: \AxTB4 p \ equivalence M \ w \ \ M \ M, w \ p\ using soundness_AxT soundness_AxB soundness_Ax4 by fast -lemma soundness\<^sub>S\<^sub>5: \\\<^sub>S\<^sub>5 p \ equivalence M \ M, w \ p\ +lemma soundness\<^sub>S\<^sub>5: \\\<^sub>S\<^sub>5 p \ equivalence M \ w \ \ M \ M, w \ p\ using soundness soundness_AxTB4 . -typedef 'i mcs\<^sub>S\<^sub>5 = - \{V :: ('i :: countable) fm set. consistent AxTB4 V \ maximal AxTB4 V}\ - using ex_mcs[where A=AxTB4 and P=equivalence] soundness_AxTB4 ex_equiv_model by fast - -abbreviation \pi\<^sub>S\<^sub>5 \ pi o Rep_mcs\<^sub>S\<^sub>5\ -abbreviation \reach\<^sub>S\<^sub>5 i V \ Abs_mcs\<^sub>S\<^sub>5 ` (reach AxTB4 i (Rep_mcs\<^sub>S\<^sub>5 V))\ - -lemma mcs\<^sub>S\<^sub>5_equiv: - assumes \consistent AxTB4 V\ \maximal AxTB4 V\ - shows \(Kripke pi (reach AxTB4), V \ p) = (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5, Abs_mcs\<^sub>S\<^sub>5 V \ p)\ - using assms by (induct p arbitrary: V) (simp_all add: Abs_mcs\<^sub>S\<^sub>5_inverse) - -lemma mcs\<^sub>S\<^sub>5_reflexive: \reflexive (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5)\ - unfolding reflexive_def -proof (intro allI) - fix i :: \'i :: countable\ and V :: \'i mcs\<^sub>S\<^sub>5\ - let ?V = \Rep_mcs\<^sub>S\<^sub>5 V\ - have \consistent AxTB4 ?V\ \maximal AxTB4 ?V\ - using Rep_mcs\<^sub>S\<^sub>5 by blast+ - then have \?V \ reach AxTB4 i ?V\ - using AxT_reflexive[where V=\?V\] by presburger - then have \V \ reach\<^sub>S\<^sub>5 i V\ - using Rep_mcs\<^sub>S\<^sub>5_inverse by (metis (no_types, lifting) image_iff) - then show \V \ \ (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5) i V\ - by simp -qed - -lemma mcs\<^sub>S\<^sub>5_symmetric: \symmetric (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5)\ - unfolding symmetric_def -proof (intro allI) - fix i :: \'i :: countable\ and V W :: \'i mcs\<^sub>S\<^sub>5\ - let ?V = \Rep_mcs\<^sub>S\<^sub>5 V\ and ?W = \Rep_mcs\<^sub>S\<^sub>5 W\ +lemma imply_completeness_S5: + assumes valid: \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. + equivalence M \ (\q \ G. M, w \ q) \ M, w \ p\ + shows \\qs. set qs \ G \ (AxTB4 \ imply qs p)\ +proof (rule ccontr) + assume \\qs. set qs \ G \ AxTB4 \ imply qs p\ + then have *: \\qs. set qs \ G \ \ AxTB4 \ imply ((\<^bold>\ p) # qs) \<^bold>\\ + using K_Boole by blast - have \consistent AxTB4 ?V\ \maximal AxTB4 ?V\ - \consistent AxTB4 ?W\ \maximal AxTB4 ?W\ - using Rep_mcs\<^sub>S\<^sub>5 by blast+ - then have \?V \ reach AxTB4 i ?W \ ?W \ reach AxTB4 i ?V\ - using AxB_symmetric[where V=\?V\ and W=\?W\] by presburger - then have \V \ reach\<^sub>S\<^sub>5 i W \ W \ reach\<^sub>S\<^sub>5 i V\ - unfolding image_iff using Rep_mcs\<^sub>S\<^sub>5_cases Rep_mcs\<^sub>S\<^sub>5_inverse - by (metis (mono_tags, lifting) mem_Collect_eq) - then show \(W \ \ (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5) i V) = (V \ \ (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5) i W)\ - by simp -qed - -lemma mcs\<^sub>S\<^sub>5_transitive: \transitive (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5)\ - unfolding transitive_def -proof safe - fix i :: \'i :: countable\ and V W U :: \'i mcs\<^sub>S\<^sub>5\ - let ?V = \Rep_mcs\<^sub>S\<^sub>5 V\ and ?W = \Rep_mcs\<^sub>S\<^sub>5 W\ and ?U = \Rep_mcs\<^sub>S\<^sub>5 U\ - - assume \W \ \ (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5) i V\ \U \ \ (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5) i W\ - then have \?W \ reach AxTB4 i ?V\ \?U \ reach AxTB4 i ?W\ - using Abs_mcs\<^sub>S\<^sub>5_inverse by auto - moreover have \consistent AxTB4 ?V\ \maximal AxTB4 ?V\ - using Rep_mcs\<^sub>S\<^sub>5 by blast+ - ultimately have \?U \ reach AxTB4 i ?V\ - using Ax4_transitive[where V=\?V\] by presburger - then have \U \ reach\<^sub>S\<^sub>5 i V\ - using Rep_mcs\<^sub>S\<^sub>5_inverse by (metis (no_types, lifting) image_iff) - then show \U \ \ (Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5) i V\ - by simp -qed - -lemma imply_completeness_S5: - assumes valid: \\(M :: ('i :: countable, 'i mcs\<^sub>S\<^sub>5) kripke) w. - equivalence M \ list_all (\q. M, w \ q) G \ M, w \ p\ - shows \AxTB4 \ imply G p\ -proof (rule K_Boole, rule ccontr) - assume \\ AxTB4 \ imply ((\<^bold>\ p) # G) \<^bold>\\ - - let ?S = \set ((\<^bold>\ p) # G)\ + let ?S = \{\<^bold>\ p} \ G\ let ?V = \Extend AxTB4 ?S from_nat\ - let ?M = \Kripke pi (reach AxTB4)\ - let ?V\<^sub>S\<^sub>5 = \Abs_mcs\<^sub>S\<^sub>5 ?V\ - let ?M\<^sub>S\<^sub>5 = \Kripke pi\<^sub>S\<^sub>5 reach\<^sub>S\<^sub>5\ + let ?M = \Kripke (mcss AxTB4) pi (reach AxTB4)\ have \consistent AxTB4 ?S\ - unfolding consistent_def using \\ AxTB4 \ imply ((\<^bold>\ p) # G) \<^bold>\\ K_imply_weaken by blast + using * by (metis (no_types, lifting) K_imply_Cons consistent_def inconsistent_subset) then have - \?M, ?V \ (\<^bold>\ p)\ \list_all (\p. ?M, ?V \ p) G\ + \?M, ?V \ (\<^bold>\ p)\ \\q \ G. ?M, ?V \ q\ \consistent AxTB4 ?V\ \maximal AxTB4 ?V\ using canonical_model unfolding list_all_def by fastforce+ - then have \?M\<^sub>S\<^sub>5, ?V\<^sub>S\<^sub>5 \ (\<^bold>\ p)\ \list_all (\p. ?M\<^sub>S\<^sub>5, ?V\<^sub>S\<^sub>5 \ p) G\ - using \consistent AxTB4 ?V\ \maximal AxTB4 ?V\ mcs\<^sub>S\<^sub>5_equiv - unfolding list_all_def by blast+ - moreover have \equivalence ?M\<^sub>S\<^sub>5\ - using mcs\<^sub>S\<^sub>5_reflexive mcs\<^sub>S\<^sub>5_symmetric mcs\<^sub>S\<^sub>5_transitive by fast - ultimately have \?M\<^sub>S\<^sub>5, ?V\<^sub>S\<^sub>5 \ p\ - using valid by blast + moreover have \equivalence ?M\ + by (simp add: mcs\<^sub>T_reflexive mcs\<^sub>K\<^sub>B_symmetric mcs\<^sub>K\<^sub>4_transitive) + ultimately have \?M, ?V \ p\ + using valid by auto then show False - using \?M\<^sub>S\<^sub>5, ?V\<^sub>S\<^sub>5 \ (\<^bold>\ p)\ by simp + using \?M, ?V \ (\<^bold>\ p)\ by simp qed lemma completeness\<^sub>S\<^sub>5: - assumes \\(M :: ('i, ('i :: countable) mcs\<^sub>S\<^sub>5) kripke) w. equivalence M \ M, w \ p\ + assumes \\(M :: ('i :: countable, 'i fm set) kripke). \w \ \ M. equivalence M \ M, w \ p\ shows \\\<^sub>S\<^sub>5 p\ - using assms imply_completeness_S5[where G=\[]\] by auto + using assms imply_completeness_S5[where G=\{}\] by auto -abbreviation \valid\<^sub>S\<^sub>5 p \ \(M :: (nat, nat mcs\<^sub>S\<^sub>5) kripke) w. equivalence M \ M, w \ p\ +abbreviation \valid\<^sub>S\<^sub>5 p \ \(M :: (nat, nat fm set) kripke). \w \ \ M. equivalence M \ M, w \ p\ theorem main\<^sub>S\<^sub>5: \valid\<^sub>S\<^sub>5 p \ \\<^sub>S\<^sub>5 p\ using soundness\<^sub>S\<^sub>5 completeness\<^sub>S\<^sub>5 by fast corollary - assumes \equivalence M\ + assumes \equivalence M\ \w \ \ M\ shows \valid\<^sub>S\<^sub>5 p \ M, w \ p\ using assms soundness\<^sub>S\<^sub>5 completeness\<^sub>S\<^sub>5 by fast subsection \Traditional formulation\ inductive SystemS5' :: \'i fm \ bool\ ("\\<^sub>S\<^sub>5'' _" [50] 50) where A1': \tautology p \ \\<^sub>S\<^sub>5' p\ | A2': \\\<^sub>S\<^sub>5' (K i (p \<^bold>\ q) \<^bold>\ K i p \<^bold>\ K i q)\ | AT': \\\<^sub>S\<^sub>5' (K i p \<^bold>\ p)\ | A5': \\\<^sub>S\<^sub>5' (\<^bold>\ K i p \<^bold>\ K i (\<^bold>\ K i p))\ | R1': \\\<^sub>S\<^sub>5' p \ \\<^sub>S\<^sub>5' (p \<^bold>\ q) \ \\<^sub>S\<^sub>5' q\ | R2': \\\<^sub>S\<^sub>5' p \ \\<^sub>S\<^sub>5' K i p\ lemma S5'_trans: \\\<^sub>S\<^sub>5' ((p \<^bold>\ q) \<^bold>\ (q \<^bold>\ r) \<^bold>\ p \<^bold>\ r)\ by (simp add: A1') lemma S5'_L: \\\<^sub>S\<^sub>5' (p \<^bold>\ L i p)\ proof - have \\\<^sub>S\<^sub>5' (K i (\<^bold>\ p) \<^bold>\ \<^bold>\ p)\ using AT' by fast moreover have \\\<^sub>S\<^sub>5' ((P \<^bold>\ \<^bold>\ Q) \<^bold>\ Q \<^bold>\ \<^bold>\ P)\ for P Q :: \'i fm\ using A1' by force ultimately show ?thesis using R1' by blast qed lemma S5'_B: \\\<^sub>S\<^sub>5' (p \<^bold>\ K i (L i p))\ using A5' S5'_L R1' S5'_trans by metis lemma S5'_KL: \\\<^sub>S\<^sub>5' (K i p \<^bold>\ L i p)\ by (meson AT' R1' S5'_L S5'_trans) lemma S5'_map_K: assumes \\\<^sub>S\<^sub>5' (p \<^bold>\ q)\ shows \\\<^sub>S\<^sub>5' (K i p \<^bold>\ K i q)\ proof - note \\\<^sub>S\<^sub>5' (p \<^bold>\ q)\ then have \\\<^sub>S\<^sub>5' K i (p \<^bold>\ q)\ using R2' by fast moreover have \\\<^sub>S\<^sub>5' (K i (p \<^bold>\ q) \<^bold>\ K i p \<^bold>\ K i q)\ using A2' by fast ultimately show ?thesis using R1' by fast qed lemma S5'_map_L: assumes \\\<^sub>S\<^sub>5' (p \<^bold>\ q)\ shows \\\<^sub>S\<^sub>5' (L i p \<^bold>\ L i q)\ using assms by (metis R1' S5'_map_K S5'_trans) lemma S5'_L_dual: \\\<^sub>S\<^sub>5' (\<^bold>\ L i (\<^bold>\ p) \<^bold>\ K i p)\ proof - have \\\<^sub>S\<^sub>5' (K i p \<^bold>\ K i p)\ \\\<^sub>S\<^sub>5' (\<^bold>\ \<^bold>\ p \<^bold>\ p)\ by (simp_all add: A1') then have \\\<^sub>S\<^sub>5' (K i (\<^bold>\ \<^bold>\ p) \<^bold>\ K i p)\ by (simp add: S5'_map_K) moreover have \\\<^sub>S\<^sub>5' ((P \<^bold>\ Q) \<^bold>\ (\<^bold>\ \<^bold>\ P \<^bold>\ Q))\ for P Q :: \'i fm\ by (simp add: A1') ultimately show \\\<^sub>S\<^sub>5' (\<^bold>\ \<^bold>\ K i (\<^bold>\ \<^bold>\ p) \<^bold>\ K i p)\ using R1' by blast qed lemma S5'_4: \\\<^sub>S\<^sub>5' (K i p \<^bold>\ K i (K i p))\ proof - have \\\<^sub>S\<^sub>5' (L i (\<^bold>\ p) \<^bold>\ K i (L i (\<^bold>\ p)))\ using A5' by fast moreover have \\\<^sub>S\<^sub>5' ((P \<^bold>\ Q) \<^bold>\ \<^bold>\ Q \<^bold>\ \<^bold>\ P)\ for P Q :: \'i fm\ using A1' by force ultimately have \\\<^sub>S\<^sub>5' (\<^bold>\ K i (L i (\<^bold>\ p)) \<^bold>\ \<^bold>\ L i (\<^bold>\ p))\ using R1' by fast then have \\\<^sub>S\<^sub>5' (L i (K i (\<^bold>\ \<^bold>\ p)) \<^bold>\ \<^bold>\ L i (\<^bold>\ p))\ by blast moreover have \\\<^sub>S\<^sub>5' (p \<^bold>\ \<^bold>\ \<^bold>\ p)\ by (simp add: A1') ultimately have \\\<^sub>S\<^sub>5' (L i (K i p) \<^bold>\ \<^bold>\ L i (\<^bold>\ p))\ by (metis (no_types, opaque_lifting) R1' S5'_map_K S5'_trans) then have \\\<^sub>S\<^sub>5' (L i (K i p) \<^bold>\ K i p)\ by (meson S5'_L_dual R1' S5'_trans) then show ?thesis by (metis A2' R1' R2' S5'_B S5'_trans) qed lemma S5_S5': \\\<^sub>S\<^sub>5 p \ \\<^sub>S\<^sub>5' p\ proof (induct p rule: AK.induct) case (A2 i p q) have \\\<^sub>S\<^sub>5' (K i (p \<^bold>\ q) \<^bold>\ K i p \<^bold>\ K i q)\ using A2' . moreover have \\\<^sub>S\<^sub>5' ((P \<^bold>\ Q \<^bold>\ R) \<^bold>\ (Q \<^bold>\ P \<^bold>\ R))\ for P Q R :: \'i fm\ by (simp add: A1') ultimately show ?case using R1' by blast next case (Ax p) then show ?case using AT' S5'_B S5'_4 by (metis Ax4.cases AxB.cases AxT.cases) qed (meson SystemS5'.intros)+ lemma S5'_S5: fixes p :: \('i :: countable) fm\ shows \\\<^sub>S\<^sub>5' p \ \\<^sub>S\<^sub>5 p\ proof (induct p rule: SystemS5'.induct) case (AT' i p) then show ?case by (simp add: Ax AxT.intros) next case (A5' i p) then show ?case using completeness\<^sub>S\<^sub>5 neg_introspection by fast qed (meson AK.intros K_A2')+ theorem main\<^sub>S\<^sub>5': \valid\<^sub>S\<^sub>5 p \ \\<^sub>S\<^sub>5' p\ using main\<^sub>S\<^sub>5 S5_S5' S5'_S5 by blast section \Acknowledgements\ text \ The formalization is inspired by Berghofer's formalization of Henkin-style completeness. \<^item> Stefan Berghofer: First-Order Logic According to Fitting. \<^url>\https://www.isa-afp.org/entries/FOL-Fitting.shtml\ \ end