diff --git a/thys/FOL_Axiomatic/FOL_Axiomatic.thy b/thys/FOL_Axiomatic/FOL_Axiomatic.thy --- a/thys/FOL_Axiomatic/FOL_Axiomatic.thy +++ b/thys/FOL_Axiomatic/FOL_Axiomatic.thy @@ -1,802 +1,802 @@ (* File: FOL_Axiomatic.thy Author: Asta Halkjær From This work is a formalization of the soundness and completeness of an axiomatic system for first-order logic. The proof system is based on System Q1 by Smullyan and the completeness proof follows his textbook "First-Order Logic" (Springer-Verlag 1968). The completeness proof is in the Henkin style where a consistent set is extended to a maximal consistent set using Lindenbaum's construction and Henkin witnesses are added during the construction to ensure saturation as well. The resulting set is a Hintikka set which, by the model existence theorem, is satisfiable in the Herbrand universe. *) theory FOL_Axiomatic imports "HOL-Library.Countable" begin section \Syntax\ datatype (params_tm: 'f) tm = Var nat (\\<^bold>#\) | Fun 'f \'f tm list\ (\\<^bold>\\) abbreviation Const (\\<^bold>\\) where \\<^bold>\a \ \<^bold>\a []\ datatype (params_fm: 'f, 'p) fm = Falsity (\\<^bold>\\) | Pre 'p \'f tm list\ (\\<^bold>\\) | Imp \('f, 'p) fm\ \('f, 'p) fm\ (infixr \\<^bold>\\ 25) | Uni \('f, 'p) fm\ (\\<^bold>\\) abbreviation Neg (\\<^bold>\ _\ [40] 40) where \\<^bold>\ p \ p \<^bold>\ \<^bold>\\ term \\<^bold>\(\<^bold>\ \<^bold>\ \<^bold>\''P'' [\<^bold>\''f'' [\<^bold>#0]])\ section \Semantics\ definition shift :: \(nat \ 'a) \ nat \ 'a \ nat \ 'a\ (\_\_:_\\ [90, 0, 0] 91) where \E\n:x\ = (\m. if m < n then E m else if m = n then x else E (m-1))\ primrec semantics_tm (\\_, _\\) where - \\E, F\(\<^bold>#n) = E n\ -| \\E, F\(\<^bold>\f ts) = F f (map \E, F\ ts)\ + \\E, F\ (\<^bold>#n) = E n\ +| \\E, F\ (\<^bold>\f ts) = F f (map \E, F\ ts)\ primrec semantics_fm (\\_, _, _\\) where \\_, _, _\ \<^bold>\ = False\ | \\E, F, G\ (\<^bold>\P ts) = G P (map \E, F\ ts)\ | \\E, F, G\ (p \<^bold>\ q) = (\E, F, G\ p \ \E, F, G\ q)\ | \\E, F, G\ (\<^bold>\p) = (\x. \E\0:x\, F, G\ p)\ proposition \\E, F, G\ (\<^bold>\(\<^bold>\P [\<^bold># 0]) \<^bold>\ \<^bold>\P [\<^bold>\a])\ by (simp add: shift_def) section \Operations\ subsection \Shift\ lemma shift_eq [simp]: \n = m \ (E\n:x\) m = x\ by (simp add: shift_def) lemma shift_gt [simp]: \m < n \ (E\n:x\) m = E m\ by (simp add: shift_def) lemma shift_lt [simp]: \n < m \ (E\n:x\) m = E (m-1)\ by (simp add: shift_def) lemma shift_commute [simp]: \E\n:y\\0:x\ = E\0:x\\n+1:y\\ proof fix m show \(E\n:y\\0:x\) m = (E\0:x\\n+1:y\) m\ unfolding shift_def by (cases m) simp_all qed subsection \Parameters\ abbreviation \params S \ \p \ S. params_fm p\ lemma upd_params_tm [simp]: \f \ params_tm t \ \E, F(f := x)\ t = \E, F\ t\ by (induct t) (auto cong: map_cong) lemma upd_params_fm [simp]: \f \ params_fm p \ \E, F(f := x), G\ p = \E, F, G\ p\ by (induct p arbitrary: E) (auto cong: map_cong) lemma finite_params_tm [simp]: \finite (params_tm t)\ by (induct t) simp_all lemma finite_params_fm [simp]: \finite (params_fm p)\ by (induct p) simp_all subsection \Instantiation\ primrec lift_tm (\\<^bold>\\) where \\<^bold>\(\<^bold>#n) = \<^bold>#(n+1)\ | \\<^bold>\(\<^bold>\f ts) = \<^bold>\f (map \<^bold>\ ts)\ primrec inst_tm (\_'\_'/_'\\ [90, 0, 0] 91) where \(\<^bold>#n)\s/m\ = (if n < m then \<^bold>#n else if n = m then s else \<^bold>#(n-1))\ | \(\<^bold>\f ts)\s/m\ = \<^bold>\f (map (\t. t\s/m\) ts)\ primrec inst_fm (\_'\_'/_'\\ [90, 0, 0] 91) where \\<^bold>\\_/_\ = \<^bold>\\ | \(\<^bold>\P ts)\s/m\ = \<^bold>\P (map (\t. t\s/m\) ts)\ | \(p \<^bold>\ q)\s/m\ = (p\s/m\ \<^bold>\ q\s/m\)\ | \(\<^bold>\p)\s/m\ = \<^bold>\(p\\<^bold>\s/m+1\)\ lemma lift_lemma [simp]: \\E\0:x\, F\ (\<^bold>\t) = \E, F\ t\ by (induct t) (auto cong: map_cong) lemma inst_tm_semantics [simp]: \\E, F\ (t\s/m\) = \E\m:\E, F\ s\, F\ t\ by (induct t) (auto cong: map_cong) lemma inst_fm_semantics [simp]: \\E, F, G\ (p\t/m\) = \E\m:\E, F\ t\, F, G\ p\ by (induct p arbitrary: E m t) (auto cong: map_cong) subsection \Size\ text \The built-in \size\ is not invariant under substitution.\ primrec size_fm where \size_fm \<^bold>\ = 1\ | \size_fm (\<^bold>\_ _) = 1\ | \size_fm (p \<^bold>\ q) = 1 + size_fm p + size_fm q\ | \size_fm (\<^bold>\p) = 1 + size_fm p\ lemma size_inst_fm [simp]: \size_fm (p\t/m\) = size_fm p\ by (induct p arbitrary: m t) auto section \Propositional Semantics\ primrec boolean where \boolean _ _ \<^bold>\ = False\ | \boolean G _ (\<^bold>\P ts) = G P ts\ | \boolean G A (p \<^bold>\ q) = (boolean G A p \ boolean G A q)\ | \boolean _ A (\<^bold>\p) = A (\<^bold>\p)\ abbreviation \tautology p \ \G A. boolean G A p\ proposition \tautology (\<^bold>\(\<^bold>\P [\<^bold>#0]) \<^bold>\ \<^bold>\(\<^bold>\P [\<^bold>#0]))\ by simp lemma boolean_semantics: \boolean (\a. G a \ map \E, F\) \E, F, G\ = \E, F, G\\ proof fix p show \boolean (\a. G a \ map \E, F\) \E, F, G\ p = \E, F, G\ p\ by (induct p) simp_all qed lemma tautology: \tautology p \ \E, F, G\ p\ using boolean_semantics by metis proposition \\p. (\E F G. \E, F, G\ p) \ \ tautology p\ by (metis boolean.simps(4) fm.simps(36) semantics_fm.simps(1,3,4)) section \Calculus\ text \Adapted from System Q1 by Smullyan in First-Order Logic (1968)\ inductive Axiomatic (\\ _\ [20] 20) where TA: \tautology p \ \ p\ | IA: \\ \<^bold>\p \<^bold>\ p\t/0\\ | MP: \\ p \<^bold>\ q \ \ p \ \ q\ | GR: \\ q \<^bold>\ p\\<^bold>\a/0\ \ a \ params {p, q} \ \ q \<^bold>\ \<^bold>\p\ lemmas TA[simp] MP[trans, dest] GR[intro] text \We simulate assumptions on the lhs of \\\ with a chain of implications on the rhs.\ primrec imply (infixr \\<^bold>\\ 26) where \([] \<^bold>\ q) = q\ | \(p # ps \<^bold>\ q) = (p \<^bold>\ ps \<^bold>\ q)\ abbreviation Axiomatic_assms (\_ \ _\ [20, 20] 20) where \ps \ q \ \ ps \<^bold>\ q\ section \Soundness\ theorem soundness: \\ p \ \E, F, G\ p\ proof (induct p arbitrary: F rule: Axiomatic.induct) case (GR q p a) moreover from this have \\E, F(a := x), G\ (q \<^bold>\ p\\<^bold>\a/0\)\ for x by blast ultimately show ?case by fastforce qed (auto simp: tautology) corollary \\ (\ \<^bold>\)\ using soundness by fastforce section \Derived Rules\ lemma AS: \\ (p \<^bold>\ q \<^bold>\ r) \<^bold>\ (p \<^bold>\ q) \<^bold>\ p \<^bold>\ r\ by auto lemma AK: \\ q \<^bold>\ p \<^bold>\ q\ by auto lemma Neg: \\ \<^bold>\ \<^bold>\ p \<^bold>\ p\ by auto lemma contraposition: \\ (p \<^bold>\ q) \<^bold>\ \<^bold>\ q \<^bold>\ \<^bold>\ p\ \\ (\<^bold>\ q \<^bold>\ \<^bold>\ p) \<^bold>\ p \<^bold>\ q\ by (auto intro: TA) lemma GR': \\ \<^bold>\ p\\<^bold>\a/0\ \<^bold>\ q \ a \ params {p, q} \ \ \<^bold>\ \<^bold>\p \<^bold>\ q\ proof - assume *: \\ \<^bold>\ p\\<^bold>\a/0\ \<^bold>\ q\ and a: \a \ params {p, q}\ have \\ \<^bold>\ q \<^bold>\ \<^bold>\ \<^bold>\ p\\<^bold>\a/0\\ using * contraposition(1) by fast then have \\ \<^bold>\ q \<^bold>\ p\\<^bold>\a/0\\ by (meson AK AS MP Neg) then have \\ \<^bold>\ q \<^bold>\ \<^bold>\p\ using a by auto then have \\ \<^bold>\ \<^bold>\p \<^bold>\ \<^bold>\ \<^bold>\ q\ using contraposition(1) by fast then show ?thesis by (meson AK AS MP Neg) qed lemma Imp3: \\ (p \<^bold>\ q \<^bold>\ r) \<^bold>\ ((s \<^bold>\ p) \<^bold>\ (s \<^bold>\ q) \<^bold>\ s \<^bold>\ r)\ by auto lemma imply_ImpE: \\ ps \<^bold>\ p \<^bold>\ ps \<^bold>\ (p \<^bold>\ q) \<^bold>\ ps \<^bold>\ q\ by (induct ps) (auto intro: Imp3 MP) lemma MP' [trans, dest]: \ps \ p \<^bold>\ q \ ps \ p \ ps \ q\ using imply_ImpE by fast lemma imply_Cons [intro]: \ps \ q \ p # ps \ q\ by (auto intro: MP AK) lemma imply_head [intro]: \p # ps \ p\ proof (induct ps) case (Cons q ps) then show ?case by (metis AK MP' imply.simps(1-2)) qed auto lemma imply_lift_Imp [simp]: assumes \\ p \<^bold>\ q\ shows \\ p \<^bold>\ ps \<^bold>\ q\ using assms MP MP' imply_head by (metis imply.simps(2)) lemma add_imply [simp]: \\ q \ ps \ q\ using MP imply_head by (auto simp del: TA) lemma imply_mem [simp]: \p \ set ps \ ps \ p\ proof (induct ps) case (Cons q ps) then show ?case by (metis imply_Cons imply_head set_ConsD) qed simp lemma deduct1: \ps \ p \<^bold>\ q \ p # ps \ q\ by (meson MP' imply_Cons imply_head) lemma imply_append [iff]: \(ps @ qs \<^bold>\ r) = (ps \<^bold>\ qs \<^bold>\ r)\ by (induct ps) simp_all lemma imply_swap_append: \ps @ qs \ r \ qs @ ps \ r\ proof (induct qs arbitrary: ps) case (Cons q qs) then show ?case by (metis deduct1 imply.simps(2) imply_append) qed simp lemma deduct2: \p # ps \ q \ ps \ p \<^bold>\ q\ by (metis imply.simps(1-2) imply_append imply_swap_append) lemmas deduct [iff] = deduct1 deduct2 lemma cut [trans, dest]: \p # ps \ r \ q # ps \ p \ q # ps \ r\ by (meson MP' deduct(2) imply_Cons) lemma Boole: \(\<^bold>\ p) # ps \ \<^bold>\ \ ps \ p\ by (meson MP' Neg add_imply deduct(2)) lemma imply_weaken: \ps \ q \ set ps \ set ps' \ ps' \ q\ proof (induct ps arbitrary: q) case (Cons p ps) then show ?case by (metis MP' deduct(2) imply_mem insert_subset list.simps(15)) qed simp section \Consistent\ definition \consistent S \ \S'. set S' \ S \ (S' \ \<^bold>\)\ lemma UN_finite_bound: assumes \finite A\ and \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 split_list: assumes \x \ set A\ shows \set (x # removeAll x A) = set A \ x \ set (removeAll x A)\ using assms by auto lemma imply_params_fm: \params_fm (ps \<^bold>\ q) = params_fm q \ (\p \ set ps. params_fm p)\ by (induct ps) auto lemma inconsistent_fm: assumes \consistent S\ and \\ consistent ({p} \ S)\ obtains S' where \set S' \ S\ and \p # S' \ \<^bold>\\ proof - obtain S' where S': \set S' \ {p} \ S\ \p \ set S'\ \S' \ \<^bold>\\ using assms unfolding consistent_def by blast then obtain S'' where S'': \set (p # S'') = set S'\ \p \ set S''\ using split_list by metis then have \p # S'' \ \<^bold>\\ using \S' \ \<^bold>\\ imply_weaken by blast then show ?thesis using that S'' S'(1) by (metis Diff_insert_absorb Diff_subset_conv list.simps(15)) qed lemma consistent_add_witness: assumes \consistent S\ and \(\<^bold>\ \<^bold>\p) \ S\ and \a \ params S\ shows \consistent ({\<^bold>\ p\\<^bold>\a/0\} \ S)\ unfolding consistent_def proof assume \\S'. set S' \ {\<^bold>\ p\\<^bold>\a/0\} \ S \ (S' \ \<^bold>\)\ then obtain S' where \set S' \ S\ and \(\<^bold>\ p\\<^bold>\a/0\) # S' \ \<^bold>\\ using assms inconsistent_fm unfolding consistent_def by metis then have \\ \<^bold>\ p\\<^bold>\a/0\ \<^bold>\ S' \<^bold>\ \<^bold>\\ by simp moreover have \a \ params_fm p\ using assms(2-3) by auto moreover have \\p \ set S'. a \ params_fm p\ using \set S' \ S\ assms(3) by auto then have \a \ params_fm (S' \<^bold>\ \<^bold>\)\ by (simp add: imply_params_fm) ultimately have \\ \<^bold>\ \<^bold>\p \<^bold>\ S' \<^bold>\ \<^bold>\\ using GR' by fast then have \(\<^bold>\ \<^bold>\p) # S' \ \<^bold>\\ by simp moreover have \set ((\<^bold>\ \<^bold>\p) # S') \ S\ using \set S' \ S\ assms(2) by simp ultimately show False using assms(1) unfolding consistent_def by blast qed lemma consistent_add_instance: assumes \consistent S\ and \\<^bold>\p \ S\ shows \consistent ({p\t/0\} \ S)\ unfolding consistent_def proof assume \\S'. set S' \ {p\t/0\} \ S \ (S' \ \<^bold>\)\ then obtain S' where \set S' \ S\ and \p\t/0\ # S' \ \<^bold>\\ using assms inconsistent_fm unfolding consistent_def by blast moreover have \\ \<^bold>\p \<^bold>\ p\t/0\\ using IA by blast ultimately have \\<^bold>\p # S' \ \<^bold>\\ by (meson add_imply cut deduct(1)) moreover have \set ((\<^bold>\p) # S') \ S\ using \set S' \ S\ assms(2) by simp ultimately show False using assms(1) unfolding consistent_def by blast qed section \Extension\ fun witness where \witness used (\<^bold>\ \<^bold>\p) = {\<^bold>\ p\\<^bold>\(SOME a. a \ used) []/0\}\ | \witness _ _ = {}\ primrec extend where \extend S f 0 = S\ | \extend S f (Suc n) = (let Sn = extend S f n in if consistent ({f n} \ Sn) then witness (params ({f n} \ Sn)) (f n) \ {f n} \ Sn else Sn)\ definition \Extend S f \ \n. extend S f n\ lemma Extend_subset: \S \ Extend S f\ unfolding Extend_def by (metis Union_upper extend.simps(1) range_eqI) lemma extend_bound: \(\n \ m. extend S f n) = extend S f m\ by (induct m) (simp_all add: atMost_Suc Let_def) lemma finite_params_witness [simp]: \finite (params (witness used p))\ by (induct used p rule: witness.induct) simp_all lemma finite_params_extend [simp]: \finite (params S) \ finite (params (extend S f n))\ by (induct n) (simp_all add: Let_def) lemma consistent_witness: fixes p :: \('f, 'p) fm\ assumes \consistent S\ and \p \ S\ and \params S \ used\ and \finite used\ and \infinite (UNIV :: 'f set)\ shows \consistent (witness used p \ S)\ using assms proof (induct used p rule: witness.induct) case (1 used p) moreover have \\a. a \ used\ using 1(4-) ex_new_if_finite by blast ultimately obtain a where a: \witness used (\<^bold>\ \<^bold>\p) = {\<^bold>\ p\\<^bold>\a/0\}\ and \a \ used\ by (metis someI_ex witness.simps(1)) then have \a \ params S\ using 1(3) by fast then show ?case using 1(1-2) a(1) consistent_add_witness by metis qed (auto simp: assms) lemma consistent_extend: fixes f :: \nat \ ('f, 'p) fm\ assumes \consistent S\ and \finite (params S)\ and \infinite (UNIV :: 'f set)\ shows \consistent (extend S f n)\ using assms proof (induct n) case (Suc n) then show ?case using consistent_witness[where S=\{f n} \ _\] by (auto simp: Let_def) qed simp lemma consistent_Extend: fixes f :: \nat \ ('f, 'p) fm\ assumes \consistent S\ and \finite (params S)\ and \infinite (UNIV :: 'f set)\ shows \consistent (Extend S f)\ proof (rule ccontr) assume \\ consistent (Extend S f)\ then obtain S' where \S' \ \<^bold>\\ and \set S' \ Extend S f\ unfolding consistent_def by blast then obtain m where \set S' \ (\n \ m. extend S f n)\ unfolding Extend_def using UN_finite_bound by (metis List.finite_set) then have \set S' \ extend S f m\ using extend_bound by blast moreover have \consistent (extend S f m)\ using assms consistent_extend by blast ultimately show False unfolding consistent_def using \S' \ \<^bold>\\ by blast qed section \Maximal\ definition \maximal S \ \p. p \ S \ \ consistent ({p} \ S)\ lemma maximal_exactly_one: assumes \consistent H\ and \maximal H\ shows \p \ H \ (\<^bold>\ p) \ H\ proof assume \p \ H\ show \(\<^bold>\ p) \ H\ proof assume \(\<^bold>\ p) \ H\ then have \set [p, \<^bold>\ p] \ H\ using \p \ H\ by simp moreover have \[p, \<^bold>\ p] \ \<^bold>\\ by blast ultimately show False using \consistent H\ unfolding consistent_def by blast qed next assume \(\<^bold>\ p) \ H\ then have \\ consistent ({\<^bold>\ p} \ H)\ using \maximal H\ unfolding maximal_def by blast then obtain S where \set S \ H\ \(\<^bold>\ p) # S \ \<^bold>\\ using \consistent H\ inconsistent_fm by blast then have \S \ p\ using Boole by blast have \consistent ({p} \ H)\ unfolding consistent_def proof assume \\S'. set S' \ {p} \ H \ (S' \ \<^bold>\)\ then obtain S' where \set S' \ H\ and \p # S' \ \<^bold>\\ using assms inconsistent_fm unfolding consistent_def by blast then have \S @ S' \ \<^bold>\\ using \S \ p\ by (metis MP' add_imply imply.simps(2) imply_append) moreover have \set (S @ S') \ H\ using \set S \ H\ \set S' \ H\ by simp ultimately show False using \consistent H\ unfolding consistent_def by blast qed then show \p \ H\ using \maximal H\ unfolding maximal_def by blast qed lemma maximal_Extend: assumes \surj f\ shows \maximal (Extend S f)\ proof (rule ccontr) assume \\ maximal (Extend S f)\ then obtain p where \p \ Extend S f\ and \consistent ({p} \ Extend S f)\ unfolding maximal_def using assms consistent_Extend by blast obtain k where k: \f k = p\ using \surj f\ unfolding surj_def by metis then have \p \ extend S f (Suc k)\ using \p \ Extend S f\ unfolding Extend_def by blast then have \\ consistent ({p} \ extend S f k)\ using k by (auto simp: Let_def) moreover have \{p} \ extend S f k \ {p} \ Extend S f\ unfolding Extend_def by blast ultimately have \\ consistent ({p} \ Extend S f)\ unfolding consistent_def by auto then show False using \consistent ({p} \ Extend S f)\ by blast qed section \Saturation\ definition \saturated S \ \p. (\<^bold>\ \<^bold>\p) \ S \ (\a. (\<^bold>\ p\\<^bold>\a/0\) \ S)\ lemma saturated_Extend: assumes \consistent (Extend S f)\ and \surj f\ shows \saturated (Extend S f)\ proof (rule ccontr) assume \\ saturated (Extend S f)\ then obtain p where p: \(\<^bold>\ \<^bold>\p) \ Extend S f\ \\a. (\<^bold>\ p\\<^bold>\a/0\) \ Extend S f\ unfolding saturated_def by blast obtain k where k: \f k = (\<^bold>\ \<^bold>\p)\ using \surj f\ unfolding surj_def by metis have \extend S f k \ Extend S f\ unfolding Extend_def by auto then have \consistent ({\<^bold>\ \<^bold>\p} \ extend S f k)\ using assms(1) p(1) unfolding consistent_def by blast then have \\a. extend S f (Suc k) = {\<^bold>\ p\\<^bold>\a/0\} \ {\<^bold>\ \<^bold>\p} \ extend S f k\ using k by (auto simp: Let_def) moreover have \extend S f (Suc k) \ Extend S f\ unfolding Extend_def by blast ultimately show False using p(2) by blast qed section \Hintikka\ locale Hintikka = fixes H :: \('f, 'p) fm set\ assumes NoFalsity: \\<^bold>\ \ H\ and ImpP: \(p \<^bold>\ q) \ H \ p \ H \ q \ H\ and ImpN: \(p \<^bold>\ q) \ H \ p \ H \ q \ H\ and UniP: \\<^bold>\p \ H \ \t. p\t/0\ \ H\ and UniN: \\<^bold>\p \ H \ \a. p\\<^bold>\a/0\ \ H\ subsection \Model Existence\ abbreviation (input) \sat H P ts \ Pre P ts \ H\ abbreviation hmodel (\\_\\) where \\H\ \ \\<^bold>#, \<^bold>\, sat H\\ lemma semantics_tm_id [simp]: \\\<^bold>#, \<^bold>\\ t = t\ by (induct t) (auto cong: map_cong) lemma semantics_tm_id_map [simp]: \map \\<^bold>#, \<^bold>\\ ts = ts\ by (auto cong: map_cong) theorem Hintikka_model: assumes \Hintikka H\ shows \p \ H \ \H\ p\ proof (induct p rule: wf_induct[where r=\measure size_fm\]) case 1 then show ?case .. next case (2 x) show \x \ H \ \H\ x\ proof (cases x; safe) case Falsity assume \\<^bold>\ \ H\ then have False using assms Hintikka.NoFalsity by fast then show \\H\ \<^bold>\\ .. next case Falsity assume \\H\ \<^bold>\\ then have False by simp then show \\<^bold>\ \ H\ .. next case (Pre P ts) assume \\<^bold>\P ts \ H\ then show \\H\ (\<^bold>\P ts)\ by simp next case (Pre P ts) assume \\H\ (\<^bold>\P ts)\ then have \sat H P ts\ by simp then show \\<^bold>\P ts \ H\ by simp next case (Imp p q) assume \(p \<^bold>\ q) \ H\ then have \p \ H \ q \ H\ using assms Hintikka.ImpP by blast then have \\ \H\ p \ \H\ q\ using 2 Imp by simp then show \\H\ (p \<^bold>\ q)\ by simp next case (Imp p q) assume \\H\ (p \<^bold>\ q)\ then have \\ \H\ p \ \H\ q\ by simp then have \p \ H \ q \ H\ using 2 Imp by simp then show \(p \<^bold>\ q) \ H\ using assms Hintikka.ImpN by blast next case (Uni p) assume \\<^bold>\p \ H\ then have \\t. p\t/0\ \ H\ using assms Hintikka.UniP by metis then have \\t. \H\ (p\t/0\)\ using 2 Uni by simp then show \\H\ (\<^bold>\p)\ by simp next case (Uni p) assume \\H\ (\<^bold>\p)\ then have \\t. \H\ (p\t/0\)\ by simp then have \\t. p\t/0\ \ H\ using 2 Uni by simp then show \\<^bold>\p \ H\ using assms Hintikka.UniN by blast qed qed subsection \Maximal Consistent Sets are Hintikka Sets\ lemma inconsistent_head: assumes \consistent S\ and \maximal S\ and \p \ S\ obtains S' where \set S' \ S\ and \p # S' \ \<^bold>\\ using assms inconsistent_fm unfolding consistent_def maximal_def by metis lemma inconsistent_parts [simp]: assumes \ps \ \<^bold>\\ and \set ps \ S\ shows \\ consistent S\ using assms unfolding consistent_def by blast lemma Hintikka_Extend: fixes H :: \('f, 'p) fm set\ assumes \consistent H\ and \maximal H\ and \saturated H\ and \infinite (UNIV :: 'f set)\ shows \Hintikka H\ proof show \\<^bold>\ \ H\ proof assume \\<^bold>\ \ H\ moreover have \[\<^bold>\] \ \<^bold>\\ by blast ultimately have \\ consistent H\ using inconsistent_parts[where ps=\[\<^bold>\]\] by simp then show False using \consistent H\ .. qed next fix p q assume *: \(p \<^bold>\ q) \ H\ show \p \ H \ q \ H\ proof safe assume \q \ H\ then obtain Hq' where Hq': \q # Hq' \ \<^bold>\\ \set Hq' \ H\ using assms inconsistent_head by metis assume \p \ H\ then have \(\<^bold>\ p) \ H\ using assms maximal_exactly_one by blast then obtain Hp' where Hp': \(\<^bold>\ p) # Hp' \ \<^bold>\\ \set Hp' \ H\ using assms inconsistent_head by metis let ?H' = \Hp' @ Hq'\ have H': \set ?H' = set Hp' \ set Hq'\ by simp then have \set Hp' \ set ?H'\ and \set Hq' \ set ?H'\ by blast+ then have \(\<^bold>\ p) # ?H' \ \<^bold>\\ and \q # ?H' \ \<^bold>\\ using Hp'(1) Hq'(1) deduct imply_weaken by metis+ then have \(p \<^bold>\ q) # ?H' \ \<^bold>\\ using Boole imply_Cons imply_head MP' cut by metis moreover have \set ((p \<^bold>\ q) # ?H') \ H\ using \q \ H\ *(1) H' Hp'(2) Hq'(2) by auto ultimately show False using assms unfolding consistent_def by blast qed next fix p q assume *: \(p \<^bold>\ q) \ H\ show \p \ H \ q \ H\ proof (safe, rule ccontr) assume \p \ H\ then obtain H' where S': \p # H' \ \<^bold>\\ \set H' \ H\ using assms inconsistent_head by metis moreover have \(\<^bold>\ (p \<^bold>\ q)) # H' \ p\ by auto ultimately have \(\<^bold>\ (p \<^bold>\ q)) # H' \ \<^bold>\\ by blast moreover have \set ((\<^bold>\ (p \<^bold>\ q)) # H') \ H\ using *(1) S'(2) assms maximal_exactly_one by auto ultimately show False using assms unfolding consistent_def by blast next assume \q \ H\ then have \(\<^bold>\ q) \ H\ using assms maximal_exactly_one by blast then obtain H' where H': \(\<^bold>\ q) # H' \ \<^bold>\\ \set H' \ H\ using assms inconsistent_head by metis moreover have \(\<^bold>\ (p \<^bold>\ q)) # H' \ \<^bold>\ q\ by auto ultimately have \(\<^bold>\ (p \<^bold>\ q)) # H' \ \<^bold>\\ by blast moreover have \set ((\<^bold>\ (p \<^bold>\ q)) # H') \ H\ using *(1) H'(2) assms maximal_exactly_one by auto ultimately show False using assms unfolding consistent_def by blast qed next fix p assume \\<^bold>\p \ H\ then show \\t. p\t/0\ \ H\ using assms consistent_add_instance unfolding maximal_def by blast next fix p assume \\<^bold>\p \ H\ then show \\a. p\\<^bold>\a/0\ \ H\ using assms maximal_exactly_one unfolding saturated_def by fast qed section \Countable Formulas\ instance tm :: (countable) countable by countable_datatype instance fm :: (countable, countable) countable by countable_datatype section \Completeness\ lemma imply_completeness: fixes p :: \('f :: countable, 'p :: countable) fm\ assumes \\(E :: _ \ 'f tm) F G. Ball X \E, F, G\ \ \E, F, G\ p\ and \finite (params X)\ and \infinite (UNIV :: 'f set)\ shows \\ps. set ps \ X \ (ps \ p)\ proof (rule ccontr) assume \\ps. set ps \ X \ (ps \ p)\ then have *: \\ps. set ps \ X \ ((\<^bold>\ p) # ps \ \<^bold>\)\ using Boole by blast let ?S = \{\<^bold>\ p} \ X\ let ?H = \Extend ?S from_nat\ have \consistent ?S\ using * by (metis consistent_def imply_Cons inconsistent_fm) moreover have \finite (params ?S)\ using assms by simp ultimately have \consistent ?H\ and \maximal ?H\ using assms(3) consistent_Extend maximal_Extend surj_from_nat by blast+ moreover from this have \saturated ?H\ using saturated_Extend by fastforce ultimately have \Hintikka ?H\ using assms(3) Hintikka_Extend by blast have \\?H\ p\ if \p \ ?S\ for p using that Extend_subset Hintikka_model \Hintikka ?H\ by blast then have \\?H\ (\<^bold>\ p)\ and \\q \ X. \?H\ q\ by blast+ moreover from this have \\?H\ p\ using assms(1) by blast ultimately show False by simp qed theorem completeness: fixes p :: \(nat, nat) fm\ assumes \\(E :: nat \ nat tm) F G. \E, F, G\ p\ shows \\ p\ using assms imply_completeness[where X=\{}\] by auto section \Main Result\ abbreviation valid :: \(nat, nat) fm \ bool\ where \valid p \ \(E :: nat \ nat tm) F G. \E, F, G\ p\ theorem main: \valid p \ (\ p)\ using completeness soundness by blast end