diff --git a/thys/SuperCalc/superposition.thy b/thys/SuperCalc/superposition.thy --- a/thys/SuperCalc/superposition.thy +++ b/thys/SuperCalc/superposition.thy @@ -1,7657 +1,7655 @@ theory superposition (* N. Peltier - http://membres-lig.imag.fr/peltier/ *) imports Main terms equational_clausal_logic well_founded_continued "HOL-Library.Multiset" multisets_continued begin section \Definition of the Superposition Calculus\ subsection \Extended Clauses\ text \An extended clause is a clause associated with a set of terms. The intended meaning is that the terms occurring in the attached set are assumed to be in normal form: any application of the superposition rule on these terms is therefore useless and can be blocked. Initially the set of irreducible terms attached to each clause is empty. At each inference step, new terms can be added or deleted from this set.\ datatype 'a eclause = Ecl "'a clause" "'a trm set" fun subst_ecl where "(subst_ecl (Ecl C S) \) = (Ecl (subst_cl C \) { s. (\t. (s = (subst t \) \ t \ S)) })" fun cl_ecl where "(cl_ecl (Ecl C X)) = C" fun trms_ecl where "(trms_ecl (Ecl C X)) = X" definition renaming_cl where "renaming_cl C D = (\ \. (renaming \ (vars_of_cl (cl_ecl C))) \ D = (subst_ecl C \))" definition closed_under_renaming where "closed_under_renaming S = (\C D. (C \ S) \ (renaming_cl C D) \ (D \ S))" definition variable_disjoint where "(variable_disjoint C D) = ((vars_of_cl (cl_ecl C)) \ (vars_of_cl (cl_ecl D)) = {})" subsection \Orders and Selection Functions\ text \We assume that the set of variables is infinite (so that shared variables can be renamed away) and that the following objects are given: (i) A term ordering that is total on ground terms, well-founded and closed under contextual embedding and substitutions. This ordering is used as usual to orient equations and to restrict the application of the replacement rule. (ii) A selection function, mapping each clause to a (possibly empty) set of negative literals. We assume that this selection function is closed under renamings. (iii) A function mapping every extended clause to an order on positions, which contains the (reversed) prefix ordering. This order allows one to control the order in which the subterms are rewritten (terms occurring at minimal positions are considered with the highest priority). (iv) A function @{term "filter_trms"} that allows one to filter away some of the terms attached to a given extended clause (it can be used for instance to remove terms if the set becomes too big). The standard superposition calculus corresponds to the case where this function always returns the empty set.\ locale basic_superposition = fixes trm_ord :: "('a trm \ 'a trm) set" fixes sel :: "'a clause \ 'a clause" fixes pos_ord :: "'a eclause \ 'a trm \ (position \ position) set" fixes vars :: "'a set" fixes filter_trms :: "'a clause \ 'a trm set \ 'a trm set" assumes trm_ord_wf : "(wf trm_ord)" and trm_ord_ground_total : "(\x y. ((ground_term x) \ (ground_term y) \ x \ y \ ((x,y) \ trm_ord \ (y,x) \ trm_ord)))" and trm_ord_trans : "trans trm_ord" and trm_ord_irrefl : "irrefl trm_ord" and trm_ord_reduction_left : "\x1 x2 y. (x1,x2) \ trm_ord \ ((Comb x1 y),(Comb x2 y)) \ trm_ord" and trm_ord_reduction_right : "\x1 x2 y. (x1,x2) \ trm_ord \ ((Comb y x1),(Comb y x2)) \ trm_ord" and trm_ord_subterm : "\x1 x2. (x1,(Comb x1 x2)) \ trm_ord \ (x2,(Comb x1 x2)) \ trm_ord" and trm_ord_subst : "\s x y. (x,y) \ trm_ord \ ( (subst x s),(subst y s)) \ trm_ord" and pos_ord_irrefl : "(\x y. (irrefl (pos_ord x y)))" and pos_ord_trans : "(\x. (trans (pos_ord x y)))" and pos_ord_prefix : "\x y p q r. ((q,p) \ (pos_ord x y) \ ((append q r),p) \ (pos_ord x y))" and pos_ord_nil : "\x y p . (p \ Nil) \ (p,Nil) \ (pos_ord x y)" and sel_neg: "(\x. ( (sel (cl_ecl x)) \ (cl_ecl x)) \ (\y \ sel (cl_ecl x). (negative_literal y)))" and sel_renaming: "\\ C. ((renaming \ (vars_of_cl C)) \ sel C = {} \ sel (subst_cl C \) = {})" and infinite_vars: "\ (finite vars)" and filter_trms_inclusion: "filter_trms C E \ E" begin text \We provide some functions to decompose a literal in a way that is compatible with the ordering and establish some basic properties.\ definition orient_lit :: "'a literal \ 'a trm \ 'a trm \ sign \ bool" where "(orient_lit L u v s) = ((( (L = (Pos (Eq u v))) \ (L = (Pos (Eq v u)))) \ ((u,v) \ trm_ord) \ (s = pos)) \ (( (L = (Neg (Eq u v))) \ (L = (Neg (Eq v u)))) \ ((u,v) \ trm_ord) \ (s = neg)))" definition orient_lit_inst :: "'a literal \ 'a trm \ 'a trm \ sign \ 'a subst \ bool" where "(orient_lit_inst L u v s \) = ((( (L = (Pos (Eq u v))) \ (L = (Pos (Eq v u)))) \ (((subst u \),(subst v \)) \ trm_ord) \ (s = pos)) \ (( (L = (Neg (Eq u v))) \ (L = (Neg (Eq v u)))) \ (((subst u \),(subst v \)) \ trm_ord) \ (s = neg)))" lemma lift_orient_lit: assumes "orient_lit_inst L t s p \" shows "orient_lit (subst_lit L \) (subst t \) (subst s \) p" unfolding orient_lit_inst_def orient_lit_def using assms orient_lit_inst_def by auto lemma orient_lit_vars: assumes "orient_lit L t s p" shows "vars_of t \ vars_of_lit L \ vars_of s \ vars_of_lit L" proof - have "p = neg \ p = pos" using sign.exhaust by auto then show ?thesis proof assume "p = neg" from this and assms(1) have "(L = Neg (Eq t s)) \ (L = (Neg (Eq s t)))" unfolding orient_lit_def by auto then show ?thesis proof assume "L = Neg (Eq t s)" then have "vars_of_lit L = vars_of t \ vars_of s" by simp from this show ?thesis by simp next assume "L = Neg (Eq s t)" then have "vars_of_lit L = vars_of s \ vars_of t" by simp from this show ?thesis by simp qed next assume "p = pos" from this and assms(1) have "(L = Pos (Eq t s)) \ (L = (Pos (Eq s t)))" unfolding orient_lit_def by auto then show ?thesis proof assume "L = Pos (Eq t s)" then have "vars_of_lit L = vars_of t \ vars_of s" by simp from this show ?thesis by simp next assume "L = Pos (Eq s t)" then have "vars_of_lit L = vars_of s \ vars_of t" by simp from this show ?thesis by simp qed qed qed lemma orient_lit_inst_vars: assumes "orient_lit_inst L t s p \" shows "vars_of t \ vars_of_lit L \ vars_of s \ vars_of_lit L" proof - have "p = neg \ p = pos" using sign.exhaust by auto then show ?thesis proof assume "p = neg" from this and assms(1) have "(L = Neg (Eq t s)) \ (L = (Neg (Eq s t)))" unfolding orient_lit_inst_def by auto then show ?thesis proof assume "L = Neg (Eq t s)" then have "vars_of_lit L = vars_of t \ vars_of s" by simp from this show ?thesis by simp next assume "L = Neg (Eq s t)" then have "vars_of_lit L = vars_of s \ vars_of t" by simp from this show ?thesis by simp qed next assume "p = pos" from this and assms(1) have "(L = Pos (Eq t s)) \ (L = (Pos (Eq s t)))" unfolding orient_lit_inst_def by auto then show ?thesis proof assume "L = Pos (Eq t s)" then have "vars_of_lit L = vars_of t \ vars_of s" by simp from this show ?thesis by simp next assume "L = Pos (Eq s t)" then have "vars_of_lit L = vars_of s \ vars_of t" by simp from this show ?thesis by simp qed qed qed lemma orient_lit_inst_coincide: assumes "orient_lit_inst L1 t s polarity \" assumes "coincide_on \ \ (vars_of_lit L1)" shows "orient_lit_inst L1 t s polarity \" proof - have "polarity = pos \ polarity = neg" using sign.exhaust by auto then show ?thesis proof assume "polarity = pos" from this and assms(1) have "L1 = Pos (Eq t s) \ L1 = Pos (Eq s t)" and "( (subst t \), (subst s \)) \ trm_ord" unfolding orient_lit_inst_def by auto from \L1 = Pos (Eq t s) \ L1 = Pos (Eq s t)\ have "vars_of t \ vars_of_lit L1" and "vars_of s \ vars_of_lit L1" by auto from \vars_of t \ vars_of_lit L1\ and assms(2) have "coincide_on \ \ (vars_of t)" unfolding coincide_on_def by auto from \vars_of s \ vars_of_lit L1\ and assms(2) have "coincide_on \ \ (vars_of s)" unfolding coincide_on_def by auto from \( (subst t \), (subst s \)) \ trm_ord\ and \coincide_on \ \ (vars_of t)\ and \coincide_on \ \ (vars_of s)\ assms(2) have "( (subst t \), (subst s \)) \ trm_ord" using coincide_on_term by metis from this and \polarity = pos\ and \L1 = Pos (Eq t s) \ L1 = Pos (Eq s t)\ show ?thesis unfolding orient_lit_inst_def by auto next assume "polarity = neg" from this and assms(1) have "L1 = Neg (Eq t s) \ L1 = Neg (Eq s t)" and "( (subst t \), (subst s \)) \ trm_ord" unfolding orient_lit_inst_def by auto from \L1 = Neg (Eq t s) \ L1 = Neg (Eq s t)\ have "vars_of t \ vars_of_lit L1" and "vars_of s \ vars_of_lit L1" by auto from \vars_of t \ vars_of_lit L1\ and assms(2) have "coincide_on \ \ (vars_of t)" unfolding coincide_on_def by auto from \vars_of s \ vars_of_lit L1\ and assms(2) have "coincide_on \ \ (vars_of s)" unfolding coincide_on_def by auto from \( (subst t \), (subst s \)) \ trm_ord\ and \coincide_on \ \ (vars_of t)\ and \coincide_on \ \ (vars_of s)\ assms(2) have "( (subst t \), (subst s \)) \ trm_ord" using coincide_on_term by metis from this and \polarity = neg\ and \L1 = Neg (Eq t s) \ L1 = Neg (Eq s t)\ show ?thesis unfolding orient_lit_inst_def by auto qed qed lemma orient_lit_inst_subterms: assumes "orient_lit_inst L t s polarity \" assumes "u \ subterms_of t" shows "u \ subterms_of_lit L" proof - have "polarity = pos \ polarity = neg" using sign.exhaust by auto then show ?thesis proof assume "polarity = pos" from this and assms(1) have "L = (Pos (Eq t s)) \ L = (Pos (Eq s t))" unfolding orient_lit_inst_def by auto then show ?thesis proof assume "L = (Pos (Eq t s))" from this and assms(2) show ?thesis by simp next assume "L = (Pos (Eq s t))" from this and assms(2) show ?thesis by simp qed next assume "polarity = neg" from this and assms(1) have "L = (Neg (Eq t s)) \ L = (Neg (Eq s t))" unfolding orient_lit_inst_def by auto then show ?thesis proof assume "L = (Neg (Eq t s))" from this and assms(2) show ?thesis by simp next assume "L = (Neg (Eq s t))" from this and assms(2) show ?thesis by simp qed qed qed subsection \Clause Ordering\ text \Clauses and extended clauses are ordered by transforming them into multisets of multisets of terms. To avoid any problem with the merging of identical literals, the multiset is assigned to a pair clause-substitution rather than to an instantiated clause.\ text \We first map every literal to a multiset of terms, using the usual conventions and then define the multisets associated with clauses and extended clauses.\ fun mset_lit :: "'a literal \ 'a trm multiset" where "mset_lit (Pos (Eq t s)) = {# t,s #}" | "mset_lit (Neg (Eq t s)) = {# t,t,s,s #}" fun mset_cl where "mset_cl (C,\) = {# (mset_lit (subst_lit x \)). x \# (mset_set C) #}" fun mset_ecl where "mset_ecl (C,\) = {# (mset_lit (subst_lit x \)). x \# (mset_set (cl_ecl C)) #}" lemma mset_ecl_and_mset_lit: assumes "L \ (cl_ecl C)" assumes "finite (cl_ecl C)" shows "(mset_lit (subst_lit L \)) \# (mset_ecl (C,\))" proof - from assms(1) assms(2) have "L \# (mset_set (cl_ecl C))" by (simp) then show ?thesis proof - have f1: "mset_set (cl_ecl C) - {#L#} + {#L#} = mset_set (cl_ecl C)" by (meson \L \# mset_set (cl_ecl C)\ insert_DiffM2) have "count {#mset_lit (subst_lit L \)#} (mset_lit (subst_lit L \)) = 1" by simp then show ?thesis by (metis (no_types, lifting) f1 image_mset_add_mset insert_iff mset_ecl.simps set_mset_add_mset_insert union_mset_add_mset_right) qed qed lemma ecl_ord_coincide: assumes "coincide_on \ \' (vars_of_cl (cl_ecl C))" shows "mset_ecl (C,\) = mset_ecl (C,\')" proof - have "\x. (x \ (cl_ecl C) \ ((subst_lit x \) = (subst_lit x \')))" proof ((rule allI),(rule impI)) fix x assume "x \ (cl_ecl C)" from this have "vars_of_lit x \ (vars_of_cl (cl_ecl C))" by auto from this and assms(1) have "coincide_on \ \' (vars_of_lit x)" unfolding coincide_on_def by auto from this show "((subst_lit x \) = (subst_lit x \'))" by (simp add: coincide_on_lit) qed then show ?thesis using equal_image_mset [of "cl_ecl C" "\x. (mset_lit (subst_lit x \))" "\x. (mset_lit (subst_lit x \'))"] by (metis mset_ecl.simps) qed text \Literal and clause orderings are obtained as usual as the multiset extensions of the term ordering.\ definition lit_ord :: "('a literal \ 'a literal) set" where "lit_ord = { (x,y). (((mset_lit x),(mset_lit y)) \ (mult trm_ord)) }" lemma mult_trm_ord_trans: shows "trans (mult trm_ord)" by (metis (no_types, lifting) mult_def transI transitive_closure_trans(2)) lemma lit_ord_trans: shows "trans lit_ord" by (metis (no_types, lifting) basic_superposition.lit_ord_def basic_superposition_axioms case_prodD case_prodI mem_Collect_eq mult_def transI transitive_closure_trans(2)) lemma lit_ord_wf: shows "wf lit_ord" proof - from trm_ord_wf have "wf (mult trm_ord)" using wf_mult by auto then show ?thesis using lit_ord_def and measure_wf [of "(mult trm_ord)" lit_ord mset_lit] by blast qed definition ecl_ord :: "(('a eclause \ 'a subst) \ ('a eclause \ 'a subst)) set" where "ecl_ord = { (x,y). (((mset_ecl x),(mset_ecl y)) \ (mult (mult trm_ord))) }" definition ecl_ord_eq :: "(('a eclause \ 'a subst) \ ('a eclause \ 'a subst)) set" where "ecl_ord_eq = ecl_ord \ { (x,y). ((mset_ecl x) = (mset_ecl y)) }" definition cl_ord :: "(('a clause \ 'a subst) \ ('a clause \ 'a subst)) set" where "cl_ord = { (x,y). (((mset_cl x),(mset_cl y)) \ (mult (mult trm_ord))) }" definition cl_ord_eq :: "(('a clause \ 'a subst) \ ('a clause \ 'a subst)) set" where "cl_ord_eq = cl_ord \ { (x,y). (mset_cl x) = (mset_cl y) }" lemma mult_mult_trm_ord_trans: shows "trans (mult (mult trm_ord))" by (metis (no_types, lifting) mult_def transI transitive_closure_trans(2)) lemma ecl_ord_trans: shows "trans ecl_ord" by (metis (no_types, lifting) basic_superposition.ecl_ord_def basic_superposition_axioms case_prodD case_prodI mem_Collect_eq mult_def transI transitive_closure_trans(2)) lemma cl_ord_trans: shows "trans cl_ord" by (metis (no_types, lifting) basic_superposition.cl_ord_def basic_superposition_axioms case_prodD case_prodI mem_Collect_eq mult_def transI transitive_closure_trans(2)) lemma cl_ord_eq_trans: shows "trans cl_ord_eq" proof - have "\r. trans r = (\p pa pb. ((p::'a literal set \ ('a \ 'a trm) list, pa) \ r \ (pa, pb) \ r) \ (p, pb) \ r)" by (simp add: trans_def) then obtain pp :: "(('a literal set \ ('a \ 'a trm) list) \ 'a literal set \ ('a \ 'a trm) list) set \ 'a literal set \ ('a \ 'a trm) list" and ppa :: "(('a literal set \ ('a \ 'a trm) list) \ 'a literal set \ ('a \ 'a trm) list) set \ 'a literal set \ ('a \ 'a trm) list" and ppb :: "(('a literal set \ ('a \ 'a trm) list) \ 'a literal set \ ('a \ 'a trm) list) set \ 'a literal set \ ('a \ 'a trm) list" where f1: "\r. (\ trans r \ (\p pa pb. (p, pa) \ r \ (pa, pb) \ r \ (p, pb) \ r)) \ (trans r \ (pp r, ppa r) \ r \ (ppa r, ppb r) \ r \ (pp r, ppb r) \ r)" by (metis (no_types)) have f2: "trans {(p, pa). (mset_cl p, mset_cl pa) \ mult (mult trm_ord)}" using cl_ord_def cl_ord_trans by presburger { assume "\ (case (ppa (cl_ord \ {(p, pa). mset_cl p = mset_cl pa}), ppb (cl_ord \ {(p, pa). mset_cl p = mset_cl pa})) of (p, pa) \ mset_cl p = mset_cl pa)" { assume "(ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). (mset_cl pa, mset_cl p) \ mult (mult trm_ord)}" moreover { assume "(ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). (mset_cl pa, mset_cl p) \ mult (mult trm_ord)} \ (mset_cl (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})), mset_cl (ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}))) \ mult (mult trm_ord)" then have "(ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). (mset_cl pa, mset_cl p) \ mult (mult trm_ord)} \ mset_cl (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ mset_cl (ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}))" by force then have "((pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). (mset_cl pa, mset_cl p) \ mult (mult trm_ord)} \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). mset_cl pa = mset_cl p}) \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). (mset_cl pa, mset_cl p) \ mult (mult trm_ord)} \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). mset_cl pa = mset_cl p}" using f2 f1 by blast } ultimately have "(mset_cl (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})), mset_cl (ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}))) \ mult (mult trm_ord) \ ((pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). (mset_cl pa, mset_cl p) \ mult (mult trm_ord)} \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). mset_cl pa = mset_cl p}) \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). (mset_cl pa, mset_cl p) \ mult (mult trm_ord)} \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). mset_cl pa = mset_cl p}" by fastforce } then have "(mset_cl (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})), mset_cl (ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}))) \ mult (mult trm_ord) \ ((pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). (mset_cl pa, mset_cl p) \ mult (mult trm_ord)} \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). mset_cl pa = mset_cl p}) \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). (mset_cl pa, mset_cl p) \ mult (mult trm_ord)} \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ {(pa, p). mset_cl pa = mset_cl p} \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ cl_ord \ {(pa, p). mset_cl pa = mset_cl p} \ (ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ cl_ord \ {(pa, p). mset_cl pa = mset_cl p} \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ cl_ord \ {(pa, p). mset_cl pa = mset_cl p}" using cl_ord_def by auto } then have "(pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ cl_ord \ {(pa, p). mset_cl pa = mset_cl p} \ (ppa (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ cl_ord \ {(pa, p). mset_cl pa = mset_cl p} \ (pp (cl_ord \ {(pa, p). mset_cl pa = mset_cl p}), ppb (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})) \ cl_ord \ {(pa, p). mset_cl pa = mset_cl p}" using cl_ord_def by force then have "trans (cl_ord \ {(pa, p). mset_cl pa = mset_cl p})" using f1 by meson from this show ?thesis unfolding cl_ord_eq_def by auto qed lemma wf_ecl_ord: shows "wf ecl_ord" proof - have "wf (mult trm_ord)" using trm_ord_wf and wf_mult by auto then have "wf (mult (mult trm_ord))" using wf_mult by auto thus ?thesis using ecl_ord_def and measure_wf [of "(mult (mult trm_ord))" ecl_ord mset_ecl] by blast qed definition maximal_literal :: "'a literal \ 'a clause \ bool" where "(maximal_literal L C) = (\x. (x \ C \ (L,x) \ lit_ord))" definition eligible_literal where "(eligible_literal L C \) = (L \ sel (cl_ecl C) \ (sel(cl_ecl C) = {} \ (maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \))))" definition strictly_maximal_literal where "strictly_maximal_literal C L \ = (\x \ (cl_ecl C) - { L }. ( (subst_lit x \),(subst_lit L \)) \ lit_ord)" definition lower_or_eq where "lower_or_eq t s = ((t = s) \ ((t,s) \ trm_ord))" lemma eligible_literal_coincide: assumes "coincide_on \ \' (vars_of_cl (cl_ecl C))" assumes "eligible_literal L C \" assumes "L \ (cl_ecl C)" shows "eligible_literal L C \'" proof - from assms(2) have "L \ sel (cl_ecl C) \ (sel (cl_ecl C) = {} \ maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \))" unfolding eligible_literal_def by auto then show ?thesis proof assume "L \ sel (cl_ecl C)" then show ?thesis unfolding eligible_literal_def by auto next assume "sel (cl_ecl C) = {} \ maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)" then have "sel (cl_ecl C) = {}" and "maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)" by auto from assms(1) have "(subst_cl (cl_ecl C) \) = (subst_cl (cl_ecl C) \')" using coincide_on_cl by blast from assms(3) and assms(1) have "coincide_on \ \' (vars_of_lit L)" unfolding coincide_on_def by auto from this have "(subst_lit L \) = (subst_lit L \')" using coincide_on_lit by auto from this and \(subst_cl (cl_ecl C) \) = (subst_cl (cl_ecl C) \')\ and \maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)\ have "maximal_literal (subst_lit L \') (subst_cl (cl_ecl C) \')" by auto from this and \sel (cl_ecl C) = {}\ show ?thesis unfolding eligible_literal_def by auto qed qed text \The next definition extends the ordering to substitutions.\ definition lower_on where "lower_on \ \ V = (\x \ V. (lower_or_eq (subst (Var x) \) ( (subst (Var x) \))))" text \We now establish some properties of the ordering relations.\ lemma lower_or_eq_monotonic: assumes "lower_or_eq t1 s1" assumes "lower_or_eq t2 s2" shows "lower_or_eq (Comb t1 t2) (Comb s1 s2)" unfolding lower_or_eq_def using trm_ord_reduction_left trm_ord_reduction_right by (metis assms(1) assms(2) lower_or_eq_def trm_ord_trans transD) lemma lower_on_term: shows "\ \ \. lower_on \ \ (vars_of t) \ (lower_or_eq (subst t \) (subst t \))" proof (induction t) case (Var x) from this show ?case unfolding lower_on_def by auto next case (Const x) show ?case unfolding lower_or_eq_def by auto next case (Comb t1 t2) show "\ \ \. lower_on \ \ (vars_of (Comb t1 t2)) \ (lower_or_eq (subst (Comb t1 t2) \) (subst (Comb t1 t2) \))" proof - fix \ \ assume "lower_on \ \ (vars_of (Comb t1 t2))" from this have "lower_on \ \ (vars_of t1)" and "lower_on \ \ (vars_of t2)" unfolding lower_on_def by auto from \lower_on \ \ (vars_of t1)\ have "lower_or_eq (subst t1 \) (subst t1 \)" using Comb.IH by auto from \lower_on \ \ (vars_of t2)\ have "lower_or_eq (subst t2 \) (subst t2 \)" using Comb.IH by auto from \lower_or_eq (subst t1 \) (subst t1 \)\ \lower_or_eq (subst t2 \) (subst t2 \)\ show "lower_or_eq (subst (Comb t1 t2) \) (subst (Comb t1 t2) \)" using lower_or_eq_monotonic by auto qed qed lemma diff_substs_yield_diff_trms: assumes "(subst (Var x) \) \ (subst (Var x) \)" shows "(x \ vars_of t) \ (subst t \) \ (subst t \)" proof (induction t) case (Var y) show "(x \ vars_of (Var y)) \ (subst (Var y) \) \ (subst (Var y) \)" proof - assume "(x \ vars_of (Var y))" from \(x \ vars_of (Var y))\ have "x = y" by auto from this and assms(1) show "(subst (Var y) \) \ (subst (Var y) \)" by auto qed next case (Const y) show " (x \ vars_of (Const y)) \ (subst (Const y) \) \ (subst (Const y) \)" proof (rule ccontr) from \(x \ vars_of (Const y))\ show False by auto qed next case (Comb t1 t2) show " (x \ vars_of (Comb t1 t2)) \ (subst (Comb t1 t2) \) \ (subst (Comb t1 t2) \)" proof - assume "(x \ vars_of (Comb t1 t2))" from \x \ vars_of (Comb t1 t2)\ have "x \ vars_of t1 \ x \ vars_of t2" by auto then show "(subst (Comb t1 t2) \) \ (subst (Comb t1 t2) \)" proof assume "x \ vars_of t1" from this have "(subst t1 \) \ (subst t1 \)" using Comb.IH by auto then show ?thesis by auto next assume "x \ vars_of t2" from this have "(subst t2 \) \ (subst t2 \)" using Comb.IH by auto then show ?thesis by auto qed qed qed lemma lower_subst_yields_lower_trms: assumes "lower_on \ \ (vars_of t)" assumes "((subst (Var x) \),(subst (Var x) \)) \ trm_ord" assumes "(x \ vars_of t)" shows "((subst t \),(subst t \)) \ trm_ord" proof - from assms(1) have "lower_or_eq (subst t \) (subst t \)" using lower_on_term by auto from assms(2) have "(subst (Var x) \) \ (subst (Var x) \)" using trm_ord_irrefl irrefl_def by fastforce from this and assms(3) have "(subst t \) \ (subst t \)" using diff_substs_yield_diff_trms by fastforce from this and \lower_or_eq (subst t \) (subst t \)\ show ?thesis unfolding lower_or_eq_def by auto qed lemma lower_on_lit: assumes "lower_on \ \ (vars_of_lit L)" assumes "((subst (Var x) \),(subst (Var x) \)) \ trm_ord" assumes "x \ vars_of_lit L" shows "((subst_lit L \), (subst_lit L \)) \ lit_ord" proof - obtain t s where def_l: "L = Pos (Eq t s) | L = (Neg (Eq t s))" by (metis mset_lit.cases) from this have "vars_of t \ vars_of_lit L" and "vars_of s \ vars_of_lit L" by auto from \vars_of s \ vars_of_lit L\ and assms(1) have "lower_on \ \ (vars_of s)" unfolding lower_on_def by auto from def_l have def_ms_l: "mset_lit L = {# t,s #} \ mset_lit L = {# t,t,s,s #}" by auto from this have "t \# (mset_lit L)" and "s \# (mset_lit L)" by auto from def_l have "mset_lit (subst_lit L \) = {# (subst u \). u \# (mset_lit L) #}" by auto from def_l have "mset_lit (subst_lit L \) = {# (subst u \). u \# (mset_lit L) #}" by auto from \lower_on \ \ (vars_of s)\ have "lower_or_eq (subst s \) (subst s \)" using lower_on_term by auto let ?L = "mset_lit L" let ?M1 = "mset_lit (subst_lit L \)" let ?M2 = "mset_lit (subst_lit L \)" from \vars_of t \ vars_of_lit L\ and assms(1) have "lower_on \ \ (vars_of t)" unfolding lower_on_def by auto from \vars_of s \ vars_of_lit L\ and assms(1) have "lower_on \ \ (vars_of s)" unfolding lower_on_def by auto have all_lower: "\u. (u \# (mset_lit L) \ (((subst u \), (subst u \)) \ trm_ord \ (subst u \) = (subst u \)))" proof (rule allI,rule impI) fix u assume "u \# (mset_lit L)" have "u = t \ u = s" proof (cases) assume "mset_lit L = {# t,s #}" from this and \u \# (mset_lit L)\ show ?thesis by (simp add: count_single insert_DiffM2 insert_noteq_member not_gr0) next assume "\mset_lit L = {# t,s #}" from this and def_ms_l have "mset_lit L = {# t,t,s,s #}" by auto from this and \u \# (mset_lit L)\ show ?thesis using not_gr0 by fastforce qed then show "(((subst u \), (subst u \)) \ trm_ord \ (subst u \) = (subst u \))" proof assume "u = t" from \lower_on \ \ (vars_of t)\ have "lower_or_eq (subst t \) (subst t \)" using lower_on_term by auto from this show ?thesis unfolding lower_or_eq_def using \u = t\ by auto next assume "u = s" from \lower_on \ \ (vars_of s)\ have "lower_or_eq (subst s \) (subst s \)" using lower_on_term by auto from this show ?thesis unfolding lower_or_eq_def using \u = s\ by auto qed qed have sl_exists: "\u. (u \# (mset_lit L) \ ((subst u \), (subst u \)) \ trm_ord)" proof - from \x \ vars_of_lit L\ and def_l have "x \ vars_of t \ x \ vars_of s" by auto then show ?thesis proof assume "x \ vars_of t" from this and \lower_on \ \ (vars_of t)\ assms(1) assms(2) have "( (subst t \),(subst t \)) \ trm_ord" using lower_subst_yields_lower_trms by auto from this and \t \# (mset_lit L)\ show ?thesis by auto next assume "x \ vars_of s" from this and \lower_on \ \ (vars_of s)\ assms(1) assms(2) have "( (subst s \),(subst s \)) \ trm_ord" using lower_subst_yields_lower_trms by auto from this and \s \# (mset_lit L)\ show ?thesis by auto qed qed from all_lower sl_exists and \mset_lit (subst_lit L \) = {# (subst u \). u \# (mset_lit L) #}\ \mset_lit (subst_lit L \) = {# (subst u \). u \# (mset_lit L) #}\ have "(?M1,?M2) \ (mult trm_ord)" using trm_ord_irrefl image_mset_ordering [of ?M1 "\x. (subst x \)" ?L ?M2 "\x. (subst x \)" trm_ord] by blast from this show ?thesis unfolding lit_ord_def by auto qed lemma lower_on_lit_eq: assumes "lower_on \ \ (vars_of_lit L)" shows "((subst_lit L \) = (subst_lit L \)) \ ((subst_lit L \), (subst_lit L \)) \ lit_ord" proof (cases) assume "coincide_on \ \ (vars_of_lit L)" then show ?thesis using coincide_on_lit by auto next assume "\coincide_on \ \ (vars_of_lit L)" then obtain x where "x \ vars_of_lit L" and "(subst (Var x) \) \ (subst (Var x) \)" unfolding coincide_on_def by auto from \x \ vars_of_lit L\ assms(1) \(subst (Var x) \) \ (subst (Var x) \)\and assms(1) have "((subst (Var x) \),(subst (Var x) \)) \ trm_ord" unfolding lower_on_def lower_or_eq_def by auto from this assms(1) \x \ vars_of_lit L\ show ?thesis using lower_on_lit by auto qed lemma lower_on_cl: assumes "lower_on \ \ (vars_of_cl (cl_ecl C))" assumes "((subst (Var x) \),(subst (Var x) \)) \ trm_ord" assumes "x \ vars_of_cl (cl_ecl C)" assumes "finite (cl_ecl C)" shows "((C,\), (C, \)) \ ecl_ord" proof - let ?M1 = "mset_ecl (C,\)" let ?M2 = "mset_ecl (C,\)" let ?M = "(mset_set (cl_ecl C))" let ?f1 = "\x. (mset_lit (subst_lit x \))" let ?f2 = "\x. (mset_lit (subst_lit x \))" have "?M1 = {# (?f1 u). u \# ?M #}" using mset_ecl.simps by blast have "?M2 = {# (?f2 u). u \# ?M #}" using mset_ecl.simps by blast have i: "\u. (u \# ?M \ (((?f1 u), (?f2 u)) \ (mult trm_ord) \ (?f1 u) = (?f2 u)))" proof ((rule allI),(rule impI)) fix u assume "u \# ?M" from this have "u \ (cl_ecl C)" using count_mset_set(3) by (simp add: assms(4)) from this and assms(1) have "lower_on \ \ (vars_of_lit u)" unfolding lower_on_def by auto then have "((subst_lit u \) = (subst_lit u \)) \ ((subst_lit u \), (subst_lit u \)) \ lit_ord" using lower_on_lit_eq by blast from this show "(((?f1 u), (?f2 u)) \ (mult trm_ord) \ (?f1 u) = (?f2 u))" unfolding lit_ord_def by auto qed have "irrefl (mult trm_ord)" by (simp add: irreflI trm_ord_wf wf_mult) have ii: "\u. (u \# ?M \ ((?f1 u), (?f2 u)) \ (mult trm_ord))" proof - from \x \ vars_of_cl (cl_ecl C)\ obtain u where "u \ (cl_ecl C)" and "x \ vars_of_lit u" by auto from assms(4) \u \ (cl_ecl C)\ have "u \# ?M" by auto from \u \ (cl_ecl C)\ and assms(1) have "lower_on \ \ (vars_of_lit u)" unfolding lower_on_def by auto from \x \ vars_of_lit u\ \lower_on \ \ (vars_of_lit u)\ assms(2) have "((subst_lit u \), (subst_lit u \)) \ lit_ord" using lower_on_lit by blast from this \u \# ?M\ have "(u \# ?M \ ((?f1 u), (?f2 u)) \ (mult trm_ord))" unfolding lit_ord_def by auto then show ?thesis by auto qed from i ii \?M1 = {# (?f1 u). u \# ?M #}\ \?M2 = {# (?f2 u). u \# ?M #}\ \irrefl (mult trm_ord)\ have "(?M1,?M2) \ (mult (mult trm_ord))" using image_mset_ordering [of ?M1 ?f1 ?M ?M2 ?f2 "(mult trm_ord)" ] by auto then show ?thesis unfolding ecl_ord_def by auto qed lemma subterm_trm_ord : shows "\ t s. subterm t p s \ p \ [] \ (s,t) \ trm_ord" proof (induction p) case Nil from \Nil \ []\ show ?case by auto next case (Cons i q) from \subterm t (i # q) s\ obtain t1 t2 where "t = (Comb t1 t2)" using subterm.elims(2) by blast have "i = Left \ i = Right" using indices.exhaust by blast then show "(s,t) \ trm_ord" proof assume "i = Left" from this and \t = Comb t1 t2\ and \subterm t (i # q) s\ have "subterm t1 q s" by auto show ?thesis proof (cases) assume "q = Nil" from this and \subterm t1 q s\ have "t1 = s" by auto from this and \t = Comb t1 t2\ show ?case using trm_ord_subterm by auto next assume "q \ Nil" from this and \subterm t1 q s\ have "(s,t1) \ trm_ord" using Cons.IH by auto from \t = Comb t1 t2\ have "(t1,t) \ trm_ord" using trm_ord_subterm by auto from this and \(s,t1) \ trm_ord\ show ?case using trm_ord_trans unfolding trans_def by metis qed next assume "i = Right" from this and \t = Comb t1 t2\ and \subterm t (i # q) s\ have "subterm t2 q s" by auto show ?thesis proof (cases) assume "q = Nil" from this and \subterm t2 q s\ have "t2 = s" by auto from this and \t = Comb t1 t2\ show ?case using trm_ord_subterm by auto next assume "q \ Nil" from this and \subterm t2 q s\ have "(s,t2) \ trm_ord" using Cons.IH by auto from \t = Comb t1 t2\ have "(t2,t) \ trm_ord" using trm_ord_subterm by auto from this and \(s,t2) \ trm_ord\ show ?case using trm_ord_trans unfolding trans_def by metis qed qed qed lemma subterm_trm_ord_eq : assumes "subterm t p s" shows "s = t \ (s,t) \ trm_ord" proof (cases) assume "p = Nil" from this and assms(1) show ?thesis by auto next assume "p \ Nil" from this and assms(1) show ?thesis using subterm_trm_ord by auto qed lemma subterms_of_trm_ord_eq : assumes "s \ subterms_of t" shows "s = t \ (s,t) \ trm_ord" proof - from assms(1) obtain p where "subterm t p s" using occurs_in_def by auto from this show ?thesis using subterm_trm_ord_eq by auto qed lemma subt_trm_ord: shows "t \ s \ (t,s) \ trm_ord" proof (induction s) case (Var x) show ?case proof assume "t \ Var x" then show "(t,Var x) \ trm_ord" by auto qed case (Const x) show ?case proof assume "t \ Const x" then show "(t,Const x) \ trm_ord" by auto qed case (Comb t1 t2) show ?case proof assume "t \ Comb t1 t2" show "(t, Comb t1 t2) \ trm_ord" proof (rule ccontr) assume "(t, Comb t1 t2) \ trm_ord" then have i: "t \ t1" using trm_ord_subterm by auto from \(t, Comb t1 t2) \ trm_ord\ have ii: "t \ t2" using trm_ord_subterm by auto from i ii and \t \ Comb t1 t2\ have "t \ t1 \ t \ t2" by auto from this and \(t, Comb t1 t2) \ trm_ord\ show False using Comb.IH trm_ord_subterm trm_ord_trans trans_def by metis qed qed qed lemma trm_ord_vars: assumes "(t,s) \ trm_ord" shows "vars_of t \ vars_of s" proof (rule ccontr) assume "\vars_of t \ vars_of s" then obtain x where "x \ vars_of t" and "x \ vars_of s" by auto let ?\ = "[(x,s)]" from assms have "((subst t ?\),(subst s ?\)) \ trm_ord" using trm_ord_subst by auto let ?\ = "[]" let ?V = "vars_of s" have "subst s ?\ = s" by simp have "subst (Var x) ?\ = s" by simp have "coincide_on ?\ ?\ ?V" proof (rule ccontr) assume "\ coincide_on ?\ ?\ ?V" then obtain y where "y \ ?V" "subst (Var y) ?\ \ subst (Var y) ?\" unfolding coincide_on_def by auto from \subst (Var y) ?\ \ subst (Var y) ?\\ have "y = x" by (metis assoc.simps(2) subst.simps(1)) from this and \x \ vars_of s\ \y \ ?V\ show False by auto qed from this and \subst s ?\ = s\ have "subst s ?\ = s" using coincide_on_term by metis from \x \ vars_of t\ have "(Var x) \ t" using \(subst t [(x, s)], subst s [(x, s)]) \ trm_ord\ \subst s [(x, s)] = s\ trm_ord_wf vars_iff_occseq by fastforce from this have "((Var x), t) \ trm_ord" using subt_trm_ord by auto from this and assms(1) have "(Var x,s) \trm_ord" using trm_ord_trans trans_def by metis from this have "((subst (Var x) ?\),(subst s ?\)) \ trm_ord" using trm_ord_subst by metis from this and \subst s ?\ = s\ \subst (Var x) ?\ = s\ have "(s,s) \ trm_ord" by auto from this show False using trm_ord_irrefl irrefl_def by metis qed lemma lower_on_ground: assumes "lower_on \ \ V" assumes "ground_on V \" shows "ground_on V \" proof (rule ccontr) assume "\ ground_on V \" from this obtain x where "x \ V" and "vars_of (subst (Var x) \) \ {}" unfolding ground_on_def ground_term_def by metis from assms(1) \x \ V\ have "(subst (Var x) \) = (subst (Var x) \) \ ((subst (Var x) \),(subst (Var x) \)) \ trm_ord" unfolding lower_on_def lower_or_eq_def by metis from this have "vars_of (subst (Var x) \) \ vars_of (subst (Var x) \)" using trm_ord_vars by auto from this and \vars_of (subst (Var x) \) \ {}\ have "vars_of (subst (Var x) \) \ {}" by auto from this and \x \ V\ and assms(2) show False unfolding ground_on_def ground_term_def by metis qed lemma replacement_monotonic : shows "\ t s. ((subst v \), (subst u \)) \ trm_ord \ subterm t p u \ replace_subterm t p v s \ ((subst s \), (subst t \)) \ trm_ord" proof (induction p) case Nil from \subterm t Nil u\ have "t = u" by auto from \replace_subterm t Nil v s\ have "s = v" by auto from \t = u\ and \s = v\ and \((subst v \), (subst u \)) \ trm_ord\ show ?case by auto next case (Cons i q) from \subterm t (i # q) u\ obtain t1 t2 where "t = (Comb t1 t2)" using subterm.elims(2) by blast have "i = Left \ i = Right" using indices.exhaust by blast then show "((subst s \), (subst t \)) \ trm_ord" proof assume "i = Left" from this and \t = Comb t1 t2\ and \subterm t (i # q) u\ have "subterm t1 q u" by auto from \i = Left\ and \t = Comb t1 t2\ and \replace_subterm t (i # q) v s\ obtain t1' where "replace_subterm t1 q v t1'" and "s = Comb t1' t2" by auto from \((subst v \), (subst u \)) \ trm_ord\ and \subterm t1 q u\ and \replace_subterm t1 q v t1'\ have "((subst t1' \), (subst t1 \)) \ trm_ord" using Cons.IH Cons.prems(1) by blast from this and \t = (Comb t1 t2)\ and \s = (Comb t1' t2)\ show "((subst s \), (subst t \)) \ trm_ord" by (simp add: trm_ord_reduction_left) next assume "i = Right" from this and \t = Comb t1 t2\ and \subterm t (i # q) u\ have "subterm t2 q u" by auto from \i = Right\ and \t = Comb t1 t2\ and \replace_subterm t (i # q) v s\ obtain t2' where "replace_subterm t2 q v t2'" and "s = Comb t1 t2'" by auto from \((subst v \), (subst u \)) \ trm_ord\ and \subterm t2 q u\ and \replace_subterm t2 q v t2'\ have "((subst t2' \), (subst t2 \)) \ trm_ord" using Cons.IH Cons.prems(2) by blast from this and \t = (Comb t1 t2)\ and \s = (Comb t1 t2')\ show "((subst s \), (subst t \)) \ trm_ord" by (simp add: trm_ord_reduction_right) qed qed lemma mset_lit_subst: shows "(mset_lit (subst_lit L \)) = {# (subst x \). x \# (mset_lit L) #}" proof - have "positive_literal L \ negative_literal L" using negative_literal.simps(2) positive_literal.elims(3) by blast then show ?thesis proof assume "positive_literal L" then obtain t s where "L = Pos (Eq t s)" by (metis equation.exhaust positive_literal.elims(2)) from this show ?thesis by auto next assume "negative_literal L" then obtain t s where "L = Neg (Eq t s)" by (metis equation.exhaust negative_literal.elims(2)) from this show ?thesis by auto qed qed lemma lit_ord_irrefl: shows "(L,L) \ lit_ord" by (simp add: lit_ord_wf) lemma lit_ord_subst: assumes "(L,M) \ lit_ord" shows "((subst_lit L \), (subst_lit M \)) \ lit_ord" proof - let ?f = "\x. (subst x \)" have i: "\ t s. ((t,s) \ trm_ord \ ((?f t), (?f s)) \ trm_ord)" using trm_ord_subst by auto from assms(1) have ii: "( (mset_lit L),(mset_lit M)) \ (mult trm_ord)" unfolding lit_ord_def by auto let ?L = "{# (?f x). x \# (mset_lit L) #}" let ?M = "{# (?f x). x \# (mset_lit M) #}" from i and ii have iii: "( ?L,?M ) \ (mult trm_ord)" using monotonic_fun_mult by metis have l: "?L = (mset_lit (subst_lit L \))" using mset_lit_subst by auto have m: "?M = (mset_lit (subst_lit M \))" using mset_lit_subst by auto from l m iii show ?thesis unfolding lit_ord_def by auto qed lemma args_are_strictly_lower: assumes "is_compound t" shows "(lhs t,t) \ trm_ord \ (rhs t,t) \ trm_ord" by (metis assms is_compound.elims(2) lhs.simps(1) rhs.simps(1) trm_ord_subterm) lemma mset_subst: assumes "C' = subst_cl D \" assumes "\ \ \ \ \" assumes "finite D" shows "mset_cl (C',\) = mset_cl (D,\) \ (mset_cl (C',\),mset_cl (D,\)) \ (mult (mult trm_ord))" proof - let ?f = "\x. (subst_lit x \)" let ?g = "\x. (mset_lit (subst_lit x \))" let ?h = "\x. (mset_lit (subst_lit x \))" have i: "\x \ D. ( (?g (?f x)) = (?h x))" proof fix x have "(subst_lit (subst_lit x \) \) = (subst_lit x (\ \ \))" using composition_of_substs_lit by auto from assms(2) have "(subst_lit x \) = (subst_lit x (\ \ \))" using subst_eq_lit by auto from this \(subst_lit (subst_lit x \) \) = (subst_lit x (\ \ \))\ show "(?g (?f x)) = (?h x)" by auto qed from assms(3) have "mset_set (?f ` D) \# {# (?f x). x \# mset_set (D) #}" using mset_set_mset_image by auto from this have ii: "{# (?g x). x \# mset_set (?f ` D) #} \# {# (?g x). x \# {# (?f x). x \# mset_set (D) #} #}" using image_mset_subseteq_mono by auto have "{# (?g x). x \# {# (?f x). x \# mset_set (D) #} #} = {# (?g (?f x)). x \# mset_set D #}" using mset_image_comp [of ?g ?f ] by auto from this and ii have iii: "{# (?g x). x \# mset_set (?f ` D) #} \# {# (?g (?f x)). x \# mset_set D #}" by auto from i have "{# (?g (?f x)). x \# (mset_set D) #} = {# (?h x). x \# (mset_set D) #} " using equal_image_mset [of D "\x. (?g (?f x))"] by auto from this and iii have "{# (?g x). x \# mset_set (?f ` D) #} \# {# (?h x). x \# mset_set D #}" by auto from this have iv: "{# (?g x). x \# mset_set (?f ` D) #} \# mset_cl (D,\)" by auto from assms(1) have "((\x. subst_lit x \) ` D) = C'" by auto from this and iv have "{#mset_lit (subst_lit x \). x \# mset_set C' #} \# mset_cl (D, \)" by auto from this have "mset_cl (C', \) \# mset_cl (D, \)" by auto from this show ?thesis using multiset_order_inclusion_eq mult_trm_ord_trans by auto qed lemma max_lit_exists: shows "finite C \ C \ {} \ ground_clause C \ (\L. (L \ C \ (maximal_literal L C)))" proof (induction rule: finite.induct) case emptyI show ?case by simp next fix A :: "'a clause" and a:: "'a literal" assume "finite A" assume hyp_ind: "A \ {} \ ground_clause A \ (\L. (L \ A \ (maximal_literal L A)))" show "(insert a A) \ {} \ ground_clause (insert a A) \ (\L. (L \ (insert a A) \ (maximal_literal L (insert a A))))" proof ((rule impI)+) assume "insert a A \ {}" assume "ground_clause (insert a A)" show "(\L. (L \ (insert a A) \ (maximal_literal L (insert a A))))" proof (cases) assume "A = {}" then have "a \ (insert a A) \ (maximal_literal a (insert a A))" unfolding maximal_literal_def using lit_ord_irrefl by auto then show ?thesis by auto next assume "A \ {}" have "insert a A = {a} \ A" by auto then have "vars_of_cl (insert a A) = vars_of_cl A \ vars_of_lit a" by auto from this and \ground_clause (insert a A)\ have "vars_of_lit a = {}" and "ground_clause A" by auto from \ground_clause A\ and \A \ {}\ and hyp_ind obtain b where "b \ A" and "maximal_literal b A" by auto show ?thesis proof (cases) assume "maximal_literal a A" then have "maximal_literal a (insert a A)" using lit_ord_wf maximal_literal_def by auto then show ?thesis by auto next assume "\maximal_literal a A" then obtain a' where "a' \ A" and "(a,a') \ lit_ord" unfolding maximal_literal_def by auto from \a' \ A\ and \maximal_literal b A\ have "(b,a') \ lit_ord" unfolding maximal_literal_def by auto from this and \(a,a') \ lit_ord\ have "(b,a) \ lit_ord" unfolding lit_ord_def using mult_def trancl_trans by fastforce from this and \maximal_literal b A\ have "maximal_literal b (insert a A)" unfolding maximal_literal_def by simp from this and \b \ A\ show ?thesis by auto qed qed qed qed text \We deduce that a clause contains at least one eligible literal.\ lemma eligible_lit_exists: assumes "finite (cl_ecl C)" assumes "(cl_ecl C) \ {}" assumes "(ground_clause (subst_cl (cl_ecl C) \))" shows "\L. ((eligible_literal L C \) \ (L \ (cl_ecl C)))" proof (cases) assume "sel (cl_ecl C) = {}" let ?C = "(subst_cl (cl_ecl C) \)" have "finite ?C" by (simp add: assms(1)) have "?C \ {}" proof - from assms(2) obtain L where "L \ (cl_ecl C)" by auto from this have "(subst_lit L \) \ ?C" by auto from this show "?C \ {}" by auto qed from \finite ?C\ \?C \ {}\ assms(3) obtain L where "L \ ?C" "maximal_literal L ?C" using max_lit_exists by metis from \L \ ?C\ obtain L' where "L' \ (cl_ecl C)" and "L = (subst_lit L' \)" by auto from \L' \ (cl_ecl C)\ \L = (subst_lit L' \)\ \maximal_literal L ?C\ \sel (cl_ecl C) = {}\ show ?thesis unfolding eligible_literal_def by metis next assume "sel (cl_ecl C) \ {}" then obtain L where "L \ sel (cl_ecl C)" by auto from this show ?thesis unfolding eligible_literal_def using sel_neg by blast qed text \The following lemmata provide various ways of proving that literals are ordered, depending on the relations between the terms they contain.\ lemma lit_ord_dominating_term: assumes "(s1,s2) \ trm_ord \ (s1,t2) \ trm_ord" assumes "orient_lit x1 s1 t1 p1" assumes "orient_lit x2 s2 t2 p2" assumes "vars_of_lit x1 = {}" assumes "vars_of_lit x2 = {}" shows "(x1,x2) \ lit_ord" proof - from \vars_of_lit x1 = {}\ and \orient_lit x1 s1 t1 p1\ have "vars_of t1 = {}" and "vars_of s1 = {}" and "\(s1,t1) \ trm_ord" unfolding orient_lit_def by auto from assms(5) and \orient_lit x2 s2 t2 p2\ have "vars_of t2 = {}" and "vars_of s2 = {}" and "\(s2,t2) \ trm_ord" unfolding orient_lit_def by auto from \vars_of t1 = {}\ and \vars_of s1 = {}\ and \\(s1,t1) \ trm_ord\ have o1: "t1 = s1 \ (t1,s1) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from \vars_of t2 = {}\ and \vars_of s2 = {}\ and \\(s2,t2) \ trm_ord\ have o2: "t2 = s2 \ (t2,s2) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from \\(s2,t2) \ trm_ord\ and assms(1) have "(s1,s2) \ trm_ord" by (metis assms(1) o2 trm_ord_trans transE) let ?m1 = "mset_lit x1" let ?m2 = "mset_lit x2" from assms(1) and o1 and o2 have "(t1,s2) \ trm_ord" using trm_ord_trans trans_def by metis from this and \(s1,s2) \ trm_ord\ have s2max: "\x. (x \# {# t1,t1,s1,s1 #} \ (x,s2) \ trm_ord)" by auto have "{# s2 #} \# {# t2,t2,s2,s2 #}" by simp from \{# s2 #} \# {# t2,t2,s2,s2 #}\ have "( {# s2 #}, {# t2,t2,s2,s2 #} ) \ mult trm_ord" using trm_ord_trans multiset_order_inclusion [of "{# s2 #}" "{# t2,t2,s2,s2 #}" "trm_ord"] by auto have "p1 = neg \ p1 = pos" using sign.exhaust by auto then show ?thesis proof assume "p1 = neg" from this and \orient_lit x1 s1 t1 p1\ have "x1 = (Neg (Eq t1 s1)) \ x1 = (Neg (Eq s1 t1))" using orient_lit_def by blast from this have m1: "?m1 = {# t1,t1,s1,s1 #}" using mset_lit.simps by auto have "p2 = neg \ p2 = pos" using sign.exhaust by auto then show ?thesis proof assume "p2 = neg" from this and \orient_lit x2 s2 t2 p2\ have "x2 = (Neg (Eq t2 s2)) \ x2 = (Neg (Eq s2 t2))" using orient_lit_def by blast from this have m2: "?m2 = {# t2,t2,s2,s2 #}" using mset_lit.simps by auto from s2max have "({# t1,t1,s1,s1 #}, {# s2 #}) \ mult trm_ord" using mult1_def_lemma [of "{# s2 #}" "{#}" s2 "{# t1,t1,s1,s1 #}" "{# t1,t1,s1,s1 #}" trm_ord] mult_def by auto from \( {# s2 #}, {# t2,t2,s2,s2 #} ) \ mult trm_ord\ and \({# t1,t1,s1,s1 #}, {# s2 #}) \ mult trm_ord\ have "( {# t1,t1,s1,s1 #}, {# t2,t2,s2,s2 #} ) \ mult trm_ord" using mult_trm_ord_trans unfolding trans_def by blast from this and m1 and m2 show ?thesis using lit_ord_def by auto next assume "p2 = pos" from this and \orient_lit x2 s2 t2 p2\ have "x2 = (Pos (Eq t2 s2)) \ x2 = (Pos (Eq s2 t2))" using orient_lit_def by blast from this have m2: "?m2 = {# t2,s2 #}" using mset_lit.simps by auto from s2max have "({# t1,t1,s1,s1 #}, {# s2 #}) \ mult trm_ord" using mult1_def_lemma [of "{# s2 #}" "{#}" s2 "{# t1,t1,s1,s1 #}" "{# t1,t1,s1,s1 #}" trm_ord] mult_def by auto from this and \( {# s2 #}, {# t2,t2,s2,s2 #} ) \ mult trm_ord\ have "({# t1,t1,s1,s1 #}, {# t2,s2 #}) \ mult trm_ord" using mset_ordering_add1 [of "{# t1,t1,s1,s1 #}" " {# s2 #}" trm_ord t2] by (auto) from this and m1 and m2 show ?thesis using lit_ord_def by auto qed next assume "p1 = pos" from this and \orient_lit x1 s1 t1 p1\ have "x1 = (Pos (Eq t1 s1)) \ x1 = (Pos (Eq s1 t1))" using orient_lit_def by blast from this have m1: "?m1 = {# t1,s1 #}" using mset_lit.simps by auto have "p2 = neg \ p2 = pos" using sign.exhaust by auto then show ?thesis proof assume "p2 = neg" from this and \orient_lit x2 s2 t2 p2\ have "x2 = (Neg (Eq t2 s2)) \ x2 = (Neg (Eq s2 t2))" using orient_lit_def by blast from this have m2: "?m2 = {# t2,t2,s2,s2 #}" using mset_lit.simps by auto from s2max have "({# t1,s1 #}, {# s2 #}) \ mult trm_ord" using mult1_def_lemma [of "{# s2 #}" "{#}" s2 "{# t1,s1 #}" "{# t1,s1 #}" trm_ord] mult_def by auto from \( {# s2 #}, {# t2,t2,s2,s2 #} ) \ mult trm_ord\ and \({# t1,s1 #}, {# s2 #}) \ mult trm_ord\ have "( {# t1,s1 #}, {# t2,t2,s2,s2 #} ) \ mult trm_ord" using mult_trm_ord_trans unfolding trans_def by blast from this and m1 and m2 show ?thesis using lit_ord_def by auto next assume "p2 = pos" from this and \orient_lit x2 s2 t2 p2\ have "x2 = (Pos (Eq t2 s2)) \ x2 = (Pos (Eq s2 t2))" using orient_lit_def by blast from this have m2: "?m2 = {# t2,s2 #}" using mset_lit.simps by auto from s2max have "({# t1,s1 #}, {# s2 #}) \ mult trm_ord" using mult1_def_lemma [of "{# s2 #}" "{#}" s2 "{# t1,s1 #}" "{# t1,s1 #}" trm_ord] mult_def by auto from this have "({# t1,s1 #}, {# t2,s2 #}) \ mult trm_ord" using mset_ordering_add1 [of "{# t1,s1 #}" " {# s2 #}" trm_ord t2] by auto from this and m1 and m2 show ?thesis using lit_ord_def by auto qed qed qed lemma lit_ord_neg_lit_lhs: assumes "orient_lit x1 s t1 pos" assumes "orient_lit x2 s t2 neg" assumes "vars_of_lit x1 = {}" assumes "vars_of_lit x2 = {}" shows "(x1,x2) \ lit_ord" proof - from assms(3) and assms(1) have "vars_of t1 = {}" and "vars_of s = {}" and "\(s,t1) \ trm_ord" unfolding orient_lit_def by auto from assms(4) and assms(2) have "vars_of t2 = {}" and "\(s,t2) \ trm_ord" unfolding orient_lit_def by auto from \vars_of t1 = {}\ and \vars_of s = {}\ and \\(s,t1) \ trm_ord\ have o1: "t1 = s \ (t1,s) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from \vars_of t2 = {}\ and \vars_of s = {}\ and \\(s,t2) \ trm_ord\ have o2: "t2 = s \ (t2,s) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto let ?m1 = "mset_lit x1" let ?m2 = "mset_lit x2" from \orient_lit x1 s t1 pos\ have "x1 = (Pos (Eq t1 s)) \ x1 = (Pos (Eq s t1))" using orient_lit_def by blast from this have m1: "?m1 = {# t1,s #}" using mset_lit.simps by auto from \orient_lit x2 s t2 neg\ have "x2 = (Neg (Eq t2 s)) \ x2 = (Neg (Eq s t2))" using orient_lit_def by blast from this have m2: "?m2 = {# t2,t2,s,s #}" using mset_lit.simps by auto show ?thesis proof (cases) assume "t1 = s" have "({# s,s #}, {# t2,s,s #}) \ mult trm_ord" using mult1_def_lemma [of "{# t2,s,s #}" "{# s,s #}" t2 "{# s,s #}" "{#}" trm_ord] mult_def by auto then have "({# s,s #}, {# t2,t2,s,s #}) \ mult trm_ord" using mset_ordering_add1 [of "{# s,s #}" " {# t2,s,s #}" trm_ord t2] by auto from this and \t1 = s\ and m1 and m2 show ?thesis using lit_ord_def by auto next assume "t1 \ s" from this and o1 have "(t1,s) \ trm_ord" by auto from this have smax: "\x. (x \# {# t1 #} \ (x,s) \ trm_ord)" by auto from smax have "({# t1,s #}, {# s,s #}) \ mult trm_ord" using mult1_def_lemma [of "{# s,s #}" "{# s #}" s "{# t1,s #}" "{# t1 #}" trm_ord] mult_def by auto from this have "({# t1,s #}, {# t2,s,s #}) \ mult trm_ord" using mset_ordering_add1 [of "{# t1,s #}" " {# s,s #}" trm_ord t2] by auto from this have "({# t1,s #}, {# t2,t2,s,s #}) \ mult trm_ord" using mset_ordering_add1 [of "{# t1,s #}" " {# t2,s,s #}" trm_ord t2] by auto from this and m1 and m2 show ?thesis using lit_ord_def by auto qed qed lemma lit_ord_neg_lit_rhs: assumes "orient_lit x1 s t1 pos" assumes "orient_lit x2 t2 s neg" assumes "vars_of_lit x1 = {}" assumes "vars_of_lit x2 = {}" shows "(x1,x2) \ lit_ord" proof - from assms(3) and assms(1) have "vars_of t1 = {}" and "vars_of s = {}" and "\(s,t1) \ trm_ord" unfolding orient_lit_def by auto from assms(4) and assms(2) have "vars_of t2 = {}" and "\(t2,s) \ trm_ord" unfolding orient_lit_def by auto from \vars_of t1 = {}\ and \vars_of s = {}\ and \\(s,t1) \ trm_ord\ have o1: "t1 = s \ (t1,s) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from \vars_of t2 = {}\ and \vars_of s = {}\ and \\(t2,s) \ trm_ord\ have o2: "t2 = s \ (s,t2) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto let ?m1 = "mset_lit x1" let ?m2 = "mset_lit x2" from \orient_lit x1 s t1 pos\ have "x1 = (Pos (Eq t1 s)) \ x1 = (Pos (Eq s t1))" using orient_lit_def by blast from this have m1: "?m1 = {# t1,s #}" using mset_lit.simps by auto from \orient_lit x2 t2 s neg\ have "x2 = (Neg (Eq t2 s)) \ x2 = (Neg (Eq s t2))" using orient_lit_def by blast from this have m2: "?m2 = {# t2,t2,s,s #}" using mset_lit.simps by auto show ?thesis proof (cases) assume "t1 = s" have "({# s,s #}, {# t2,s,s #}) \ mult trm_ord" using mult1_def_lemma [of "{# t2,s,s #}" "{# s,s #}" t2 "{# s,s #}" "{#}" trm_ord] mult_def by auto then have "({# s,s #}, {# t2,t2,s,s #}) \ mult trm_ord" using mset_ordering_add1 [of "{# s,s #}" " {# t2,s,s #}" trm_ord t2] by auto from this and \t1 = s\ and m1 and m2 show ?thesis using lit_ord_def by auto next assume "t1 \ s" from this and o1 have "(t1,s) \ trm_ord" by auto from this have smax: "\x. (x \# {# t1 #} \ (x,s) \ trm_ord)" by auto from smax have "({# t1,s #}, {# s,s #}) \ mult trm_ord" using mult1_def_lemma [of "{# s,s #}" "{# s #}" s "{# t1,s #}" "{# t1 #}" trm_ord] mult_def by auto from this have "({# t1,s #}, {# t2,s,s #}) \ mult trm_ord" using mset_ordering_add1 [of "{# t1,s #}" " {# s,s #}" trm_ord t2] by auto from this have "({# t1,s #}, {# t2,t2,s,s #}) \ mult trm_ord" using mset_ordering_add1 [of "{# t1,s #}" " {# t2,s,s #}" trm_ord t2] by auto from this and m1 and m2 show ?thesis using lit_ord_def by auto qed qed lemma lit_ord_rhs: assumes "(t1,t2) \ trm_ord" assumes "orient_lit x1 s t1 p" assumes "orient_lit x2 s t2 p" assumes "vars_of_lit x1 = {}" assumes "vars_of_lit x2 = {}" shows "(x1,x2) \ lit_ord" proof - from assms(2) and assms(4) have "vars_of t1 = {}" and "vars_of s = {}" and "\(s,t1) \ trm_ord" unfolding orient_lit_def by auto from assms(3) and assms(5) have "vars_of t2 = {}" and "\(s,t2) \ trm_ord" unfolding orient_lit_def by auto from \vars_of t1 = {}\ and \vars_of s = {}\ and \\(s,t1) \ trm_ord\ have o1: "t1 = s \ (t1,s) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from \vars_of t2 = {}\ and \vars_of s = {}\ and \\(s,t2) \ trm_ord\ have o2: "t2 = s \ (t2,s) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto let ?m1 = "mset_lit x1" let ?m2 = "mset_lit x2" have "p = pos \ p = neg" using sign.exhaust by auto then show ?thesis proof assume "p = pos" from this and \orient_lit x1 s t1 p\ have "x1 = (Pos (Eq t1 s)) \ x1 = (Pos (Eq s t1))" using orient_lit_def by blast from this have m1: "?m1 = {# t1,s #}" using mset_lit.simps by auto from \p = pos\ and \orient_lit x2 s t2 p\ have "x2 = (Pos (Eq t2 s)) \ x2 = (Pos (Eq s t2))" using orient_lit_def by blast from this have m2: "?m2 = {# t2,s #}" using mset_lit.simps by auto from assms(1) have "(\b. b \# {#t1#} \ (b, t2) \ trm_ord)" by auto then have "({# t1,s #}, {# t2,s #}) \ mult trm_ord" using mult1_def_lemma [of "{# t2,s #}" "{# s #}" t2 "{# t1,s #}" "{# t1 #}" trm_ord] mult_def by auto from this and m1 and m2 show ?thesis using lit_ord_def by auto next assume "p = neg" from this and \orient_lit x1 s t1 p\ have "x1 = (Neg (Eq t1 s)) \ x1 = (Neg (Eq s t1))" using orient_lit_def by blast from this have m1: "?m1 = {# t1,t1,s,s #}" using mset_lit.simps by auto from \p = neg\ and \orient_lit x2 s t2 p\ have "x2 = (Neg (Eq t2 s)) \ x2 = (Neg (Eq s t2))" using orient_lit_def by blast from this have m2: "?m2 = {# t2,t2,s,s #}" using mset_lit.simps by auto from assms(1) have max: "(\b. b \# {#t1,t1#} \ (b, t2) \ trm_ord)" by auto have i: "{# t2,s,s #} = {# s,s,t2 #}" by (simp add: add.commute add.left_commute) have ii: "{# t1,t1,s,s #} = {# s,s,t1,t1 #}" by (simp add: add.commute add.left_commute) from i and ii and max have "({# t1,t1,s,s #}, {# t2,s,s #}) \ mult trm_ord" using mult1_def_lemma [of "{# t2,s,s #}" "{# s,s #}" t2 "{# t1,t1,s,s #}" "{# t1,t1 #}" trm_ord] mult_def by auto then have "({# t1,t1,s,s #}, {# t2,t2,s,s #}) \ mult trm_ord" using mset_ordering_add1 [of "{# t1,t1,s,s #}" " {# t2,s,s #}" trm_ord t2] by auto from this and m1 and m2 show ?thesis using lit_ord_def by auto qed qed text \We show that the replacement of a term by an equivalent term preserves the semantics.\ lemma trm_rep_preserves_eq_semantics: assumes "fo_interpretation I" assumes "(I (subst t1 \) (subst t2 \))" assumes "(validate_ground_eq I (subst_equation (Eq t1 s) \))" shows "(validate_ground_eq I (subst_equation (Eq t2 s) \))" proof - from assms(1) have "transitive I" and "symmetric I" unfolding fo_interpretation_def congruence_def equivalence_relation_def by auto have "(subst_equation (Eq t1 s) \) = (Eq (subst t1 \) (subst s \))" by simp from this and assms(3) have "I (subst t1 \) (subst s \)" by simp from this and assms(2) and \transitive I\ and \symmetric I\ have "I (subst t2 \) (subst s \)" unfolding transitive_def symmetric_def by metis have "(subst_equation (Eq t2 s) \) = (Eq (subst t2 \) (subst s \))" by simp from this and \I (subst t2 \) (subst s \)\ show ?thesis by simp qed lemma trm_rep_preserves_lit_semantics: assumes "fo_interpretation I" assumes "(I (subst t1 \) (subst t2 \))" assumes "orient_lit_inst L t1 s polarity \'" assumes "\(validate_ground_lit I (subst_lit L \))" shows "\validate_ground_lit I (subst_lit (mk_lit polarity (Eq t2 s)) \)" proof - from assms(1) have "transitive I" and "symmetric I" unfolding fo_interpretation_def congruence_def equivalence_relation_def by auto have "polarity = pos \ polarity = neg" using sign.exhaust by auto then show ?thesis proof assume "polarity = pos" from this have mk: "(mk_lit polarity (Eq t2 s)) = (Pos (Eq t2 s))" by auto from \polarity = pos\ and assms(3) have "L = (Pos (Eq t1 s)) \ L = (Pos (Eq s t1))" unfolding orient_lit_inst_def by auto then show ?thesis proof assume "L = (Pos (Eq t1 s))" from this and assms(4) have "\I (subst t1 \) (subst s \)" by simp from this and assms(2) and \transitive I\ and \symmetric I\ have "\I (subst t2 \) (subst s \)" unfolding transitive_def symmetric_def by metis from this and mk show ?thesis by simp next assume "L = (Pos (Eq s t1))" from this and assms(4) have "\I (subst s \) (subst t1 \)" by simp from this and assms(2) and \transitive I\ and \symmetric I\ have "\I (subst t2 \) (subst s \)" unfolding transitive_def symmetric_def by metis from this and mk show ?thesis by simp qed next assume "polarity = neg" from this have mk: "(mk_lit polarity (Eq t2 s)) = (Neg (Eq t2 s))" by auto from \polarity = neg\ and assms(3) have "L = (Neg (Eq t1 s)) \ L = (Neg (Eq s t1))" unfolding orient_lit_inst_def by auto then show ?thesis proof assume "L = (Neg (Eq t1 s))" from this and assms(4) have "I (subst t1 \) (subst s \)" by simp from this and assms(2) and \transitive I\ and \symmetric I\ have "I (subst t2 \) (subst s \)" unfolding transitive_def symmetric_def by metis from this and mk show ?thesis by simp next assume "L = (Neg (Eq s t1))" from this and assms(4) have "I (subst s \) (subst t1 \)" by simp from this and assms(2) and \transitive I\ and \symmetric I\ have "I (subst t2 \) (subst s \)" unfolding transitive_def symmetric_def by metis from this and mk show ?thesis by simp qed qed qed lemma subterms_dominated : assumes "maximal_literal L C" assumes "orient_lit L t s p" assumes "u \ subterms_of_cl C" assumes "vars_of_lit L = {}" assumes "vars_of_cl C = {}" shows "u = t \ (u,t) \ trm_ord" proof (rule ccontr) assume neg_h: "\(u = t \ (u,t) \ trm_ord)" from assms(5) and assms(3) have "vars_of u = {}" using subterm_vars by blast from \vars_of_lit L = {}\ and \orient_lit L t s p\ have "vars_of s = {}" and "vars_of t = {}" and "\(t,s) \ trm_ord" unfolding orient_lit_def by auto from assms(3) obtain L' where "u \ subterms_of_lit L'" and "L' \ C" by auto from assms(5) and \L' \ C\ have "vars_of_lit L' = {}" using vars_of_cl.simps by auto from \u \ subterms_of_lit L'\ obtain t' s' p' where "orient_lit L' t' s' p'" and "u \ subterms_of t' \ subterms_of s'" unfolding orient_lit_def by (metis Un_commute mset_lit.cases subterms_of_eq.simps subterms_of_lit.simps(1) subterms_of_lit.simps(2) trm_ord_wf wf_asym) from \u \ subterms_of t' \ subterms_of s'\ have "u \ subterms_of t' \ u \ subterms_of s'" by auto then show False proof assume "u \ subterms_of t'" from this have "u = t' \ (u,t') \ trm_ord" using subterms_of_trm_ord_eq [of u t'] by auto from neg_h and \vars_of u = {}\ and \vars_of t = {}\ have "(t,u) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from this and \u = t' \ (u,t') \ trm_ord\ have "(t,t') \ trm_ord" using trm_ord_trans unfolding trans_def by metis from this and \vars_of_lit L' = {}\ and assms(4) and \orient_lit L t s p\ and \orient_lit L' t' s' p'\ have "(L,L') \ lit_ord" using lit_ord_dominating_term by blast from this and assms(1) and \L' \ C\ show False unfolding maximal_literal_def by auto next assume "u \ subterms_of s'" from this have "u = s' \ (u,s') \ trm_ord" using subterms_of_trm_ord_eq [of u s'] by auto from neg_h and \vars_of u = {}\ and \vars_of t = {}\ have "(t,u) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from this and \u = s' \ (u,s') \ trm_ord\ have "(t,s') \ trm_ord" using trm_ord_trans unfolding trans_def by metis from this and \vars_of_lit L' = {}\ and assms(4) and \orient_lit L t s p\ and \orient_lit L' t' s' p'\ have "(L,L') \ lit_ord" using lit_ord_dominating_term by blast from this and assms(1) and \L' \ C\ show False unfolding maximal_literal_def by auto qed qed text \A term dominates an expression if the expression contains no strictly greater subterm:\ fun dominate_eq:: "'a trm \ 'a equation \ bool" where "(dominate_eq t (Eq u v)) = ((t,u) \ trm_ord \ (t,v) \ trm_ord)" fun dominate_lit:: "'a trm \ 'a literal \ bool" where "(dominate_lit t (Pos e)) = (dominate_eq t e)" | "(dominate_lit t (Neg e)) = (dominate_eq t e)" definition dominate_cl:: "'a trm \ 'a clause \ bool" where "(dominate_cl t C) = (\x \ C. (dominate_lit t x))" definition no_disequation_in_cl:: "'a trm \ 'a clause \ bool" where "(no_disequation_in_cl t C) = (\u v. (Neg (Eq u v) \ C \ (u \ t \ v \ t)))" definition no_taut_eq_in_cl:: "'a trm \ 'a clause \ bool" where "(no_taut_eq_in_cl t C) = (Pos (Eq t t) \ C)" definition eq_occurs_in_cl where "(eq_occurs_in_cl t s C \) = (\L t' s'. (L \ C) \ (orient_lit_inst L t' s' pos \) \ (t = subst t' \) \ (s = subst s' \))" subsection \Inference Rules\ text \We now define the rules of the superposition calculus. Standard superposition is a refinement of the paramodulation rule based on the following ideas: (i) the replacement of a term by a bigger term is forbidden; (ii) the replacement can be performed only in the maximal term of a maximal (or selected) literal; (iii) replacement of variables is forbidden. Our definition imposes additional conditions on the positions on which the replacements are allowed: any superposition inference inside a term occurring in the set attached to the extended clause is blocked. \ text \We consider two different kinds of inferences: ground or first-order. Ground inferences are those needed for completeness, first-order inferences are those actually used by theorem provers. For conciseness, these two notions of inferences are defined simultaneously, and a parameter is added to the corresponding functions to determine whether the inference is ground or first-order.\ datatype inferences = Ground | FirstOrder text \The following function checks whether a given substitution is a unifier of two terms. If the inference is first-order then the unifier must be maximal.\ -definition ck_unifier -where - "(ck_unifier t s \ type) = - (if (type = FirstOrder) then (MGU \ t s) else (Unifier \ t s))" +definition ck_unifier where + "ck_unifier t s \ type \ (if type = FirstOrder then IMGU \ t s else Unifier \ t s)" lemma ck_unifier_thm: assumes "ck_unifier t s \ k" shows "(subst t \) = (subst s \)" -by (metis assms MGU_is_Unifier ck_unifier_def Unifier_def) + by (metis assms IMGU_iff_Idem_and_MGU MGU_is_Unifier ck_unifier_def Unifier_def) lemma subst_preserve_ck_unifier: assumes "ck_unifier t s \ k" shows "ck_unifier t s (comp \ \) Ground" proof - let ?\' = "(comp \ \)" from assms have "(subst t \) = (subst s \)" using ck_unifier_thm by auto then have "(subst t ?\') = (subst s ?\')" by simp then show ?thesis unfolding ck_unifier_def Unifier_def by auto qed text \The following function checks whether a given term is allowed to be reduced according to the strategy described above, i.e., that it does not occur in the set of terms associated with the clause (we do not assume that the set of irreducible terms is closed under subterm thus we use the function @{term "occurs_in"} instead of a mere membership test.\ definition allowed_redex where "allowed_redex t C \ = (\ (\s \ (trms_ecl C). (occurs_in (subst t \) (subst s \))))" text \The following function allows one to compute the set of irreducible terms attached to the conclusion of an inference. The computation depends on the type of the considered inference: for ground inferences the entire set of irreducible terms is kept. For first-order inferences, the function @{term "filter_trms"} is called to remove some of the terms (see also the function @{term "dom_trms"} below).\ definition get_trms where "get_trms C E t = (if (t = FirstOrder) then (filter_trms C E) else E)" text \The following definition provides the conditions that allow one to propagate irreducible terms from the parent clauses to the conclusion. A term can be propagated if it is strictly lower than a term occurring in the derived clause, or if it occurs in a negative literal of the derived clause. Note that this condition is slightly more restrictive than that of the basic superposition calculus, because maximal terms occurring in maximal positive literals cannot be kept in the set of irreducible terms. However in our definition, terms can be propagated even if they do not occur in the parent clause or in the conclusion. Extended clauses whose set of irreducible terms fulfills this property are called well-constrained.\ definition dom_trm where "dom_trm t C = (\ L u v p. (L \ C \ (decompose_literal L u v p) \ (( (p = neg \ t = u) \ (t,u) \ trm_ord))))" lemma dom_trm_lemma: assumes "dom_trm t C" shows "\ u. (u \ (subterms_of_cl C) \ (u = t \ (t,u) \ trm_ord))" proof - from assms(1) obtain L u v p where "L \ C" "decompose_literal L u v p" "(u = t \ (t,u) \ trm_ord)" unfolding dom_trm_def by blast from \decompose_literal L u v p\ have "u \ subterms_of_lit L" unfolding decompose_literal_def decompose_equation_def using root_subterm by force from this and \L \ C\ have "u \ (subterms_of_cl C)" by auto from this and \(u = t \ (t,u) \ trm_ord)\ show ?thesis by auto qed definition dom_trms where "dom_trms C E = { x. (x \ E) \ (dom_trm x C) }" lemma dom_trms_subset: shows "(dom_trms C E ) \ E" unfolding dom_trms_def by auto lemma dom_trm_vars: assumes "dom_trm t C" shows "vars_of t \ vars_of_cl C" proof - from assms obtain L u v p where "L \ C" "decompose_literal L u v p" "t = u \ (t,u) \ trm_ord" unfolding dom_trm_def by auto from \t = u \ (t,u) \ trm_ord\ have "vars_of t \ vars_of u" using trm_ord_vars by blast from this and \decompose_literal L u v p\ have "vars_of t \ vars_of_lit L" using decompose_literal_vars by blast from this show ?thesis using \L \ C\ by auto qed definition well_constrained where "well_constrained C = (\y. (y \ trms_ecl C \ dom_trm y (cl_ecl C)))" text \The next function allows one to check that a set of terms is in normal form. The argument @{term "f"} denotes the function mapping a term to its normal form (we do not assume that @{term "f"} is compatible with the term structure at this point).\ definition all_trms_irreducible where "(all_trms_irreducible E f) = (\x y. (x \ E \ occurs_in y x \ (f y) = y))" paragraph \Superposition\ text \We now define the superposition rule. Note that we assume that the parent clauses are variable-disjoint, but we do not explicitly rename them at this point, thus for completeness we will have to assume that the clause sets are closed under renaming. During the application of the rule, all the terms occurring at a position that is lower than that of the reduced term can be added in the set of irreducible terms attached to the conclusion (the intuition is that we assume that the terms occurring at minimal positions are reduced first). In particular, every proper subterm of the reduced term @{term "u'"} is added in the set of irreducible terms, thus every application of the superposition rule in a term introduced by unification will be blocked. Clause @{term "P1"} is the ``into'' clause and clause @{term "P2"} is the ``from'' clause.\ definition superposition :: "'a eclause \ 'a eclause \ 'a eclause \ 'a subst \ inferences \ 'a clause \ bool" where "(superposition P1 P2 C \ k C') = (\L t s u v M p Cl_P1 Cl_P2 Cl_C polarity t' u' L' trms_C. (L \ Cl_P1) \ (M \ Cl_P2) \ (eligible_literal L P1 \) \ (eligible_literal M P2 \) \ (variable_disjoint P1 P2) \ (Cl_P1 = (cl_ecl P1)) \ (Cl_P2 = (cl_ecl P2)) \ (\ is_a_variable u') \ (allowed_redex u' P1 \) \ trms_C = (get_trms Cl_C (dom_trms Cl_C (subst_set ((trms_ecl P1) \ (trms_ecl P2) \ { r. \ q. (q,p) \ (pos_ord P1 t) \ (subterm t q r) }) \)) k) \ (C = (Ecl Cl_C trms_C)) \ (orient_lit_inst M u v pos \) \ (orient_lit_inst L t s polarity \) \ ((subst u \) \ (subst v \)) \ (subterm t p u') \ (ck_unifier u' u \ k) \ (replace_subterm t p v t') \ ((k = FirstOrder) \ ( (subst_lit M \),(subst_lit L \)) \ lit_ord) \ ((k = FirstOrder) \ (strictly_maximal_literal P2 M \)) \ (L' = mk_lit polarity (Eq t' s)) \ (Cl_C = (subst_cl C' \)) \ (C' = (Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )))" paragraph \Reflexion\ text \We now define the Reflexion rule, which deletes contradictory literals (after unification). All the terms occurring in these literals can be added into the set of irreducible terms (intuitively, we can assume that these terms have been normalized before applying the rule). It is sufficient to add the term @{term "t"}, since every term occurring in the considered literal is a subterm of @{term "t"} (after unification).\ definition reflexion :: "'a eclause \ 'a eclause \ 'a subst \ inferences \ 'a clause \ bool" where "(reflexion P C \ k C') = (\L1 t s Cl_P Cl_C trms_C. (eligible_literal L1 P \) \ (L1 \ (cl_ecl P)) \ (Cl_C = (cl_ecl C)) \ (Cl_P = (cl_ecl P)) \ (orient_lit_inst L1 t s neg \) \ (ck_unifier t s \ k) \ (C = (Ecl Cl_C trms_C)) \ trms_C = (get_trms Cl_C (dom_trms Cl_C (subst_set ( (trms_ecl P) \ { t } ) \)) k) \ (Cl_C = (subst_cl C' \)) \ (C' = ((Cl_P - { L1 }) )))" paragraph \Factorization\ text \We now define the equational factorization rule, which merges two equations sharing the same left-hand side (after unification), if the right-hand sides are equivalent. Here, contrarily to the previous rule, the term @{term "t"} cannot be added into the set of irreducible terms, because we cannot assume that this term is in normal form (e.g., the application of the equational factorization rule may yield a new rewrite rule of left-hand side @{term "t"}). However, all proper subterms of @{term "t"} can be added.\ definition factorization :: "'a eclause \ 'a eclause \ 'a subst \ inferences \ 'a clause \ bool" where "(factorization P C \ k C') = (\L1 L2 L' t s u v Cl_P Cl_C trms_C. (eligible_literal L1 P \) \ (L1 \ (cl_ecl P)) \ (L2 \ (cl_ecl P) - { L1 }) \ (Cl_C = (cl_ecl C)) \ (Cl_P = (cl_ecl P)) \ (orient_lit_inst L1 t s pos \) \ (orient_lit_inst L2 u v pos \) \ ((subst t \) \ (subst s \)) \ ((subst t \) \ (subst v \)) \ (ck_unifier t u \ k) \ (L' = Neg (Eq s v)) \ (C = (Ecl Cl_C trms_C) \ trms_C = (get_trms Cl_C (dom_trms Cl_C (subst_set ( (trms_ecl P) \ (proper_subterms_of t) ) \))) k) \ (Cl_C = (subst_cl C' \)) \ (C' = ( (Cl_P - { L2 }) \ { L' } )))" subsection \Derivations\ text \We now define the set of derivable clauses by induction. Note that redundancy criteria are not taken into account at this point. Our definition of derivations also covers renaming.\ definition derivable :: "'a eclause \ 'a eclause set \ 'a eclause set \ 'a subst \ inferences \ 'a clause \ bool" where "(derivable C P S \ k C') = ((\P1 P2. (P1 \ S \ P2 \ S \ P = { P1, P2 } \ superposition P1 P2 C \ k C')) \ (\P1. (P1 \ S \ P = { P1 } \ factorization P1 C \ k C')) \ (\P1. (P1 \ S \ P = { P1 } \ reflexion P1 C \ k C')))" lemma derivable_premisses: assumes "derivable C P S \ k C'" shows "P \ S" using assms derivable_def by auto inductive derivable_ecl :: "'a eclause \ 'a eclause set \ bool" where init [simp, intro!]: "C \ S \ (derivable_ecl C S)" | rn [simp, intro!]: "(derivable_ecl C S) \ (renaming_cl C D) \ (derivable_ecl D S)" | deriv [simp, intro!]: "(\x. (x \ P \ (derivable_ecl x S))) \ (derivable C P S' \ FirstOrder C') \ (derivable_ecl C S)" text \We define a notion of instance by associating clauses with ground substitutions.\ definition instances:: "'a eclause set \ ('a eclause \ 'a subst) set" where "instances S = { x. \C \. (C \ S \ (ground_clause (subst_cl (cl_ecl C) \)) \ x = ( C,\))}" definition clset_instances:: "('a eclause \ 'a subst) set \ 'a clause set" where "clset_instances S = { C. \x. (x \ S \ C = (subst_cl (cl_ecl (fst x)) (snd x))) }" definition grounding_set where "grounding_set S \ = (\x. (x \ S \ (ground_clause (subst_cl (cl_ecl x) \))))" section \Soundness\ text \In this section, we prove that the conclusion of every inference rule is a logical consequence of the premises. Thus a clause set is unsatisfiable if the empty clause is derivable. For each rule, we first prove that all ground instances of the conclusion are entailed by the corresponding instances of the parent clauses, then we lift the result to first-order clauses. The proof is standard and straightforward, but note that we also prove that the derived clauses are finite and well-constrained.\ lemma cannot_validate_contradictary_literals : assumes "l = Neg (Eq t t)" assumes "fo_interpretation I" shows "\ (validate_ground_lit I l)" proof - from assms(2) have "congruence I" unfolding fo_interpretation_def by auto then have "I t t" unfolding congruence_def reflexive_def equivalence_relation_def by auto from this and assms(1) show ?thesis by auto qed lemma ground_reflexion_is_sound : assumes "finite (cl_ecl C)" assumes "reflexion C D \ k C'" assumes "(ground_clause (subst_cl (cl_ecl D) \))" shows "clause_entails_clause (subst_cl (subst_cl (cl_ecl C) \) \) (subst_cl (cl_ecl D) \)" proof (rule ccontr) let ?C = "(cl_ecl C)" let ?D = "(cl_ecl D)" let ?C' = "(subst_cl (subst_cl (cl_ecl C) \) \)" let ?D' = "(subst_cl (cl_ecl D) \)" assume "\ (clause_entails_clause ?C' ?D')" then obtain I where "validate_clause I ?C'" and "\ (validate_clause I ?D')" "fo_interpretation I" unfolding clause_entails_clause_def by auto from assms(2) obtain L1 t s where "?D = (subst_cl (?C - { L1 }) \)" and "orient_lit_inst L1 t s neg \" and "ck_unifier t s \ k" using reflexion_def [of C D \ k] by auto from assms(1) have "finite (subst_cl (subst_cl ?C \) \)" by auto then obtain \ where i: "ground_clause (subst_cl (subst_cl (subst_cl ?C \) \) \)" using ground_instance_exists [of "(subst_cl (subst_cl ?C \) \)"] by auto let ?CC = "(subst_cl (subst_cl (subst_cl ?C \) \) \)" let ?\'' = "comp \ \" let ?\' = "comp ?\'' \" have "?CC = (subst_cl (subst_cl ?C ?\'') \)" using composition_of_substs_cl [of ?C] by auto then have "?CC = (subst_cl ?C ?\')" using composition_of_substs_cl [of ?C] by auto from \validate_clause I (subst_cl (subst_cl (cl_ecl C) \) \)\ have "validate_ground_clause I ?CC" using i validate_clause.simps by blast then obtain l' where "l' \ ?CC" and "validate_ground_lit I l'" by auto from \l' \ ?CC\ and \?CC = (subst_cl ?C ?\')\ obtain l where "l \ ?C" and "l' = (subst_lit l ?\')" using subst_cl.simps by blast have "subst_lit l \ \ ?D" proof (rule ccontr) assume "subst_lit l \ \ ?D" from this and \?D = (subst_cl (?C - { L1 }) \)\ and \l \ ?C\ have "l = L1" by auto from this and \orient_lit_inst L1 t s neg \\ have "l = (Neg (Eq t s)) \ l = (Neg (Eq s t))" unfolding orient_lit_inst_def by auto from \ck_unifier t s \ k\ have "subst t \ = subst s \" using ck_unifier_thm by auto then have "subst (subst (subst t \) \) \ = subst (subst (subst s \) \) \" by auto then have "(subst t ?\') = subst s ?\'" by auto from this and \l = (Neg (Eq t s)) \ l = (Neg (Eq s t))\ have "(subst_lit l ?\') = (Neg (Eq (subst t ?\') (subst t ?\')))" by auto from this and \fo_interpretation I\ have "\ (validate_ground_lit I (subst_lit l ?\'))" using cannot_validate_contradictary_literals [of "(subst_lit l ?\')" "(subst t ?\')" I] by auto from this and \l' = subst_lit l ?\'\ and \validate_ground_lit I l'\ show False by auto qed from \subst_lit l \ \ ?D\ and \l' = subst_lit l ?\'\ have "l' \ (subst_cl (subst_cl ?D \) \)" using subst_cl.simps composition_of_substs_lit mem_Collect_eq by (metis (mono_tags, lifting)) from this and \validate_ground_lit I l'\ have "validate_ground_clause I (subst_cl (subst_cl ?D \) \)" by auto from \ground_clause (subst_cl ?D \)\ have "(subst_cl ?D \) = (subst_cl (subst_cl ?D \) \)" using substs_preserve_ground_clause [of "(subst_cl ?D \)" \] by blast from this and \validate_ground_clause I (subst_cl (subst_cl ?D \) \)\ have "validate_ground_clause I (subst_cl ?D \)" by force from this and assms(3) and \\ validate_clause I (subst_cl (cl_ecl D) \)\ show False using substs_preserve_ground_clause validate_clause.elims(3) by metis qed lemma reflexion_is_sound : assumes "finite (cl_ecl C)" assumes "reflexion C D \ k C'" shows "clause_entails_clause (cl_ecl C) (cl_ecl D)" proof (rule ccontr) let ?C = "(cl_ecl C)" let ?D = "(cl_ecl D)" assume "\ (clause_entails_clause ?C ?D)" then obtain I where "validate_clause I ?C" and "\ (validate_clause I ?D)" "fo_interpretation I" unfolding clause_entails_clause_def by auto from \\ (validate_clause I ?D)\ obtain \ where D_false: "\ (validate_ground_clause I (subst_cl ?D \))" and "(ground_clause (subst_cl ?D \))" by auto have "validate_clause I (subst_cl (subst_cl ?C \) \)" using \validate_clause I (cl_ecl C)\ instances_are_entailed by blast from this and assms(1) and assms(2) have "validate_clause I (subst_cl ?D \)" using ground_reflexion_is_sound unfolding clause_entails_clause_def using \fo_interpretation I\ \ground_clause (subst_cl (cl_ecl D) \)\ by blast from this and D_false show False by (metis \ground_clause (subst_cl (cl_ecl D) \)\ substs_preserve_ground_clause validate_clause.elims(1)) qed lemma orient_lit_semantics_pos : assumes "fo_interpretation I" assumes "orient_lit_inst l u v pos \" assumes "validate_ground_lit I (subst_lit l \)" shows "I (subst u \) (subst v \)" proof - let ?u = "subst u \" let ?v = "subst v \" from assms(2) have "l = (Pos (Eq u v)) \ l = (Pos (Eq v u))" using orient_lit_inst_def by auto from this and assms(3) have "validate_ground_eq I (Eq ?u ?v) \ validate_ground_eq I (Eq ?v ?u)" by auto then have "I ?u ?v \ I ?v ?u" by auto from this and \fo_interpretation I\ show "I ?u ?v" unfolding fo_interpretation_def congruence_def equivalence_relation_def symmetric_def by blast qed lemma orient_lit_semantics_neg : assumes "fo_interpretation I" assumes "orient_lit_inst l u v neg \" assumes "validate_ground_lit I (subst_lit l \)" shows "\I (subst u \) (subst v \)" proof - let ?u = "subst u \" let ?v = "subst v \" from assms(2) have "l = (Neg (Eq u v)) \ l = (Neg (Eq v u))" using orient_lit_inst_def by auto from this and assms(3) have "\validate_ground_eq I (Eq ?u ?v) \ \validate_ground_eq I (Eq ?v ?u)" by auto then have "\I ?u ?v \ \I ?v ?u" by auto from this and \fo_interpretation I\ show "\I ?u ?v" unfolding fo_interpretation_def congruence_def equivalence_relation_def symmetric_def by blast qed lemma orient_lit_semantics_replacement : assumes "fo_interpretation I" assumes "orient_lit_inst l u v polarity \" assumes "validate_ground_lit I (subst_lit l \)" assumes "I (subst u \) (subst u' \)" shows "validate_ground_lit I (subst_lit (mk_lit polarity (Eq u' v)) \)" proof - from assms(2) obtain e where "l = Pos e \ l = Neg e" and "e = Eq u v \ e = Eq v u" unfolding orient_lit_inst_def by auto have "polarity = pos \ polarity = neg" using sign.exhaust by blast then show ?thesis proof assume "polarity = pos" from this and assms(1) and assms(2) and \validate_ground_lit I (subst_lit l \)\ have "I (subst u \) (subst v \)" using orient_lit_semantics_pos by auto from this and assms(1) and \I (subst u \) (subst u' \)\ have "I (subst u' \) (subst v \)" unfolding fo_interpretation_def congruence_def equivalence_relation_def symmetric_def transitive_def by blast from this and \polarity = pos\ show ?thesis by auto next assume "polarity = neg" from this and assms(1) and assms(2) and \validate_ground_lit I (subst_lit l \)\ have "\I (subst u \) (subst v \)" using orient_lit_semantics_neg by blast from this and assms(1) and \I (subst u \) (subst u' \)\ have "\I (subst u' \) (subst v \)" unfolding fo_interpretation_def congruence_def equivalence_relation_def symmetric_def transitive_def by blast from this and \polarity = neg\ show ?thesis by auto qed qed lemma ground_factorization_is_sound : assumes "finite (cl_ecl C)" assumes "factorization C D \ k C'" assumes "(ground_clause (subst_cl (cl_ecl D) \))" shows "clause_entails_clause (subst_cl (subst_cl (cl_ecl C) \) \) (subst_cl (cl_ecl D) \)" proof (rule ccontr) let ?C = "(cl_ecl C)" let ?D = "(cl_ecl D)" assume "\ clause_entails_clause (subst_cl (subst_cl (cl_ecl C) \) \) (subst_cl (cl_ecl D) \)" then obtain I where "validate_clause I (subst_cl (subst_cl (cl_ecl C) \) \)" and "\ (validate_clause I (subst_cl (cl_ecl D) \))" and "fo_interpretation I" unfolding clause_entails_clause_def by auto from assms(2) obtain L1 L2 L' t s u v where "orient_lit_inst L1 t s pos \" and "orient_lit_inst L2 u v pos \" and "ck_unifier t u \ k" and "L' = Neg (Eq s v)" and "(?D = (subst_cl ( (?C - { L2 }) \ { L' } )) \)" and "L1 \ L2" and "L1 \ ?C" using factorization_def by auto from assms(1) have "finite (subst_cl (subst_cl ?C \) \)" by auto then obtain \ where i: "ground_clause (subst_cl (subst_cl (subst_cl ?C \) \) \)" using ground_instance_exists [of "(subst_cl (subst_cl ?C \) \)"] by auto let ?CC = "(subst_cl (subst_cl (subst_cl ?C \) \) \)" let ?\'' = "comp \ \" let ?\' = "comp ?\'' \" have "?CC = (subst_cl (subst_cl ?C ?\'') \)" using composition_of_substs_cl [of ?C] by auto then have "?CC = (subst_cl ?C ?\')" using composition_of_substs_cl [of ?C] by auto from \validate_clause I (subst_cl (subst_cl (cl_ecl C) \) \)\ have "validate_ground_clause I ?CC" using i validate_clause.simps by blast then obtain l' where "l' \ ?CC" and "validate_ground_lit I l'" by auto from \l' \ ?CC\ and \?CC = (subst_cl ?C ?\')\ obtain l where "l \ ?C" and "l' = (subst_lit l ?\')" using subst_cl.simps by blast from \\ validate_clause I (subst_cl (cl_ecl D) \)\ have "\ validate_ground_clause I (subst_cl ?D \)" using assms(3) substs_preserve_ground_clause validate_clause.elims(3) by metis from \ground_clause (subst_cl ?D \)\ have "(subst_cl ?D \) = (subst_cl (subst_cl ?D \) \)" using substs_preserve_ground_clause [of "(subst_cl ?D \)" \] by blast from this and \\ validate_ground_clause I (subst_cl ?D \)\ have "\ validate_ground_clause I (subst_cl (subst_cl ?D \) \)" by force from \(?D = (subst_cl ( (?C - { L2 }) \ { L' } )) \)\ have "(subst_lit L' \) \ ?D" by auto then have "(subst_lit (subst_lit (subst_lit L' \) \) \) \ (subst_cl (subst_cl ?D \) \)" by auto from this and \\ validate_ground_clause I (subst_cl (subst_cl ?D \) \)\ have "\validate_ground_lit I (subst_lit (subst_lit (subst_lit L' \) \) \)" by auto from this and \L'= Neg (Eq s v)\ have "I (subst (subst (subst s \) \) \) (subst (subst (subst v \) \) \)" by auto from this have "I (subst s ?\') (subst v ?\')" by simp have "subst_lit l \ \ ?D" proof (rule ccontr) assume "subst_lit l \ \ ?D" from this and \(?D = (subst_cl ( (?C - { L2 }) \ { L' } )) \)\ and \l \ ?C\ have "l = L2" by auto from \ck_unifier t u \ k\ have "subst t \ = subst u \" using ck_unifier_thm by auto then have "subst (subst (subst t \) \) \ = subst (subst (subst u \) \) \" by auto then have "(subst t ?\') = subst u ?\'" by auto from \validate_ground_lit I l'\ and \l' = (subst_lit l ?\')\ have "validate_ground_lit I (subst_lit l ?\')" by auto from this and \fo_interpretation I\ and \l = L2\ and \orient_lit_inst L2 u v pos \\ have "I (subst u ?\') (subst v ?\')" using orient_lit_semantics_pos by blast from this and \fo_interpretation I\ and \I (subst s ?\') (subst v ?\')\ have "I (subst u ?\') (subst s ?\')" unfolding fo_interpretation_def congruence_def equivalence_relation_def symmetric_def transitive_def by blast from this and \(subst t ?\') = subst u ?\'\ have "I (subst t ?\') (subst s ?\')" by auto from this have "validate_ground_eq I (subst_equation (Eq t s) ?\')" by auto from \I (subst t ?\') (subst s ?\')\ and \fo_interpretation I\ have "I (subst s ?\') (subst t ?\')" unfolding fo_interpretation_def congruence_def equivalence_relation_def symmetric_def by auto from this have "validate_ground_eq I (subst_equation (Eq s t) ?\')" by auto from \orient_lit_inst L1 t s pos \\ have "L1 = (Pos (Eq t s)) \ L1 = (Pos (Eq s t))" unfolding orient_lit_inst_def by auto from this and \validate_ground_eq I (subst_equation (Eq s t) ?\')\ and \validate_ground_eq I (subst_equation (Eq t s) ?\')\ have "validate_ground_lit I (subst_lit L1 ?\')" by auto from \L1 \ ?C\ and \?D = (subst_cl ( (?C - { L2 }) \ { L' } )) \\ and \L1 \ L2\ have "(subst_lit L1 \) \ ?D" by auto then have "(subst_lit (subst_lit (subst_lit L1 \) \) \) \ (subst_cl (subst_cl ?D \) \)" by auto then have "(subst_lit L1 ?\') \ (subst_cl (subst_cl ?D \) \)" using composition_of_substs_lit by metis from this and \validate_ground_lit I (subst_lit L1 ?\')\ and \\ validate_ground_clause I (subst_cl (subst_cl ?D \) \)\ show False by auto qed from \subst_lit l \ \ ?D\ and \l' = subst_lit l ?\'\ have "l' \ (subst_cl (subst_cl ?D \) \)" using subst_cl.simps composition_of_substs_lit mem_Collect_eq by (metis (mono_tags, lifting)) from this and \validate_ground_lit I l'\ have "validate_ground_clause I (subst_cl (subst_cl ?D \) \)" by auto from this and \\ validate_ground_clause I (subst_cl (subst_cl ?D \) \)\ show False by blast qed lemma factorization_is_sound : assumes "finite (cl_ecl C)" assumes "factorization C D \ k C'" shows "clause_entails_clause (cl_ecl C) (cl_ecl D)" proof (rule ccontr) let ?C = "(cl_ecl C)" let ?D = "(cl_ecl D)" assume "\ (clause_entails_clause ?C ?D)" then obtain I where "validate_clause I ?C" and "\ (validate_clause I ?D)" "fo_interpretation I" unfolding clause_entails_clause_def by auto from \\ (validate_clause I ?D)\ obtain \ where D_false: "\ (validate_ground_clause I (subst_cl ?D \))" and "(ground_clause (subst_cl ?D \))" by auto have "validate_clause I (subst_cl (subst_cl ?C \) \)" using \validate_clause I (cl_ecl C)\ instances_are_entailed by blast from this and assms(1) and assms(2) have "validate_clause I (subst_cl ?D \)" using ground_factorization_is_sound unfolding clause_entails_clause_def using \fo_interpretation I\ \ground_clause (subst_cl (cl_ecl D) \)\ by blast from this and D_false show False by (metis \ground_clause (subst_cl (cl_ecl D) \)\ substs_preserve_ground_clause validate_clause.elims(1)) qed lemma ground_superposition_is_sound : assumes "finite (cl_ecl P1)" assumes "finite (cl_ecl P2)" assumes "superposition P1 P2 C \ k C'" assumes "(ground_clause (subst_cl (cl_ecl C) \))" shows "set_entails_clause { (subst_cl (subst_cl (cl_ecl P1) \) \), (subst_cl (subst_cl (cl_ecl P2) \) \) } (subst_cl (cl_ecl C) \)" proof (rule ccontr) let ?P1 = "(cl_ecl P1)" let ?P2 = "(cl_ecl P2)" let ?C = "(cl_ecl C)" assume "\ set_entails_clause { (subst_cl (subst_cl (cl_ecl P1) \) \), (subst_cl (subst_cl (cl_ecl P2) \) \) } (subst_cl (cl_ecl C) \)" then obtain I where "validate_clause I (subst_cl (subst_cl (cl_ecl P1) \) \)" and "validate_clause I (subst_cl (subst_cl (cl_ecl P2) \) \)" and "\ (validate_clause I (subst_cl (cl_ecl C) \))" and "fo_interpretation I" unfolding set_entails_clause_def by (meson insert_iff validate_clause_set.elims(2)) from assms(3) obtain t s u v M p polarity t' u' L L' where "orient_lit_inst M u v pos \" and "orient_lit_inst L t s polarity \" and "subterm t p u'" and "ck_unifier u' u \ k" and "replace_subterm t p v t'" and "L' = mk_lit polarity (Eq t' s)" and "?C = (subst_cl ((?P1 - { L }) \ ((?P2 - { M }) \ { L' } )) \)" using superposition_def by auto let ?P1' = "(subst_cl (subst_cl ?P1 \) \)" let ?P2' = "(subst_cl (subst_cl ?P2 \) \)" from assms(1) have "finite ?P1'" by simp from assms(2) have "finite ?P2'" by simp let ?vars = "(vars_of_cl ?P1') \ (vars_of_cl ?P2')" from \finite ?P1'\ have "finite (vars_of_cl ?P1')" using set_of_variables_is_finite_cl [of ?P1'] by auto from \finite ?P2'\ have "finite (vars_of_cl ?P2')" using set_of_variables_is_finite_cl [of ?P2'] by auto from \finite (vars_of_cl ?P1')\ and \finite (vars_of_cl ?P2')\ have "finite ?vars" by auto then obtain \ where "ground_on ?vars \" using ground_subst_exists by blast then have "ground_on (vars_of_cl ?P1') \" unfolding ground_on_def by auto then have "ground_clause (subst_cl (subst_cl (subst_cl ?P1 \) \) \)" using ground_substs_yield_ground_clause [of "(subst_cl (subst_cl ?P1 \) \)" \] by auto from \ground_on ?vars \\ have "ground_on (vars_of_cl ?P2') \" unfolding ground_on_def by auto then have "ground_clause (subst_cl (subst_cl (subst_cl ?P2 \) \) \)" using ground_substs_yield_ground_clause [of "(subst_cl (subst_cl ?P2 \) \)" \] by auto let ?P1'' = "(subst_cl ?P1' \)" let ?P2'' = "(subst_cl ?P2' \)" let ?\'' = "comp \ \" let ?\' = "comp ?\'' \" have "?P1'' = (subst_cl (subst_cl ?P1 ?\'') \)" using composition_of_substs_cl [of ?P1] by auto then have "?P1'' = (subst_cl ?P1 ?\')" using composition_of_substs_cl [of ?P1] by auto from \ground_clause (subst_cl (subst_cl (subst_cl (cl_ecl P1) \) \) \)\ and \validate_clause I (subst_cl (subst_cl (cl_ecl P1) \) \)\ have "validate_ground_clause I ?P1''" using validate_clause.simps by blast then obtain l1' where "l1' \ ?P1''" and "validate_ground_lit I l1'" by auto have "?P2'' = (subst_cl (subst_cl ?P2 ?\'') \)" using composition_of_substs_cl [of ?P2] by auto then have "?P2'' = (subst_cl ?P2 ?\')" using composition_of_substs_cl [of ?P2] by auto from \ground_clause (subst_cl (subst_cl (subst_cl (cl_ecl P2) \) \) \)\ \validate_clause I (subst_cl (subst_cl (cl_ecl P2) \) \)\ have "validate_ground_clause I ?P2''" using validate_clause.simps by blast then obtain l2' where "l2' \ ?P2''" and "validate_ground_lit I l2'" by auto from \l1' \ ?P1''\ and \?P1'' = (subst_cl ?P1 ?\')\ obtain l1 where "l1 \ ?P1" and "l1' = (subst_lit l1 ?\')" using subst_cl.simps by blast from \l2' \ ?P2''\ and \?P2'' = (subst_cl ?P2 ?\')\ obtain l2 where "l2 \ ?P2" and "l2' = (subst_lit l2 ?\')" using subst_cl.simps by blast let ?C' = "(subst_cl (subst_cl ?C \) \)" from \ground_clause (subst_cl ?C \)\ have "(subst_cl ?C \) = (subst_cl (subst_cl ?C \) \)" using substs_preserve_ground_clause [of "(subst_cl ?C \)" \] by blast from \\ validate_clause I (subst_cl (cl_ecl C) \)\ have "\ validate_ground_clause I ?C'" by (metis assms(4) substs_preserve_ground_clause validate_clause.simps) have "l1 = L" proof (rule ccontr) assume "l1 \ L" from this and \l1 \ ?P1\ and \?C = (subst_cl ((?P1 - { L }) \ ((?P2 - { M }) \ { L' } )) \)\ have "(subst_lit l1 \) \ ?C" by auto from this have "(subst_lit (subst_lit (subst_lit l1 \) \) \) \ ?C'" by auto from this and \l1' = (subst_lit l1 ?\')\ have "l1' \ ?C'" by (simp add: composition_of_substs_lit) from this and \validate_ground_lit I l1'\ have "validate_ground_clause I ?C'" by auto from this and \\ validate_ground_clause I (subst_cl (subst_cl (cl_ecl C) \) \)\ show False by auto qed have "l2 = M" proof (rule ccontr) assume "l2 \ M" from this and \l2 \ ?P2\ and \?C = (subst_cl ((?P1 - { L }) \ ((?P2 - { M }) \ { L' } )) \)\ have "(subst_lit l2 \) \ ?C" by auto from this have "(subst_lit (subst_lit (subst_lit l2 \) \) \) \ ?C'" by auto from this and \l2' = (subst_lit l2 ?\')\ have "l2' \ ?C'" by (simp add: composition_of_substs_lit) from this and \validate_ground_lit I l2'\ have "validate_ground_clause I ?C'" by auto from this and \\ validate_ground_clause I (subst_cl (subst_cl (cl_ecl C) \) \)\ show False by auto qed from \orient_lit_inst M u v pos \\ and \l2 = M\ and \fo_interpretation I\ and \validate_ground_lit I l2'\ and \l2' = (subst_lit l2 ?\')\ have "I (subst u ?\') (subst v ?\')" using orient_lit_semantics_pos by blast from \subterm t p u'\ have "subterm (subst t ?\') p (subst u' ?\')" using substs_preserve_subterms [of t p u'] by metis from \ck_unifier u' u \ k\ have "(subst u \) = (subst u' \)" using ck_unifier_thm [of u' u \ k] by auto from this have "(subst (subst (subst u \) \) \) = (subst (subst (subst u' \) \ ) \)" by auto from this have "(subst u ?\') = (subst u' ?\')" using composition_of_substs by auto from \(subst u ?\') = (subst u' ?\')\ and \I (subst u ?\') (subst v ?\')\ have "I (subst u' ?\') (subst v ?\')" by auto from \subterm t p u'\ and \I (subst u' ?\') (subst v ?\')\ and \fo_interpretation I\ and \replace_subterm t p v t'\ have "I (subst t ?\') (subst t' ?\')" unfolding fo_interpretation_def using replacement_preserves_congruences [of I u' ?\' v t p t'] by auto from \l1 = L\ and \fo_interpretation I\ and \validate_ground_lit I l1'\ and \l1' = (subst_lit l1 ?\')\ and \orient_lit_inst L t s polarity \\ and \I (subst t ?\') (subst t' ?\')\ and \L' = mk_lit polarity (Eq t' s)\ have "validate_ground_lit I (subst_lit L' ?\')" using orient_lit_semantics_replacement [of I L t s polarity \ ?\' t'] by blast from \?C = (subst_cl ((?P1 - { L }) \ ((?P2 - { M }) \ { L' } )) \)\ have "subst_lit L' \ \ ?C" by auto then have "subst_lit (subst_lit (subst_lit L' \) \) \ \ ?C'" by auto then have "subst_lit L' ?\' \ ?C'" by (simp add: composition_of_substs_lit) from this and \validate_ground_lit I (subst_lit L' ?\')\ and \\validate_ground_clause I ?C'\ show False by auto qed lemma superposition_is_sound : assumes "finite (cl_ecl P1)" assumes "finite (cl_ecl P2)" assumes "superposition P1 P2 C \ k C'" shows "set_entails_clause { cl_ecl P1, cl_ecl P2 } (cl_ecl C)" proof (rule ccontr) let ?P1 = "(cl_ecl P1)" let ?P2 = "(cl_ecl P2)" let ?C = "(cl_ecl C)" assume "\ set_entails_clause { cl_ecl P1, cl_ecl P2 } (cl_ecl C)" then obtain I where "validate_clause I ?P1" and "validate_clause I ?P2" and "\ (validate_clause I ?C)" and "fo_interpretation I" unfolding set_entails_clause_def by (meson insert_iff validate_clause_set.elims(2)) from \\ (validate_clause I ?C)\ obtain \ where "\ (validate_ground_clause I (subst_cl ?C \))" and "(ground_clause (subst_cl ?C \))" by auto have P1_true: "validate_clause I (subst_cl (subst_cl ?P1 \) \)" using \validate_clause I (cl_ecl P1)\ instances_are_entailed by blast have P2_true: "validate_clause I (subst_cl (subst_cl ?P2 \) \)" using \validate_clause I (cl_ecl P2)\ instances_are_entailed by blast have "\ (validate_clause I (subst_cl ?C \))" by (metis \\ validate_ground_clause I (subst_cl (cl_ecl C) \)\ \ground_clause (subst_cl (cl_ecl C) \)\ substs_preserve_ground_clause validate_clause.elims(1)) let ?S = "{ (subst_cl (subst_cl (cl_ecl P1) \) \), (subst_cl (subst_cl (cl_ecl P2) \) \) }" from P1_true and P2_true have "validate_clause_set I ?S" by (metis insert_iff singletonD validate_clause_set.elims(3)) from this and \\ (validate_clause I (subst_cl ?C \))\ \fo_interpretation I\ have "\ set_entails_clause ?S (subst_cl (cl_ecl C) \)" using set_entails_clause_def by blast from this and assms(1) and assms(2) and assms(3) and \(ground_clause (subst_cl ?C \))\ show False using ground_superposition_is_sound by auto qed lemma superposition_preserves_finiteness: assumes "finite (cl_ecl P1)" assumes "finite (cl_ecl P2)" assumes "superposition P1 P2 C \ k C'" shows "finite (cl_ecl C) \ (finite C')" proof - from assms(3) obtain L M L' where def_C: "(cl_ecl C) = (subst_cl (((cl_ecl P1) - { L }) \ (((cl_ecl P2) - { M }) \ { L' } )) \)" and def_C': "C' = (((cl_ecl P1) - { L }) \ (((cl_ecl P2) - { M }) \ { L' } ))" using superposition_def by auto from assms(1) and assms(2) have "finite (((cl_ecl P1) - { L }) \ (((cl_ecl P2) - { M }) \ { L' } ))" by auto from this and def_C def_C' show ?thesis using substs_preserve_finiteness by auto qed lemma reflexion_preserves_finiteness: assumes "finite (cl_ecl P1)" assumes "reflexion P1 C \ k C'" shows "finite (cl_ecl C) \ (finite C')" proof - from assms(2) obtain L1 where def_C: "(cl_ecl C) = (subst_cl ((cl_ecl P1) - { L1 }) \)" and def_C': "C' = ((cl_ecl P1) - { L1 })" using reflexion_def by auto from assms(1) have "finite ((cl_ecl P1) - { L1 })" by auto from this and def_C def_C' show ?thesis using substs_preserve_finiteness by auto qed lemma factorization_preserves_finiteness: assumes "finite (cl_ecl P1)" assumes "factorization P1 C \ k C'" shows "finite (cl_ecl C) \ (finite C')" proof - from assms(2) obtain L2 L' where def_C: "(cl_ecl C) = (subst_cl ( ((cl_ecl P1) - { L2 }) \ { L' } ) \)" and def_C': "C' = ( ((cl_ecl P1) - { L2 }) \ { L' } )" using factorization_def by auto from assms(1) have "(finite (((cl_ecl P1) - { L2 }) \ { L' }))" by auto from this and def_C def_C' show ?thesis using substs_preserve_finiteness by auto qed lemma derivable_clauses_are_finite: assumes "derivable C P S \ k C'" assumes "\x \ P. (finite (cl_ecl x))" shows "finite (cl_ecl C) \ (finite C')" proof (rule ccontr) assume hyp: "\ (finite (cl_ecl C) \ (finite C'))" have not_sup: "\ (\P1 P2. (P1 \ P \ P2 \ P \ superposition P1 P2 C \ k C'))" proof assume "(\P1 P2. (P1 \ P \ P2 \ P \ superposition P1 P2 C \ k C'))" then obtain P1 P2 where "P1 \ P" "P2 \ P" "superposition P1 P2 C \ k C'" by auto from \P1 \ P\ and assms(2) have "finite (cl_ecl P1)" by auto from \P2 \ P\ and assms(2) have "finite (cl_ecl P2)" by auto from \(finite (cl_ecl P1))\ and \(finite (cl_ecl P2))\ and \superposition P1 P2 C \ k C'\ have "finite (cl_ecl C) \ (finite C')" using superposition_preserves_finiteness [of P1 P2 C \] by auto then show False using hyp by auto qed have not_ref: "\ (\P1. (P1 \ P \ reflexion P1 C \ k C'))" proof assume "(\P1. (P1 \ P \ reflexion P1 C \ k C'))" then obtain P1 where "P1 \ P" "reflexion P1 C \ k C'" by auto from \P1 \ P\ and assms(2) have "finite (cl_ecl P1)" by auto from \(finite (cl_ecl P1))\ and \reflexion P1 C \ k C'\ have "finite (cl_ecl C) \ (finite C')" using reflexion_preserves_finiteness [of P1 C \] by auto then show False using hyp by auto qed have not_fact: "\ (\P1. (P1 \ P \ factorization P1 C \ k C'))" proof assume "(\P1. (P1 \ P \ factorization P1 C \ k C'))" then obtain P1 where "P1 \ P" " factorization P1 C \ k C'" by auto from \P1 \ P\ and assms(2) have "finite (cl_ecl P1)" by auto from \(finite (cl_ecl P1))\ and \ factorization P1 C \ k C'\ have "finite (cl_ecl C) \ (finite C')" using factorization_preserves_finiteness [of P1 C \] by auto then show False using hyp by auto qed from not_sup not_ref not_fact and assms(1) show False unfolding derivable_def by blast qed lemma derivable_clauses_lemma: assumes "derivable C P S \ k C'" shows "((cl_ecl C) = (subst_cl C' \))" proof (rule ccontr) assume hyp: "\ ((cl_ecl C) = (subst_cl C' \))" have not_sup: "\ (\P1 P2. (P1 \ S \ P2 \ S \ superposition P1 P2 C \ k C'))" proof assume "(\P1 P2. (P1 \ S \ P2 \ S \ superposition P1 P2 C \ k C'))" then obtain P1 P2 where "P1 \ S" "P2 \ S" "superposition P1 P2 C \ k C'" by auto from \superposition P1 P2 C \ k C'\ obtain Cl_C Cl_P1 L Cl_P2 M L' T where "Cl_C = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \)" "(C' = (Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } ))" "C = (Ecl Cl_C T)" unfolding superposition_def by blast from \Cl_C = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \)\ \(C' = (Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } ))\ \C = (Ecl Cl_C T)\ hyp show False by auto qed have not_ref: "\ (\P1. (P1 \ S \ reflexion P1 C \ k C'))" proof assume "(\P1. (P1 \ S \ reflexion P1 C \ k C'))" then obtain P1 where "P1 \ S" "reflexion P1 C \ k C'" by auto from \reflexion P1 C \ k C'\ obtain T Cl_C Cl_P L1 where "C = (Ecl Cl_C T)" "Cl_C = (subst_cl ((Cl_P - { L1 }) )) \" "(C' = ((Cl_P - { L1 }) ))" unfolding reflexion_def by blast from \Cl_C = (subst_cl ((Cl_P - { L1 }) )) \\ \(C' = ((Cl_P - { L1 }) ))\ \C = (Ecl Cl_C T)\ hyp show False by auto qed have not_fact: "\ (\P1. (P1 \ S \ factorization P1 C \ k C'))" proof assume "(\P1. (P1 \ S \ factorization P1 C \ k C'))" then obtain P1 where "P1 \ S" "factorization P1 C \ k C'" by auto from \factorization P1 C \ k C'\ obtain T Cl_C Cl_P L' L2 where "C = (Ecl Cl_C T)" "Cl_C = (subst_cl ( (Cl_P - { L2 }) \ { L' } )) \" "C' = ( (Cl_P - { L2 }) \ { L' } )" unfolding factorization_def by blast from \Cl_C = (subst_cl ( (Cl_P - { L2 }) \ { L' } )) \\ \C' = ( (Cl_P - { L2 }) \ { L' } )\ \C = (Ecl Cl_C T)\ hyp show False by auto qed from not_sup not_ref not_fact and assms(1) show False unfolding derivable_def by blast qed lemma substs_preserves_decompose_literal: assumes "decompose_literal L t s polarity" shows "decompose_literal (subst_lit L \) (subst t \) (subst s \) polarity" proof - let ?L = "(subst_lit L \)" let ?t = "(subst t \)" let ?s = "(subst s \)" have "polarity = pos \ polarity = neg" using sign.exhaust by auto then show ?thesis proof assume "polarity = pos" from this and assms(1) have "L = Pos (Eq t s) \ L = Pos (Eq s t)" unfolding decompose_literal_def decompose_equation_def by auto from \L = Pos (Eq t s) \ L = Pos (Eq s t)\ have "?L = Pos (Eq ?t ?s) \ ?L = Pos (Eq ?s ?t)" by auto from this \polarity = pos\ show ?thesis unfolding decompose_literal_def decompose_equation_def by auto next assume "polarity = neg" from this and assms(1) have "L = Neg (Eq t s) \ L = Neg (Eq s t)" unfolding decompose_literal_def decompose_equation_def by auto from this \polarity = neg\ show ?thesis unfolding decompose_literal_def decompose_equation_def by auto qed qed lemma substs_preserve_dom_trm: assumes "dom_trm t C" shows "dom_trm (subst t \) (subst_cl C \)" proof - let ?t = "(subst t \)" from assms(1) have "(\ L u v p. (L \ C \ (decompose_literal L u v p) \ (( (p = neg \ t = u) \ (t,u) \ trm_ord))))" unfolding dom_trm_def by auto from this obtain L u v p where "L \ C" "decompose_literal L u v p" "(( (p = neg \ t = u) \ (t,u) \ trm_ord))" unfolding dom_trm_def by blast let ?u = "(subst u \)" from \L \ C\ have "(subst_lit L \) \ (subst_cl C \)" by auto from \decompose_literal L u v p\ have "decompose_literal (subst_lit L \) (subst u \) (subst v \) p" using substs_preserves_decompose_literal by metis from \(( (p = neg \ t = u) \ (t,u) \ trm_ord))\ have "(( (p = neg \ ?t = ?u) \ (?t,?u) \ trm_ord))" using trm_ord_subst by auto from this \(subst_lit L \) \ (subst_cl C \)\ \decompose_literal (subst_lit L \) (subst u \) (subst v \) p\ show "dom_trm (subst t \) (subst_cl C \)" unfolding dom_trm_def by auto qed lemma substs_preserve_well_constrainedness: assumes "well_constrained C" shows "well_constrained (subst_ecl C \)" proof (rule ccontr) assume "\?thesis" from this obtain y where "y \ trms_ecl (subst_ecl C \)" and "\ dom_trm y (cl_ecl (subst_ecl C \))" unfolding well_constrained_def by auto obtain Cl_C T where "C = (Ecl Cl_C T)" using "eclause.exhaust" by auto from this have "(subst_ecl C \) = (Ecl (subst_cl Cl_C \) (subst_set T \))" by auto from this have "(cl_ecl (subst_ecl C \) = (subst_cl Cl_C \))" and "trms_ecl (subst_ecl C \) = (subst_set T \)" by auto from \(cl_ecl (subst_ecl C \) = (subst_cl Cl_C \))\ \C = (Ecl Cl_C T)\ have "(cl_ecl (subst_ecl C \) = (subst_cl (cl_ecl C) \))" by auto from \y \ trms_ecl (subst_ecl C \)\ \C = (Ecl Cl_C T)\ obtain z where "z \ T" and "y = (subst z \)" by auto from \z \ T\ assms(1) \C = (Ecl Cl_C T)\ have "dom_trm z (cl_ecl C)" unfolding well_constrained_def by auto from this have "dom_trm (subst z \) (subst_cl (cl_ecl C) \)" using substs_preserve_dom_trm by auto from this \y = (subst z \)\ have "dom_trm y (subst_cl (cl_ecl C) \)" by auto from this \(cl_ecl (subst_ecl C \) = (subst_cl (cl_ecl C) \))\ \\ dom_trm y (cl_ecl (subst_ecl C \))\ show False by auto qed lemma ck_trms_sound: assumes "T = get_trms D (dom_trms C E) k" shows "T \ (dom_trms C E)" proof (cases) assume "k = FirstOrder" from this and assms have "T = filter_trms D (dom_trms C E)" unfolding get_trms_def by auto from this show ?thesis using filter_trms_inclusion by blast next assume "k \ FirstOrder" from this and assms have "T = (dom_trms C E)" unfolding get_trms_def by auto from this show ?thesis using filter_trms_inclusion by blast qed lemma derivable_clauses_are_well_constrained: assumes "derivable C P S \ k C'" shows "well_constrained C" proof (rule ccontr) assume hyp: "\ well_constrained C" then obtain y where "y \ trms_ecl C" and "\ dom_trm y (cl_ecl C)" unfolding well_constrained_def by auto have not_sup: "\ (\P1 P2. (P1 \ S \ P2 \ S \ superposition P1 P2 C \ k C'))" proof assume "(\P1 P2. (P1 \ S \ P2 \ S \ superposition P1 P2 C \ k C'))" then obtain P1 P2 where "P1 \ S" "P2 \ S" "superposition P1 P2 C \ k C'" by auto from \superposition P1 P2 C \ k C'\ obtain Cl_C T E where "T = (get_trms Cl_C (dom_trms Cl_C (subst_set E \)) k)" "Cl_C = (subst_cl C' \)" "C = (Ecl Cl_C T)" unfolding superposition_def by blast from \T = (get_trms Cl_C (dom_trms Cl_C (subst_set E \)) k)\ have "T \(dom_trms Cl_C (subst_set E \))" using ck_trms_sound by metis from this and \y \ trms_ecl C\ and \C = (Ecl Cl_C T)\ have "y \ (dom_trms (cl_ecl C) (subst_set E \))" by auto from this and \\ dom_trm y (cl_ecl C)\ show False unfolding dom_trms_def by auto qed have not_ref: "\ (\P1. (P1 \ S \ reflexion P1 C \ k C'))" proof assume "(\P1. (P1 \ S \ reflexion P1 C \ k C'))" then obtain P1 where "P1 \ S" "reflexion P1 C \ k C'" by auto from \reflexion P1 C \ k C'\ obtain T Cl_C E where "T = (get_trms Cl_C (dom_trms Cl_C (subst_set E \)) k)" "Cl_C = (subst_cl C' \)" "C = (Ecl Cl_C T)" unfolding reflexion_def by blast from \T = (get_trms Cl_C (dom_trms Cl_C (subst_set E \)) k)\ have "T \(dom_trms Cl_C (subst_set E \))" using ck_trms_sound by metis from this and \y \ trms_ecl C\ and \C = (Ecl Cl_C T)\ have "y \ (dom_trms (cl_ecl C) (subst_set E \))" by auto from this and \\ dom_trm y (cl_ecl C)\ show False unfolding dom_trms_def by auto qed have not_fact: "\ (\P1. (P1 \ S \ factorization P1 C \ k C'))" proof assume "(\P1. (P1 \ S \ factorization P1 C \ k C'))" then obtain P1 where "P1 \ S" "factorization P1 C \ k C'" by auto from \factorization P1 C \ k C'\ obtain T Cl_C E where "T = (get_trms Cl_C (dom_trms Cl_C (subst_set E \)) k)" "Cl_C = (subst_cl C' \)" "C = (Ecl Cl_C T)" unfolding factorization_def by blast from \T = (get_trms Cl_C (dom_trms Cl_C (subst_set E \)) k)\ have "T \(dom_trms Cl_C (subst_set E \))" using ck_trms_sound by metis from this and \y \ trms_ecl C\ and \C = (Ecl Cl_C T)\ have "y \ (dom_trms (cl_ecl C) (subst_set E \))" by auto from this and \\ dom_trm y (cl_ecl C)\ show False unfolding dom_trms_def by auto qed from not_sup not_ref not_fact and assms(1) show False unfolding derivable_def by blast qed lemma derivable_clauses_are_entailed: assumes "derivable C P S \ k C'" assumes "validate_clause_set I (cl_ecl ` P)" assumes "fo_interpretation I" assumes "\x \ P. (finite (cl_ecl x))" shows "validate_clause I (cl_ecl C)" proof (rule ccontr) assume "\validate_clause I (cl_ecl C)" have not_sup: "\ (\P1 P2. (P1 \ S \ P2 \ S \ P = { P1, P2 } \ superposition P1 P2 C \ k C'))" proof assume "(\P1 P2. (P1 \ S \ P2 \ S \ P = { P1, P2 } \ superposition P1 P2 C \ k C'))" from this obtain P1 P2 where "P1 \ P" "P2 \ P" and "superposition P1 P2 C \ k C'" by auto from \P1 \ P\ and assms(2) have "validate_clause I (cl_ecl P1)" by auto from \P2 \ P\ and assms(2) have "validate_clause I (cl_ecl P2)" by auto from assms(4) and \P1 \ P\ have "finite (cl_ecl P1)" by auto from assms(4) and \P2 \ P\ have "finite (cl_ecl P2)" by auto from assms(3) and \finite (cl_ecl P1)\ and \finite (cl_ecl P2)\ and \superposition P1 P2 C \ k C'\ have "set_entails_clause { (cl_ecl P1), (cl_ecl P2) } (cl_ecl C)" using superposition_is_sound by blast from this and assms(3) and \validate_clause I (cl_ecl P1)\ and \validate_clause I (cl_ecl P2)\ have "validate_clause I (cl_ecl C)" using set_entails_clause_def [of "{ (cl_ecl P1), (cl_ecl P2) }" "cl_ecl C"] by auto from this and \\validate_clause I (cl_ecl C)\ show False by auto qed have not_fact: "\ (\P1. (P1 \ S \ P = { P1 } \ factorization P1 C \ k C'))" proof assume "(\P1. (P1 \ S \ P = { P1 } \ factorization P1 C \ k C'))" from this obtain P1 where "P1 \ P" and "factorization P1 C \ k C'" by auto from \P1 \ P\ and assms(2) have "validate_clause I (cl_ecl P1)" by auto from assms(4) and \P1 \ P\ have "finite (cl_ecl P1)" by auto from assms(3) and \finite (cl_ecl P1)\ and \factorization P1 C \ k C'\ have "clause_entails_clause (cl_ecl P1) (cl_ecl C)" using factorization_is_sound by auto from this and assms(3) and \validate_clause I (cl_ecl P1)\ have "validate_clause I (cl_ecl C)" unfolding clause_entails_clause_def by auto from this and \\validate_clause I (cl_ecl C)\ show False by auto qed have not_ref: "\ (\P1. (P1 \ S \ P = { P1 } \ reflexion P1 C \ k C'))" proof assume "(\P1. (P1 \ S \ P = { P1 } \ reflexion P1 C \ k C'))" from this obtain P1 where "P1 \ P" and "reflexion P1 C \ k C'" by auto from \P1 \ P\ and assms(2) have "validate_clause I (cl_ecl P1)" by auto from assms(4) and \P1 \ P\ have "finite (cl_ecl P1)" by auto from assms(3) and \finite (cl_ecl P1)\ and \reflexion P1 C \ k C'\ have "clause_entails_clause (cl_ecl P1) (cl_ecl C)" using reflexion_is_sound by auto from this and assms(3) and \validate_clause I (cl_ecl P1)\ have "validate_clause I (cl_ecl C)" unfolding clause_entails_clause_def by auto from this and \\validate_clause I (cl_ecl C)\ show False by auto qed from not_sup not_fact not_ref and assms(1) show False unfolding derivable_def by blast qed lemma all_derived_clauses_are_finite: shows "derivable_ecl C S \ \x \ S. (finite (cl_ecl x)) \ finite (cl_ecl C)" proof (induction rule: derivable_ecl.induct) fix C :: "'a eclause" fix S assume "C \ S" assume "\x \ S. (finite (cl_ecl x))" from this \C \ S\ show "finite (cl_ecl C)" by auto next fix C S fix D :: "'a eclause" assume "derivable_ecl C S" assume "\x \ S. (finite (cl_ecl x))" assume hyp_ind: "\x \ S. (finite (cl_ecl x)) \ finite (cl_ecl C)" "(renaming_cl C D)" from \(renaming_cl C D)\ obtain \ where "D = (subst_ecl C \)" unfolding renaming_cl_def by auto obtain C_Cl T where "C = (Ecl C_Cl T)" using "eclause.exhaust" by auto from this and \D = (subst_ecl C \)\ have "(cl_ecl D) = (subst_cl (cl_ecl C) \)" by auto from this hyp_ind \\x \ S. (finite (cl_ecl x))\ show "finite (cl_ecl D)" using substs_preserve_finiteness by auto next fix P S C S' \ C' assume h: "\x. x \ P \ derivable_ecl x S \ ((\x\S. finite (cl_ecl x)) \ finite (cl_ecl x))" assume "derivable C P S' \ FirstOrder C'" assume "\x\S. finite (cl_ecl x)" from h and \\x\S. finite (cl_ecl x)\ have "\x \ P. (finite (cl_ecl x))" by metis from this and \derivable C P S' \ FirstOrder C'\ show "finite (cl_ecl C)" using derivable_clauses_are_finite by auto qed lemma all_derived_clauses_are_wellconstrained: shows "derivable_ecl C S \ \x \ S. (well_constrained x) \ well_constrained C" proof (induction rule: derivable_ecl.induct) fix C :: "'a eclause" fix S assume "C \ S" assume "\x \ S. (well_constrained x)" from this \C \ S\ show "well_constrained C" by auto next fix C S fix D :: "'a eclause" assume "derivable_ecl C S" assume "\x \ S. (well_constrained x)" assume hyp_ind: "\x \ S. (well_constrained x) \ well_constrained C" "(renaming_cl C D)" from \\x \ S. (well_constrained x)\ and hyp_ind have "well_constrained C" by auto from \(renaming_cl C D)\ obtain \ where "D = (subst_ecl C \)" unfolding renaming_cl_def by auto from this and \well_constrained C\ show "well_constrained D" using substs_preserve_well_constrainedness by auto next fix P S C S' \ C' assume "\x. x \ P \ derivable_ecl x S \ (Ball S well_constrained \ well_constrained x)" assume "derivable C P S' \ FirstOrder C'" assume "Ball S well_constrained" from \derivable C P S' \ FirstOrder C'\ show "well_constrained C" using derivable_clauses_are_well_constrained by auto qed lemma SOUNDNESS: shows "derivable_ecl C S \ \x \ S. (finite (cl_ecl x)) \ set_entails_clause (cl_ecl ` S) (cl_ecl C)" proof (induction rule: derivable_ecl.induct) fix C :: "'a eclause" fix S assume "C \ S" assume "\x \ S. (finite (cl_ecl x))" from \C \ S\ show "set_entails_clause (cl_ecl ` S) (cl_ecl C)" unfolding set_entails_clause_def by auto next fix C S fix D :: "'a eclause" assume "derivable_ecl C S" assume "\x \ S. (finite (cl_ecl x))" assume hyp_ind: "\x \ S. (finite (cl_ecl x)) \ set_entails_clause (cl_ecl ` S) (cl_ecl C)" assume "(renaming_cl C D)" from \(renaming_cl C D)\ obtain \ where "D = (subst_ecl C \)" unfolding renaming_cl_def by auto obtain C_Cl T where "C = (Ecl C_Cl T)" using "eclause.exhaust" by auto from this and \D = (subst_ecl C \)\ have "(cl_ecl D) = (subst_cl (cl_ecl C) \)" by auto show "set_entails_clause (cl_ecl ` S) (cl_ecl D)" proof (rule ccontr) assume "\?thesis" from this obtain I where "fo_interpretation I" and i: "validate_clause_set I (cl_ecl `S)" "\validate_clause I (cl_ecl D)" unfolding set_entails_clause_def by auto from \\validate_clause I (cl_ecl D)\ and \(cl_ecl D) = (subst_cl (cl_ecl C) \)\ have "\validate_clause I (cl_ecl C)" using instances_are_entailed by metis from this and \fo_interpretation I\ i have "\set_entails_clause (cl_ecl ` S) (cl_ecl C)" unfolding set_entails_clause_def by auto from this and \\x \ S. (finite (cl_ecl x))\ hyp_ind show False by auto qed next fix P S C S' \ C' assume h: "\x. x \ P \ derivable_ecl x S \ ((\x\S. finite (cl_ecl x)) \ set_entails_clause (cl_ecl ` S) (cl_ecl x))" assume "derivable C P S' \ FirstOrder C'" assume "\x\S. finite (cl_ecl x)" from h and \\x\S. finite (cl_ecl x)\ have i: "\x \ P. set_entails_clause (cl_ecl ` S) (cl_ecl x)" by metis show "set_entails_clause (cl_ecl ` S) (cl_ecl C)" proof (rule ccontr) assume "\?thesis" from this obtain I where "fo_interpretation I" and ii: "validate_clause_set I (cl_ecl `S)" "\validate_clause I (cl_ecl C)" unfolding set_entails_clause_def by auto from h \\x\S. finite (cl_ecl x)\ have "(\x\P. finite (cl_ecl x))" using all_derived_clauses_are_finite by metis from \fo_interpretation I\ i and ii have "\x \ P. (validate_clause I (cl_ecl x))" unfolding set_entails_clause_def by auto from this have "validate_clause_set I (cl_ecl ` P)" by auto from this and \(\x\P. finite (cl_ecl x))\ \fo_interpretation I\ \derivable C P S' \ FirstOrder C'\ have "validate_clause I (cl_ecl C)" using derivable_clauses_are_entailed [of C P S' \ FirstOrder C' I] by blast from this and \\validate_clause I (cl_ecl C)\ show False by auto qed qed lemma REFUTABLE_SETS_ARE_UNSAT: assumes "\x \ S. (finite (cl_ecl x))" assumes "derivable_ecl C S" assumes "(cl_ecl C = {})" shows "\ (satisfiable_clause_set (cl_ecl ` S))" proof assume "(satisfiable_clause_set (cl_ecl ` S))" then obtain I where "fo_interpretation I" and model: "validate_clause_set I (cl_ecl ` S)" unfolding satisfiable_clause_set_def [of "cl_ecl ` S"] by blast from assms(1) assms(2) have "set_entails_clause (cl_ecl ` S) (cl_ecl C)" using SOUNDNESS by metis from this \fo_interpretation I\ and model have "validate_clause I (cl_ecl C)" unfolding set_entails_clause_def by auto from this and assms(3) show False by auto qed section \Redundancy Criteria and Saturated Sets\ text \We define redundancy criteria. We use similar notions as in the Bachmair and Ganzinger paper, the only difference is that we have to handle the sets of irreducible terms associated with the clauses. Indeed, to ensure completeness, we must guarantee that all the terms that are irreducible in the entailing clauses are also irreducible in the entailed one (otherwise some needed inferences could be blocked due the irreducibility condition, as in the basic superposition calculus). Of course, if the attached sets of terms are empty, then this condition trivially holds and the definition collapses to the usual one. We introduce the following relation:\ definition subterms_inclusion :: "'a trm set \ 'a trm set \ bool" where "subterms_inclusion E1 E2 = (\x1 \ E1. \x2 \ E2. (occurs_in x1 x2))" lemma subterms_inclusion_refl: shows "subterms_inclusion E E" proof (rule ccontr) assume "\subterms_inclusion E E" from this obtain x1 where "x1 \ E" and "\ occurs_in x1 x1" unfolding subterms_inclusion_def by force from \\ occurs_in x1 x1\ have "\ (\p. subterm x1 p x1)" unfolding occurs_in_def by auto from this have "\subterm x1 Nil x1" by metis from this show False by auto qed lemma subterms_inclusion_subset: assumes "subterms_inclusion E1 E2" assumes "E2 \ E2'" shows "subterms_inclusion E1 E2'" by (meson assms(1) assms(2) basic_superposition.subterms_inclusion_def basic_superposition_axioms subsetD) lemma set_inclusion_preserve_normalization: assumes "all_trms_irreducible E f" assumes "E' \ E" shows "all_trms_irreducible E' f" by (meson all_trms_irreducible_def assms(1) assms(2) subsetD) lemma subterms_inclusion_preserves_normalization: assumes "all_trms_irreducible E f" assumes "subterms_inclusion E' E" shows "all_trms_irreducible E' f" by (meson all_trms_irreducible_def assms(1) assms(2) occur_in_subterm subterms_inclusion_def) text \We define two notions of redundancy, the first one is for inferences: any derivable clause must be entailed by a set of clauses that are strictly smaller than one of the premises.\ definition redundant_inference :: "'a eclause \ 'a eclause set \ 'a eclause set \ 'a subst \ bool" where "(redundant_inference C S P \) = (\S'. (S' \ (instances S) \ (set_entails_clause (clset_instances S') (cl_ecl C)) \ (\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl C))) \ (\x \ S'. \D' \ P. (((fst x),(snd x)),(D',\)) \ ecl_ord)))" text \The second one is the usual notion for clauses: a clause is redundant if it is entailed by smaller (or equal) clauses.\ definition redundant_clause :: "'a eclause \ 'a eclause set \ 'a subst \ 'a clause \ bool" where "(redundant_clause C S \ C') = (\S'. (S' \ (instances S) \ (set_entails_clause (clset_instances S') (cl_ecl C)) \ (\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl C))) \ (\x \ S'. ( ((mset_ecl ((fst x),(snd x))),(mset_cl (C',\))) \ (mult (mult trm_ord)) \ (mset_ecl ((fst x),(snd x))) = mset_cl (C',\)))))" text \Note that according to the definition above, an extended clause is always redundant w.r.t.\ a clause obtained from the initial one by adding in the attached set of terms a subterm of a term that already occurs in this set. This remark is important because explicitly adding such subterms in the attached set may prune the search space, due to the fact that the containing term can be removed at some point when calling the function @{term "dom_trm"}. Adding the subterm explicitly is thus useful in this case. In practice, the simplest solution may be to assume that the set of irreducible terms is closed under subterm. Of course, a clause is also redundant w.r.t.\ any clause obtained by removing terms in the attached set. In particular, terms can be safely removed from the set of irreducible terms of the entailing clauses if needed to make a given clause redundant.\ lemma self_redundant_clause: assumes "C \ S" assumes "C' = (cl_ecl C)" assumes "ground_clause (subst_cl (cl_ecl C) \)" shows "redundant_clause (subst_ecl C \) S \ C'" proof - obtain Cl_C and T where "C = Ecl Cl_C T" using eclause.exhaust by auto from this have "cl_ecl C = Cl_C" and "trms_ecl C = T" by auto let ?Cl_C = "subst_cl Cl_C \" let ?T = "subst_set T \" let ?C = "subst_ecl C \" from \C = Ecl Cl_C T\ have "?C = (Ecl ?Cl_C ?T)" by auto from this have "cl_ecl ?C = ?Cl_C" and "trms_ecl ?C = ?T" by auto let ?S = "{ (C,\) }" from assms(1) assms(3) have i: "?S \ (instances S)" unfolding instances_def by auto from \cl_ecl C = Cl_C\ have "clset_instances ?S = { ?Cl_C }" unfolding clset_instances_def by auto from this and \cl_ecl ?C = ?Cl_C\ have ii: "set_entails_clause (clset_instances ?S) (cl_ecl ?C)" using set_entails_clause_member by force have iii: "(\x \ ?S. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl ?C)))" proof fix x assume "x \ ?S" from this have "x = (C,\)" by auto from this \C = Ecl Cl_C T\ have "subst_set (trms_ecl (fst x)) (snd x) = ?T" by auto from this and \trms_ecl ?C = ?T\ have "subst_set (trms_ecl (fst x)) (snd x) = (trms_ecl ?C)" by auto from this show "( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl ?C))" using subterms_inclusion_refl by auto qed have iv: "(\x \ ?S. ( ((mset_ecl ((fst x),(snd x))),(mset_cl (C',\))) \ (mult (mult trm_ord)) \ (mset_ecl ((fst x),(snd x))) = mset_cl (C',\)))" proof fix x assume "x \ ?S" from this have "x = (C,\)" by auto from this \C = Ecl Cl_C T\ have "(mset_ecl ((fst x),(snd x))) = (mset_ecl (C,\))" by auto from this \C' = (cl_ecl C)\ have "(mset_ecl ((fst x),(snd x))) = mset_cl (C',\)" by auto from this show "( ((mset_ecl ((fst x),(snd x))),(mset_cl (C',\))) \ (mult (mult trm_ord)) \ (mset_ecl ((fst x),(snd x))) = mset_cl (C',\))" by auto qed from i ii iii iv show ?thesis unfolding redundant_clause_def by metis qed definition trms_subsumes where "trms_subsumes C D \ = ( (subst_cl (cl_ecl C) \) = (cl_ecl D) \ ((subst_set (trms_ecl C) \) \ trms_ecl D))" definition inference_closed where "inference_closed S = (\ P C' D \. (derivable D P S \ FirstOrder C') \ (D \ S))" text \Various notions of saturatedness are defined, depending on the kind of inferences that are considered and on the redundancy criterion.\ text \The first definition is the weakest one: all ground inferences must be redundant (this definition is used for the completeness proof to make it the most general).\ definition ground_inference_saturated :: "'a eclause set \ bool" where "(ground_inference_saturated S) = (\ C P \ C'. (derivable C P S \ Ground C') \ (ground_clause (cl_ecl C)) \ (grounding_set P \) \ (redundant_inference C S P \))" text \The second one states that every ground instance of a first-order inference must be redundant.\ definition inference_saturated :: "'a eclause set \ bool" where "(inference_saturated S) = (\ C P \ C' D \ \. (derivable C P S \ Ground C') \ (ground_clause (cl_ecl C)) \ (grounding_set P \) \ (derivable D P S \ FirstOrder C') \ (trms_subsumes D C \) \ (\ \ \ \ \) \ (redundant_inference (subst_ecl D \) S P \))" text \The last definition is the most restrictive one: every derivable clause must be redundant.\ definition clause_saturated :: "'a eclause set \ bool" where "(clause_saturated S) = (\ C P \ C' D \ \. (derivable C P S \ Ground C') \ (ground_clause (cl_ecl C)) \ (derivable D P S \ FirstOrder C') \ (trms_subsumes D C \) \ (\ \ \ \ \) \ (redundant_clause (subst_ecl D \) S \ C'))" text \We now relate these various notions, so that the forthcoming completeness proof applies to all of them. To this purpose, we have to show that the conclusion of a (ground) inference rule is always strictly smaller than one of the premises.\ lemma conclusion_is_smaller_than_premisses: assumes "derivable C P S \ Ground C'" assumes "\x \ S. (finite (cl_ecl x))" assumes "grounding_set P \" shows "\ D. (D \ P \ (( (mset_cl (C',\)), (mset_ecl (D,\))) \ (mult (mult trm_ord))))" proof (rule ccontr) assume hyp: "\ (\ D. (D \ P \ (( (mset_cl (C',\)), (mset_ecl (D,\))) \ (mult (mult trm_ord)))))" from assms(1) have "P \ S" unfolding derivable_def by auto have not_sup: "\ (\P1 P2. (P1 \ P \ P2 \ P \ superposition P1 P2 C \ Ground C'))" proof assume "(\P1 P2. (P1 \ P \ P2 \ P \ superposition P1 P2 C \ Ground C'))" then obtain P1 P2 where "P1 \ P" "P2 \ P" "superposition P1 P2 C \ Ground C'" by auto from \superposition P1 P2 C \ Ground C'\ obtain L t s u v M L' polarity u' p t' Cl_C NT where "M \ (cl_ecl P2)" "L \ (cl_ecl P1)" "orient_lit_inst M u v pos \" "orient_lit_inst L t s polarity \" "subterm t p u'" "ck_unifier u' u \ Ground" "replace_subterm t p v t'" "L' = mk_lit polarity (Eq t' s)" "(C = (Ecl Cl_C NT))" "(subst u \) \ (subst v \)" "( (subst_lit M \),(subst_lit L \)) \ lit_ord" "strictly_maximal_literal P2 M \" "Cl_C = (subst_cl (((cl_ecl P1) - { L }) \ (((cl_ecl P2) - { M }) \ { L' } )) \)" "C' = (((cl_ecl P1) - { L }) \ (((cl_ecl P2) - { M }) \ { L' } ))" unfolding superposition_def by blast from \P1 \ P\ and assms(2) and \P \ S\ have "finite (cl_ecl P1)" by auto from \P2 \ P\ and assms(2) and \P \ S\ have "finite (cl_ecl P2)" by auto from assms(3) and \P2 \ P\ have "ground_clause (subst_cl (cl_ecl P2) \)" unfolding grounding_set_def by auto from this have "vars_of_cl (subst_cl (cl_ecl P2) \) = {}" by auto from \M \ (cl_ecl P2)\have "(subst_lit M \) \ (subst_cl (cl_ecl P2) \)" by auto from this and \vars_of_cl (subst_cl (cl_ecl P2) \) = {}\ have "vars_of_lit (subst_lit M \) = {}" by auto from \orient_lit_inst M u v pos \\ have "orient_lit (subst_lit M \) (subst u \) (subst v \) pos" using lift_orient_lit by auto from this and \vars_of_lit (subst_lit M \) = {}\ have "vars_of (subst u \) = {}" using orient_lit_vars by blast from \orient_lit (subst_lit M \) (subst u \) (subst v \) pos\ and \vars_of_lit (subst_lit M \) = {}\ have "vars_of (subst v \) = {}" using orient_lit_vars by blast from \orient_lit (subst_lit M \) (subst u \) (subst v \) pos\ have "((subst u \),(subst v \)) \ trm_ord" unfolding orient_lit_def by auto from this and \(subst u \) \ (subst v \)\ and \vars_of (subst u \) = {}\ \vars_of (subst v \) = {}\ have "((subst v \),(subst u \)) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by blast from assms(3) and \P1 \ P\ have "ground_clause (subst_cl (cl_ecl P1) \)" unfolding grounding_set_def by auto from this have "vars_of_cl (subst_cl (cl_ecl P1) \) = {}" by auto from \L \ (cl_ecl P1)\have "(subst_lit L \) \ (subst_cl (cl_ecl P1) \)" by auto from this and \vars_of_cl (subst_cl (cl_ecl P1) \) = {}\ have "vars_of_lit (subst_lit L \) = {}" by auto from \orient_lit_inst L t s polarity \\ have "orient_lit (subst_lit L \) (subst t \) (subst s \) polarity" using lift_orient_lit by auto from this and \vars_of_lit (subst_lit L \) = {}\ have "vars_of (subst t \) = {}" using orient_lit_vars by blast from \orient_lit (subst_lit L \) (subst t \) (subst s \) polarity\ and \vars_of_lit (subst_lit L \) = {}\ have "vars_of (subst s \) = {}" using orient_lit_vars by blast let ?mC1 = "mset_ecl (P1, \)" let ?mC2 = "mset_ecl (C, \)" from \L \ (cl_ecl P1)\ \finite (cl_ecl P1)\ have "mset_set (cl_ecl P1) = mset_set ((cl_ecl P1)-{ L }) + mset_set { L }" using split_mset_set [of "cl_ecl P1" "cl_ecl P1 - { L }" "{ L }"] by blast from this have d1: "{# (mset_lit (subst_lit x \)). x \# (mset_set (cl_ecl P1)) #} = {# (mset_lit (subst_lit x \)). x \# (mset_set ((cl_ecl P1) - { L })) #} + {# (mset_lit (subst_lit x \)). x \# (mset_set { L }) #}" using split_image_mset by auto let ?C = "(((cl_ecl P1) - { L }) \ (((cl_ecl P2) - { M }) \ { L' } ))" from \finite (cl_ecl P1)\ \finite (cl_ecl P2)\ have "finite ?C" by auto let ?C' = "?C - ( (cl_ecl P1) - { L })" from \finite ?C\ have "finite ?C'" by auto have "?C = ( (cl_ecl P1) - { L }) \ ?C'" by auto from \finite (cl_ecl P1)\ \finite ?C'\ have "mset_set ?C = mset_set ((cl_ecl P1)-{ L }) + mset_set ?C'" using split_mset_set [of "?C" "cl_ecl P1 - { L }" "?C'"] by blast from this have d2: "{# (mset_lit (subst_lit x \)). x \# (mset_set ?C) #} = {# (mset_lit (subst_lit x \)). x \# (mset_set ((cl_ecl P1) - { L })) #} + {# (mset_lit (subst_lit x \)). x \# (mset_set ?C') #}" using split_image_mset by auto have "{# (mset_lit (subst_lit x \)). x \# (mset_set { L }) #} \ {#}" by auto let ?K = "{# (mset_lit (subst_lit x \)). x \# (mset_set ?C') #}" let ?J = "{# (mset_lit (subst_lit x \)). x \# (mset_set { L }) #}" have "(\k \ set_mset ?K. \j \ set_mset ?J. (k, j) \ (mult trm_ord))" proof fix k assume "k \ set_mset ?K" from this have "k \# ?K" by auto from this obtain M' where "M' \# (mset_set ?C')" and "k = (mset_lit (subst_lit M' \))" using image_mset_thm [of "?K" "\x. (mset_lit (subst_lit x \))" "(mset_set ?C')"] by metis from \M' \# (mset_set ?C')\and \finite ?C'\ have "M' \ ?C'" by auto have "L \# (mset_set { L })" by auto from this have "(mset_lit (subst_lit L \) \# ?J)" by auto from this have "(mset_lit (subst_lit L \) \ set_mset ?J)" by auto have "{# (mset_lit (subst_lit x \)). x \# (mset_set { L }) #} \ {#}" by auto show "\j \ set_mset ?J. (k, j) \ (mult trm_ord)" proof (cases) assume "M' \ (cl_ecl P2) - { M }" from this and \strictly_maximal_literal P2 M \\ have "((subst_lit M' \),(subst_lit M \)) \ lit_ord" unfolding strictly_maximal_literal_def by metis from this and \( (subst_lit M \),(subst_lit L \)) \ lit_ord\ have "((subst_lit M' \),(subst_lit L \)) \ lit_ord" using lit_ord_trans unfolding trans_def by metis from this have "((mset_lit (subst_lit M' \)), (mset_lit (subst_lit L \))) \ (mult trm_ord)" unfolding lit_ord_def by auto from \(mset_lit (subst_lit L \) \ set_mset ?J)\ this \((mset_lit (subst_lit M' \)), (mset_lit (subst_lit L \))) \ (mult trm_ord)\ and \k = (mset_lit (subst_lit M' \))\ show ?thesis by blast next assume "M' \ (cl_ecl P2) - { M }" from this and \M' \ ?C'\ have "M' = L'" by auto from \subterm t p u'\ have "subterm (subst t \) p (subst u' \)" using substs_preserve_subterms by blast from \ck_unifier u' u \ Ground\ have "(subst u \) = (subst u' \)" unfolding ck_unifier_def Unifier_def by auto from this and \((subst v \),(subst u \)) \ trm_ord\ have "((subst v \),(subst u' \)) \ trm_ord" by auto from this \subterm t p u'\ \replace_subterm t p v t'\ have "((subst t' \),(subst t \)) \ trm_ord" using replacement_monotonic by auto have "polarity = pos \ polarity = neg" using sign.exhaust by auto then have "((subst_lit L' \),(subst_lit L \)) \ lit_ord" proof assume "polarity = pos" from this and \orient_lit_inst L t s polarity \\ have i: "(mset_lit (subst_lit L \)) = {# (subst s \) #} + {# (subst t \) #}" unfolding orient_lit_inst_def using add.commute by force from \L' = mk_lit polarity (Eq t' s)\ \polarity = pos\ have ii: "(mset_lit (subst_lit L' \)) = {# (subst s \) #} + {# (subst t' \) #}" using add.commute by force have "{# (subst t \) #} \ {#}" by auto have "(\k' \ set_mset {# (subst t' \) #}. \j' \ set_mset {# (subst t \) #}. (k', j') \ (trm_ord))" proof fix k' assume "k' \set_mset {# (subst t' \) #}" from this have "k' = (subst t' \)" by auto have "(subst t \) \ set_mset {# (subst t \) #}" by auto from this \k' = (subst t' \)\ and \((subst t' \),(subst t \)) \ trm_ord\ show "\j' \ set_mset {# (subst t \) #}. (k', j') \ (trm_ord)" by auto qed from i ii \((subst t' \),(subst t \)) \ trm_ord\ have "(mset_lit (subst_lit L' \),(mset_lit (subst_lit L \))) \ (mult trm_ord)" by (metis one_step_implies_mult empty_iff insert_iff set_mset_add_mset_insert set_mset_empty) from this show ?thesis unfolding lit_ord_def by auto next assume "polarity = neg" from this and \orient_lit_inst L t s polarity \\ have i: "(mset_lit (subst_lit L \)) = {# (subst s \), (subst s \) #} + {# (subst t \), (subst t \) #}" unfolding orient_lit_inst_def by auto from \L' = mk_lit polarity (Eq t' s)\ \polarity = neg\ have "subst_lit L' \ = (Neg (Eq (subst t' \) (subst s \)))" by auto from this have "(mset_lit (subst_lit L' \)) = {# (subst t' \), (subst t' \), (subst s \), (subst s \) #}" by auto from this have ii: "(mset_lit (subst_lit L' \)) = {# (subst s \), (subst s \) #} + {# (subst t' \), (subst t' \) #}" by (simp add: add.commute add.left_commute) have "{# (subst t \), (subst t \) #} \ {#}" by auto have "(\k' \ set_mset {# (subst t' \),(subst t' \) #}. \j' \ set_mset {# (subst t \),(subst t \) #}. (k', j') \ (trm_ord))" proof fix k' assume "k' \set_mset {# (subst t' \),(subst t' \) #}" from this have "k' = (subst t' \)" by auto have "(subst t \) \ set_mset {# (subst t \),(subst t \) #}" by auto from this \k' = (subst t' \)\ and \((subst t' \),(subst t \)) \ trm_ord\ show "\j' \ set_mset {# (subst t \),(subst t \) #}. (k', j') \ (trm_ord)" by auto qed from this i ii \{# (subst t \), (subst t \) #} \ {#}\ have "(mset_lit (subst_lit L' \), (mset_lit (subst_lit L \))) \ (mult trm_ord)" using one_step_implies_mult [of "{# (subst t \), (subst t \) #}" "{# (subst t' \),(subst t' \) #}" trm_ord "{# (subst s \),(subst s \) #}"] trm_ord_trans by auto from this show ?thesis unfolding lit_ord_def by auto qed from this and \(mset_lit (subst_lit L \) \ set_mset ?J)\ \k = (mset_lit (subst_lit M' \))\ \M' = L'\ show ?thesis unfolding lit_ord_def by auto qed qed from this d1 d2 have o: " ({#mset_lit (subst_lit x \). x \# mset_set ?C #}, {#mset_lit (subst_lit x \). x \# mset_set (cl_ecl P1)#}) \ mult (mult trm_ord)" using mult_trm_ord_trans one_step_implies_mult [of "{# (mset_lit (subst_lit x \)). x \# (mset_set { L }) #}" "{# (mset_lit (subst_lit x \)). x \# (mset_set ?C') #}" "mult trm_ord" "{# (mset_lit (subst_lit x \)). x \# (mset_set ((cl_ecl P1) - { L })) #} " ] by auto from this \C' = (((cl_ecl P1) - { L }) \ (((cl_ecl P2) - { M }) \ { L' } ))\ and \P1 \ P\ and hyp show False by auto qed have not_ref: "\ (\P1. (P1 \ P \ reflexion P1 C \ Ground C'))" proof assume "(\P1. (P1 \ P \ reflexion P1 C \ Ground C'))" then obtain P1 where "P1 \ P" "reflexion P1 C \ Ground C'" by auto from \reflexion P1 C \ Ground C'\ obtain L1 t s Cl_C Cl_P where "(eligible_literal L1 P1 \)" "(L1 \ (cl_ecl P1))" "(Cl_C = (cl_ecl C))" "(Cl_P = (cl_ecl P1))" "(orient_lit_inst L1 t s neg \)" "(ck_unifier t s \ Ground)" "(Cl_C = (subst_cl ((Cl_P - { L1 }) )) \)" "(C' = ((Cl_P - { L1 }) ))" unfolding reflexion_def by blast from \P1 \ P\ and assms(2) and \P \ S\ have "finite (cl_ecl P1)" by auto let ?mC1 = "mset_ecl (P1, \)" let ?mC2 = "mset_ecl (C, \)" from \L1 \ (cl_ecl P1)\ \finite (cl_ecl P1)\ have "mset_set (cl_ecl P1) = mset_set ((cl_ecl P1)-{ L1 }) + mset_set { L1 }" using split_mset_set [of "cl_ecl P1" "cl_ecl P1 - { L1 }" "{ L1 }"] by blast from this have d1: "{# (mset_lit (subst_lit x \)). x \# (mset_set (cl_ecl P1)) #} = {# (mset_lit (subst_lit x \)). x \# (mset_set ((cl_ecl P1) - { L1 })) #} + {# (mset_lit (subst_lit x \)). x \# (mset_set { L1 }) #}" using split_image_mset by auto let ?C = "((cl_ecl P1) - { L1 })" from \finite (cl_ecl P1)\ have "finite ?C" by auto let ?C' = "{}" have "finite ?C'" by auto have "?C = ( (cl_ecl P1) - { L1 }) \ ?C'" by auto from \finite (cl_ecl P1)\ \finite ?C'\ have "mset_set ?C = mset_set ((cl_ecl P1)-{ L1 }) + mset_set ?C'" using split_mset_set [of "?C" "cl_ecl P1 - { L1 }" "?C'"] by blast from this have d2: "{# (mset_lit (subst_lit x \)). x \# (mset_set ?C) #} = {# (mset_lit (subst_lit x \)). x \# (mset_set ((cl_ecl P1) - { L1 })) #} + {# (mset_lit (subst_lit x \)). x \# (mset_set ?C') #}" using split_image_mset by auto have "{# (mset_lit (subst_lit x \)). x \# (mset_set { L1 }) #} \ {#}" by auto let ?K = "{# (mset_lit (subst_lit x \)). x \# (mset_set ?C') #}" let ?J = "{# (mset_lit (subst_lit x \)). x \# (mset_set { L1 }) #}" have "(\k \ set_mset ?K. \j \ set_mset ?J. (k, j) \ (mult trm_ord))" by auto from this d1 d2 \{# (mset_lit (subst_lit x \)). x \# (mset_set { L1 }) #} \ {#}\ have o: " ({#mset_lit (subst_lit x \). x \# mset_set ?C #}, {#mset_lit (subst_lit x \). x \# mset_set (cl_ecl P1)#}) \ mult (mult trm_ord)" using mult_trm_ord_trans one_step_implies_mult [of "{# (mset_lit (subst_lit x \)). x \# (mset_set { L1 }) #}" "{# (mset_lit (subst_lit x \)). x \# (mset_set ?C') #}" "mult trm_ord" "{# (mset_lit (subst_lit x \)). x \# (mset_set ((cl_ecl P1) - { L1 })) #} " ] by auto from this \Cl_P = (cl_ecl P1)\ \C' = ((Cl_P - { L1 }) )\ and \P1 \ P\ and hyp show False by auto qed have not_fact: "\ (\P1. (P1 \ P \ factorization P1 C \ Ground C'))" proof assume "(\P1. (P1 \ P \ factorization P1 C \ Ground C'))" then obtain P1 where "P1 \ P" "factorization P1 C \ Ground C'" by auto from \factorization P1 C \ Ground C'\ obtain L1 L2 L' t s u v Cl_P Cl_C where "(eligible_literal L1 P1 \)" "(L1 \ (cl_ecl P1))" "(L2 \ (cl_ecl P1) - { L1 })" "(Cl_C = (cl_ecl C))" "(Cl_P = (cl_ecl P1))" "(orient_lit_inst L1 t s pos \)" "(orient_lit_inst L2 u v pos \)" "((subst t \) \ (subst s \))" "((subst t \) \ (subst v \))" "(ck_unifier t u \ Ground)" "(L' = Neg (Eq s v))" "(Cl_C = (subst_cl ( (Cl_P - { L2 }) \ { L' } )) \)" "(C' = ( (Cl_P - { L2 }) \ { L' } ))" unfolding factorization_def by blast from \P1 \ P\ and assms(2) and \P \ S\ have "finite (cl_ecl P1)" by auto from assms(3) and \P1 \ P\ have "ground_clause (subst_cl (cl_ecl P1) \)" unfolding grounding_set_def by auto from this have "vars_of_cl (subst_cl (cl_ecl P1) \) = {}" by auto from \L1 \ (cl_ecl P1)\have "(subst_lit L1 \) \ (subst_cl (cl_ecl P1) \)" by auto from this and \vars_of_cl (subst_cl (cl_ecl P1) \) = {}\ have "vars_of_lit (subst_lit L1 \) = {}" by auto from \orient_lit_inst L1 t s pos \\ have "orient_lit (subst_lit L1 \) (subst t \) (subst s \) pos" using lift_orient_lit by auto from this and \vars_of_lit (subst_lit L1 \) = {}\ have "vars_of (subst t \) = {}" using orient_lit_vars by blast from \orient_lit (subst_lit L1 \) (subst t \) (subst s \) pos\ and \vars_of_lit (subst_lit L1 \) = {}\ have "vars_of (subst s \) = {}" using orient_lit_vars by blast from \(L2 \ (cl_ecl P1) - { L1 })\ have "L2 \ (cl_ecl P1)" by auto from \L2 \ (cl_ecl P1)\ have "(subst_lit L2 \) \ (subst_cl (cl_ecl P1) \)" by auto from this and \vars_of_cl (subst_cl (cl_ecl P1) \) = {}\ have "vars_of_lit (subst_lit L2 \) = {}" by auto from \orient_lit_inst L2 u v pos \\ have "orient_lit (subst_lit L2 \) (subst u \) (subst v \) pos" using lift_orient_lit by auto from this and \vars_of_lit (subst_lit L2 \) = {}\ have "vars_of (subst u \) = {}" using orient_lit_vars by blast from \orient_lit (subst_lit L2 \) (subst u \) (subst v \) pos\ and \vars_of_lit (subst_lit L2 \) = {}\ have "vars_of (subst v \) = {}" using orient_lit_vars by blast from \ck_unifier t u \ Ground\ have "(subst t \) = (subst u \)" unfolding ck_unifier_def Unifier_def by auto from \orient_lit (subst_lit L1 \) (subst t \) (subst s \) pos\ have "((subst t \),(subst s \)) \ trm_ord" unfolding orient_lit_def by auto from this and \(subst t \) \ (subst s \)\ and \vars_of (subst t \) = {}\ \vars_of (subst s \) = {}\ have "((subst s \),(subst t \)) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by blast from this and \(subst t \) = (subst u \)\ have "((subst s \),(subst u \)) \ trm_ord" by auto from \orient_lit (subst_lit L2 \) (subst u \) (subst v \) pos\ have "((subst u \),(subst v \)) \ trm_ord" unfolding orient_lit_def by auto from this and \(subst t \) \ (subst v \)\ and \(subst t \) = (subst u \)\ and \vars_of (subst u \) = {}\ \vars_of (subst v \) = {}\ have "((subst v \),(subst u \)) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by metis let ?mC1 = "mset_ecl (P1, \)" let ?mC2 = "mset_ecl (C, \)" from \L2 \ (cl_ecl P1)\ \finite (cl_ecl P1)\ have "mset_set (cl_ecl P1) = mset_set ((cl_ecl P1)-{ L2 }) + mset_set { L2 }" using split_mset_set [of "cl_ecl P1" "cl_ecl P1 - { L2 }" "{ L2 }"] by blast from this have d1: "{# (mset_lit (subst_lit x \)). x \# (mset_set (cl_ecl P1)) #} = {# (mset_lit (subst_lit x \)). x \# (mset_set ((cl_ecl P1) - { L2 })) #} + {# (mset_lit (subst_lit x \)). x \# (mset_set { L2 }) #}" using split_image_mset by auto let ?C = "(((cl_ecl P1) - { L2 }) \ { L' } )" from \finite (cl_ecl P1)\ have "finite ?C" by auto let ?C' = "?C - ( (cl_ecl P1) - { L2 })" from \finite ?C\ have "finite ?C'" by auto have "?C = ( (cl_ecl P1) - { L2 }) \ ?C'" by auto from \finite (cl_ecl P1)\ \finite ?C'\ have "mset_set ?C = mset_set ((cl_ecl P1)-{ L2 }) + mset_set ?C'" using split_mset_set [of "?C" "cl_ecl P1 - { L2 }" "?C'"] by blast from this have d2: "{# (mset_lit (subst_lit x \)). x \# (mset_set ?C) #} = {# (mset_lit (subst_lit x \)). x \# (mset_set ((cl_ecl P1) - { L2 })) #} + {# (mset_lit (subst_lit x \)). x \# (mset_set ?C') #}" using split_image_mset by auto have "{# (mset_lit (subst_lit x \)). x \# (mset_set { L2 }) #} \ {#}" by auto let ?K = "{# (mset_lit (subst_lit x \)). x \# (mset_set ?C') #}" let ?J = "{# (mset_lit (subst_lit x \)). x \# (mset_set { L2 }) #}" have "(\k \ set_mset ?K. \j \ set_mset ?J. (k, j) \ (mult trm_ord))" proof fix k assume "k \ set_mset ?K" from this have "k \# ?K" by simp from this obtain M' where "M' \# (mset_set ?C')" and "k = (mset_lit (subst_lit M' \))" using image_mset_thm [of "?K" "\x. (mset_lit (subst_lit x \))" "(mset_set ?C')"] by metis from \M' \# (mset_set ?C')\and \finite ?C'\ have "M' \ ?C'" by auto have "L2 \# (mset_set { L2 })" by auto from this have "(mset_lit (subst_lit L2 \) \# ?J)" by auto from this have "(mset_lit (subst_lit L2 \) \ set_mset ?J)" by auto have "{# (mset_lit (subst_lit x \)). x \# (mset_set { L2 }) #} \ {#}" by auto show "\j \ set_mset ?J. (k, j) \ (mult trm_ord)" proof - from \M' \ ?C'\ have "M' = L'" by auto from \orient_lit_inst L2 u v pos \\ have i: "(mset_lit (subst_lit L2 \)) = {#} + {# (subst u \), (subst v \) #}" unfolding orient_lit_inst_def using add.commute by force from \L' = Neg (Eq s v)\ have ii: "(mset_lit (subst_lit L' \)) = {#} + {# (subst s \), (subst s \), (subst v \), (subst v \) #}" by force have "{# (subst u \), (subst v \) #} \ {#}" by auto have "(\k' \ set_mset {# (subst s \), (subst s \), (subst v \), (subst v \) #}. \j' \ set_mset {# (subst u \), (subst v \) #}. (k', j') \ (trm_ord))" proof fix k' assume nh: "k' \set_mset {# (subst s \), (subst s \), (subst v \), (subst v \) #}" have "(subst u \) \ set_mset {# (subst u \), (subst v \) #}" by auto from nh have "k' = (subst s \) \ k' = (subst v \)" by auto then show "\j' \ set_mset {# (subst u \), (subst v \) #}. (k', j') \ (trm_ord)" proof assume "k' = (subst s \)" from this and \((subst s \),(subst u \)) \ trm_ord\ and \(subst u \) \ set_mset {# (subst u \), (subst v \) #}\ show ?thesis by auto next assume "k' = (subst v \)" from this and \((subst v \),(subst u \)) \ trm_ord\ and \(subst u \) \ set_mset {# (subst u \), (subst v \) #}\ show ?thesis by auto qed qed from this i ii \{# (subst u \), (subst v \) #} \ {#}\ have "(mset_lit (subst_lit L' \), (mset_lit (subst_lit L2 \))) \ (mult trm_ord)" using one_step_implies_mult [of "{# (subst u \), (subst v \) #}" "{# (subst s \),(subst s \), (subst v \),(subst v \) #}" trm_ord "{#}"] trm_ord_trans by metis from this \M' = L'\ \k = (mset_lit (subst_lit M' \))\ show ?thesis by auto qed qed from this d1 d2 \{# (mset_lit (subst_lit x \)). x \# (mset_set { L2 }) #} \ {#}\ have o: " ({#mset_lit (subst_lit x \). x \# mset_set ?C #}, {#mset_lit (subst_lit x \). x \# mset_set (cl_ecl P1)#}) \ mult (mult trm_ord)" using mult_trm_ord_trans one_step_implies_mult [of "{# (mset_lit (subst_lit x \)). x \# (mset_set { L2 }) #}" "{# (mset_lit (subst_lit x \)). x \# (mset_set ?C') #}" "mult trm_ord" "{# (mset_lit (subst_lit x \)). x \# (mset_set ((cl_ecl P1) - { L2 })) #} " ] by metis from this \(Cl_P = (cl_ecl P1))\ \C' = ( (Cl_P - { L2 }) \ { L' } )\ and \P1 \ P\ and hyp show False by auto qed from not_sup not_ref not_fact and assms(1) show False unfolding derivable_def by blast qed lemma redundant_inference_clause: assumes "redundant_clause E S \ C'" assumes "derivable C P S \ Ground C'" assumes "grounding_set P \" assumes "\x \ S. (finite (cl_ecl x))" shows "redundant_inference E S P \" proof - from assms(1) obtain S' where "S' \ (instances S)" "(set_entails_clause (clset_instances S') (cl_ecl E))" "(\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl E)))" "(\x \ S'. ( ((mset_ecl ((fst x),(snd x))),(mset_cl (C',\))) \ (mult (mult trm_ord)) \ (mset_ecl ((fst x),(snd x))) = mset_cl (C',\)))" unfolding redundant_clause_def by auto from assms(3) assms(4) \derivable C P S \ Ground C'\ obtain D where "D \ P" "(( (mset_cl (C',\)), (mset_ecl (D,\))) \ (mult (mult trm_ord)))" using conclusion_is_smaller_than_premisses by blast have "(\x \ S'. \D' \ P. (((fst x),(snd x)),(D',\)) \ ecl_ord)" proof fix x assume "x \ S'" from this and \(\x \ S'. ( ((mset_ecl ((fst x),(snd x))),(mset_cl (C',\))) \ (mult (mult trm_ord)) \ (mset_ecl ((fst x),(snd x))) = mset_cl (C',\)))\ have "((mset_ecl ((fst x),(snd x))),(mset_cl (C',\))) \ (mult (mult trm_ord)) \ (mset_ecl ((fst x),(snd x))) = mset_cl (C',\)" by auto then have "(((fst x),(snd x)),(D,\)) \ ecl_ord" proof assume "((mset_ecl ((fst x),(snd x))),(mset_cl (C',\))) \ (mult (mult trm_ord))" from this and \(( (mset_cl (C',\)), (mset_ecl (D,\))) \ (mult (mult trm_ord)))\ have "((mset_ecl ((fst x),(snd x))),(mset_ecl (D,\))) \ (mult (mult trm_ord))" using mult_mult_trm_ord_trans unfolding trans_def by metis from this show ?thesis unfolding ecl_ord_def by auto next assume "(mset_ecl ((fst x),(snd x))) = mset_cl (C',\)" from this and \(( (mset_cl (C',\)), (mset_ecl (D,\))) \ (mult (mult trm_ord)))\ have "((mset_ecl ((fst x),(snd x))),(mset_ecl (D,\))) \ (mult (mult trm_ord))" by auto from this show ?thesis unfolding ecl_ord_def by auto qed from this and \D \ P\ show "\D' \ P. (((fst x),(snd x)),(D',\)) \ ecl_ord" by auto qed from this and \S' \ (instances S)\ and \(set_entails_clause (clset_instances S') (cl_ecl E))\ and \(\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl E)))\ show ?thesis unfolding redundant_inference_def by auto qed lemma clause_saturated_and_inference_saturated: assumes "clause_saturated S" assumes "\x \ S. (finite (cl_ecl x))" shows "inference_saturated S" proof (rule ccontr) assume "\ inference_saturated S" then obtain C P \ C' D \ \ where "derivable C P S \ Ground C'" "ground_clause (cl_ecl C)" "derivable D P S \ FirstOrder C'" "trms_subsumes D C \" "\ \ \ \ \" "grounding_set P \" "\redundant_inference (subst_ecl D \) S P \" unfolding inference_saturated_def by blast from assms(2) \grounding_set P \\ \derivable C P S \ Ground C'\ \\redundant_inference (subst_ecl D \) S P \\ have "\redundant_clause (subst_ecl D \) S \ C'" using redundant_inference_clause by blast from assms(1) have "\C P \ C' D \ \. (derivable C P S \ Ground C') \ (ground_clause (cl_ecl C)) \ (derivable D P S \ FirstOrder C') \ (trms_subsumes D C \) \ (\ \ \ \ \) \ (redundant_clause (subst_ecl D \) S \ C')" unfolding clause_saturated_def by blast from this and \derivable C P S \ Ground C'\ \ground_clause (cl_ecl C)\ \derivable D P S \ FirstOrder C'\ \trms_subsumes D C \\ \\ \ \ \ \\ assms(1) have "redundant_clause (subst_ecl D \) S \ C'" by auto from this and \\redundant_clause (subst_ecl D \) S \ C'\ show False by auto qed section \Refutational Completeness\ text \We prove that our variant of the superposition calculus is complete under the redundancy criteria defined above. This is done as usual, by constructing a model of every saturated set not containing the empty clause.\ subsection \Model Construction\ text \We associate as usual every set of extended clauses with an interpretation. The interpretation is constructed in such a way that it is a model of the set of clauses if the latter is saturated and does not contain the empty clause. The interpretation is constructed by defining directly a normalization function mapping every term to its normal form, i.e., to the minimal equivalent term. Note that we do not consider sets of rewrite rules explicitly.\ text \The next function associates every normalization function with the corresponding interpretation (two terms are in relation if they share the same normal form). The obtained relation is an interpretation if the normalization function is compatible with the term combination operator.\ definition same_values :: "('a trm \ 'a trm) \ 'a trm \ 'a trm \ bool" where "(same_values f) = (\x y. (f x) = (f y))" definition value_is_compatible_with_structure :: "('a trm \ 'a trm) \ bool" where "(value_is_compatible_with_structure f) = (\ t s. (f (Comb t s)) = (f (Comb (f t) (f s))))" lemma same_values_fo_int: assumes "value_is_compatible_with_structure f" shows "fo_interpretation (same_values f)" proof - let ?I = "(same_values f)" have ref: "reflexive ?I" unfolding same_values_def reflexive_def by simp have sym: "symmetric ?I" unfolding same_values_def symmetric_def by auto have trans: "transitive ?I" unfolding same_values_def transitive_def by auto from assms(1) have comp: "compatible_with_structure ?I" unfolding same_values_def compatible_with_structure_def value_is_compatible_with_structure_def [of f] by metis from ref trans sym comp have "congruence ?I" unfolding congruence_def equivalence_relation_def by auto then show ?thesis unfolding fo_interpretation_def by auto qed text \The normalization function is defined by mapping each term to a set of pairs. Intuitively, the second element of each pair represents the right hand side of a rule that can be used to rewrite the considered term, and the first element of the pair denotes its normal form. The value of the term is the first component of the pair with the smallest second component.\ text \The following function returns the set of values for which the second component is minimal. We then prove that this set is non-empty and define a function returning an arbitrary chosen element.\ definition min_trms :: "('a trm \ 'a trm) set \ 'a trm set" where "(min_trms E) = ({ x. (\ pair. (pair \ E \ (\pair' \ E. (snd pair',snd pair) \ trm_ord)) \ x = fst pair) })" lemma min_trms_not_empty: assumes "E \ {}" shows "min_trms E \ {}" proof - from assms(1) obtain x where "x \ E" by auto let ?pair_ordering = "{ (x,y). ((snd x),(snd y)) \ trm_ord }" from trm_ord_wf have "wf ?pair_ordering" using measure_wf by auto from this \x \ E\ obtain y where "y \ E" and "\z. (z,y) \ ?pair_ordering \ (z \ E)" using wfE_min [of ?pair_ordering ] by metis from this have "fst y \ min_trms E" unfolding min_trms_def by blast then show ?thesis by auto qed definition get_min :: "'a trm \ ('a trm \ 'a trm) set \ 'a trm" where "(get_min t E) = (if ((min_trms E) = {}) then t else (SOME x. (x \ min_trms E)))" text \We now define the normalization function. The definition is tuned to make the termination proof straightforward. We will reformulate it afterward to get a simpler definition. We first test whether a subterm of the considered term is reducible. If this is the case then the value can be obtained by applying recursively the function on each subterm, and then again on the term obtained by combining the obtained normal forms. If not, then we collect all possible pairs (as explained above), and we use the one with the minimal second component. These pairs can be interpreted as rewrite rules, giving the value of the considered term: the second component is the right-hand side of the rule and the first component is the normal form of the right-hand side. As usual, such rewrite rules are obtained from ground clauses that have a strictly positive maximal literal, no selected literals, and that are not validated by the constructed interpretation. \ function trm_rep:: "'a trm \ ('a eclause set \ 'a trm)" where "(trm_rep t) = (\S. (if ((is_compound t) \ ((lhs t),t) \ trm_ord \ ((rhs t),t) \ trm_ord \ ( ((lhs t,t) \ trm_ord \ (trm_rep (lhs t) S) \ (lhs t)) \ ((rhs t,t) \ trm_ord \(trm_rep (rhs t) S) \ (rhs t)))) then (if ((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord then (trm_rep (Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)) S) else t) else (get_min t { pair. \z CC C' C s L L' \ t' s'. pair = (z,s) \ CC \ S \ (t \ (subst_set (trms_ecl CC) \)) \ (\x. (\x' \ (trms_ecl CC). occurs_in x (subst x' \)) \ ( (x,t) \ trm_ord \ (trm_rep x S) = x)) \ (C' = (cl_ecl CC)) \ (s,t) \ trm_ord \ ((s,t) \ trm_ord \ (z = trm_rep s S)) \ (orient_lit_inst L' t' s' pos \) \ (sel C') = {} \ (L' \ C') \ (maximal_literal L C) \ (L = (subst_lit L' \)) \ (C = (subst_cl C' \)) \ (ground_clause C) \ (t = (subst t' \)) \ (s = (subst s' \)) \ (finite C') \ (\ L u v. (L \ C \ orient_lit L u v pos \ (u,t) \ trm_ord \ (v,t) \ trm_ord \ (trm_rep u S) \ (trm_rep v S)) \ (L \ C \ orient_lit L u v neg \ (u,t) \ trm_ord \ (v,t) \ trm_ord \ (trm_rep u S) = (trm_rep v S))) \ (\s''. ( (eq_occurs_in_cl t s'' (C'- { L' }) \) \ (s'',t) \ trm_ord \ (s,t) \ trm_ord \ (trm_rep s S) \ (trm_rep s'' S))) })))" by auto termination apply (relation "trm_ord") by auto (simp add: trm_ord_wf) text \We now introduce a few shorthands and rewrite the previous definition into an equivalent simpler form. The key point is to prove that a term is always greater than its normal form.\ definition subterm_reduction_aux:: "'a eclause set \ 'a trm \ 'a trm" where "subterm_reduction_aux S t = (if ((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord then (trm_rep (Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)) S) else t)" definition subterm_reduction:: "'a eclause set \ 'a trm \ 'a trm" where "subterm_reduction S t = (trm_rep (Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)) S)" definition maximal_literal_is_unique where "(maximal_literal_is_unique t s C' L' S \) = (\s''. ( (eq_occurs_in_cl t s'' (C'- { L' }) \) \ (s'',t) \ trm_ord \ (s,t) \ trm_ord \ (trm_rep s S) \ (trm_rep s'' S)))" definition smaller_lits_are_false where "(smaller_lits_are_false t C S) = (\ L u v. (L \ C \ orient_lit L u v pos \ (u,t) \ trm_ord \ (v,t) \ trm_ord \ (trm_rep u S) \ (trm_rep v S)) \ (L \ C \ orient_lit L u v neg \ (u,t) \ trm_ord \ (v,t) \ trm_ord \ (trm_rep u S) = (trm_rep v S)))" definition int_clset where "int_clset S = (same_values (\x. (trm_rep x S)))" lemma smaller_lits_are_false_if_cl_not_valid: assumes "\(validate_ground_clause (int_clset S) C)" shows "smaller_lits_are_false t C S" proof (rule ccontr) let ?I = "int_clset S" assume "\smaller_lits_are_false t C S" from this obtain L u v where "L \ C" and "(orient_lit L u v pos \ (trm_rep u S) = (trm_rep v S)) \ (orient_lit L u v neg \ (trm_rep u S) \ (trm_rep v S))" unfolding smaller_lits_are_false_def by blast then have "(orient_lit L u v pos \ (trm_rep u S) = (trm_rep v S)) \ (orient_lit L u v neg \ (trm_rep u S) \ (trm_rep v S))" by blast then show False proof assume c_pos: "(orient_lit L u v pos \ (trm_rep u S) = (trm_rep v S))" then have "orient_lit L u v pos" by blast from c_pos have "(trm_rep u S) = (trm_rep v S)" by blast from \orient_lit L u v pos\ have "L = (Pos (Eq u v)) \ L = (Pos (Eq v u))" unfolding orient_lit_def by auto from this and \(trm_rep u S) = (trm_rep v S)\ have "validate_ground_lit ?I L" using validate_ground_lit.simps(1) validate_ground_eq.simps unfolding same_values_def int_clset_def by (metis (mono_tags, lifting)) from this and \L \ C\ and assms show False unfolding int_clset_def using validate_ground_clause.simps by blast next assume c_neg: "(orient_lit L u v neg \ (trm_rep u S) \ (trm_rep v S))" then have "orient_lit L u v neg" by blast from c_neg have "(trm_rep u S) \ (trm_rep v S)" by blast from \orient_lit L u v neg\ have "L = (Neg (Eq u v)) \ L = (Neg (Eq v u))" unfolding orient_lit_def by auto from this and \(trm_rep u S) \ (trm_rep v S)\ have "validate_ground_lit ?I L" using validate_ground_lit.simps(2) validate_ground_eq.simps unfolding same_values_def int_clset_def by (metis (mono_tags, lifting)) from this and \L \ C\ and assms show False unfolding int_clset_def using validate_ground_clause.simps by blast qed qed text \The following function states that all instances of the terms attached to a clause are in normal form w.r.t.\ the interpretation associated with @{term "S"}, up to some maximal term @{term "t"}\ definition trms_irreducible where "trms_irreducible CC \ S t = (\x. (\x' \ (trms_ecl CC). occurs_in x (subst x' \)) \ ( (x,t) \ trm_ord \ (trm_rep x S) = x))" lemma trms_irreducible_lemma: assumes "all_trms_irreducible (subst_set (trms_ecl C) \) (\t. trm_rep t S)" shows "trms_irreducible C \ S t" proof (rule ccontr) assume "\trms_irreducible C \ S t" from this obtain x where "\x'\ trms_ecl C. occurs_in x (subst x' \)" and "trm_rep x S \ x" unfolding trms_irreducible_def by blast from \\x'\ trms_ecl C. occurs_in x (subst x' \)\ obtain x' where "x' \ trms_ecl C" and "occurs_in x (subst x' \)" by blast from \x' \ trms_ecl C\ have "(subst x' \) \ (subst_set (trms_ecl C) \)" by auto from this and assms(1) \occurs_in x (subst x' \)\ have "trm_rep x S = x" unfolding all_trms_irreducible_def by metis from this and \trm_rep x S \ x\ show False by blast qed text \The following predicate states that a term @{term "z"} is the normal form of the right-hand side of a rule of left-hand side @{term "t"}. It is used to define the set of possible values for term @{term "t"}. The actual value is that corresponding to the smallest right-hand side.\ definition candidate_values where "(candidate_values z CC C' C s L L' \ t' s' t S) = (CC \ S \ (t \ (subst_set (trms_ecl CC) \)) \ (trms_irreducible CC \ S t) \ (C' = (cl_ecl CC)) \ (s,t) \ trm_ord \ ((s,t) \ trm_ord \ (z = trm_rep s S)) \ (orient_lit_inst L' t' s' pos \) \ (sel C' = {}) \ (L' \ C') \ (maximal_literal L C) \ (L = (subst_lit L' \)) \ (C = (subst_cl C' \)) \ (ground_clause C) \ (t = (subst t' \)) \ (s = (subst s' \)) \ (finite C') \ (smaller_lits_are_false t C S) \ (maximal_literal_is_unique t s C' L' S \))" definition set_of_candidate_values:: "'a eclause set \ 'a trm \ ('a trm \ 'a trm) set" where "set_of_candidate_values S t = { pair. \z CC C' C s L L' \ t' s'. pair = (z,s) \ (candidate_values z CC C' C s L L' \ t' s' t S) }" definition subterm_reduction_applicable_aux:: "'a eclause set \ 'a trm \ bool" where "subterm_reduction_applicable_aux S t = (is_compound t \ (lhs t,t) \ trm_ord \ (rhs t,t) \ trm_ord \ ( ((lhs t,t) \ trm_ord \ (trm_rep (lhs t) S) \ (lhs t)) \ ((rhs t,t) \ trm_ord \(trm_rep (rhs t) S) \ (rhs t))))" definition subterm_reduction_applicable:: "'a eclause set \ 'a trm \ bool" where "subterm_reduction_applicable S t = (is_compound t \ ((trm_rep (lhs t) S) \ (lhs t) \ (trm_rep (rhs t) S) \ (rhs t)))" lemma trm_rep_is_lower_aux: assumes "\y. (y,t) \ trm_ord \ (y \ (trm_rep y S) \ ((trm_rep y S),y) \ trm_ord)" assumes "(subterm_reduction_applicable S t)" shows "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord" proof - have "(lhs t,t) \ trm_ord" using \subterm_reduction_applicable S t\ args_are_strictly_lower subterm_reduction_applicable_def by blast have "(rhs t,t) \ trm_ord" using \subterm_reduction_applicable S t\ args_are_strictly_lower subterm_reduction_applicable_def by blast from assms(1) and \(lhs t,t) \ trm_ord\ have l: "( (lhs t \ (trm_rep (lhs t) S)) \ ((trm_rep (lhs t) S), (lhs t)) \ trm_ord)" by metis from assms(1) and \(rhs t,t) \ trm_ord\ have r: "(rhs t \ (trm_rep (rhs t) S) \ ((trm_rep (rhs t) S), (rhs t)) \ trm_ord)" by metis from \subterm_reduction_applicable S t\ have "((trm_rep (lhs t) S) \ (lhs t) \ (trm_rep (rhs t) S) \ (rhs t))" unfolding subterm_reduction_applicable_def [of S t] by blast then show ?thesis proof assume "(trm_rep (lhs t) S) \ (lhs t)" from this and l have "((trm_rep (lhs t) S), (lhs t)) \ trm_ord" by metis from this have i: "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),(Comb (lhs t) (trm_rep (rhs t) S))) \ trm_ord" using trm_ord_reduction_left by blast show "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord" proof (cases) assume "(trm_rep (rhs t) S) = (rhs t)" from this and \((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),(Comb (lhs t) (trm_rep (rhs t) S))) \ trm_ord\ show "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord" by (metis assms(2) is_compound.elims(2) lhs.simps(1) rhs.simps(1) subterm_reduction_applicable_def) next assume "(trm_rep (rhs t) S) \ (rhs t)" from this and r have "((trm_rep (rhs t) S), (rhs t)) \ trm_ord" by metis from this have "((Comb (lhs t) (trm_rep (rhs t) S)), ((Comb (lhs t) (rhs t)))) \ trm_ord" using trm_ord_reduction_right by blast from this and i show "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord" by (metis assms(2) is_compound.elims(2) lhs.simps(1) rhs.simps(1) subterm_reduction_applicable_def trm_ord_trans transE) qed next assume "(trm_rep (rhs t) S) \ (rhs t)" from this and r have "((trm_rep (rhs t) S), (rhs t)) \ trm_ord" by metis from this have i: "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),(Comb (trm_rep (lhs t) S) (rhs t))) \ trm_ord" using trm_ord_reduction_right by blast show "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord" proof (cases) assume "(trm_rep (lhs t) S) = (lhs t)" from this and \((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),(Comb (trm_rep (lhs t) S) (rhs t))) \ trm_ord\ show "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord" by (metis assms(2) basic_superposition.subterm_reduction_applicable_def basic_superposition_axioms is_compound.elims(2) lhs.simps(1) rhs.simps(1)) next assume "(trm_rep (lhs t) S) \ (lhs t)" from this and l have "((trm_rep (lhs t) S), (lhs t)) \ trm_ord" by metis from this have "((Comb (trm_rep (lhs t) S) (rhs t)), ((Comb (lhs t) (rhs t)))) \ trm_ord" using trm_ord_reduction_left by blast from this and i show "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord" by (metis assms(2) basic_superposition.subterm_reduction_applicable_def basic_superposition_axioms is_compound.elims(2) lhs.simps(1) rhs.simps(1) trm_ord_trans transE) qed qed qed text \The following lemma corresponds to the initial definition of the function @{term "trm_rep"}.\ lemma trm_rep_init_def: shows "(trm_rep t) = (\S. (if (subterm_reduction_applicable_aux S t) then (subterm_reduction_aux S t) else (get_min t (set_of_candidate_values S t))))" unfolding subterm_reduction_aux_def set_of_candidate_values_def candidate_values_def subterm_reduction_applicable_aux_def maximal_literal_is_unique_def smaller_lits_are_false_def trms_irreducible_def using trm_rep.simps [of t] by force lemma trm_rep_aux_def: assumes "\y. (y,t) \ trm_ord \ (y \ (trm_rep y S) \ ((trm_rep y S),y) \ trm_ord)" shows "(trm_rep t S) = (if (subterm_reduction_applicable S t) then (subterm_reduction S t) else (get_min t (set_of_candidate_values S t)))" proof (cases) assume "subterm_reduction_applicable S t" then have "subterm_reduction_applicable_aux S t" using args_are_strictly_lower subterm_reduction_applicable_def subterm_reduction_applicable_aux_def by blast from this have "(trm_rep t S) = (subterm_reduction_aux S t)" using trm_rep_init_def [of t] by meson then have "((Comb (trm_rep (lhs t) S) (trm_rep (rhs t) S)),t) \ trm_ord" using \subterm_reduction_applicable S t\ assms trm_rep_is_lower_aux by blast then show ?thesis by (metis \trm_rep t S = subterm_reduction_aux S t\ \subterm_reduction_applicable S t\ subterm_reduction_def subterm_reduction_aux_def) next assume "\subterm_reduction_applicable S t" then have "\subterm_reduction_applicable_aux S t" using subterm_reduction_applicable_def subterm_reduction_applicable_aux_def by blast from this and \\subterm_reduction_applicable S t\ show ?thesis by (meson trm_rep_init_def) qed lemma trm_rep_is_lower: shows "(t \ (trm_rep t S)) \ (((trm_rep t S),t) \ trm_ord)" (is "?P t") proof ((rule wf_induct [of "trm_ord" "?P" "t"]),(simp add: trm_ord_wf)) next fix x assume hyp_ind: "\y. (y,x) \ trm_ord \ (?P y)" let ?v = "(Comb (trm_rep (lhs x) S) (trm_rep (rhs x) S))" show "(?P x)" proof (rule impI) assume "x \ (trm_rep x S)" show "((trm_rep x S),x) \ trm_ord" proof cases assume c1: "subterm_reduction_applicable S x" from this and hyp_ind have "(?v,x) \ trm_ord" using trm_rep_is_lower_aux by metis from c1 and hyp_ind have "(trm_rep x S) = (subterm_reduction S x)" using trm_rep_aux_def [of x S] by metis from this have "(trm_rep x S) = (trm_rep ?v S)" unfolding subterm_reduction_def by metis from \(?v,x) \ trm_ord\ and hyp_ind have "?P ?v" by metis from this and \(trm_rep x S) = (trm_rep ?v S)\ show ?thesis by (metis \(trm_rep (lhs x) S \ trm_rep (rhs x) S, x) \ trm_ord\ trm_ord_trans transE) next assume c2: "\subterm_reduction_applicable S x" from c2 and hyp_ind have "(trm_rep x S) = (get_min x (set_of_candidate_values S x))" using trm_rep_aux_def [of x S] by metis from this and \x \ (trm_rep x S)\ have "(trm_rep x S) \ (min_trms (set_of_candidate_values S x))" unfolding get_min_def by (metis (full_types) some_in_eq) then obtain pair where "pair \ (set_of_candidate_values S x)" "(trm_rep x S) = fst pair" unfolding min_trms_def by blast from \pair \ (set_of_candidate_values S x)\ have "\ CC C' C L L' \ t' s'. candidate_values (fst pair) CC C' C (snd pair) L L' \ t' s' x S" unfolding set_of_candidate_values_def by fastforce from this have "(snd pair,x) \ trm_ord" unfolding candidate_values_def by blast from \\ CC C' C L L' \ t' s'. candidate_values (fst pair) CC C' C (snd pair) L L' \ t' s' x S\ have "((snd pair, x) \ trm_ord \ fst pair = trm_rep (snd pair) S)" unfolding candidate_values_def by blast from \(snd pair,x) \ trm_ord\ \((snd pair, x) \ trm_ord \ fst pair = trm_rep (snd pair) S)\ have "fst pair = trm_rep (snd pair) S" by blast from \(snd pair,x) \ trm_ord\ and hyp_ind have "(?P (snd pair))" by blast from this and \fst pair = (trm_rep (snd pair) S)\ have "fst pair = snd pair \ (fst pair,snd pair) \ trm_ord" by metis from this and \(trm_rep x S) = fst pair\ and \(snd pair,x) \ trm_ord\ \x \ (trm_rep x S)\ show ?thesis by (metis trm_ord_trans transD) qed qed qed lemma trm_rep_is_lower_subt_red: assumes "(subterm_reduction_applicable S x)" shows "((trm_rep x S),x) \ trm_ord" proof - let ?v = "(Comb (trm_rep (lhs x) S) (trm_rep (rhs x) S))" from assms(1) have "(?v,x) \ trm_ord" using trm_rep_is_lower_aux trm_rep_is_lower by metis from assms(1) have "(trm_rep x S) = (subterm_reduction S x)" using trm_rep_aux_def [of x S] trm_rep_is_lower by metis from this have "(trm_rep x S) = (trm_rep ?v S)" unfolding subterm_reduction_def by metis have "?v = trm_rep ?v S \ (trm_rep ?v S,?v) \ trm_ord" using trm_rep_is_lower by blast from this and \(trm_rep x S) = (trm_rep ?v S)\ show "(((trm_rep x S),x) \ trm_ord)" by (metis \(trm_rep (lhs x) S \ trm_rep (rhs x) S, x) \ trm_ord\ trm_ord_trans transE) qed lemma trm_rep_is_lower_root_red: assumes "\(subterm_reduction_applicable S x)" assumes "min_trms (set_of_candidate_values S x) \ {}" shows "(((trm_rep x S),x) \ trm_ord)" proof - from assms(1) have "(trm_rep x S) = (get_min x (set_of_candidate_values S x))" using trm_rep_aux_def [of x S] trm_rep_is_lower by metis from this and assms(2) have "(trm_rep x S) \ (min_trms (set_of_candidate_values S x))" unfolding get_min_def by (metis (full_types) some_in_eq) then obtain pair where "pair \ (set_of_candidate_values S x)" and "(trm_rep x S) = fst pair" unfolding min_trms_def by blast from \pair \ (set_of_candidate_values S x)\ have "\ CC C' C L L' \ t' s'. candidate_values (fst pair) CC C' C (snd pair) L L' \ t' s' x S" unfolding set_of_candidate_values_def by fastforce from this have "(snd pair,x) \ trm_ord" unfolding candidate_values_def by blast from \\ CC C' C L L' \ t' s'. candidate_values (fst pair) CC C' C (snd pair) L L' \ t' s' x S\ have "((snd pair, x) \ trm_ord \ fst pair = trm_rep (snd pair) S)" unfolding candidate_values_def by blast from \(snd pair,x) \ trm_ord\ and \((snd pair, x) \ trm_ord \ fst pair = trm_rep (snd pair) S)\ have "fst pair = trm_rep (snd pair) S" by blast have "snd pair = trm_rep (snd pair) S \ (trm_rep (snd pair) S,snd pair) \ trm_ord" using trm_rep_is_lower by blast from this and \(snd pair,x) \ trm_ord\ have "(trm_rep (snd pair) S,x) \ trm_ord" using trm_ord_trans trans_def by metis from this and \(trm_rep x S) = fst pair\ and \fst pair = trm_rep (snd pair) S\ show ?thesis by metis qed text \Finally, the next lemma gives a simpler and more convenient definition of the function @{term "trm_rep"}.\ lemma trm_rep_simp_def: shows "(trm_rep t S) = (if (subterm_reduction_applicable S t) then (subterm_reduction S t) else (get_min t (set_of_candidate_values S t)))" using trm_rep_is_lower trm_rep_aux_def by blast text \We now establish some useful properties of the normalization function.\ lemma trm_rep_involutive: shows "(trm_rep (trm_rep t S) S) = (trm_rep t S)" (is "?P t") proof ((rule wf_induct [of "trm_ord" "?P" "t"]),(simp add: trm_ord_wf)) next fix x assume hyp_ind: "\y. (y,x) \ trm_ord \ (?P y)" let ?v = "(Comb (trm_rep (lhs x) S) (trm_rep (rhs x) S))" show "(?P x)" proof cases assume c1: "subterm_reduction_applicable S x" from this and hyp_ind have "(?v,x) \ trm_ord" using trm_rep_is_lower_aux trm_rep_is_lower by metis from this hyp_ind have "(trm_rep (trm_rep ?v S) S) = (trm_rep ?v S)" using trm_rep_aux_def [of x S] by metis from c1 have "trm_rep x S = trm_rep ?v S" using trm_rep_simp_def [of x S] unfolding subterm_reduction_def by metis from this and \(trm_rep (trm_rep ?v S) S) = (trm_rep ?v S)\ \trm_rep x S = trm_rep ?v S\ show ?thesis by metis next assume c2: "\subterm_reduction_applicable S x" from c2 have "(trm_rep x S) = (get_min x (set_of_candidate_values S x))" using trm_rep_simp_def [of x S] by metis show ?thesis proof (rule ccontr) assume "(trm_rep (trm_rep x S) S) \ (trm_rep x S)" from this have "x \ (trm_rep x S)" by metis from c2 and \x \ (trm_rep x S)\ have "(trm_rep x S) \ (min_trms (set_of_candidate_values S x))" using trm_rep_simp_def [of x S] unfolding get_min_def by (metis (full_types) some_in_eq) then obtain pair where "pair \ (set_of_candidate_values S x)" and "(trm_rep x S) = fst pair" unfolding min_trms_def by blast from \pair \ (set_of_candidate_values S x)\ have i: "\ CC C' C L L' \ t' s'. candidate_values (fst pair) CC C' C (snd pair) L L' \ t' s' x S" unfolding set_of_candidate_values_def by fastforce from this have "(snd pair,x) \ trm_ord" unfolding candidate_values_def by blast from i have "((snd pair, x) \ trm_ord \ fst pair = trm_rep (snd pair) S)" unfolding candidate_values_def by blast from \(snd pair,x) \ trm_ord\ and \((snd pair, x) \ trm_ord \ fst pair = trm_rep (snd pair) S)\ have "fst pair = trm_rep (snd pair) S" by blast from \(snd pair,x) \ trm_ord\ and hyp_ind have "(?P (snd pair))" by blast from this and \fst pair = (trm_rep (snd pair) S)\ and \(trm_rep x S) = fst pair\ and \(trm_rep (trm_rep x S) S) \ (trm_rep x S)\ show False by metis qed qed qed text \The following predicate is true if all proper subterms are in normal form.\ definition root_term :: "'a eclause set \ 'a trm \ bool" where "(root_term S t) = ((trm_rep t S) = (get_min t (set_of_candidate_values S t)))" text \The following function checks that the considered term contains a subterm that can be reduced.\ definition st_red :: "'a eclause set \ 'a trm \ bool" where "(st_red S t) = (\ t' p. ( (subterm t p t') \ (root_term S t') \ (trm_rep t' S \ t')))" lemma red_arg_implies_red_trm : assumes "st_red S t1" assumes "t = (Comb t1 t2) \ t = (Comb t2 t1)" shows "st_red S t" proof - from assms(1) obtain t' p where "subterm t1 p t'" and "root_term S t'" and "trm_rep t' S \ t'" unfolding st_red_def by blast from \subterm t1 p t'\ and assms(2) obtain q where "subterm t q t'" by (metis subterm.simps(4) subterm.simps(5)) from this and \root_term S t'\ and \trm_rep t' S \ t'\ show ?thesis unfolding st_red_def by blast qed lemma subterms_of_irred_trms_are_irred: "(trm_rep t S) \ t \ st_red S t" (is "(?P t)") proof ((rule wf_induct [of "trm_ord" "?P" "t"]),(simp add: trm_ord_wf)) next fix x assume hyp_ind: "\y. (y,x) \ trm_ord \ (?P y)" show "(?P x)" proof (rule impI) assume "(trm_rep x S) \ x" show "st_red S x" proof (rule ccontr) assume neg_h: "\st_red S x" have i: " \subterm_reduction_applicable S x" proof assume decomp_case: "subterm_reduction_applicable S x" then obtain x1 x2 where "x = (Comb x1 x2)" using is_compound.elims(2) unfolding subterm_reduction_applicable_def by blast from this and decomp_case have "((trm_rep x1 S) \ x1 \ (trm_rep x2 S) \ x2)" using lhs.simps(1) rhs.simps(1) unfolding subterm_reduction_applicable_def by metis then show False proof assume "(trm_rep x1 S) \ x1" from \x = (Comb x1 x2)\ and trm_ord_subterm have "(x1,x) \ trm_ord" by auto from this and hyp_ind and \(trm_rep x1 S) \ x1\ have "st_red S x1" by blast from this and neg_h and \x = (Comb x1 x2)\ show False using red_arg_implies_red_trm [of S x1 x x2] by blast next assume "(trm_rep x2 S) \ x2" from \x = (Comb x1 x2)\ and trm_ord_subterm have "(x2,x) \ trm_ord" by auto from this and hyp_ind and \(trm_rep x2 S) \ x2\ have "st_red S x2" by metis from this and neg_h and \x = (Comb x1 x2)\ show False using red_arg_implies_red_trm [of S x2 x x1] by blast qed qed then have "(trm_rep x S) = (get_min x (set_of_candidate_values S x))" using trm_rep_simp_def by metis then have "root_term S x" unfolding root_term_def by blast have "subterm x [] x" by auto from this and \root_term S x\ and \(trm_rep x S) \ x\have "st_red S x" unfolding st_red_def by blast from this and neg_h show False by auto qed qed qed lemma trm_rep_compatible_with_structure: shows "value_is_compatible_with_structure (\x. trm_rep x S)" proof (rule ccontr) assume "\value_is_compatible_with_structure (\x. trm_rep x S)" from this obtain t s where neg_h:"trm_rep (Comb t s) S \ (trm_rep (Comb (trm_rep t S) (trm_rep s S)) S)" unfolding value_is_compatible_with_structure_def by blast from this have "(trm_rep t S) \ t \ (trm_rep s S) \ s" by metis from this have "subterm_reduction_applicable S (Comb t s)" unfolding subterm_reduction_applicable_def by (metis is_compound.simps(3) lhs.simps(1) rhs.simps(1)) from this have "(trm_rep (Comb t s) S) = (subterm_reduction S (Comb t s))" using trm_rep_simp_def by metis from this and neg_h show False unfolding subterm_reduction_def by (metis lhs.simps(1) rhs.simps(1)) qed text \The following function checks that a position can be reduced, taking into account the order on positions associated with the considered clause and term. A term is reducible when all terms occurring at smaller positions are irreducible.\ definition minimal_redex where "minimal_redex p t C S t' = (\q s. ((q,p) \ (pos_ord C t') \ (subterm t q s) \ (trm_rep s S = s)))" text \The next function checks that a given clause contains two equations with the same left-hand side and whose right-hand sides are equivalent in a given interpretation. If no such equations exist then it is clear that the maximal literal is necessarily unique.\ definition equivalent_eq_exists where "equivalent_eq_exists t s C I \ L1 = (\L\ C - { L1 }. \ u v. (orient_lit_inst L u v pos \) \ ((subst t \) = (subst u \)) \ (I (subst s \) (subst v \)))" lemma maximal_literal_is_unique_lemma: assumes "\equivalent_eq_exists t s C (int_clset S) \ L1" shows "maximal_literal_is_unique (subst t \) (subst s \) C L1 S \" proof (rule ccontr) let ?t = "(subst t \)" let ?s = "(subst s \)" let ?L = "(subst_lit L \)" let ?C = "(subst_cl C \)" assume "\(maximal_literal_is_unique ?t ?s C L1 S \)" from this obtain s'' where "(eq_occurs_in_cl ?t s'' (C- { L1 }) \)" and "(trm_rep ?s S) = (trm_rep s'' S)" unfolding maximal_literal_is_unique_def by blast from \(eq_occurs_in_cl ?t s'' (C- { L1 }) \)\ obtain L' t' s' where "L' \ (C- { L1 })" and "orient_lit_inst L' t' s' pos \" and "(subst t' \) = ?t" and "s'' = subst s' \" unfolding eq_occurs_in_cl_def by auto from \s'' = subst s' \\ and \(trm_rep ?s S) = (trm_rep s'' S)\ have "(trm_rep ?s S) = (trm_rep (subst s' \) S)" by blast from \L' \ (C- { L1 })\ \orient_lit_inst L' t' s' pos \\ \(subst t' \) = ?t\ \(trm_rep ?s S) = (trm_rep (subst s' \) S)\ have "equivalent_eq_exists t s C (int_clset S) \ L1" unfolding equivalent_eq_exists_def same_values_def int_clset_def by metis from this and assms(1) show False by blast qed lemma max_pos_lit_dominates_cl: assumes "maximal_literal (subst_lit L \) (subst_cl C \)" assumes "orient_lit_inst L t s pos \" assumes "L' \ C - { L }" assumes "\equivalent_eq_exists t s C I \ L" assumes "vars_of_lit (subst_lit L \) = {}" assumes "vars_of_lit (subst_lit L' \) = {}" assumes "fo_interpretation I" shows "((subst_lit L' \),(subst_lit L \)) \ lit_ord" proof - let ?L' = "(subst_lit L' \)" let ?L = "(subst_lit L \)" let ?t = "(subst t \)" let ?s = "(subst s \)" from assms(2) have "(?t,?s) \ trm_ord" unfolding orient_lit_inst_def by auto obtain u' v' where "L' = (Pos (Eq u' v')) \ L' = (Neg (Eq u' v'))" using literal.exhaust equation.exhaust by metis from this obtain polarity u v where "orient_lit_inst L' u v polarity \" and "((subst u \),(subst v \)) \ trm_ord" using trm_ord_trans trm_ord_irrefl unfolding trans_def irrefl_def orient_lit_inst_def by metis let ?u = "(subst u \)" let ?v = "(subst v \)" from \orient_lit_inst L' u v polarity \\ have "orient_lit ?L' ?u ?v polarity" using lift_orient_lit by auto from \orient_lit_inst L t s pos \\ have "orient_lit ?L ?t ?s pos" using lift_orient_lit by auto from assms(6) and \orient_lit ?L' ?u ?v polarity\ have "vars_of ?u \ {}" using orient_lit_vars by metis from assms(6) and \orient_lit ?L' ?u ?v polarity\ have "vars_of ?v \ {}" using orient_lit_vars by metis from assms(5) and \orient_lit ?L ?t ?s pos\ have "vars_of ?t \ {}" using orient_lit_vars by metis from assms(5) and \orient_lit ?L ?t ?s pos\ have "vars_of ?s \ {}" using orient_lit_vars by metis from assms(1) and \L' \ C - { L }\ have "(?L,?L') \ lit_ord" unfolding maximal_literal_def by auto from this and \orient_lit ?L ?t ?s pos\ \orient_lit ?L' ?u ?v polarity\ and assms(5) assms(6) have "(?t,?u) \ trm_ord" using lit_ord_dominating_term by metis from this and \vars_of ?t \ {}\ \vars_of ?u \ {}\ have "?u = ?t \ (?u,?t) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by blast from \(?u,?v) \ trm_ord\ and \vars_of ?u \ {}\ \vars_of ?v \ {}\ have "?u = ?v \ (?v,?u) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by blast from \(?t,?s) \ trm_ord\ and \vars_of ?t \ {}\ \vars_of ?s \ {}\ have "?t = ?s \ (?s,?t) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by blast from \vars_of ?v \ {}\ \vars_of ?s \ {}\ have "?v = ?s \ (?v,?s) \ trm_ord \ (?s,?v) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by blast show ?thesis proof (cases) assume "(?u,?t) \ trm_ord" from this and \?u = ?v \ (?v,?u) \ trm_ord\ have "(?v,?t) \ trm_ord" using trm_ord_trans unfolding trans_def by auto from this and \(?u,?t) \ trm_ord\ and \orient_lit ?L ?t ?s pos\ \orient_lit ?L' ?u ?v polarity\ assms(5) assms(6) show ?thesis using lit_ord_dominating_term by metis next assume "(?u,?t) \ trm_ord" from this and \?u = ?t \ (?u,?t) \ trm_ord\ have "?u = ?t" by auto have "polarity = pos" proof (rule ccontr) assume "polarity \ pos" then have "polarity = neg" using sign.exhaust by auto from this and \?u = ?t\ and \orient_lit ?L ?t ?s pos\ \orient_lit ?L' ?u ?v polarity\ assms(5) assms(6) have "(?L,?L') \ lit_ord" using lit_ord_neg_lit_lhs by auto from this and \(?L,?L') \ lit_ord\ show False by auto qed have "?v \ ?s" proof assume "?v = ?s" from this assms(7) have "I ?s ?v" unfolding fo_interpretation_def congruence_def equivalence_relation_def reflexive_def by auto from this and \orient_lit_inst L' u v polarity \\ \polarity = pos\ \?u = ?t\ and assms(3) have "equivalent_eq_exists t s C I \ L" unfolding equivalent_eq_exists_def by metis from this and assms(4) show False by auto qed have "(?s,?v) \ trm_ord" proof assume "(?s,?v) \ trm_ord" from this and \?u = ?t\ and \orient_lit ?L ?t ?s pos\ \orient_lit ?L' ?u ?v polarity\ and \polarity=pos\ assms(5) assms(6) have "(?L,?L') \ lit_ord" using lit_ord_rhs by auto from this and \(?L,?L') \ lit_ord\show False by auto qed from this and \?v \ ?s\ and \?v = ?s \ (?v,?s) \ trm_ord \ (?s,?v) \ trm_ord\ have "(?v,?s) \ trm_ord" by metis from this and \?u = ?t\ and \orient_lit ?L ?t ?s pos\ \orient_lit ?L' ?u ?v polarity\ and \polarity=pos\ assms(5) assms(6) show "(?L',?L) \ lit_ord" using lit_ord_rhs by auto qed qed lemma if_all_smaller_are_false_then_cl_not_valid: assumes "(smaller_lits_are_false (subst t \) (subst_cl C \) S)" assumes "(fo_interpretation (same_values (\t. (trm_rep t S))))" assumes "orient_lit_inst L1 t s pos \" assumes "maximal_literal (subst_lit L1 \) (subst_cl C \)" assumes "ground_clause (subst_cl C \)" assumes "(subst_lit L1 \) \ (subst_cl C \)" assumes "\equivalent_eq_exists t s C (same_values (\t. (trm_rep t S))) \ L1" assumes "trm_rep (subst t \) S = trm_rep (subst s \) S" shows "(\ validate_ground_clause (same_values (\t. (trm_rep t S))) (subst_cl ( C - { L1 } ) \))" proof let ?I = "(same_values (\t. (trm_rep t S)))" assume "validate_ground_clause ?I (subst_cl ( C - { L1 } ) \)" then obtain L where "L \ (subst_cl ( C - { L1 } ) \)" and "validate_ground_lit ?I L" using validate_ground_clause.simps [of ?I "(subst_cl ( C - { L1 } ) \)"] by blast from \L \ (subst_cl ( C - { L1 } ) \)\ obtain L' where "L' \ C - { L1 }" and "L = (subst_lit L' \)" by auto from \L' \ C - { L1 }\ and \L = (subst_lit L' \)\ have "L \(subst_cl C \)" by auto from \L \(subst_cl C \)\ and assms(5) have "vars_of_lit L = {}" by auto from assms(6) and assms(5) have "vars_of_lit (subst_lit L1 \) = {}" by auto obtain u v polarity where o: "orient_lit_inst L' u v polarity \" and "((subst u \), (subst v \)) \ trm_ord" unfolding orient_lit_inst_def using literal.exhaust equation.exhaust trm_ord_trans trm_ord_irrefl unfolding trans_def irrefl_def by metis from o and \L = (subst_lit L' \)\ have o' : "orient_lit L (subst u \) (subst v \) polarity" using lift_orient_lit by auto from o' and \vars_of_lit L = {}\ have "vars_of (subst u \) = {}" using orient_lit_vars by blast from o' and \vars_of_lit L = {}\ have "vars_of (subst v \) = {}" using orient_lit_vars by blast from assms(3) have o1 : "orient_lit (subst_lit L1 \) (subst t \) (subst s \) pos" using lift_orient_lit [of L1 t s pos \] by auto from o1 and \vars_of_lit (subst_lit L1 \) = {}\ have "vars_of (subst t \) = {}" using orient_lit_vars by blast have "polarity = pos \ polarity = neg" using sign.exhaust by auto then show False proof assume "polarity = pos" from this and o and assms(2) and \validate_ground_lit ?I L\ and \L = (subst_lit L' \)\ have "(trm_rep (subst u \) S) = (trm_rep (subst v \) S)" using orient_lit_semantics_pos [of ?I ] unfolding same_values_def by metis from assms(4) and \L \(subst_cl C \)\ have "((subst_lit L1 \),L) \ lit_ord" unfolding maximal_literal_def by blast from this and o' and o1 and \polarity=pos\ and \vars_of_lit L = {}\ and \L = (subst_lit L' \)\ and \vars_of_lit (subst_lit L1 \) = {}\ have "(subst t \, subst u \) \ trm_ord" and "(subst t \, subst v \) \ trm_ord" using lit_ord_dominating_term [of "subst t \" "subst u \" "subst v \" "subst_lit L1 \" "subst s \" pos] by auto show ?thesis proof (cases) assume "(subst t \) = (subst u \)" from this and assms(8) and \(trm_rep (subst u \) S) = (trm_rep (subst v \) S)\ have "(trm_rep (subst s \) S) = (trm_rep (subst v \) S)" by metis from this o and \L' \ C - { L1 }\ \polarity = pos\ \(subst t \) = (subst u \)\ assms(7) show False unfolding equivalent_eq_exists_def same_values_def by blast next assume "(subst t \) \ (subst u \)" from this and \(subst t \, subst u \) \ trm_ord\ and \vars_of (subst t \) = {}\ and \vars_of (subst u \) = {}\ have "(subst u \, subst t \) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from this and \(subst u \, subst v \) \ trm_ord\ and \vars_of (subst v \) = {}\ and \vars_of (subst t \) = {}\ have "(subst v \, subst t \) \ trm_ord" using trm_ord_ground_total trm_ord_trans unfolding ground_term_def trans_def by metis from \polarity = pos\ and o' and assms(1) and \L \(subst_cl C \)\ and \L = (subst_lit L' \)\ and \((subst u \), subst t \) \ trm_ord\ and \((subst v \), subst t \) \ trm_ord\ have "trm_rep (subst u \) S \ trm_rep (subst v \) S" unfolding smaller_lits_are_false_def by metis from this and \trm_rep (subst u \) S = trm_rep (subst v \) S\ show False by blast qed next assume "polarity = neg" from this and o and assms(2) and \validate_ground_lit ?I L\ and \L = (subst_lit L' \)\ have "(trm_rep (subst u \) S) \ (trm_rep (subst v \) S)" using orient_lit_semantics_neg [of ?I ] unfolding same_values_def by metis from assms(4) and \L \(subst_cl C \)\ have "((subst_lit L1 \),L) \ lit_ord" unfolding maximal_literal_def by blast from this and o' and o1 and \vars_of_lit L = {}\ and \L = (subst_lit L' \)\ and \vars_of_lit (subst_lit L1 \) = {}\ have "(subst t \, subst u \) \ trm_ord" and "(subst t \, subst v \) \ trm_ord" using lit_ord_dominating_term [of "subst t \" "subst u \" "subst v \" "subst_lit L1 \" "subst s \" pos] by auto from \((subst_lit L1 \),L) \ lit_ord\ and o' and o1 and \polarity=neg\ and \vars_of_lit L = {}\ and \L = (subst_lit L' \)\ and \vars_of_lit (subst_lit L1 \) = {}\ have "subst t \ \ subst u \" using lit_ord_neg_lit_lhs by auto from this and \(subst t \, subst u \) \ trm_ord\ and \vars_of (subst t \) = {}\ and \vars_of (subst u \) = {}\ have "(subst u \, subst t \) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from this and \(subst u \, subst v \) \ trm_ord\ and \vars_of (subst v \) = {}\ and \vars_of (subst t \) = {}\ have "(subst v \, subst t \) \ trm_ord" using trm_ord_ground_total trm_ord_trans unfolding ground_term_def trans_def by metis from \polarity = neg\ and o' and assms(1) and \L \(subst_cl C \)\ and \L = (subst_lit L' \)\ and \((subst u \), subst t \) \ trm_ord\ and \((subst v \), subst t \) \ trm_ord\ have "trm_rep (subst u \) S = trm_rep (subst v \) S" unfolding smaller_lits_are_false_def by metis from this and \trm_rep (subst u \) S \ trm_rep (subst v \) S\ show False by blast qed qed text \We introduce the notion of a reduction, which is a ground superposition inference satisfying some additional conditions: (i) the ``from'' clause is smaller than the ``into'' clause; (ii) its ``body'' (i.e., the part of the clause without the equation involved in the rule) is false in a given interpretation and strictly smaller than the involved equation.\ definition reduction where "(reduction L1 C \' t s polarity L2 u u' p v D I S \) = ( (D \ S) \ (eligible_literal L1 C \') \ (eligible_literal L2 D \') \ ground_clause (subst_cl (cl_ecl D) \') \ (minimal_redex p (subst t \) C S t) \ (coincide_on \ \' (vars_of_cl (cl_ecl C))) \ (allowed_redex u' C \) \ (\ is_a_variable u') \ (L1 \ (cl_ecl C)) \ (L2 \ (cl_ecl D)) \ (orient_lit_inst L1 t s polarity \') \ (orient_lit_inst L2 u v pos \') \ (subst u \') \ (subst v \') \ (subst u' \') = (subst u \') \ (\ validate_ground_clause I (subst_cl ( (cl_ecl D) - { L2 } ) \')) \ ( (subst_lit L2 \'),(subst_lit L1 \')) \ lit_ord \ (\x \ (cl_ecl D) - { L2 }. ( (subst_lit x \'),(subst_lit L2 \')) \ lit_ord) \ (all_trms_irreducible (subst_set (trms_ecl D) \') (\t. (trm_rep t S))) \ (I (subst u \') (subst v \')) \ (subterm t p u'))" text \The next lemma states that the rules used to evaluate terms can be renamed so that they share no variable with the clause in which the term occurs.\ lemma candidate_values_renaming: assumes "(candidate_values z CC C' C s L L' \ t' s' t S)" assumes "finite C'" assumes "finite (cl_ecl (D:: 'a eclause))" assumes "closed_under_renaming S" assumes "Ball S well_constrained" shows "\ CC_bis C'_bis L'_bis \_bis t'_bis s'_bis t_bis. (candidate_values z CC_bis C'_bis C s L L'_bis \_bis t'_bis s'_bis t S) \ (vars_of_cl (cl_ecl D)) \ vars_of_cl (cl_ecl CC_bis) = {}" proof - from assms(2) have "finite (vars_of_cl C')" using set_of_variables_is_finite_cl by auto from assms(3) have "finite (vars_of_cl (cl_ecl D))" using set_of_variables_is_finite_cl by auto from infinite_vars have "\ (finite vars)" by auto from \finite (vars_of_cl C')\ \finite (vars_of_cl (cl_ecl D))\ and \\ (finite vars)\ obtain \ where "renaming \ (vars_of_cl C')" and "((subst_codomain \ (vars_of_cl C')) \ (vars_of_cl (cl_ecl D))) = {}" using renaming_exists [of vars ] by meson from this \finite (vars_of_cl C')\ obtain \' where i: "(\ x \ (vars_of_cl C'). (subst (subst (Var x) \ ) \') = (Var x))" using renamings_admit_inverse by blast obtain CC_bis where "CC_bis = (subst_ecl CC \)" by auto obtain C'_bis where "C'_bis = (subst_cl C' \)" by auto from assms(1) have "C' = (cl_ecl CC)" unfolding candidate_values_def by metis from this obtain T where "CC = (Ecl C' T)" by (metis cl_ecl.simps trms_ecl.elims) from this and \CC_bis = (subst_ecl CC \)\ and \C'_bis = (subst_cl C' \)\ have "C'_bis = (cl_ecl CC_bis)" by auto from assms(1) have "CC \ S" unfolding candidate_values_def by metis from assms(1) have "(s,t) \ trm_ord" unfolding candidate_values_def by metis from assms(1) have "((s,t) \ trm_ord \ (z = trm_rep s S))" unfolding candidate_values_def by metis from assms(1) have "(maximal_literal L C)" unfolding candidate_values_def by metis from assms(1) have "(ground_clause C)" unfolding candidate_values_def by metis from assms(1) have "L' \ C'" unfolding candidate_values_def by metis from assms(1) have "L = (subst_lit L' \)" unfolding candidate_values_def by metis from assms(1) have "(smaller_lits_are_false t C S)" unfolding candidate_values_def by metis from assms(1) have "C = (subst_cl C' \)" unfolding candidate_values_def by metis from assms(1) have "(orient_lit_inst L' t' s' pos \)" unfolding candidate_values_def by metis from assms(1) have "(trms_irreducible CC \ S t)" unfolding candidate_values_def by metis from assms(1) have "t = subst t' \" unfolding candidate_values_def by metis from assms(1) have "s = subst s' \" unfolding candidate_values_def by metis from \CC \ S\ and assms(4) and \renaming \ (vars_of_cl C')\ and \C' = (cl_ecl CC)\ \CC_bis = (subst_ecl CC \)\ have "CC_bis \ S" unfolding closed_under_renaming_def renaming_cl_def by auto from assms(1) have "(sel C' = {})" unfolding candidate_values_def by metis from this and \renaming \ (vars_of_cl C')\ \C' = (cl_ecl CC)\ \C'_bis = (subst_cl C' \)\ have "sel C'_bis = {}" using sel_renaming by auto obtain L'_bis where "L'_bis = (subst_lit L' \)" by auto from this and \L' \ C'\ \C'_bis = (subst_cl C' \)\ have "L'_bis \ C'_bis" by auto let ?\ = "(comp (comp \ \') \)" let ?\' = "(comp \' \)" have "coincide_on \ ?\ (vars_of_cl C')" proof (rule ccontr) assume "\coincide_on \ ?\ (vars_of_cl C')" from this obtain x where "(subst (Var x) \) \ (subst (Var x) ?\)" and "x \ vars_of_cl C'" unfolding coincide_on_def by auto from \x \ vars_of_cl C'\ i have "(subst (subst (Var x) \ ) \') = (Var x)" by blast from this and \(subst (Var x) \) \ (subst (Var x) ?\)\ show False by simp qed from \L' \ C'\ have "vars_of_lit L' \ vars_of_cl C'" by auto from this and \coincide_on \ ?\ (vars_of_cl C')\ have "coincide_on \ ?\ (vars_of_lit L')" unfolding coincide_on_def by auto from this and \L = (subst_lit L' \)\ have "L = (subst_lit L' ?\)" using coincide_on_lit by auto have "subst_lit L' ?\ = subst_lit (subst_lit L' \) ?\'" by (simp add: coincide_on_def coincide_on_lit composition_of_substs_lit) from this and \L = (subst_lit L' ?\)\ and \L'_bis = (subst_lit L' \)\ have "L = (subst_lit L'_bis ?\')" by auto from \coincide_on \ ?\ (vars_of_cl C')\ and \C = (subst_cl C' \)\ have "C = subst_cl C' ?\" using coincide_on_cl by blast have "subst_cl C' ?\ = subst_cl (subst_cl C' \) ?\'" by (metis composition_of_substs_cl) from this and \C = subst_cl C' ?\\ and \C'_bis = (subst_cl C' \)\ have "C = subst_cl C'_bis ?\'" by auto from \(finite C')\ and \C'_bis = (subst_cl C' \)\ have "finite C'_bis" by auto have "t \ (subst_set (trms_ecl CC_bis) ?\')" proof assume "t \ (subst_set (trms_ecl CC_bis) ?\')" from this obtain u where "u \ (trms_ecl CC_bis)" and "t = (subst u ?\')" by auto from \u \ (trms_ecl CC_bis)\ and \CC_bis = (subst_ecl CC \)\ obtain v where "v \ trms_ecl CC" and "u = (subst v \)" using \CC = Ecl C' T\ by auto from \u = (subst v \)\ \t = (subst u ?\')\ have "subst v ?\ = t" by simp from \v \ trms_ecl CC\ \CC \ S\ assms(5) have "dom_trm v (cl_ecl CC)" unfolding Ball_def well_constrained_def by metis from this have "vars_of v \ vars_of_cl (cl_ecl CC)" using dom_trm_vars by auto from this and \C' = (cl_ecl CC)\ \coincide_on \ ?\ (vars_of_cl C')\ have "coincide_on \ ?\ (vars_of v)" unfolding coincide_on_def by auto from this and \subst v ?\ = t\ have "(subst v \) = t" using coincide_on_term by metis from this and \v \ trms_ecl CC\ have "(t \ (subst_set (trms_ecl CC) \))" by auto from this and assms(1) show False unfolding candidate_values_def by metis qed have "(trms_irreducible CC_bis ?\' S t)" proof (rule ccontr) assume "\(trms_irreducible CC_bis ?\' S t)" then obtain x x' where "x' \ trms_ecl CC_bis" "occurs_in x (subst x' ?\')" "(x,t) \ trm_ord" "(trm_rep x S) \ x" using trms_irreducible_def by blast from \x' \ (trms_ecl CC_bis)\ and \CC_bis = (subst_ecl CC \)\ obtain v where "v \ trms_ecl CC" and "x' = (subst v \)" using \CC = Ecl C' T\ by auto from \occurs_in x (subst x' ?\')\ \x' = subst v \\ have "occurs_in x (subst v ?\)" by simp from \v \ trms_ecl CC\ \CC \ S\ assms(5) have "dom_trm v (cl_ecl CC)" unfolding Ball_def well_constrained_def by auto from this have "vars_of v \ vars_of_cl (cl_ecl CC)" using dom_trm_vars by auto from this and \C' = (cl_ecl CC)\ \coincide_on \ ?\ (vars_of_cl C')\ have "coincide_on \ ?\ (vars_of v)" unfolding coincide_on_def by auto from this have "(subst v \) = (subst v ?\)" using coincide_on_term by metis from this and \occurs_in x (subst v ?\)\ have "occurs_in x (subst v \)" by auto from this and \v \ trms_ecl CC\ and \(x,t) \ trm_ord\ \(trm_rep x S) \ x\ and \(trms_irreducible CC \ S t)\ show False unfolding trms_irreducible_def by metis qed obtain t'_bis where "t'_bis = (subst t' \)" by auto obtain s'_bis where "s'_bis = (subst s' \)" by auto from \(orient_lit_inst L' t' s' pos \)\ have "vars_of t' \ vars_of_lit L'" using orient_lit_inst_vars by auto from this \coincide_on \ ?\ (vars_of_lit L')\ have "coincide_on \ ?\ (vars_of t')" unfolding coincide_on_def by blast from this have "subst t' ?\ = subst t' \" using coincide_on_term by metis from this \t'_bis = (subst t' \)\ have "subst t'_bis ?\' = subst t' \" by simp from \(orient_lit_inst L' t' s' pos \)\ have "vars_of s' \ vars_of_lit L'" using orient_lit_inst_vars by auto from this \coincide_on \ ?\ (vars_of_lit L')\ have "coincide_on \ ?\ (vars_of s')" unfolding coincide_on_def by blast from this have "subst s' ?\ = subst s' \" using coincide_on_term by metis from this \s'_bis = (subst s' \)\ have "subst s'_bis ?\' = subst s' \" by simp have "(orient_lit_inst L'_bis t'_bis s'_bis pos ?\')" proof - from \(orient_lit_inst L' t' s' pos \)\ have "((subst t' \),(subst s' \)) \ trm_ord" using orient_lit_inst_def by simp from \(orient_lit_inst L' t' s' pos \)\ have "L' = (Pos (Eq t' s')) \ L' = (Pos (Eq s' t'))" by (simp add: orient_lit_inst_def) from this \L'_bis = (subst_lit L' \)\ \t'_bis = (subst t' \)\ \s'_bis = (subst s' \)\ have "L'_bis = (Pos (Eq t'_bis s'_bis)) \ L'_bis = (Pos (Eq s'_bis t'_bis))" by auto from \subst s'_bis ?\' = subst s' \\ and \subst t'_bis ?\' = subst t' \\ and \((subst t' \),(subst s' \)) \ trm_ord\ have "((subst t'_bis ?\'),(subst s'_bis ?\')) \ trm_ord" by auto from this and \L'_bis = (Pos (Eq t'_bis s'_bis)) \ L'_bis = (Pos (Eq s'_bis t'_bis))\ show ?thesis unfolding orient_lit_inst_def by auto qed have "(maximal_literal_is_unique t s C'_bis L'_bis S ?\')" proof (rule ccontr) assume "\(maximal_literal_is_unique t s C'_bis L'_bis S ?\')" from this obtain s'' where "(eq_occurs_in_cl t s'' (C'_bis- { L'_bis }) ?\')" "(s'',t) \ trm_ord" "(s,t) \ trm_ord" "(trm_rep s S) = (trm_rep s'' S)" unfolding maximal_literal_is_unique_def by metis from \(eq_occurs_in_cl t s'' (C'_bis- { L'_bis }) ?\')\ obtain M u v where "M \ C'_bis - { L'_bis }" "orient_lit_inst M u v pos ?\'" "t = (subst u ?\')" "s'' = (subst v ?\')" unfolding eq_occurs_in_cl_def by blast from \M \ C'_bis - { L'_bis }\ and \C'_bis = (subst_cl C' \)\ and \L'_bis = (subst_lit L' \)\ obtain M' where "M' \ C'- { L' }" and "subst_lit M' \ = M" by auto from \orient_lit_inst M u v pos ?\'\ obtain e where "M = (Pos e)" unfolding orient_lit_inst_def by auto from this and \orient_lit_inst M u v pos ?\'\ have "e = (Eq u v) \ e = (Eq v u)" unfolding orient_lit_inst_def by auto from \orient_lit_inst M u v pos ?\'\ have "( (subst u ?\'),(subst v ?\')) \ trm_ord" unfolding orient_lit_inst_def by auto from \M = (Pos e)\ and \subst_lit M' \ = M\ obtain e' where "(subst_equation e' \) = e" and "M' = (Pos e')" by (metis (no_types, opaque_lifting) subst_lit.simps(1) subst_lit.simps(2) atom.simps(1) literal.distinct(1) positive_literal.elims(1)) from \e = (Eq u v) \ e = (Eq v u)\ and \(subst_equation e' \) = e\ obtain u' v' where "e' = (Eq u' v') \ (e' = (Eq v' u'))" and "(subst u' \) = u" and "(subst v' \) = v" by (metis subst_equation.simps equation.inject subterms_of_eq.cases) from \( (subst u ?\'),(subst v ?\')) \ trm_ord\ \(subst u' \) = u\ \(subst v' \) = v\ have "( (subst u' ?\),(subst v' ?\)) \ trm_ord" by simp from this and \M' = (Pos e')\ and \e' = (Eq u' v') \ (e' = (Eq v' u'))\ have "orient_lit_inst M' u' v' pos ?\" unfolding orient_lit_inst_def by auto from \M' \ C' - { L' }\ have "vars_of_lit M' \ vars_of_cl C'" by auto from this and \coincide_on \ ?\ (vars_of_cl C')\ have "coincide_on \ ?\ (vars_of_lit M')" unfolding coincide_on_def by auto from this have "coincide_on ?\ \ (vars_of_lit M')" using coincide_sym by auto from this and \orient_lit_inst M' u' v' pos ?\\ have "orient_lit_inst M' u' v' pos \" using orient_lit_inst_coincide by auto from \orient_lit_inst M' u' v' pos ?\\ have "vars_of u' \ vars_of_lit M'" and "vars_of v' \ vars_of_lit M'" using orient_lit_inst_vars by auto from \vars_of u' \ vars_of_lit M'\ and \coincide_on ?\ \ (vars_of_lit M')\ have "coincide_on ?\ \ (vars_of u')" unfolding coincide_on_def by auto from this have "subst u' ?\ = subst u' \" using coincide_on_term by metis from this and \(subst u' \) = u\ have "subst u ?\' = subst u' \" by simp from \vars_of v' \ vars_of_lit M'\ and \coincide_on ?\ \ (vars_of_lit M')\ have "coincide_on ?\ \ (vars_of v')" unfolding coincide_on_def by auto from this have "subst v' ?\ = subst v' \" using coincide_on_term by metis from this and \(subst v' \) = v\ have "subst v ?\' = subst v' \" by simp from \subst v ?\' = subst v' \\ \s'' = (subst v ?\')\ have "s'' = (subst v' \)" by auto from \subst u ?\' = subst u' \\ \t = (subst u ?\')\ have "t = (subst u' \)" by auto from \s'' = (subst v' \)\ \t = (subst u' \)\ \orient_lit_inst M' u' v' pos \\ \M' \ C' - { L'}\ have "eq_occurs_in_cl t s'' (C'- { L' }) \" unfolding eq_occurs_in_cl_def by auto from this and \(s'',t) \ trm_ord\ and \(s,t) \ trm_ord\ and \(trm_rep s S) = (trm_rep s'' S)\ have "\(maximal_literal_is_unique t s C' L' S \)" unfolding maximal_literal_is_unique_def by blast from this and assms(1) show False unfolding candidate_values_def by metis qed from \t'_bis = (subst t' \)\ and \t = subst t' \\ have "t = subst t'_bis (comp \' \)" using \subst t'_bis (comp \' \) = subst t' \\ by auto from \s'_bis = (subst s' \)\ and \s = subst s' \\ have "s = subst s'_bis (comp \' \)" using \subst s'_bis (comp \' \) = subst s' \\ by auto from \CC_bis \ S\ \t \ subst_set (trms_ecl CC_bis) (comp \' \)\ \trms_irreducible CC_bis (comp \' \) S t\ \C'_bis = cl_ecl CC_bis\ \(s, t) \ trm_ord\ \((s, t) \ trm_ord \ z = trm_rep s S)\ \orient_lit_inst L'_bis t'_bis s'_bis pos (comp \' \)\ \sel C'_bis = {}\ \L'_bis \ C'_bis\ \maximal_literal L C\ \L = subst_lit L'_bis (comp \' \)\ \C = subst_cl C'_bis (comp \' \)\ \ground_clause C\ \t = subst t'_bis (comp \' \)\ \s = subst s'_bis (comp \' \)\ \finite C'_bis\ \smaller_lits_are_false t C S\ \maximal_literal_is_unique t s C'_bis L'_bis S (comp \' \)\ have "(candidate_values z CC_bis C'_bis C s L L'_bis ?\' t'_bis s'_bis t S)" unfolding candidate_values_def by metis have "vars_of_cl (cl_ecl D) \ (vars_of_cl (cl_ecl CC_bis)) = {}" proof (rule ccontr) assume "vars_of_cl (cl_ecl D) \ (vars_of_cl (cl_ecl CC_bis)) \ {}" from this and \C'_bis = (cl_ecl CC_bis)\ obtain x where "x \ vars_of_cl C'_bis" and "x \ vars_of_cl (cl_ecl D)" by auto from \x \ vars_of_cl C'_bis\ obtain N where "N \ C'_bis" and "x \ vars_of_lit N" by auto from \N \ C'_bis\ and \C'_bis = (subst_cl C' \)\ obtain N' where "N' \ C'" and "N = subst_lit N' \" by auto from \x \ vars_of_lit N\ obtain e where "N = (Pos e) \ (N = (Neg e))" and "x \ vars_of_eq e" by (metis negative_literal.elims(2) negative_literal.elims(3) vars_of_lit.simps(1) vars_of_lit.simps(2)) from \N = (Pos e) \ (N = (Neg e))\ and \N = subst_lit N' \\ obtain e' where "N' = (Pos e') \ (N' = (Neg e'))" and "e = subst_equation e' \" by (metis subst_lit.elims atom.simps(1) atom.simps(2)) from \x \ vars_of_eq e\ obtain u v where "e = (Eq u v)" and "x \ vars_of u \ vars_of v" by (metis equation.exhaust vars_of_eq.simps) from \e = (Eq u v)\ and \e = subst_equation e' \\ obtain u' v' where "e' = (Eq u' v')" "u = (subst u' \)" and "v = (subst v' \)" by (metis subst_equation.simps equation.exhaust equation.inject) from \x \ vars_of u \ vars_of v\ have "x \ vars_of u \ x \ vars_of v" by auto then show False proof assume "x \ vars_of u" from this and \u = (subst u' \)\ obtain y where "y \ vars_of u'" and "x \ vars_of (subst (Var y) \)" using vars_of_instances [of u' \] by auto from \y \ vars_of u'\ and \e' = (Eq u' v')\ have "y \ vars_of_eq e'" by auto from this and \N' = (Pos e') \ (N' = (Neg e'))\ have "y \ vars_of_lit N'" by auto from this and \N' \ C'\ have "y \ vars_of_cl C'" by auto from this and \renaming \ (vars_of_cl C')\ have "is_a_variable (subst (Var y) \)" unfolding renaming_def by auto from this and \x \ vars_of (subst (Var y) \)\ have "Var x = (subst (Var y) \)" by (metis is_a_variable.elims(2) singletonD vars_of.simps(1)) from this and \y \ vars_of_cl C'\ have "x \ (subst_codomain \ (vars_of_cl C'))" unfolding subst_codomain_def by auto from this and \x \ vars_of_cl (cl_ecl D)\ and \((subst_codomain \ (vars_of_cl C')) \ (vars_of_cl (cl_ecl D))) = {}\ show False by auto next assume "x \ vars_of v" from this and \v = (subst v' \)\ obtain y where "y \ vars_of v'" and "x \ vars_of (subst (Var y) \)" using vars_of_instances [of v' \] by auto from \y \ vars_of v'\ and \e' = (Eq u' v')\ have "y \ vars_of_eq e'" by auto from this and \N' = (Pos e') \ (N' = (Neg e'))\ have "y \ vars_of_lit N'" by auto from this and \N' \ C'\ have "y \ vars_of_cl C'" by auto from this and \renaming \ (vars_of_cl C')\ have "is_a_variable (subst (Var y) \)" unfolding renaming_def by auto from this and \x \ vars_of (subst (Var y) \)\ have "Var x = (subst (Var y) \)" by (metis is_a_variable.elims(2) singletonD vars_of.simps(1)) from this and \y \ vars_of_cl C'\ have "x \ (subst_codomain \ (vars_of_cl C'))" unfolding subst_codomain_def by auto from this and \x \ vars_of_cl (cl_ecl D)\ and \((subst_codomain \ (vars_of_cl C')) \ (vars_of_cl (cl_ecl D))) = {}\ show False by auto qed qed from this and \(candidate_values z CC_bis C'_bis C s L L'_bis ?\' t'_bis s'_bis t S)\ show ?thesis by blast qed lemma pos_ord_acyclic: shows "acyclic (pos_ord C t)" by (simp add: acyclic_irrefl pos_ord_irrefl pos_ord_trans) definition proper_subterm_red where "proper_subterm_red t S \ = (\p s. (subterm t p s \ p \ Nil \ (trm_rep (subst s \) S \ (subst s \))))" text \The following lemma states that if an eligible term in a clause instance is not in normal form, then the clause instance must be reducible (according to the previous definition of @{term "reduction"}). This is the key lemma for proving completeness. Note that we assume that the considered substitution is in normal form, so that the reduction cannot occur inside a variable. We also rename the clause used for the reduction, to ensure that it shares no variable with the provided clause. The proof requires an additional hypothesis in the case where the reducible term occurs at the root position in an eligible term of a positive literal, see the first hypothesis below and function @{term "equivalent_eq_exists"}.\ lemma reduction_exists: assumes "polarity = neg \ \ equivalent_eq_exists t s (cl_ecl C) (int_clset S) \ L1 \ proper_subterm_red t S \" assumes "\x y. (( x \ vars_of_cl (cl_ecl C)) \ (occurs_in y (subst (Var x) \)) \ trm_rep y S = y)" assumes "eligible_literal L1 C \" assumes "(trm_rep (subst t \) S) \ (subst t \)" assumes "L1 \ (cl_ecl C)" assumes "(orient_lit_inst L1 t s polarity \)" assumes "\x \ S. finite (cl_ecl x)" assumes "ground_clause (subst_cl (cl_ecl C) \)" assumes "(fo_interpretation (same_values (\t. (trm_rep t S))))" assumes "C \ S" assumes "Ball S well_constrained" assumes "all_trms_irreducible (subst_set (trms_ecl C) \) (\t. trm_rep t S)" assumes "\ validate_ground_clause (int_clset S) (subst_cl (cl_ecl C) \)" assumes "closed_under_renaming S" shows "\\' u u' p v D L2. ((reduction L1 C \' t s polarity L2 u u' p v D (same_values (\t. (trm_rep t S))) S \) \ (variable_disjoint C D))" proof - text \The first step is to get the minimal reducible position in @{term "(subst t \)"} and the corresponding subterm @{term "v"}.\ let ?Redexes = "{ p. \v. subterm (subst t \) p v \ root_term S v \ trm_rep v S \ v }" have "?Redexes \ pos_of (subst t \)" proof fix x assume "x \ ?Redexes" then have "\v. subterm (subst t \) x v" by blast then have "position_in x (subst t \)" unfolding position_in_def by metis then show "x \ pos_of (subst t \)" by auto qed from this have "finite ?Redexes" using set_of_positions_is_finite [of "(subst t \)" ] using finite_subset by blast from assms(4) have "st_red S (subst t \)" using subterms_of_irred_trms_are_irred by blast from this obtain p' where "p' \ ?Redexes" unfolding st_red_def by blast from \finite ?Redexes\ this obtain mp where "mp \ ?Redexes" and "\p'. (p', mp) \ (pos_ord C t) \ p' \ ?Redexes" using pos_ord_acyclic [of C t] finite_proj_wf [of ?Redexes p' "pos_ord C t"] by blast have mr: "minimal_redex mp (subst t \) C S t" proof (rule ccontr) assume "\minimal_redex mp (subst t \) C S t" from this obtain p'' v' where "(p'',mp) \ (pos_ord C t)" "subterm (subst t \) p'' v'" and "trm_rep v' S \ v'" unfolding minimal_redex_def by blast show False proof (cases) assume "(root_term S v')" from this and \subterm (subst t \) p'' v'\ \trm_rep v' S \ v'\ have "p'' \?Redexes" by blast from this and \\p'. (p', mp) \ (pos_ord C t) \ p' \ ?Redexes\ and \(p'',mp) \ (pos_ord C t)\ show False by blast next assume "\(root_term S v')" from \trm_rep v' S \ v'\ have "st_red S v'" using subterms_of_irred_trms_are_irred by blast from this obtain p''' v'' where "subterm v' p''' v''" "root_term S v''" "trm_rep v'' S \ v''" unfolding st_red_def by blast from \subterm v' p''' v''\ and \subterm (subst t \) p'' v'\ have "subterm (subst t \) (append p'' p''') v''" using subterm_of_a_subterm_is_a_subterm by metis from this and \trm_rep v'' S \ v''\ \root_term S v''\ have "(append p'' p''') \ ?Redexes" by blast from this and \\p'. (p', mp) \ (pos_ord C t) \ p' \ ?Redexes\ have "(append p'' p''',mp) \ (pos_ord C t)" by blast from this and \(p'',mp) \ (pos_ord C t)\ show False using pos_ord_prefix by auto qed qed from \mp \ ?Redexes\ obtain p v where "mp=p" "subterm (subst t \) p v" and "root_term S v" and "trm_rep v S \ v" unfolding st_red_def by blast text \Second, we find the clause @{term "C2"} and substitution @{term "\"} that are used to determine the value of @{term "v"} according to the definition of @{term "trm_rep"}, and we prove that they satisfy all the desired properties. In particular, clause @{term "C2"} is renamed to ensure that it shares no variable with @{term "C"}.\ from \subterm (subst t \) p v\ have si: "(\x q1 q2. (is_a_variable x) \ (subterm (subst x \) q1 v) \ (subterm t q2 x) \ (p = (append q2 q1))) \ ((\ u. (\(is_a_variable u) \ (subterm t p u) \ (v = (subst u \)))))" using subterms_of_instances by metis let ?v = "trm_rep v S" from \trm_rep v S \ v\ and \root_term S v\ have "?v \ min_trms (set_of_candidate_values S v)" unfolding root_term_def get_min_def by (metis some_in_eq) from \?v \ min_trms (set_of_candidate_values S v)\ obtain pair where "?v = fst pair" and "pair \ (set_of_candidate_values S v)" and min_pair: "(\pair'\set_of_candidate_values S v. (snd pair', snd pair) \ trm_ord)" unfolding min_trms_def by blast from \pair \ (set_of_candidate_values S v)\ have "\z CC C' C s L L' \ t' s'. pair = (z, s) \ (candidate_values z CC C' C s L L' \ t' s' v S)" unfolding set_of_candidate_values_def [of S v] by blast from this obtain zz C2_init Cl_C2_init gr_Cl_C2 gr_rhs gr_L2 L2_init \_init lhs_init rhs_init where "pair = (zz, gr_rhs)" and "(candidate_values zz C2_init Cl_C2_init gr_Cl_C2 gr_rhs gr_L2 L2_init \_init lhs_init rhs_init v S)" by blast from assms(7) and \C \ S\ have "finite (cl_ecl C)" by auto from \(candidate_values zz C2_init Cl_C2_init gr_Cl_C2 gr_rhs gr_L2 L2_init \_init lhs_init rhs_init v S)\ have "finite Cl_C2_init" unfolding candidate_values_def by metis from assms(11) \closed_under_renaming S\ \finite Cl_C2_init\ \finite (cl_ecl C)\ \(candidate_values zz C2_init Cl_C2_init gr_Cl_C2 gr_rhs gr_L2 L2_init \_init lhs_init rhs_init v S)\ obtain C2 Cl_C2 \ L2 lhs rhs where "(candidate_values zz C2 Cl_C2 gr_Cl_C2 gr_rhs gr_L2 L2 \ lhs rhs v S)" and "(vars_of_cl (cl_ecl C) \ vars_of_cl (cl_ecl C2)) = {}" using candidate_values_renaming [of zz C2_init Cl_C2_init gr_Cl_C2 gr_rhs gr_L2 L2_init \_init lhs_init rhs_init v S C] by auto from \(candidate_values zz C2 Cl_C2 gr_Cl_C2 gr_rhs gr_L2 L2 \ lhs rhs v S)\ and \pair = (zz, gr_rhs)\ and \?v = fst pair\ have cv: "(candidate_values ?v C2 Cl_C2 gr_Cl_C2 gr_rhs gr_L2 L2 \ lhs rhs v S)" by (metis fst_conv) from cv have "C2 \ S" unfolding candidate_values_def by metis from cv have "ground_clause gr_Cl_C2" unfolding candidate_values_def by metis from assms(7) and assms(10) have "finite (vars_of_cl (cl_ecl C))" using set_of_variables_is_finite_cl by blast from cv have "smaller_lits_are_false v gr_Cl_C2 S" unfolding candidate_values_def by metis from cv have "gr_Cl_C2 = subst_cl Cl_C2 \" unfolding candidate_values_def by metis from cv have "orient_lit_inst L2 lhs rhs pos \" unfolding candidate_values_def by metis from cv have "maximal_literal gr_L2 gr_Cl_C2" unfolding candidate_values_def by metis from cv have "gr_L2 = subst_lit L2 \" unfolding candidate_values_def by metis from cv have "ground_clause gr_Cl_C2" unfolding candidate_values_def by metis from cv have "L2 \ Cl_C2" unfolding candidate_values_def by metis from this and \gr_Cl_C2 = subst_cl Cl_C2 \\ and \gr_L2 = subst_lit L2 \\ have "gr_L2 \ gr_Cl_C2" by auto from cv have "trm_rep v S = trm_rep gr_rhs S" unfolding candidate_values_def by metis from cv have "(gr_rhs, v) \ trm_ord" unfolding candidate_values_def by metis from cv have "Cl_C2 = cl_ecl C2" unfolding candidate_values_def by metis from cv have "v \ subst_set (trms_ecl C2) \" unfolding candidate_values_def by metis from cv have "sel (cl_ecl C2) = {}" unfolding candidate_values_def by metis from this and \maximal_literal gr_L2 gr_Cl_C2\ and \gr_Cl_C2 = subst_cl Cl_C2 \\ and \Cl_C2 = (cl_ecl C2)\ and \gr_L2 = subst_lit L2 \\ have "eligible_literal L2 C2 \" unfolding eligible_literal_def by auto from cv have "(gr_rhs, v) \ trm_ord" unfolding candidate_values_def by metis from cv have norm: "(\x. (\x'\ trms_ecl C2. occurs_in x (subst x' \)) \ (x, v) \ trm_ord \ trm_rep x S = x)" unfolding candidate_values_def trms_irreducible_def by metis from \ground_clause gr_Cl_C2\ and \gr_L2 \ gr_Cl_C2\ have "vars_of_lit gr_L2 = {}" by auto from cv have "v = subst lhs \" unfolding candidate_values_def by metis from cv have "gr_rhs = subst rhs \" unfolding candidate_values_def by metis let ?I = "(same_values (\t. (trm_rep t S)))" have no_fact: "\ equivalent_eq_exists lhs rhs Cl_C2 (same_values (\t. trm_rep t S)) \ L2" proof assume "equivalent_eq_exists lhs rhs Cl_C2 (same_values (\t. trm_rep t S)) \ L2" from this have "\L\Cl_C2 - {L2}.\u v. orient_lit_inst L u v pos \ \ subst lhs \ = subst u \ \ same_values (\t. trm_rep t S) (subst rhs \) (subst v \)" unfolding equivalent_eq_exists_def by blast from this obtain M where "M\Cl_C2 - {L2}" and e: "\u v. orient_lit_inst M u v pos \ \ subst lhs \ = subst u \ \ same_values (\t. trm_rep t S) (subst rhs \) (subst v \)" by blast from e obtain u' v' where "orient_lit_inst M u' v' pos \" and i: "subst lhs \ = subst u' \ \ same_values (\t. trm_rep t S) (subst rhs \) (subst v' \)" by blast from i have "subst lhs \ = subst u' \" by blast from i have "trm_rep (subst rhs \) S = trm_rep (subst v' \) S" unfolding same_values_def by blast let ?u' = "(subst u' \)" let ?v' = "(subst v' \)" from \orient_lit_inst M u' v' pos \\ have "orient_lit (subst_lit M \) ?u' ?v' pos" using lift_orient_lit by auto from \orient_lit_inst L2 lhs rhs pos \\ have "orient_lit (subst_lit L2 \) (subst lhs \) (subst rhs \) pos" using lift_orient_lit by auto from \orient_lit_inst M u' v' pos \\ and \M \ (Cl_C2 - {L2})\ and \gr_Cl_C2 = subst_cl Cl_C2 \\ have "eq_occurs_in_cl ?u' ?v' (Cl_C2 - {L2}) \" unfolding eq_occurs_in_cl_def by auto from \M\Cl_C2 - {L2}\ and \gr_Cl_C2 = subst_cl Cl_C2 \\ have "(subst_lit M \) \ (subst_cl (Cl_C2 - { L2 }) \)" by auto from \M\Cl_C2 - {L2}\ and \gr_Cl_C2 = subst_cl Cl_C2 \\ have "(subst_lit M \) \ gr_Cl_C2" by auto from \vars_of_lit gr_L2 = {}\ and \gr_L2 = subst_lit L2 \\ \orient_lit (subst_lit L2 \) (subst lhs \) (subst rhs \) pos\ have "vars_of (subst rhs \) = {}" using orient_lit_vars by blast from \ground_clause gr_Cl_C2\ and \(subst_lit M \) \ gr_Cl_C2\ have "vars_of_lit (subst_lit M \) = {}" by auto from this and \orient_lit (subst_lit M \) ?u' ?v' pos\ have "vars_of ?v' = {}" using orient_lit_vars by blast from \maximal_literal gr_L2 gr_Cl_C2\ and \(subst_lit M \) \ gr_Cl_C2\ have "(gr_L2,(subst_lit M \)) \ lit_ord" unfolding maximal_literal_def by auto from this and \orient_lit (subst_lit M \) ?u' ?v' pos\ and \orient_lit (subst_lit L2 \) (subst lhs \) (subst rhs \) pos\ and \subst lhs \ = subst u' \\ and \vars_of_lit gr_L2 = {}\ and \vars_of_lit (subst_lit M \) = {}\ and \gr_L2 = subst_lit L2 \\ have "((subst rhs \),?v') \ trm_ord" using lit_ord_rhs by auto from this and \vars_of ?v' = {}\ and \vars_of (subst rhs \) = {}\ have "?v' = (subst rhs \) \ (?v',(subst rhs \)) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by auto from this and \(gr_rhs,v) \ trm_ord\ and \gr_rhs = subst rhs \\ have "(?v',v) \ trm_ord" using trm_ord_trans unfolding trans_def by auto from cv have "maximal_literal_is_unique v gr_rhs Cl_C2 L2 S \" unfolding candidate_values_def by metis from \orient_lit_inst M u' v' pos \\ have "((subst u' \),(subst v' \)) \ trm_ord" unfolding orient_lit_inst_def by auto have "trm_rep gr_rhs S \ trm_rep (subst v' \) S" by (metis \(subst v' \, v) \ trm_ord\ \(gr_rhs, v) \ trm_ord\ \subst lhs \ = subst u' \\ \eq_occurs_in_cl (subst u' \) (subst v' \) (Cl_C2 - {L2}) \\ \maximal_literal_is_unique v gr_rhs Cl_C2 L2 S \\ \v = subst lhs \\ maximal_literal_is_unique_def) from this and \trm_rep (subst rhs \) S = trm_rep (subst v' \) S\ and \gr_rhs = (subst rhs \)\ show False by blast qed from this \gr_Cl_C2 = subst_cl Cl_C2 \\ and \gr_L2 = subst_lit L2 \\ and \smaller_lits_are_false v gr_Cl_C2 S\ and assms(9) and \orient_lit_inst L2 lhs rhs pos \\ and \maximal_literal gr_L2 gr_Cl_C2\ and \ground_clause gr_Cl_C2\ and \gr_L2 \ gr_Cl_C2\ and \v = subst lhs \\ \gr_rhs = subst rhs \\ and \trm_rep v S = trm_rep gr_rhs S\ have "(\ validate_ground_clause ?I (subst_cl ( Cl_C2 - { L2 } ) \))" using if_all_smaller_are_false_then_cl_not_valid [of lhs \ "Cl_C2" S L2 rhs] by blast text \We fuse the substitutions @{term "\"} and @{term "\"} so that the superposition rule can be applied:\ from \ground_clause (subst_cl (cl_ecl C) \)\ have "ground_on (vars_of_cl (cl_ecl C)) \" using ground_clauses_and_ground_substs by auto from \finite (vars_of_cl (cl_ecl C))\ \(vars_of_cl (cl_ecl C) \ vars_of_cl (cl_ecl C2)) = {}\ \ground_on (vars_of_cl (cl_ecl C)) \\ obtain \' where "coincide_on \' \ (vars_of_cl (cl_ecl C))" and "coincide_on \' \ (vars_of_cl (cl_ecl C2))" using combine_substs [of "(vars_of_cl (cl_ecl C))" "(vars_of_cl (cl_ecl C2))" \ \] by blast from \coincide_on \' \ (vars_of_cl (cl_ecl C))\ have "coincide_on \ \' (vars_of_cl (cl_ecl C))" using coincide_sym by auto from \coincide_on \' \ (vars_of_cl (cl_ecl C2))\ have "coincide_on \ \' (vars_of_cl (cl_ecl C2))" using coincide_sym by auto from \eligible_literal L1 C \\ \L1 \ (cl_ecl C)\ \coincide_on \ \' (vars_of_cl (cl_ecl C))\ have "eligible_literal L1 C \'" using eligible_literal_coincide by auto from \eligible_literal L2 C2 \\ \L2 \ Cl_C2\ \Cl_C2 = (cl_ecl C2)\ \coincide_on \ \' (vars_of_cl (cl_ecl C2))\ have "eligible_literal L2 C2 \'" using eligible_literal_coincide by auto from \ground_clause gr_Cl_C2\ and \gr_Cl_C2 = (subst_cl Cl_C2 \)\ have "ground_clause (subst_cl Cl_C2 \')" by (metis \Cl_C2 = cl_ecl C2\ \coincide_on \' \ (vars_of_cl (cl_ecl C2))\ coincide_on_cl) from \coincide_on \ \' (vars_of_cl (cl_ecl C))\ \L1 \ (cl_ecl C)\ have "coincide_on \ \' (vars_of_lit L1)" unfolding coincide_on_def by auto from \coincide_on \ \' (vars_of_cl (cl_ecl C2))\ \L2 \ Cl_C2\ and \Cl_C2 = (cl_ecl C2)\ have "coincide_on \ \' (vars_of_lit L2)" unfolding coincide_on_def by auto from \(orient_lit_inst L1 t s polarity \)\ and \coincide_on \ \' (vars_of_lit L1)\ have "(orient_lit_inst L1 t s polarity \')" using orient_lit_inst_coincide [of L1 t s polarity \ \'] by blast from \(orient_lit_inst L2 lhs rhs pos \)\ and \coincide_on \ \' (vars_of_lit L2)\ have "(orient_lit_inst L2 lhs rhs pos \')" using orient_lit_inst_coincide by blast text \To prove that the superposition rule is applicable, we need to show that @{term "v"} does not occur inside a variable:\ have "\(\x q1 q2. (is_a_variable x) \ (subterm (subst x \) q1 v) \ (subterm t q2 x) \ (p = (append q2 q1)))" proof assume "(\x q1 q2. (is_a_variable x) \ (subterm (subst x \) q1 v) \ (subterm t q2 x) \ (p = (append q2 q1)))" then obtain x q1 q2 where "is_a_variable x" "subterm (subst x \) q1 v" "(subterm (subst x \) q1 v)" "(subterm t q2 x)" by auto from \(subterm (subst x \) q1 v)\ have "occurs_in v (subst x \)" unfolding occurs_in_def by auto from \is_a_variable x\ obtain x' where "x = Var x'" using is_a_variable.elims(2) by blast from \subterm t q2 x\ have "x \ subterms_of t" using subterms_of.simps unfolding occurs_in_def by blast from this have "x \ subterms_of_lit L1" using assms(6) by (simp add: orient_lit_inst_subterms) from this \L1 \ (cl_ecl C)\ have "x \ subterms_of_cl (cl_ecl C)" by auto from this have "vars_of x \ vars_of_cl (cl_ecl C)" using subterm_vars by blast from this and \x = (Var x')\ have "x' \ vars_of_cl (cl_ecl C)" by auto from \x' \ vars_of_cl (cl_ecl C)\ \occurs_in v (subst x \)\ \x = Var x'\ assms(2) have "trm_rep v S = v" by blast from this and \trm_rep v S \ v\ show False by blast qed from this and si obtain u where "\ (is_a_variable u)" "(subterm t p u)" and "v = (subst u \)" by auto from \orient_lit_inst L1 t s polarity \\ have "vars_of t \ vars_of_lit L1" using orient_lit_inst_vars by auto from \subterm t p u\ have "vars_of u \ vars_of t" using vars_subterm by auto from \vars_of t \ vars_of_lit L1\ \vars_of u \ vars_of t\ \coincide_on \ \' (vars_of_lit L1)\ have "coincide_on \ \' (vars_of u)" unfolding coincide_on_def by blast from this have "subst u \ = subst u \'" using coincide_on_term by auto from \orient_lit_inst L2 lhs rhs pos \\ have "vars_of lhs \ vars_of_lit L2" and "vars_of rhs \ vars_of_lit L2" using orient_lit_inst_vars by auto from \vars_of lhs \ vars_of_lit L2\ \coincide_on \ \' (vars_of_lit L2)\ have "coincide_on \ \' (vars_of lhs)" unfolding coincide_on_def by blast from this have "subst lhs \ = subst lhs \'" using coincide_on_term by auto from \vars_of rhs \ vars_of_lit L2\ \coincide_on \ \' (vars_of_lit L2)\ have "coincide_on \ \' (vars_of rhs)" unfolding coincide_on_def by blast from this have "subst rhs \ = subst rhs \'" using coincide_on_term by auto from \trm_rep v S = trm_rep gr_rhs S\ and \v= subst lhs \\ and \gr_rhs = (subst rhs \)\ have "trm_rep (subst rhs \) S = trm_rep (subst lhs \) S" by metis from this and \subst rhs \ = subst rhs \'\ \subst lhs \ = subst lhs \'\ have "trm_rep (subst rhs \') S = trm_rep (subst lhs \') S" by metis from \subst lhs \ = subst lhs \'\ \subst u \ = subst u \'\ \v = subst u \\ and \v = subst lhs \\ have "subst u \' = subst lhs \'" by auto from \coincide_on \' \ (vars_of_cl (cl_ecl C2))\ and \Cl_C2 = (cl_ecl C2)\ have "coincide_on \' \ (vars_of_cl (Cl_C2 - { L2 }))" unfolding coincide_on_def by auto from this and \(\ validate_ground_clause ?I (subst_cl ( Cl_C2 - { L2 } ) \))\ have "(\ validate_ground_clause ?I (subst_cl ( Cl_C2 - { L2 } ) \'))" using coincide_on_cl by metis have "(\x\cl_ecl C2 - {L2}. (subst_lit x \', subst_lit L2 \') \ lit_ord)" proof fix x assume "x \cl_ecl C2 - {L2}" from \L2 \ Cl_C2\ and \gr_L2 = (subst_lit L2 \)\ \gr_Cl_C2 = (subst_cl Cl_C2 \)\ have "gr_L2 \ gr_Cl_C2" by auto from this and \ground_clause gr_Cl_C2\ have "vars_of_lit gr_L2 = {}" by auto from \x \ cl_ecl C2 - {L2}\ and \Cl_C2 = (cl_ecl C2)\ \gr_Cl_C2 = (subst_cl Cl_C2 \)\ have "(subst_lit x \) \ gr_Cl_C2" by auto from this and \ground_clause gr_Cl_C2\ have "vars_of_lit (subst_lit x \) = {}" by auto from this \x \ cl_ecl C2 - {L2}\ \maximal_literal gr_L2 gr_Cl_C2\ \Cl_C2 = cl_ecl C2\ \gr_L2 = (subst_lit L2 \)\ \gr_Cl_C2 = (subst_cl Cl_C2 \)\ \orient_lit_inst L2 lhs rhs pos \\ no_fact assms(9) \vars_of_lit gr_L2 = {}\ \vars_of_lit (subst_lit x \) = {}\ have "(subst_lit x \, subst_lit L2 \) \ lit_ord" using max_pos_lit_dominates_cl [of L2 \ Cl_C2 lhs rhs x ?I] by metis from \L2 \ Cl_C2\ have "vars_of_lit L2 \ vars_of_cl Cl_C2" by auto from this and \coincide_on \' \ (vars_of_cl (cl_ecl C2))\ and \Cl_C2 = cl_ecl C2\ have "coincide_on \' \ (vars_of_lit L2)" unfolding coincide_on_def by auto from this have "subst_lit L2 \' = subst_lit L2 \" using coincide_on_lit by auto from \x \ (cl_ecl C2) - {L2}\ have "x \ cl_ecl C2" by auto from this have "vars_of_lit x \ vars_of_cl (cl_ecl C2)" by auto from this and \coincide_on \' \ (vars_of_cl (cl_ecl C2))\ have "coincide_on \' \ (vars_of_lit x)" unfolding coincide_on_def by auto from this have "subst_lit x \' = subst_lit x \" using coincide_on_lit by auto from \(subst_lit x \, subst_lit L2 \) \ lit_ord\ \subst_lit L2 \' = subst_lit L2 \\ \subst_lit x \' = subst_lit x \\ show "(subst_lit x \',subst_lit L2 \') \ lit_ord" by metis qed have "all_trms_irreducible (subst_set (trms_ecl C2) \') (\t. trm_rep t S)" proof (rule ccontr) assume "\all_trms_irreducible (subst_set (trms_ecl C2) \') (\t. trm_rep t S)" from this obtain x y where "x \ (subst_set (trms_ecl C2) \')" and "occurs_in y x" and "trm_rep y S \ y" unfolding all_trms_irreducible_def by blast from \x \ (subst_set (trms_ecl C2) \')\ obtain x' where "x' \ trms_ecl C2" and "x = (subst x' \')" by auto from assms(11) and \x' \ (trms_ecl C2)\ and \C2 \ S\ have "dom_trm x' (cl_ecl C2)" unfolding Ball_def well_constrained_def by blast from this obtain x'' where "x'' \ subterms_of_cl (cl_ecl C2)" and "x'' = x' \ (x',x'') \ trm_ord" using dom_trm_lemma by blast from \dom_trm x' (cl_ecl C2)\ have "vars_of x' \ vars_of_cl (cl_ecl C2)" using dom_trm_vars by blast from this and \coincide_on \' \ (vars_of_cl (cl_ecl C2))\ have "coincide_on \' \ (vars_of x')" unfolding coincide_on_def by auto from this have "(subst x' \) = (subst x' \')" using coincide_on_term by metis from this and \x = (subst x' \')\ have "x = (subst x' \)" by auto from this and \x' \ trms_ecl C2\ have "x \(subst_set (trms_ecl C2) \)" by auto from \x'' \ (subterms_of_cl (cl_ecl C2))\ have "(subst x'' \) \ (subterms_of_cl (subst_cl (cl_ecl C2) \))" using subterm_cl_subst [of x'' "cl_ecl C2"] by auto from \orient_lit_inst L2 lhs rhs pos \\ \gr_rhs = (subst rhs \)\ \gr_L2 = (subst_lit L2 \)\ have "orient_lit gr_L2 (subst lhs \) gr_rhs pos" using lift_orient_lit by auto from \ground_clause gr_Cl_C2\ have "vars_of_cl gr_Cl_C2 = {}" by auto from \vars_of_lit gr_L2 = {}\ \vars_of_cl gr_Cl_C2 = {}\ \(subst x'' \) \ (subterms_of_cl (subst_cl (cl_ecl C2) \))\ \orient_lit gr_L2 (subst lhs \) gr_rhs pos\ \maximal_literal gr_L2 gr_Cl_C2\ \Cl_C2 = cl_ecl C2\ \gr_L2 = (subst_lit L2 \)\ \gr_Cl_C2 = (subst_cl Cl_C2 \)\ \v = (subst lhs \)\ \v = (subst lhs \)\ have "(subst x'' \) = v \ (((subst x'' \),v) \ trm_ord)" using subterms_dominated [of gr_L2 gr_Cl_C2 "(subst lhs \)" gr_rhs pos "subst x'' \"] by metis from \x'' = x' \ (x',x'') \ trm_ord\ \x = (subst x' \)\ have "(subst x'' \) = x \ (x,(subst x'' \)) \ trm_ord" using trm_ord_subst by metis from this and \(subst x'' \) = v \ (((subst x'' \),v) \ trm_ord)\ have "x = v \ ((x,v) \ trm_ord)" using trm_ord_trans trans_def by metis then show False proof assume "x = v" from this and \v \ subst_set (trms_ecl C2) \\ \x \ (subst_set (trms_ecl C2) \)\ show False by auto next assume "(x,v) \ trm_ord" from \occurs_in y x\ have "y = x \ (y,x) \ trm_ord" unfolding occurs_in_def using subterm_trm_ord_eq by auto from this and \(x,v) \ trm_ord\ have "(y,v) \ trm_ord" using trm_ord_trans unfolding trans_def by metis from this and norm and \trm_rep y S \ y\ and \occurs_in y x\ and \x' \ trms_ecl C2\ and \x = (subst x' \)\ show False by metis qed qed from \subterm t p u\ have "u = t \ (u,t) \ trm_ord" using subterm_trm_ord_eq by auto from assms(8) and \L1 \ (cl_ecl C)\ have "vars_of_lit (subst_lit L1 \) = {}" by auto from \coincide_on \ \' (vars_of_lit L1)\ have "(subst_lit L1 \) = (subst_lit L1 \')" using coincide_on_lit by metis from this and \vars_of_lit (subst_lit L1 \) = {}\ have "vars_of_lit (subst_lit L1 \') = {}" by auto from \coincide_on \ \' (vars_of_lit L2)\ have "(subst_lit L2 \) = (subst_lit L2 \')" using coincide_on_lit by metis from \vars_of_lit gr_L2 = {}\ \ground_clause gr_Cl_C2\ \gr_Cl_C2 = (subst_cl Cl_C2 \)\ \L2 \ Cl_C2\ have "vars_of_lit (subst_lit L2 \) = {}" by auto from this and \(subst_lit L2 \) = (subst_lit L2 \')\ have "vars_of_lit (subst_lit L2 \') = {}" by auto text \We now prove that the ``into'' clause is strictly smaller than the ``from'' clause. This is easy if the rewritten literal is negative or if the reduction does not occur at root level. Otherwise, we must use the fact that the function @{term "trm_rep"} selects the smallest right-hand side to compute the value of a term.\ have "(subst_lit L2 \', subst_lit L1 \') \ lit_ord" proof (rule ccontr) assume "\(subst_lit L2 \', subst_lit L1 \') \ lit_ord" from \orient_lit_inst L1 t s polarity \'\ have "orient_lit (subst_lit L1 \') (subst t \') (subst s \') polarity" using lift_orient_lit [of L1 t s polarity \'] by auto from \orient_lit_inst L2 lhs rhs pos \'\ have "orient_lit (subst_lit L2 \') (subst lhs \') (subst rhs \') pos" using lift_orient_lit by auto have "(u,t) \ trm_ord" proof assume "(u,t) \ trm_ord" from this have "(subst u \', subst t \') \ trm_ord" using trm_ord_subst by auto have False "subst u \' = subst lhs \'" using \(subst u \', subst t \') \ trm_ord\ \(subst_lit L2 \', subst_lit L1 \') \ lit_ord\ \subst lhs \ = subst lhs \'\ \subst u \ = subst u \'\ \subst_lit L2 \ = subst_lit L2 \'\ \gr_L2 = subst_lit L2 \\ \orient_lit (subst_lit L1 \') (subst t \') (subst s \') polarity\ \orient_lit (subst_lit L2 \') (subst lhs \') (subst rhs \') pos\ \v = subst lhs \\ \v = subst u \\ \vars_of_lit (subst_lit L1 \') = {}\ \vars_of_lit gr_L2 = {}\ lit_ord_dominating_term apply fastforce using \(subst u \', subst t \') \ trm_ord\ \(subst_lit L2 \', subst_lit L1 \') \ lit_ord\ \subst u \' = subst lhs \'\ \orient_lit (subst_lit L1 \') (subst t \') (subst s \') polarity\ \orient_lit (subst_lit L2 \') (subst lhs \') (subst rhs \') pos\ \vars_of_lit (subst_lit L1 \') = {}\ \vars_of_lit (subst_lit L2 \') = {}\ lit_ord_dominating_term by fastforce then show False by auto qed from this and \subterm t p u\ have "p = Nil" using subterm_trm_ord by auto have "\ proper_subterm_red t S \" proof assume "proper_subterm_red t S \" from this obtain p' s where "p' \ Nil" and "subterm t p' s" "trm_rep (subst s \) S \ (subst s \)" unfolding proper_subterm_red_def by blast from \p = Nil\ and \p' \ Nil\ have "(p',p) \ (pos_ord C t)" using pos_ord_nil by auto from \subterm t p' s\ have "subterm (subst t \) p' (subst s \)" by (simp add: substs_preserve_subterms) from this and \(p',p) \ (pos_ord C t)\ mr and \trm_rep (subst s \) S \ (subst s \)\ \mp=p\ show False using minimal_redex_def by blast qed from \(u,t) \ trm_ord\ and \u = t \ (u,t) \ trm_ord\ have "u = t" by auto have "polarity = pos" proof (rule ccontr) assume "polarity \ pos" then have "polarity = neg" using sign.exhaust by auto from \u = t\ have "subst t \' = subst u \'" by auto from this and \v = (subst u \)\ and \v = (subst lhs \)\ and \subst lhs \ = subst lhs \'\ and \(subst u \) = (subst u \')\ have "(subst t \') = (subst lhs \')" by auto from this and \polarity = neg\ \orient_lit (subst_lit L1 \') (subst t \') (subst s \') polarity\ and \orient_lit (subst_lit L2 \') (subst lhs \') (subst rhs \') pos\ \(subst_lit L2 \', subst_lit L1 \') \ lit_ord\ \vars_of_lit (subst_lit L1 \') = {}\ \vars_of_lit (subst_lit L2 \') = {}\ show False using lit_ord_neg_lit_lhs by auto qed from \vars_of_lit (subst_lit L1 \) = {}\ assms(6) have "vars_of (subst t \) = {}" using lift_orient_lit orient_lit_vars by blast from \vars_of_lit (subst_lit L1 \) = {}\ assms(6) have "vars_of (subst s \) = {}" using lift_orient_lit orient_lit_vars by blast have "trm_rep (subst t \) S \ trm_rep (subst s \) S" proof assume "trm_rep (subst t \) S = trm_rep (subst s \) S" from this have "validate_ground_eq ?I (Eq (subst t \) (subst s \))" unfolding same_values_def using validate_ground_eq.simps by (metis (mono_tags, lifting)) from \trm_rep (subst t \) S = trm_rep (subst s \) S\ have "validate_ground_eq ?I (Eq (subst s \) (subst t \))" unfolding same_values_def using validate_ground_eq.simps by (metis (mono_tags, lifting)) from \orient_lit_inst L1 t s polarity \\ and \polarity=pos\ have "L1 = (Pos (Eq t s)) \ L1 = (Pos (Eq s t))" unfolding orient_lit_inst_def by auto from this have "subst_lit L1 \ = (Pos (Eq (subst t \) (subst s \))) \ subst_lit L1 \ = (Pos (Eq (subst s \) (subst t \)))" by auto from this and \validate_ground_eq ?I (Eq (subst s \) (subst t \))\ and \validate_ground_eq ?I (Eq (subst t \) (subst s \))\ have "validate_ground_lit ?I (subst_lit L1 \)" using validate_ground_lit.simps(1) by metis from \L1 \ (cl_ecl C)\ have "(subst_lit L1 \) \ (subst_cl (cl_ecl C) \)" by auto from this and \validate_ground_lit ?I (subst_lit L1 \)\ have "validate_ground_clause ?I (subst_cl (cl_ecl C) \)" using validate_ground_clause.simps by metis from this and \\ validate_ground_clause (int_clset S) (subst_cl (cl_ecl C) \)\ show False unfolding int_clset_def by blast qed have cv': "(candidate_values (trm_rep (subst s \) S) C (cl_ecl C) (subst_cl (cl_ecl C) \) (subst s \) (subst_lit L1 \) L1 \ t s v S)" proof - from \polarity=pos\ and \orient_lit_inst L1 t s polarity \\ have "\negative_literal L1" unfolding orient_lit_inst_def by auto from this and \eligible_literal L1 C \\ have "sel(cl_ecl C) = {}" and "maximal_literal (subst_lit L1 \) (subst_cl (cl_ecl C) \)" using sel_neg unfolding eligible_literal_def by auto from \v = subst u \\ and \u = t\ have "v = subst t \" by auto from assms(7) \C \ S\ have "finite (cl_ecl C)" by auto have "v \ subst_set (trms_ecl C) \" proof assume "v \ subst_set (trms_ecl C) \" from this and assms(12) \subterm (subst t \) p v\ \v = subst t \\ have "trm_rep v S = v" unfolding all_trms_irreducible_def occurs_in_def by blast from this \v = subst t \\ \trm_rep (subst t \) S \ (subst t \)\ show False by blast qed from assms(13) have "smaller_lits_are_false v (subst_cl (cl_ecl C) \) S" using smaller_lits_are_false_if_cl_not_valid [of S "(subst_cl (cl_ecl C) \)" ] by blast from assms(1) \\ proper_subterm_red t S \\ \polarity=pos\ \v = subst t \\ have "maximal_literal_is_unique v (subst s \) (cl_ecl C) L1 S \" using maximal_literal_is_unique_lemma [of t s "(cl_ecl C)" S \ L1] by blast from \all_trms_irreducible (subst_set (trms_ecl C) \) (\t. trm_rep t S)\ have "trms_irreducible C \ S v" using trms_irreducible_lemma [of C \ S v] by blast have "(subst s \, subst t \) \ trm_ord" proof - from \orient_lit_inst L1 t s polarity \\ have "(subst t \, subst s \) \ trm_ord" unfolding orient_lit_inst_def by auto from \trm_rep (subst t \) S \ trm_rep (subst s \) S\ have "(subst t \) \ (subst s \)" by metis from this and \(subst t \, subst s \) \ trm_ord\ \vars_of (subst t \) = {}\ \vars_of (subst s \) = {}\ show "(subst s \, subst t \) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by metis qed from \C \ S\ \(subst s \, subst t \) \ trm_ord\ and \polarity=pos\ \orient_lit_inst L1 t s polarity \\ and \sel (cl_ecl C) = {}\ and \L1 \ cl_ecl C\ and \maximal_literal (subst_lit L1 \) (subst_cl (cl_ecl C) \)\ and \ground_clause (subst_cl (cl_ecl C) \)\ and \v = subst t \\ and \finite (cl_ecl C)\ and \v \ subst_set (trms_ecl C) \\ and \smaller_lits_are_false v (subst_cl (cl_ecl C) \) S\ and \maximal_literal_is_unique v (subst s \) (cl_ecl C) L1 S \\ and \trms_irreducible C \ S v\ show cv': "(candidate_values (trm_rep (subst s \) S) C (cl_ecl C) (subst_cl (cl_ecl C) \) (subst s \) (subst_lit L1 \) L1 \ t s v S)" unfolding candidate_values_def by blast qed from cv' have "(trm_rep (subst s \) S,(subst s \)) \ set_of_candidate_values S v" unfolding set_of_candidate_values_def by blast from this and min_pair and \pair = (zz, gr_rhs)\ have "((subst s \),gr_rhs) \ trm_ord" by (metis snd_conv) have "(subst s \) \ gr_rhs" using \trm_rep v S = trm_rep gr_rhs S\ \u = t\ \v = subst u \\ \trm_rep (subst t \) S \ trm_rep (subst s \) S\ by blast have "vars_of gr_rhs = {}" using \subst rhs \ = subst rhs \'\ \subst_lit L2 \ = subst_lit L2 \'\ \gr_L2 = subst_lit L2 \\ \gr_rhs = subst rhs \\ \orient_lit (subst_lit L2 \') (subst lhs \') (subst rhs \') pos\ \vars_of_lit gr_L2 = {}\ orient_lit_vars by fastforce from \(subst s \) \ gr_rhs\ and \vars_of (subst s \) = {}\ \vars_of gr_rhs = {}\ \((subst s \),gr_rhs) \ trm_ord\ have "(gr_rhs,(subst s \)) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by blast have "(subst_lit L2 \', subst_lit L1 \') \ lit_ord" using \(gr_rhs, subst s \) \ trm_ord\ \subst lhs \ = subst lhs \'\ \subst rhs \ = subst rhs \'\ \subst_lit L1 \ = subst_lit L1 \'\ \gr_rhs = subst rhs \\ \orient_lit (subst_lit L2 \') (subst lhs \') (subst rhs \') pos\ \polarity = pos\ \u = t\ \v = subst lhs \\ \v = subst u \\ \vars_of_lit (subst_lit L1 \') = {}\ \vars_of_lit (subst_lit L2 \') = {}\ assms(6) lit_ord_rhs lift_orient_lit by fastforce from this and \(subst_lit L2 \', subst_lit L1 \') \ lit_ord\ show False by auto qed have "trm_rep (subst u \) S \ (subst u \)" using \trm_rep v S \ v\ \v = subst u \\ by blast have "allowed_redex u C \" proof (rule ccontr) assume "\allowed_redex u C \" from this obtain ss where "ss \ trms_ecl C" and "occurs_in (subst u \) (subst ss \)" unfolding allowed_redex_def by auto from \ss \ trms_ecl C\ have "(subst ss \) \ (subst_set (trms_ecl C) \)" by auto from this and assms(12) and \occurs_in (subst u \) (subst ss \)\ \trm_rep (subst u \) S \ (subst u \)\ show False unfolding all_trms_irreducible_def by blast qed have "subst lhs \' \ subst rhs \'" using \(gr_rhs, v) \ trm_ord\ \subst lhs \ = subst lhs \'\ \subst rhs \ = subst rhs \'\ \gr_rhs = subst rhs \\ \v = subst lhs \\ trm_ord_wf by auto from this \mp=p\ \\ (is_a_variable u)\ \all_trms_irreducible (subst_set (trms_ecl C2) \') (\t. trm_rep t S)\ \(subst_lit L2 \', subst_lit L1 \') \ lit_ord\ \all_trms_irreducible (subst_set (trms_ecl C2) \') (\t. trm_rep t S)\ \(\x\cl_ecl C2 - {L2}. (subst_lit x \', subst_lit L2 \') \ lit_ord)\ \C2 \ S\ \eligible_literal L1 C \'\ \eligible_literal L2 C2 \'\ \ground_clause (subst_cl Cl_C2 \')\ \Cl_C2 = cl_ecl C2\ mr \coincide_on \ \' (vars_of_cl (cl_ecl C))\ \L1 \ cl_ecl C\ \L2 \ Cl_C2\ \orient_lit_inst L1 t s polarity \'\ \(orient_lit_inst L2 lhs rhs pos \')\ \(subterm t p u)\ \subst u \' = subst lhs \'\ \trm_rep (subst rhs \') S = trm_rep (subst lhs \') S\ \(\ validate_ground_clause ?I (subst_cl ( Cl_C2 - { L2 } ) \'))\ \allowed_redex u C \\ have "(reduction L1 C \' t s polarity L2 lhs u mp rhs C2 (same_values (\t. (trm_rep t S))) S \)" unfolding reduction_def same_values_def by metis from \vars_of_cl (cl_ecl C) \ vars_of_cl (cl_ecl C2) = {}\ have "variable_disjoint C C2" unfolding variable_disjoint_def by auto from this and \(reduction L1 C \' t s polarity L2 lhs u mp rhs C2 (same_values (\t. (trm_rep t S))) S \)\ show ?thesis by blast qed lemma subts_of_irred_trms_are_irred: assumes "trm_rep y S \ y" shows "\ x. subterm x p y \ trm_rep x S \ x" proof (induction p) case (Nil) from assms(1) show ?case by (metis subterm.simps(1)) next case (Cons i p) show "\ x. subterm x (Cons i p) y \ trm_rep x S \ x" proof fix x assume "subterm x (Cons i p) y" from this obtain x1 x2 where "x = Comb x1 x2" using subterm.elims(2) by blast have "i = Left | i = Right" using indices.exhaust by auto then show "trm_rep x S \ x" proof assume "i = Left" from this and \subterm x (Cons i p) y\ \x = Comb x1 x2\ have "subterm x1 p y" by auto from this and Cons.IH have "trm_rep x1 S \ x1" by blast from this and \x = Comb x1 x2\ have "subterm_reduction_applicable S x" unfolding subterm_reduction_applicable_def by (metis is_compound.simps(3) lhs.simps(1)) from this have "(trm_rep x S, x) \ trm_ord" using trm_rep_is_lower_subt_red by blast from this show ?thesis using trm_ord_irrefl unfolding irrefl_def by metis next assume "i = Right" from this and \subterm x (Cons i p) y\ \x = Comb x1 x2\ have "subterm x2 p y" by auto from this and Cons.IH have "trm_rep x2 S \ x2" by blast from this and \x = Comb x1 x2\ have "subterm_reduction_applicable S x" unfolding subterm_reduction_applicable_def by (metis is_compound.simps(3) rhs.simps(1)) from this have "(trm_rep x S, x) \ trm_ord" using trm_rep_is_lower_subt_red by blast from this show ?thesis using trm_ord_irrefl unfolding irrefl_def by metis qed qed qed lemma allowed_redex_coincide: assumes "allowed_redex t C \" assumes "t \ subterms_of_cl (cl_ecl C)" assumes "coincide_on \ \' (vars_of_cl (cl_ecl C))" assumes "well_constrained C" shows "allowed_redex t C \'" proof (rule ccontr) assume "\allowed_redex t C \'" from this obtain s where "s \ trms_ecl C" and "occurs_in (subst t \') (subst s \')" unfolding allowed_redex_def by auto from \s \ trms_ecl C\ and assms(4) have "vars_of s \ vars_of_cl (cl_ecl C)" using dom_trm_vars unfolding well_constrained_def by blast from this have "vars_of s \ vars_of_cl (cl_ecl C)" using subterm_vars by blast from this and assms(3) have "coincide_on \ \' (vars_of s)" unfolding coincide_on_def by auto from this have "(subst s \) = (subst s \')" using coincide_on_term by auto from assms(2) have "vars_of t \ vars_of_cl (cl_ecl C)" using subterm_vars by blast from this and assms(3) have "coincide_on \ \' (vars_of t)" unfolding coincide_on_def by auto from this have "(subst t \) = (subst t \')" using coincide_on_term by auto from this and \(subst s \) = (subst s \')\ and \occurs_in (subst t \') (subst s \')\ have "occurs_in (subst t \) (subst s \)" by auto from this and \s \ trms_ecl C\ have "\allowed_redex t C \" unfolding allowed_redex_def by auto from this and assms(1) show False by auto qed text \The next lemma states that the irreducibility of an instance of a set of terms is preserved when the substitution is replaced by its equivalent normal form.\ lemma irred_terms_and_reduced_subst: assumes "f = (\t. (trm_rep t S))" assumes "\ = (map_subst f \)" assumes "all_trms_irreducible (subst_set E \) f" assumes "I = int_clset S" assumes "equivalent_on \ \ (vars_of_cl (cl_ecl C)) I" assumes "lower_on \ \ (vars_of_cl (cl_ecl C))" assumes "E = (trms_ecl C)" assumes "\x \ S. \y. (y \ trms_ecl x \ dom_trm y (cl_ecl x))" assumes "C \ S" assumes "fo_interpretation I" shows "all_trms_irreducible (subst_set E \ ) f" proof (rule ccontr) assume "\all_trms_irreducible (subst_set E \) f" from this obtain t y where "y \ (subst_set E \)" "occurs_in t y" "f t \ t" unfolding all_trms_irreducible_def by metis from \occurs_in t y\ obtain p where "subterm y p t" unfolding occurs_in_def by auto from this and \f t \ t\ and assms(1) have "f y \ y" using subts_of_irred_trms_are_irred by blast from \y \ (subst_set E \)\ obtain z where "z \ E" and "y = (subst z \)" by auto from \z \ E\ have "(subst z \) \ (subst_set E \)" by auto have "subterm (subst z \) [] (subst z \)" by auto then have "occurs_in (subst z \) (subst z \)" unfolding occurs_in_def by blast from this and assms(3) and \(subst z \) \ (subst_set E \)\ have "f (subst z \) = (subst z \)" unfolding all_trms_irreducible_def by metis from this and \f y \ y\ and \y = (subst z \)\ have "(subst z \) \ (subst z \)" by metis from \z \ E\ and assms(7) assms(8) assms(9) have "dom_trm z (cl_ecl C)" by metis from this have "vars_of z \ vars_of_cl (cl_ecl C)" using dom_trm_vars by auto from this assms(5) have "equivalent_on \ \ (vars_of z) I" unfolding equivalent_on_def by auto from \vars_of z \ vars_of_cl (cl_ecl C)\ assms(6) have "lower_on \ \ (vars_of z)" unfolding lower_on_def by auto from \(subst z \) \ (subst z \)\ \lower_on \ \ (vars_of z)\ have "( (subst z \),(subst z \) ) \ trm_ord" using lower_on_term unfolding lower_or_eq_def by metis from this have "( (subst z \),(subst z \) ) \ trm_ord" using trm_ord_trans trm_ord_irrefl irrefl_def trans_def by metis from assms(10) \equivalent_on \ \ (vars_of z) I\ have "(I (subst z \) (subst z \))" using equivalent_on_term unfolding fo_interpretation_def by auto from this and assms(4) assms(1) \f (subst z \) = (subst z \)\ have "(subst z \) = f (subst z \)" unfolding same_values_def int_clset_def by metis from this \( (subst z \),(subst z \) ) \ trm_ord\ \(subst z \) \ (subst z \)\ assms(1) show False using trm_rep_is_lower by metis qed lemma no_valid_literal: assumes "L \ C" assumes "orient_lit_inst L t s pos \" assumes "\(validate_ground_clause (int_clset S) (subst_cl C \))" shows "trm_rep (subst t \) S \ trm_rep (subst s \) S" proof assume neg_hyp: "trm_rep (subst t \) S = trm_rep (subst s \) S" let ?I = "int_clset S" from neg_hyp have "validate_ground_eq ?I (Eq (subst t \) (subst s \))" unfolding same_values_def int_clset_def using validate_ground_eq.simps by (metis (mono_tags, lifting)) from \trm_rep (subst t \) S = trm_rep (subst s \) S\ have "validate_ground_eq ?I (Eq (subst s \) (subst t \))" unfolding same_values_def int_clset_def using validate_ground_eq.simps by (metis (mono_tags, lifting)) from \orient_lit_inst L t s pos \\ have "L = (Pos (Eq t s)) \ L = (Pos (Eq s t))" unfolding orient_lit_inst_def by auto from this have "subst_lit L \ = (Pos (Eq (subst t \) (subst s \))) \ subst_lit L \ = (Pos (Eq (subst s \) (subst t \)))" by auto from this and \validate_ground_eq ?I (Eq (subst s \) (subst t \))\ and \validate_ground_eq ?I (Eq (subst t \) (subst s \))\ have "validate_ground_lit ?I (subst_lit L \)" using validate_ground_lit.simps(1) by metis from assms(1) have "(subst_lit L \) \ (subst_cl C \)" by auto from \(subst_lit L \) \ (subst_cl C \)\ and \validate_ground_lit ?I (subst_lit L \)\ have "validate_ground_clause ?I (subst_cl C \)" using validate_ground_clause.elims(3) by blast from this and \\ validate_ground_clause ?I (subst_cl C \)\ show False by blast qed subsection \Lifting\ text \This section contains all the necessary lemmata for transforming ground inferences into first-order inferences. We show that all the necessary properties can be lifted.\ lemma lift_orient_lit_inst: assumes "orient_lit_inst L t s polarity \" assumes "subst_eq \ (comp \ \)" shows "orient_lit_inst L t s polarity \" proof - let ?\ = "(comp \ \)" have "polarity = pos \ polarity = neg" using sign.exhaust by auto then show ?thesis proof assume "polarity = pos" from this and assms(1) have "L = Pos (Eq t s) \ L = Pos (Eq s t)" and "( (subst t \), (subst s \)) \ trm_ord" unfolding orient_lit_inst_def by auto from assms(2) have "(subst t \) = (subst (subst t \) \)" by auto from assms(2) have "(subst s \) = (subst (subst s \) \)" by auto from \(subst t \) = (subst (subst t \) \)\ \(subst s \) = (subst (subst s \) \)\ \( (subst t \), (subst s \)) \ trm_ord\ have "( (subst (subst t \) \),(subst (subst s \) \)) \ trm_ord" by auto from this have "( (subst t \), (subst s \)) \ trm_ord" using trm_ord_subst by auto from this and \polarity = pos\ \L = Pos (Eq t s) \ L = Pos (Eq s t)\ show ?thesis unfolding orient_lit_inst_def by blast next assume "polarity = neg" from this and assms(1) have "L = Neg (Eq t s) \ L = Neg (Eq s t)" and "( (subst t \), (subst s \)) \ trm_ord" unfolding orient_lit_inst_def by auto from assms(2) have "(subst t \) = (subst (subst t \) \)" by auto from assms(2) have "(subst s \) = (subst (subst s \) \)" by auto from \(subst t \) = (subst (subst t \) \)\ \(subst s \) = (subst (subst s \) \)\ \( (subst t \), (subst s \)) \ trm_ord\ have "( (subst (subst t \) \),(subst (subst s \) \)) \ trm_ord" by auto from this have "( (subst t \), (subst s \)) \ trm_ord" using trm_ord_subst by auto from this and \polarity = neg\ \L = Neg (Eq t s) \ L = Neg (Eq s t)\ show ?thesis unfolding orient_lit_inst_def by blast qed qed lemma lift_maximal_literal: assumes "maximal_literal (subst_lit L \) (subst_cl C \)" shows "maximal_literal L C" proof (rule ccontr) assume "\maximal_literal L C" then obtain M where "M \ C" and "(L,M) \ lit_ord" unfolding maximal_literal_def by auto from \M \ C\ have "(subst_lit M \) \ (subst_cl C \)" by auto from \(L,M) \ lit_ord\ have "((subst_lit L \),(subst_lit M \)) \ lit_ord" using lit_ord_subst by auto from this and \(subst_lit M \) \ (subst_cl C \)\ and assms(1) show False unfolding maximal_literal_def by auto qed lemma lift_eligible_literal: assumes "eligible_literal L C \" assumes "\ \ \ \ \" shows "eligible_literal L C \" proof - from assms(1) have "(L \ sel (cl_ecl C) \ (sel(cl_ecl C) = {} \ (maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \))))" unfolding eligible_literal_def by auto then show ?thesis proof assume "L \ sel (cl_ecl C)" then show ?thesis unfolding eligible_literal_def by auto next assume "sel(cl_ecl C) = {} \ (maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \))" then have "sel (cl_ecl C) = {}" and "maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)" by auto let ?\ = "\ \ \" from assms(2) have "(subst_lit L \) = (subst_lit L ?\)" using subst_eq_lit by auto then have "(subst_lit L \) = (subst_lit (subst_lit L \) \)" using composition_of_substs_lit [of L \ \] by auto from assms(2) have "(subst_cl (cl_ecl C) \) = (subst_cl (cl_ecl C) ?\)" using subst_eq_cl [of \ "?\" "(cl_ecl C)"] by auto then have "(subst_cl (cl_ecl C) \) = (subst_cl (subst_cl (cl_ecl C) \) \)" using composition_of_substs_cl [of "cl_ecl C" \ \] by auto from \maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)\ \(subst_lit L \) = (subst_lit (subst_lit L \) \)\ \(subst_cl (cl_ecl C) \) = (subst_cl (subst_cl (cl_ecl C) \) \)\ have "maximal_literal (subst_lit (subst_lit L \) \) (subst_cl (subst_cl (cl_ecl C) \) \)" by auto from this have "maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)" using lift_maximal_literal by metis from this and \sel (cl_ecl C) = {}\ show ?thesis unfolding eligible_literal_def by auto qed qed lemma lift_allowed_redex: assumes "\ \ \ \ \" assumes "(allowed_redex u C \)" shows "(allowed_redex u C \)" proof (rule ccontr) assume "\(allowed_redex u C \)" from this obtain s where "s \ (trms_ecl C)" and "(occurs_in (subst u \) (subst s \))" unfolding allowed_redex_def by metis from \(occurs_in (subst u \) (subst s \))\ have "(occurs_in (subst (subst u \) \) (subst (subst s \) \))" using substs_preserve_occurs_in by auto from \\ \ \ \ \\ have "(subst u \) = (subst (subst u \) \)" by auto from \\ \ \ \ \\ have "(subst s \) = (subst (subst s \) \)" by auto from \(occurs_in (subst (subst u \) \) (subst (subst s \) \))\ \(subst u \) = (subst (subst u \) \)\ \(subst s \) = (subst (subst s \) \)\ have "(occurs_in (subst u \) (subst s \))" by auto from this and \s \ (trms_ecl C)\ assms(2) show False unfolding allowed_redex_def by auto qed lemma lift_decompose_literal: assumes "decompose_literal (subst_lit L \) t s polarity" assumes "subst_eq \ (comp \ \)" shows "decompose_literal (subst_lit L \) (subst t \) (subst s \) polarity" proof - let ?L = "(subst_lit L \)" let ?t' = "(subst t \)" let ?s' = "(subst s \)" let ?\ = "(comp \ \)" let ?L' = "(subst_lit ?L \)" from assms(2) have "(subst_lit L \) = (subst_lit L ?\)" using subst_eq_lit by auto from this have "(subst_lit L \) = ?L'" using composition_of_substs_lit by metis have "polarity = pos \ polarity = neg" using sign.exhaust by auto then show ?thesis proof assume "polarity = pos" from this and assms(1) have "?L = Pos (Eq t s) \ ?L = Pos (Eq s t)" unfolding decompose_literal_def decompose_equation_def by auto from \?L = Pos (Eq t s) \ ?L = Pos (Eq s t)\ have "?L' = Pos (Eq ?t' ?s') \ ?L' = Pos (Eq ?s' ?t')" by auto from this \(subst_lit L \) = ?L'\ have "(subst_lit L \) = Pos (Eq ?t' ?s') \ (subst_lit L \) = Pos (Eq ?s' ?t')" by auto from this \polarity = pos\ show ?thesis unfolding decompose_literal_def decompose_equation_def by auto next assume "polarity = neg" from this and assms(1) have "?L = Neg (Eq t s) \ ?L = Neg (Eq s t)" unfolding decompose_literal_def decompose_equation_def by auto from \?L = Neg (Eq t s) \ ?L = Neg (Eq s t)\ have "?L' = Neg (Eq ?t' ?s') \ ?L' = Neg (Eq ?s' ?t')" by auto from this and \(subst_lit L \) = ?L'\ have "(subst_lit L \) = Neg (Eq ?t' ?s') \ (subst_lit L \) = Neg (Eq ?s' ?t')" by auto from this \polarity = neg\ show ?thesis unfolding decompose_literal_def decompose_equation_def by auto qed qed lemma lift_dom_trm: assumes "dom_trm (subst t \) (subst_cl C \)" assumes "\ \ \ \ \" shows "dom_trm (subst t \) (subst_cl C \)" proof - let ?t = "(subst t \)" let ?t' = "(subst ?t \)" let ?t'' = "(subst t \)" have "?t' = (subst t (\ \ \))" by auto from assms(2) have "?t'' = (subst t (\ \ \))" by auto from this and \?t' = (subst t (\ \ \))\ have "?t' = ?t''" by metis from assms(1) have "(\ L u v p. (L \ (subst_cl C \) \ (decompose_literal L u v p) \ (( (p = neg \ ?t = u) \ (?t,u) \ trm_ord))))" unfolding dom_trm_def by auto from this obtain L u v p where "L \ (subst_cl C \)" "decompose_literal L u v p" "(( (p = neg \ ?t = u) \ (?t,u) \ trm_ord))" unfolding dom_trm_def by blast from \L \ (subst_cl C \)\ obtain L' where "L' \ C" "L = (subst_lit L' \)" by auto from this and \decompose_literal L u v p\ have "decompose_literal (subst_lit L' \) u v p" by auto from this assms(2) \L = (subst_lit L' \)\ have "decompose_literal (subst_lit L' \) (subst u \) (subst v \) p" using lift_decompose_literal [of L' \ u v p \ \] by auto let ?u = "(subst u \)" from \L' \ C\ have "(subst_lit L' \) \ (subst_cl C \)" by auto from \(( (p = neg \ ?t = u) \ (?t,u) \ trm_ord))\ have "(( (p = neg \ ?t' = ?u) \ (?t',?u) \ trm_ord))" using trm_ord_subst by auto from this and \?t' = ?t''\ have "(( (p = neg \ ?t'' = ?u) \ (?t'',?u) \ trm_ord))" by auto from this \(subst_lit L' \) \ (subst_cl C \)\ \decompose_literal (subst_lit L' \) (subst u \) (subst v \) p\ show "dom_trm (subst t \) (subst_cl C \)" unfolding dom_trm_def by auto qed lemma lift_irreducible_terms: assumes "T = get_trms C (dom_trms (subst_cl D \) (subst_set E \)) Ground" assumes "\ \ \ \ \" shows "\T'. ( (subst_set T' \) \ T \ T' = get_trms C' (dom_trms (subst_cl D \) (subst_set E \)) FirstOrder)" proof - let ?E = "(dom_trms (subst_cl D \) (subst_set E \))" let ?E' = "(dom_trms (subst_cl D \) (subst_set E \))" let ?T' = "(filter_trms C' ?E)" have "?T' = get_trms C' ?E FirstOrder" unfolding get_trms_def by auto from assms(1) have "T = ?E'" unfolding get_trms_def by auto have "(subst_set ?T' \) \ ?E'" proof fix x assume "x \ (subst_set ?T' \)" from this obtain x' where "x = (subst x' \)" and "x' \ ?T'" by auto from \x' \ ?T'\ have "x' \ ?E" using filter_trms_inclusion by auto from \x' \ ?E\ have "x' \ (subst_set E \)" and "dom_trm x' (subst_cl D \)" unfolding dom_trms_def by auto from \x' \ (subst_set E \)\ obtain y where "y \ E" and "x' = (subst y \)" by auto from \x' = (subst y \)\ and \dom_trm x' (subst_cl D \)\ have "dom_trm (subst y \) (subst_cl D \)" by auto from this assms(2) have "dom_trm (subst y \) (subst_cl D \)" using lift_dom_trm by auto from \y \ E\ have "(subst y \) \ (subst_set E \)" by auto from this and \dom_trm (subst y \) (subst_cl D \)\ have "(subst y \) \ ?E'" unfolding dom_trms_def by auto from assms(2) have "(subst y \) = (subst y (\ \ \))" by auto from this \x = (subst x' \)\ and \x' = (subst y \)\ have "x = (subst y \)" by auto from this and \(subst y \) \ ?E'\ show "x \ ?E'" by auto qed from this and \T = ?E'\ \?T' = get_trms C' ?E FirstOrder\ show ?thesis by auto qed text \We eventually deduce the following lemma, which allows one to transform ground derivations into first-order derivations.\ lemma lifting_lemma: assumes "derivable C P S \ Ground C'" shows "\ D \ \. ((derivable D P S \ FirstOrder C') \ (\ \ \ \ \) \ (trms_subsumes D C \))" proof (rule ccontr) assume hyp: "\ (\ D \ \. ((derivable D P S \ FirstOrder C') \ (\ \ \ \ \) \ (trms_subsumes D C \)))" from assms(1) have "P \ S" unfolding derivable_def by auto have not_sup: "\ (\P1 P2. (P = { P1, P2 } \ superposition P1 P2 C \ Ground C'))" proof assume "(\P1 P2. (P = { P1, P2 } \ superposition P1 P2 C \ Ground C'))" then obtain P1 P2 where "P = { P1, P2 }" "superposition P1 P2 C \ Ground C'" by auto from this obtain L t s u v M p Cl_P1 Cl_P2 Cl_C polarity t' u' L' trms_C where "L \ Cl_P1" "(M \ Cl_P2)" "(eligible_literal L P1 \)" "(eligible_literal M P2 \)" "(variable_disjoint P1 P2)" "(Cl_P1 = (cl_ecl P1))" "(Cl_P2 = (cl_ecl P2))" "(\ is_a_variable u')" "(allowed_redex u' P1 \)" "(C = (Ecl Cl_C trms_C))" "(orient_lit_inst M u v pos \)" "(orient_lit_inst L t s polarity \)" "((subst u \) \ (subst v \))" "(subterm t p u')" "(ck_unifier u' u \ Ground)" "(replace_subterm t p v t')" "(L' = mk_lit polarity (Eq t' s))" "(trms_C = get_trms Cl_C (dom_trms Cl_C (subst_set ((trms_ecl P1) \ (trms_ecl P2) \ { r. \ q. (q,p) \ (pos_ord P1 t) \ (subterm t q r) }) \)) Ground)" "(Cl_C = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \))" "(C' = (Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } ))" unfolding superposition_def get_trms_def by auto from \(ck_unifier u' u \ Ground)\ have " Unifier \ u' u" unfolding ck_unifier_def by auto from this have "(subst u' \) = (subst u \)" unfolding Unifier_def by auto from this have "unify u' u \ None" using MGU_exists by auto from this obtain \ where "unify u' u = Some \" by auto - from this have "MGU \ u' u" using unify_computes_MGU by auto - from this and \Unifier \ u' u\ obtain \ where "\ \ \ \ \" - unfolding MGU_def by auto + hence "IMGU \ u' u" by (rule unify_computes_IMGU) + with \Unifier \ u' u\ obtain \ where "\ \ \ \ \" + unfolding IMGU_iff_Idem_and_MGU MGU_def by auto from \\ \ \ \ \\ and \(eligible_literal L P1 \)\ have "eligible_literal L P1 \" using lift_eligible_literal by auto from \\ \ \ \ \\ and \(eligible_literal M P2 \)\ have "eligible_literal M P2 \" using lift_eligible_literal by auto - from \MGU \ u' u\ have "ck_unifier u' u \ FirstOrder" unfolding ck_unifier_def by auto + from \IMGU \ u' u\ have "ck_unifier u' u \ FirstOrder" unfolding ck_unifier_def by auto from \\ \ \ \ \\ have "(subst u \) = (subst (subst u \) \)" by auto from \\ \ \ \ \\ have "(subst v \) = (subst (subst v \) \)" by auto from \((subst u \) \ (subst v \))\ \(subst u \) = (subst (subst u \) \)\ \(subst v \) = (subst (subst v \) \)\ have "(subst (subst u \) \) \ (subst (subst v \) \)" by auto from this have "(subst u \) \ (subst v \)" by auto from \\ \ \ \ \\ \(allowed_redex u' P1 \)\ have "(allowed_redex u' P1 \)" using lift_allowed_redex [of \ \ \ ] by auto from \\ \ \ \ \\ \orient_lit_inst M u v pos \\ have "orient_lit_inst M u v pos \" using lift_orient_lit_inst by auto from \\ \ \ \ \\ \orient_lit_inst L t s polarity \\ have "orient_lit_inst L t s polarity \" using lift_orient_lit_inst by auto from \(Cl_C = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \))\ and \C' = (Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )\ have "(Cl_C = (subst_cl C' \))" by auto obtain E where "E = ((trms_ecl P1) \ (trms_ecl P2) \ { r. \ q. (q,p) \ (pos_ord P1 t) \ (subterm t q r) })" by auto from this and \(Cl_C = (subst_cl C' \))\ \trms_C = (get_trms Cl_C (dom_trms Cl_C (subst_set ((trms_ecl P1) \ (trms_ecl P2) \ { r. \ q. (q,p) \ (pos_ord P1 t) \ (subterm t q r) }) \)) Ground)\ have "trms_C = (get_trms Cl_C (dom_trms (subst_cl C' \) (subst_set E \)) Ground)" by auto let ?Cl_C' = "(subst_cl C' \)" from \\ \ \ \ \\ \trms_C = (get_trms Cl_C (dom_trms (subst_cl C' \) (subst_set E \)) Ground)\ obtain "\T'. ( (subst_set T' \) \ trms_C \ T' = get_trms ?Cl_C' (dom_trms (subst_cl C' \) (subst_set E \)) FirstOrder)" using lift_irreducible_terms by auto from this obtain T' where "(subst_set T' \) \ trms_C" and "T' = get_trms ?Cl_C' (dom_trms (subst_cl C' \) (subst_set E \)) FirstOrder" by auto obtain C_fo where "C_fo = (Ecl ?Cl_C' T')" by auto from \C' = (Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )\ have "(?Cl_C' = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \))" by auto from \L \ Cl_P1\ \(M \ Cl_P2)\ \(eligible_literal L P1 \)\ \(eligible_literal M P2 \)\ \(variable_disjoint P1 P2)\ \(Cl_P1 = (cl_ecl P1))\ \(Cl_P2 = (cl_ecl P2))\ \(\ is_a_variable u')\ \(allowed_redex u' P1 \)\ \(C_fo = (Ecl ?Cl_C' T'))\ \(orient_lit_inst M u v pos \)\ \(orient_lit_inst L t s polarity \)\ \((subst u \) \ (subst v \))\ \(subterm t p u')\ \(ck_unifier u' u \ FirstOrder)\ \(replace_subterm t p v t')\ \(L' = mk_lit polarity (Eq t' s))\ \T' = (get_trms ?Cl_C' (dom_trms (subst_cl C' \) (subst_set E \)) FirstOrder)\ \(?Cl_C' = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \))\ \(C' = (Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } ))\ \E = ((trms_ecl P1) \ (trms_ecl P2) \ { r. \ q. (q,p) \ (pos_ord P1 t) \ (subterm t q r) })\ have "superposition P1 P2 C_fo \ FirstOrder C'" unfolding superposition_def by blast from this \P = { P1, P2 }\ \P \ S\ have "(derivable C_fo P S \ FirstOrder C')" unfolding derivable_def by auto from \(?Cl_C' = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \))\ have i: "(subst_cl ?Cl_C' \) = (subst_cl (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \)) \" by auto have ii: "(subst_cl (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \) \) = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) (\ \ \))" using composition_of_substs_cl [of "((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } ))" ] by auto from \\ \ \ \ \\ have "(subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \) = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) (\ \ \))" using subst_eq_cl [of \ "\ \ \" "((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } ))" ] by auto from this and i ii \Cl_C = (subst_cl ((Cl_P1 - { L }) \ ((Cl_P2 - { M }) \ { L' } )) \)\ have "(subst_cl ?Cl_C' \) = Cl_C" by metis from this and \(C = (Ecl Cl_C trms_C))\ and \(C_fo = (Ecl ?Cl_C' T'))\ have "(subst_cl (cl_ecl C_fo) \) = (cl_ecl C)" by auto from \(subst_set T' \) \ trms_C\ and \(C = (Ecl Cl_C trms_C))\ and \(C_fo = (Ecl ?Cl_C' T'))\ have "(subst_set (trms_ecl C_fo) \) \ (trms_ecl C)" by auto from \(subst_cl (cl_ecl C_fo) \) = (cl_ecl C)\ \(subst_set (trms_ecl C_fo) \) \ (trms_ecl C)\ have "(trms_subsumes C_fo C \)" unfolding trms_subsumes_def by auto from this and \(derivable C_fo P S \ FirstOrder C')\ and \\ \ \ \ \\ and hyp show False by auto qed have not_fact: "\ (\P1. ({ P1 } = P \ factorization P1 C \ Ground C'))" proof assume "(\P1. ({ P1 } = P \ factorization P1 C \ Ground C'))" then obtain P1 where "P = { P1 }" "factorization P1 C \ Ground C'" by auto from this obtain L1 L2 L' t s u v Cl_P Cl_C trms_C where "eligible_literal L1 P1 \" "L1 \ (cl_ecl P1)" "L2 \ (cl_ecl P1) - { L1 }" "Cl_C = (cl_ecl C)" "(Cl_P = (cl_ecl P1))" "(orient_lit_inst L1 t s pos \)" "(orient_lit_inst L2 u v pos \)" "((subst t \) \ (subst s \))" "(subst t \) \ (subst v \)" "(ck_unifier t u \ Ground)" "(L' = Neg (Eq s v))" "C = (Ecl Cl_C trms_C)" "trms_C = (get_trms Cl_C (dom_trms Cl_C (subst_set ( (trms_ecl P1) \ (proper_subterms_of t) ) \))) Ground" "(Cl_C = (subst_cl ( (Cl_P - { L2 }) \ { L' } )) \)" "(C' = ( (Cl_P - { L2 }) \ { L' } ))" unfolding factorization_def get_trms_def by auto from \(ck_unifier t u \ Ground)\ have " Unifier \ t u" unfolding ck_unifier_def Unifier_def by auto from this have "(subst t \) = (subst u \)" unfolding Unifier_def by auto from this have "unify t u \ None" using MGU_exists by auto from this obtain \ where "unify t u = Some \" by auto - from this have "MGU \ t u" using unify_computes_MGU by auto - from this and \Unifier \ t u\ obtain \ where "\ \ \ \ \" - unfolding MGU_def by auto + hence "IMGU \ t u" by (rule unify_computes_IMGU) + with \Unifier \ t u\ obtain \ where "\ \ \ \ \" + unfolding IMGU_iff_Idem_and_MGU MGU_def by auto from \\ \ \ \ \\ and \(eligible_literal L1 P1 \)\ have "eligible_literal L1 P1 \" using lift_eligible_literal by auto - from \MGU \ t u\ have "ck_unifier t u \ FirstOrder" unfolding ck_unifier_def by auto + from \IMGU \ t u\ have "ck_unifier t u \ FirstOrder" unfolding ck_unifier_def by auto from \\ \ \ \ \\ have "(subst t \) = (subst (subst t \) \)" by auto from \\ \ \ \ \\ have "(subst s \) = (subst (subst s \) \)" by auto from \\ \ \ \ \\ have "(subst v \) = (subst (subst v \) \)" by auto from \((subst t \) \ (subst s \))\ \(subst t \) = (subst (subst t \) \)\ \(subst s \) = (subst (subst s \) \)\ have "(subst (subst t \) \) \ (subst (subst s \) \)" by auto from this have "(subst t \) \ (subst s \)" by auto from \((subst t \) \ (subst v \))\ \(subst t \) = (subst (subst t \) \)\ \(subst v \) = (subst (subst v \) \)\ have "(subst (subst t \) \) \ (subst (subst v \) \)" by auto from this have "(subst t \) \ (subst v \)" by auto from \\ \ \ \ \\ \orient_lit_inst L1 t s pos \\ have "orient_lit_inst L1 t s pos \" using lift_orient_lit_inst by auto from \\ \ \ \ \\ \orient_lit_inst L2 u v pos \\ have "orient_lit_inst L2 u v pos \" using lift_orient_lit_inst by auto from \(Cl_C = (subst_cl ( (Cl_P - { L2 }) \ { L' } )) \)\ and \C' = ( (Cl_P - { L2 }) \ { L' } )\ have "(Cl_C = (subst_cl C' \))" by auto obtain E where "E = (trms_ecl P1)" by auto from this and \(Cl_C = (subst_cl C' \))\ \trms_C = (get_trms Cl_C (dom_trms Cl_C (subst_set ( (trms_ecl P1) \ (proper_subterms_of t) ) \))) Ground\ have "trms_C = (get_trms Cl_C (dom_trms (subst_cl C' \) (subst_set (E \ (proper_subterms_of t)) \)) Ground)" by auto let ?Cl_C' = "(subst_cl C' \)" from \\ \ \ \ \\ \trms_C = (get_trms Cl_C (dom_trms (subst_cl C' \) (subst_set (E \ (proper_subterms_of t)) \)) Ground)\ have "\T'. ( (subst_set T' \) \ trms_C \ T' = get_trms ?Cl_C' (dom_trms (subst_cl C' \) (subst_set (E \ (proper_subterms_of t)) \)) FirstOrder)" using lift_irreducible_terms by blast from this obtain T' where "(subst_set T' \) \ trms_C" and "T' = get_trms ?Cl_C' (dom_trms (subst_cl C' \) (subst_set (E \ (proper_subterms_of t)) \)) FirstOrder" by auto obtain C_fo where "C_fo = (Ecl ?Cl_C' T')" by auto from \C' = ( (Cl_P - { L2 }) \ { L' } )\ have "(?Cl_C' = (subst_cl ( (Cl_P - { L2 }) \ { L' } ) \))" by auto from \C_fo = (Ecl ?Cl_C' T')\ have "?Cl_C' = (cl_ecl C_fo)" by auto have "?Cl_C' = (subst_cl C' \)" by auto from \eligible_literal L1 P1 \\ \L1 \ (cl_ecl P1)\ \L2 \ (cl_ecl P1) - { L1 }\ \?Cl_C' = (cl_ecl C_fo)\ \(Cl_P = (cl_ecl P1))\ \(orient_lit_inst L1 t s pos \)\ \(orient_lit_inst L2 u v pos \)\ \((subst t \) \ (subst s \))\ \(subst t \) \ (subst v \)\ \(ck_unifier t u \ FirstOrder)\ \(L' = Neg (Eq s v))\ \C_fo = (Ecl ?Cl_C' T')\ \T' = get_trms?Cl_C' (dom_trms ?Cl_C' (subst_set (E \ (proper_subterms_of t)) \)) FirstOrder\ \(?Cl_C' = (subst_cl ( (Cl_P - { L2 }) \ { L' } ) \))\ \C' = ( (Cl_P - { L2 }) \ { L' } )\ \E = (trms_ecl P1)\ have "factorization P1 C_fo \ FirstOrder C'" unfolding factorization_def by blast from this \P = { P1 }\ \P \ S\ have "(derivable C_fo P S \ FirstOrder C')" unfolding derivable_def by auto from \(?Cl_C' = (subst_cl ( (Cl_P - { L2 }) \ { L' } ) \))\ have i: "(subst_cl ?Cl_C' \) = (subst_cl (subst_cl ( (Cl_P - { L2 }) \ { L' } ) \)) \" by auto have ii: "(subst_cl (subst_cl ( (Cl_P - { L2 }) \ { L' } ) \) \) = (subst_cl ( (Cl_P - { L2 }) \ { L' } ) (\ \ \))" using composition_of_substs_cl [of "( (Cl_P - { L2 }) \ { L' } )" ] by auto from \\ \ \ \ \\ have "(subst_cl ( (Cl_P - { L2 }) \ { L' } ) \) = (subst_cl ( (Cl_P - { L2 }) \ { L' } ) (\ \ \))" using subst_eq_cl [of \ "\ \ \" "( (Cl_P - { L2 }) \ { L' } )" ] by auto from this and i ii \Cl_C = (subst_cl ( (Cl_P - { L2 }) \ { L' } ) \)\ have "(subst_cl ?Cl_C' \) = Cl_C" by metis from this and \(C = (Ecl Cl_C trms_C))\ and \(C_fo = (Ecl ?Cl_C' T'))\ have "(subst_cl (cl_ecl C_fo) \) = (cl_ecl C)" by auto from \(subst_set T' \) \ trms_C\ and \(C = (Ecl Cl_C trms_C))\ and \(C_fo = (Ecl ?Cl_C' T'))\ have "(subst_set (trms_ecl C_fo) \) \ (trms_ecl C)" by auto from \(subst_cl (cl_ecl C_fo) \) = (cl_ecl C)\ \(subst_set (trms_ecl C_fo) \) \ (trms_ecl C)\ have "(trms_subsumes C_fo C \)" unfolding trms_subsumes_def by auto from this and \(derivable C_fo P S \ FirstOrder C')\ and \\ \ \ \ \\ and hyp show False by auto qed have not_ref: "\ (\P1. ({ P1 } = P \ reflexion P1 C \ Ground C'))" proof assume "(\P1. ({ P1 } = P \ reflexion P1 C \ Ground C'))" then obtain P1 where "{ P1 } = P" "reflexion P1 C \ Ground C'" by auto from this obtain L1 t s Cl_P Cl_C trms_C where "(eligible_literal L1 P1 \)" "(L1 \ (cl_ecl P1))" "(Cl_C = (cl_ecl C))" "(Cl_P = (cl_ecl P1))" "(orient_lit_inst L1 t s neg \)" "(ck_unifier t s \ Ground)" "(C = (Ecl Cl_C trms_C))" "(trms_C = get_trms Cl_C (dom_trms Cl_C (subst_set ( (trms_ecl P1) \ { t } ) \)) Ground)" "(Cl_C = (subst_cl ((Cl_P - { L1 }) )) \)" "(C' = ((Cl_P - { L1 }) ))" unfolding reflexion_def get_trms_def by auto from \(ck_unifier t s \ Ground)\ have " Unifier \ t s" unfolding ck_unifier_def Unifier_def by auto from this have "(subst t \) = (subst s \)" unfolding Unifier_def by auto from this have "unify t s \ None" using MGU_exists by auto from this obtain \ where "unify t s = Some \" by auto - from this have "MGU \ t s" using unify_computes_MGU by auto - from this and \Unifier \ t s\ obtain \ where "\ \ \ \ \" - unfolding MGU_def by auto + hence "IMGU \ t s" by (rule unify_computes_IMGU) + with \Unifier \ t s\ obtain \ where "\ \ \ \ \" + unfolding IMGU_iff_Idem_and_MGU MGU_def by auto from \\ \ \ \ \\ and \(eligible_literal L1 P1 \)\ have "eligible_literal L1 P1 \" using lift_eligible_literal by auto - from \MGU \ t s\ have "ck_unifier t s \ FirstOrder" unfolding ck_unifier_def by auto + from \IMGU \ t s\ have "ck_unifier t s \ FirstOrder" unfolding ck_unifier_def by auto from \\ \ \ \ \\ \orient_lit_inst L1 t s neg \\ have "orient_lit_inst L1 t s neg \" using lift_orient_lit_inst by auto from \(Cl_C = (subst_cl ((Cl_P - { L1 }) )) \)\ and \C' = ((Cl_P - { L1 }) )\ have "(Cl_C = (subst_cl C' \))" by auto obtain E where "E = (trms_ecl P1)" by auto from this and \(Cl_C = (subst_cl C' \))\ \trms_C = (get_trms Cl_C (dom_trms Cl_C (subst_set ( (trms_ecl P1) \ { t } ) \))) Ground\ have "trms_C = (get_trms Cl_C (dom_trms (subst_cl C' \) (subst_set (E \ { t }) \)) Ground)" by auto let ?Cl_C' = "(subst_cl C' \)" from \\ \ \ \ \\ \trms_C = (get_trms Cl_C (dom_trms (subst_cl C' \) (subst_set (E \ { t }) \)) Ground)\ have "\T'. ( (subst_set T' \) \ trms_C \ T' = get_trms ?Cl_C' (dom_trms (subst_cl C' \) (subst_set (E \ { t }) \)) FirstOrder)" using lift_irreducible_terms by blast from this obtain T' where "(subst_set T' \) \ trms_C" and "T' = get_trms ?Cl_C' (dom_trms (subst_cl C' \) (subst_set (E \ { t }) \)) FirstOrder" by auto obtain C_fo where "C_fo = (Ecl ?Cl_C' T')" by auto from \C' = ((Cl_P - { L1 }) )\ have "(?Cl_C' = (subst_cl ((Cl_P - { L1 }) ) \))" by auto from \C_fo = (Ecl ?Cl_C' T')\ have "?Cl_C' = (cl_ecl C_fo)" by auto have "?Cl_C' = (subst_cl C' \)" by auto from \(eligible_literal L1 P1 \)\ \(L1 \ (cl_ecl P1))\ \?Cl_C' = (cl_ecl C_fo)\ \(Cl_P = (cl_ecl P1))\ \(orient_lit_inst L1 t s neg \)\ \(ck_unifier t s \ FirstOrder)\ \(C_fo = (Ecl ?Cl_C' T'))\ \(T' = get_trms ?Cl_C' (dom_trms (subst_cl C' \) (subst_set (E \ { t }) \)) FirstOrder)\ \(?Cl_C' = (subst_cl ((Cl_P - { L1 }) )) \)\ \(C' = ((Cl_P - { L1 }) ))\ \E = (trms_ecl P1)\ have "reflexion P1 C_fo \ FirstOrder C'" unfolding reflexion_def by metis from this \{ P1 } = P\ \P \ S\ have "(derivable C_fo P S \ FirstOrder C')" unfolding derivable_def by auto from \(?Cl_C' = (subst_cl ((Cl_P - { L1 }) )) \)\ have i: "(subst_cl ?Cl_C' \) = (subst_cl (subst_cl ((Cl_P - { L1 }) ) \)) \" by auto have ii: "(subst_cl (subst_cl ((Cl_P - { L1 }) ) \) \) = (subst_cl ((Cl_P - { L1 }) ) (\ \ \))" using composition_of_substs_cl [of "((Cl_P - { L1 }) )" ] by auto from \\ \ \ \ \\ have "(subst_cl((Cl_P - { L1 }) ) \) = (subst_cl ((Cl_P - { L1 }) ) (\ \ \))" using subst_eq_cl [of \ "\ \ \" "((Cl_P - { L1 }) )" ] by auto from this and i ii \Cl_C = (subst_cl ((Cl_P - { L1 }) ) \)\ have "(subst_cl ?Cl_C' \) = Cl_C" by metis from this and \(C = (Ecl Cl_C trms_C))\ and \(C_fo = (Ecl ?Cl_C' T'))\ have "(subst_cl (cl_ecl C_fo) \) = (cl_ecl C)" by auto from \(subst_set T' \) \ trms_C\ and \(C = (Ecl Cl_C trms_C))\ and \(C_fo = (Ecl ?Cl_C' T'))\ have "(subst_set (trms_ecl C_fo) \) \ (trms_ecl C)" by auto from \(subst_cl (cl_ecl C_fo) \) = (cl_ecl C)\ \(subst_set (trms_ecl C_fo) \) \ (trms_ecl C)\ have "(trms_subsumes C_fo C \)" unfolding trms_subsumes_def by auto from this and \(derivable C_fo P S \ FirstOrder C')\ and \\ \ \ \ \\ and hyp show False by auto qed from not_sup not_ref not_fact and assms(1) show False unfolding derivable_def by blast qed lemma trms_subsumes_and_red_inf: assumes "trms_subsumes D C \" assumes "redundant_inference (subst_ecl D \) S P \" assumes "\ \ \ \ \" shows "redundant_inference C S P \" proof - from assms(2) obtain S' where "S' \ (instances S)" "(set_entails_clause (clset_instances S') (cl_ecl (subst_ecl D \)))" "(\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl (subst_ecl D \))))" "(\x \ S'. \D' \ P. (((fst x),(snd x)),(D',\)) \ ecl_ord)" unfolding redundant_inference_def by auto from assms(1) have "(subst_cl (cl_ecl D) \) = (cl_ecl C)" unfolding trms_subsumes_def by auto obtain Cl_D T where "D = Ecl Cl_D T" using eclause.exhaust by auto from this have "(cl_ecl D) = Cl_D" and "trms_ecl D = T" by auto from \D = Ecl Cl_D T\ have "subst_ecl D \ = Ecl (subst_cl Cl_D \) (subst_set T \)" by auto from this have "(cl_ecl (subst_ecl D \)) = (subst_cl Cl_D \)" and "trms_ecl (subst_ecl D \) = (subst_set T \)" by auto from \(cl_ecl (subst_ecl D \)) = (subst_cl Cl_D \)\ and \(cl_ecl D) = Cl_D\ \(subst_cl (cl_ecl D) \) = (cl_ecl C)\ have "(cl_ecl (subst_ecl D \)) = (cl_ecl C)" by auto from this and \(set_entails_clause (clset_instances S') (cl_ecl (subst_ecl D \)))\ have "(set_entails_clause (clset_instances S') (cl_ecl C))" by auto from \trms_ecl D = T\ and \trms_ecl (subst_ecl D \) = (subst_set T \)\ have "trms_ecl (subst_ecl D \) = (subst_set (trms_ecl D) \)" by auto from assms(1) have "(subst_set (trms_ecl D) \) \ (trms_ecl C)" unfolding trms_subsumes_def by auto have ii: "(\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl C)))" proof (rule ccontr) assume "\ (\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl C)))" from this obtain x where "x \ S'" and "\( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl C))" by auto from \x \ S'\ and \(\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl (subst_ecl D \))))\ have "( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl (subst_ecl D \)))" by auto obtain E1 where "E1 = (subst_set (trms_ecl (fst x)) (snd x))" by auto obtain E2 where "E2 = (subst_set (trms_ecl D) \)" by auto obtain E2' where "E2' = (trms_ecl C)" by auto from \E2 = (subst_set (trms_ecl D) \)\ \E2' = (trms_ecl C)\ \(subst_set (trms_ecl D) \) \ (trms_ecl C)\ have "E2 \ E2'" by auto from \E1 = (subst_set (trms_ecl (fst x)) (snd x))\ \E2 = (subst_set (trms_ecl D) \)\ \trms_ecl (subst_ecl D \) = (subst_set (trms_ecl D) \)\ \( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl (subst_ecl D \)))\ have "subterms_inclusion E1 E2" by auto from this and \E2 \ E2'\ have "subterms_inclusion E1 E2'" using subterms_inclusion_subset by auto from this and \E1 = (subst_set (trms_ecl (fst x)) (snd x))\ \E2' = (trms_ecl C)\ and \\( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl C))\ show False by auto qed from this and \(set_entails_clause (clset_instances S') (cl_ecl C))\ and \(\x \ S'. \D' \ P. (((fst x),(snd x)),(D',\)) \ ecl_ord)\ and \S' \ (instances S)\ show "redundant_inference C S P \" unfolding redundant_inference_def by auto qed lemma lift_inference: assumes "inference_saturated S" shows "ground_inference_saturated S" proof (rule ccontr) assume "\ (ground_inference_saturated S)" from this obtain C P \ C' where "derivable C P S \ Ground C'" "ground_clause (cl_ecl C)" "grounding_set P \" "\redundant_inference C S P \" unfolding ground_inference_saturated_def by blast from \derivable C P S \ Ground C'\ obtain D \ \ where "derivable D P S \ FirstOrder C'" "\ \ \ \ \" "trms_subsumes D C \" using lifting_lemma by blast from \trms_subsumes D C \\ and \\redundant_inference C S P \\ \\ \ \ \ \\ have "\ redundant_inference (subst_ecl D \) S P \" using trms_subsumes_and_red_inf by auto from this and \derivable C P S \ Ground C'\ \derivable D P S \ FirstOrder C'\ \trms_subsumes D C \\ \\ \ \ \ \\ \ground_clause (cl_ecl C)\ \grounding_set P \\ \\ redundant_inference (subst_ecl D \) S P \\ assms(1) show False unfolding inference_saturated_def by auto qed lemma lift_redundant_cl : assumes "C' = subst_cl D \" assumes "redundant_clause C S \ C'" assumes "\ \ \ \ \" assumes "finite D" shows "redundant_clause C S \ D" proof - from assms(2) have "(\S'. (S' \ (instances S) \ (set_entails_clause (clset_instances S') (cl_ecl C)) \ (\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl C))) \ (\x \ S'. ( ((mset_ecl ((fst x),(snd x))),(mset_cl (C',\))) \ (mult (mult trm_ord)) \ (mset_ecl ((fst x),(snd x))) = mset_cl (C',\)))))" unfolding redundant_clause_def by auto from this obtain S' where i: "S' \ (instances S)" and ii: "(set_entails_clause (clset_instances S') (cl_ecl C))" and iii: "(\x \ S'. ( subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl C)))" and iv: "(\x \ S'. ( ((mset_ecl ((fst x),(snd x))),(mset_cl (C',\))) \ (mult (mult trm_ord)) \ (mset_ecl ((fst x),(snd x))) = mset_cl (C',\)))" by auto let ?m1 = "mset_cl (C',\)" let ?m2 = "mset_cl (D,\)" from assms(1) assms(3) assms(4) have "mset_cl (C',\) = mset_cl (D,\) \ (mset_cl (C',\),mset_cl (D,\)) \ (mult (mult trm_ord))" using mset_subst by auto from this iv have "(\x \ S'. ( ((mset_ecl ((fst x),(snd x))),(mset_cl (D,\))) \ (mult (mult trm_ord)) \ (mset_ecl ((fst x),(snd x))) = mset_cl (D,\)))" using mult_mult_trm_ord_trans unfolding trans_def by metis from this and i ii iii show ?thesis unfolding redundant_clause_def by meson qed text \We deduce the following (trivial) lemma, stating that sets that are closed under all inferences are also saturated.\ lemma inference_closed_sets_are_saturated: assumes "inference_closed S" assumes "\x \ S. (finite (cl_ecl x))" shows "clause_saturated S" proof (rule ccontr) assume "\?thesis" from this obtain C P \ C' D \ \ where "(derivable C P S \ Ground C')" "(ground_clause (cl_ecl C))" "(derivable D P S \ FirstOrder C')" "(trms_subsumes D C \)" "(\ \ \ \ \)" "\(redundant_clause (subst_ecl D \) S \ C')" unfolding clause_saturated_def by blast from \(derivable D P S \ FirstOrder C')\ assms(1) have "D \ S" unfolding inference_closed_def by auto from \derivable D P S \ FirstOrder C'\ have "(cl_ecl D) = (subst_cl C' \)" using derivable_clauses_lemma by auto from \trms_subsumes D C \\ have "(cl_ecl C) = (subst_cl (cl_ecl D) \)" unfolding trms_subsumes_def by blast from \\ \ \ \ \\ have "subst_cl (cl_ecl D) \ = subst_cl (cl_ecl D) (\ \ \)" using subst_eq_cl by blast then have "(subst_cl (cl_ecl D) \) = (subst_cl (subst_cl (cl_ecl D) \) \)" using composition_of_substs_cl [of "cl_ecl D" \ \] by auto from this and \(cl_ecl C) = (subst_cl (cl_ecl D) \)\ \(ground_clause (cl_ecl C))\ have "ground_clause (subst_cl (cl_ecl D) \)" by auto from this \D \ S\ \(cl_ecl D) = (subst_cl C' \)\ \(cl_ecl D) = (subst_cl C' \)\ have "redundant_clause (subst_ecl D \) S \ (subst_cl C' \)" using self_redundant_clause by metis from \derivable D P S \ FirstOrder C'\ have "P \ S" unfolding derivable_def by auto from this assms(2) have "\x \ P. (finite (cl_ecl x))" by auto from this \derivable D P S \ FirstOrder C'\ have "finite C'" using derivable_clauses_are_finite by auto from this \(cl_ecl D) = subst_cl C' \\ \redundant_clause (subst_ecl D \) S \ (subst_cl C' \)\ \(\ \ \ \ \)\ have "redundant_clause (subst_ecl D \) S \ C'" using lift_redundant_cl by metis from this and \\(redundant_clause (subst_ecl D \) S \ C')\ show False by auto qed subsection \Satisfiability of Saturated Sets with No Empty Clause\ text \We are now in the position to prove that the previously constructed interpretation is indeed a model of the set of extended clauses, if the latter is saturated and does not contain an extended clause with empty clausal part. More precisely, the constructed interpretation satisfies the clausal part of every extended clause whose attached set of terms is in normal form. This is the case in particular if this set is empty, hence if the clause is an input clause. Note that we do not provide any function for explicitly constructing such saturated sets, except by generating all derivable clauses (see below).\ lemma int_clset_is_a_model: assumes "ground_inference_saturated S" assumes all_finite: "\x \ S. (finite (cl_ecl x))" assumes "Ball S well_constrained" assumes all_non_empty: "\x \ S. (cl_ecl x) \ {}" assumes "closed_under_renaming S" shows "\ C \. (fst pair = C) \ \ = (snd pair) \ C \ S \ ground_clause (subst_cl (cl_ecl C) \) \ (all_trms_irreducible (subst_set (trms_ecl C) \) (\t. (trm_rep t S))) \ validate_ground_clause (same_values (\t. (trm_rep t S))) (subst_cl (cl_ecl C) \)" (is "?P pair") proof ((rule wf_induct [of "ecl_ord" "?P" "pair"]),(simp add: wf_ecl_ord)) next text \The proof is by induction and contradiction. We consider a minimal instance that is not true in the interpretation and we derive a contradiction.\ fix pair assume hyp_ind: "\y. (y,pair) \ ecl_ord \ (?P y)" let ?I = "(int_clset S)" have "fo_interpretation ?I" unfolding int_clset_def using trm_rep_compatible_with_structure same_values_fo_int by metis show "(?P pair)" proof (rule ccontr) assume "\(?P pair)" then obtain C \ where "C = (fst pair)" and "\ = (snd pair)" and "C \ S" and "ground_clause (subst_cl (cl_ecl C) \)" and "(all_trms_irreducible (subst_set (trms_ecl C) \) (\t. (trm_rep t S)))" and cm: "\validate_ground_clause (int_clset S) (subst_cl (cl_ecl C) \)" unfolding int_clset_def by metis text \First, we prove that no reduction is possible (otherwise the superposition rule applies).\ let ?nored = "(\L1 L2 D t s u' u v p polarity \'. \ ((reduction L1 C \' t s polarity L2 u u' p v D ?I S \) \ (variable_disjoint C D)))" have ?nored proof (rule ccontr) assume "\ ?nored" then obtain L1 L2 D t s u' u v p polarity \' where red: "reduction L1 C \' t s polarity L2 u u' p v D ?I S \" and "variable_disjoint C D" by blast from red have "(subst u \') \ (subst v \')" unfolding reduction_def by blast from red have "ground_clause (subst_cl (cl_ecl D) \')" unfolding reduction_def by blast from red have "(coincide_on \ \' (vars_of_cl (cl_ecl C)))" unfolding reduction_def by blast from red have "\ is_a_variable u'" unfolding reduction_def by blast from red have "D \ S" unfolding reduction_def by blast from red have el1: "(eligible_literal L1 C \')" unfolding reduction_def by blast from red have el2: "(eligible_literal L2 D \')" unfolding reduction_def by blast from red have "D \ S" unfolding reduction_def by blast from red have "(minimal_redex p (subst t \) C S t)" unfolding reduction_def by blast from red have l1: "L1 \ (cl_ecl C)" unfolding reduction_def by blast from red have l2: "L2 \ (cl_ecl D)" unfolding reduction_def by blast from red have o1: "(orient_lit_inst L1 t s polarity \')" unfolding reduction_def by blast from red have o2: "(orient_lit_inst L2 u v pos \')" unfolding reduction_def by blast from red have e:"(subst u' \') = (subst u \')" unfolding reduction_def by blast from red have "(\ validate_ground_clause ?I (subst_cl ( (cl_ecl D) - { L2 } ) \'))" unfolding reduction_def by blast from red have "(\x \ (cl_ecl D) - { L2 }. ( (subst_lit x \'),(subst_lit L2 \')) \ lit_ord)" unfolding reduction_def by blast from red have st: "(subterm t p u')" unfolding reduction_def by blast from red have "(allowed_redex u' C \)" unfolding reduction_def by blast from st have "u' \ (subterms_of t)" using occurs_in_def by auto from this and o1 have "u' \ (subterms_of_lit L1)" using orient_lit_inst_subterms by auto from this and \L1 \ (cl_ecl C)\ have "u' \ (subterms_of_cl (cl_ecl C))" by auto from this and \(allowed_redex u' C \)\ and \C \ S\ and \(coincide_on \ \' (vars_of_cl (cl_ecl C)))\ assms(3) have rte: "(allowed_redex u' C \')" using allowed_redex_coincide [of u' C \ \'] by metis from red have "( (subst_lit L2 \'),(subst_lit L1 \')) \ lit_ord" unfolding reduction_def by blast from red have "(all_trms_irreducible (subst_set (trms_ecl D) \') (\t. (trm_rep t S)))" unfolding reduction_def by blast from red have "?I (subst u \') (subst v \')" unfolding reduction_def by blast from e have t: "ck_unifier u' u \' Ground" unfolding ck_unifier_def Unifier_def by auto have "\x \ (cl_ecl D). ( (mset_lit (subst_lit x \')),(mset_lit (subst_lit L1 \'))) \ (mult trm_ord)" proof (rule ccontr) assume "\(\x \ (cl_ecl D). ( (mset_lit (subst_lit x \')),(mset_lit (subst_lit L1 \'))) \ (mult trm_ord))" from this obtain x where "x \ (cl_ecl D)" and "( (mset_lit (subst_lit x \')),(mset_lit (subst_lit L1 \'))) \ (mult trm_ord)" by auto show False proof (cases) assume "x = L2" from this and \((subst_lit L2 \'),(subst_lit L1 \')) \ lit_ord\ and \( (mset_lit (subst_lit x \')),(mset_lit (subst_lit L1 \'))) \ (mult trm_ord)\ show False unfolding lit_ord_def by auto next assume "x \ L2" from this and \x \ (cl_ecl D)\ have "x \ (cl_ecl D) - { L2 }" by auto from this and \(\x \ (cl_ecl D) - { L2 }. ( (subst_lit x \'),(subst_lit L2 \')) \ lit_ord)\ have "( (subst_lit x \'),(subst_lit L2 \')) \ lit_ord" by auto from \((mset_lit (subst_lit x \')),(mset_lit (subst_lit L1 \'))) \ (mult trm_ord)\ have "((subst_lit x \'),(subst_lit L1 \')) \ lit_ord" unfolding lit_ord_def by auto from this and \( (subst_lit x \'),(subst_lit L2 \')) \ lit_ord\ and \((subst_lit L2 \'),(subst_lit L1 \')) \ lit_ord\ show False using lit_ord_trans unfolding trans_def by blast qed qed from all_finite and \C \ S\ have "finite (cl_ecl C)" by auto from this and \L1 \ (cl_ecl C)\ have "(mset_lit (subst_lit L1 \')) \# mset_ecl (C,\')" using mset_ecl_and_mset_lit by auto from this have "(mset_lit (subst_lit L1 \')) \ (set_mset (mset_ecl (C,\')))" by simp have "\x. (x \ (set_mset (mset_ecl (D,\'))) \ (\y \ set_mset (mset_ecl (C,\')). (x,y) \ (mult trm_ord)))" proof ((rule allI),(rule impI)) fix x assume "x \ (set_mset (mset_ecl (D,\')))" then have "x \# mset_ecl (D,\')" by simp from \x \# mset_ecl (D,\')\ obtain x' where "x' \# (mset_set (cl_ecl D))" and "x = (mset_lit (subst_lit x' \'))" by auto from \x' \# (mset_set (cl_ecl D))\ have "x' \ (cl_ecl D)" using count_mset_set(3) by (fastforce simp: count_eq_zero_iff) from this and \\x \ (cl_ecl D). ( (mset_lit (subst_lit x \')),(mset_lit (subst_lit L1 \'))) \ (mult trm_ord)\ and \x = (mset_lit (subst_lit x' \'))\ have "(x,(mset_lit (subst_lit L1 \'))) \ (mult trm_ord)" by auto from \(mset_lit (subst_lit L1 \')) \ (set_mset (mset_ecl (C,\')))\ and \(x,(mset_lit (subst_lit L1 \'))) \ (mult trm_ord)\ show "(\y \ set_mset (mset_ecl (C,\')). (x,y) \ (mult trm_ord))" by auto qed from this have dom: "\I J K. J \ {#} \ (\k\set_mset K. \j\set_mset J. (k, j) \ (mult trm_ord)) \ (I + K, I + J) \ mult (mult trm_ord)" by (blast intro: one_step_implies_mult) from \(mset_lit (subst_lit L1 \')) \# mset_ecl (C,\')\ have "mset_ecl (C,\') \ {#}" by auto from \\x. (x \ (set_mset (mset_ecl (D,\'))) \ (\y \ set_mset (mset_ecl (C,\')). (x,y) \ (mult trm_ord)))\ and \mset_ecl (C,\') \ {#}\ have "({#} + mset_ecl (D, \'), {#} + mset_ecl (C, \')) \ mult (mult trm_ord)" using dom [of "(mset_ecl (C,\'))" "mset_ecl (D,\')" "{#}"] by auto from this have "(mset_ecl (D, \'), mset_ecl (C, \')) \ mult (mult trm_ord)" by auto from this have "( (D,\'), (C,\') ) \ ecl_ord" unfolding ecl_ord_def by auto from st obtain t' where rt: "(replace_subterm t p v t')" using replace_subterm_is_a_function by blast from st obtain R Cl_R nt_R L' Cl_C Cl_D where ntr: "nt_R = (dom_trms Cl_R (subst_set ((trms_ecl C) \ (trms_ecl D) \ { r. \ q. (q,p) \ (pos_ord C t) \ (subterm t q r) }) \'))" and r: "R = Ecl Cl_R nt_R" and clc: "Cl_C = (cl_ecl C)" and cld: "Cl_D = (cl_ecl D)" and clr: "Cl_R = (subst_cl ((Cl_C - { L1 }) \ ((Cl_D - { L2 }) \ { L' } )) \')" and l': "L' = mk_lit polarity (Eq t' s)" by auto from \orient_lit_inst L1 t s polarity \'\ have "vars_of t \ vars_of_lit L1" using orient_lit_inst_vars by auto from \L1 \ (cl_ecl C)\ have "vars_of_lit L1 \ vars_of_cl (cl_ecl C)" by auto from this and \vars_of t \ vars_of_lit L1\ have "vars_of t \vars_of_cl (cl_ecl C)" by auto from this and \coincide_on \ \' (vars_of_cl (cl_ecl C))\ have "coincide_on \ \' (vars_of t)" unfolding coincide_on_def by auto from this have "subst t \ = subst t \'" using coincide_on_term by auto from \(\x \ (cl_ecl D) - { L2 }. ( (subst_lit x \'),(subst_lit L2 \')) \ lit_ord)\ have "strictly_maximal_literal D L2 \'" unfolding strictly_maximal_literal_def by metis from ntr have "nt_R = get_trms Cl_R (dom_trms Cl_R (subst_set ((trms_ecl C) \ (trms_ecl D) \ { r. \ q. (q,p) \ (pos_ord C t) \ (subterm t q r) }) \')) Ground" unfolding get_trms_def by auto from this \(subst u \') \ (subst v \')\ \\ is_a_variable u'\ l1 l2 el1 el2 \variable_disjoint C D\ rte r o1 o2 t st rt l' clr ntr clr clc cld \R = Ecl Cl_R nt_R\ \( (subst_lit L2 \'),(subst_lit L1 \')) \ lit_ord\ \strictly_maximal_literal D L2 \' \ have "superposition C D R \' Ground ((Cl_C - { L1 }) \ ((Cl_D - { L2 }) \ { L' } ))" unfolding superposition_def by blast from l2 have "(subst_lit L2 \') \ (subst_cl (cl_ecl D) \')" by auto from this and \ground_clause (subst_cl (cl_ecl D) \')\ have "vars_of_lit (subst_lit L2 \') = {}" by auto from this and o2 have "vars_of (subst v \') = {}" unfolding orient_lit_inst_def using vars_of_lit.simps vars_of_eq.simps by force from \(coincide_on \ \' (vars_of_cl (cl_ecl C)))\ have "(subst_cl (cl_ecl C) \') = (subst_cl (cl_ecl C) \)" using coincide_on_cl [of \ \' "(cl_ecl C)"] by auto from this and \ground_clause (subst_cl (cl_ecl C) \)\ have "ground_clause (subst_cl (cl_ecl C) \')" using coincide_on_cl by auto from l1 have "(subst_lit L1 \') \ (subst_cl (cl_ecl C) \')" by auto from this and \ground_clause (subst_cl (cl_ecl C) \')\ have "vars_of_lit (subst_lit L1 \') = {}" by auto from this and o1 have "vars_of (subst t \') = {}" unfolding orient_lit_inst_def using vars_of_lit.simps vars_of_eq.simps by force from \vars_of_lit (subst_lit L1 \') = {}\ and o1 have "vars_of (subst s \') = {}" unfolding orient_lit_inst_def using vars_of_lit.simps vars_of_eq.simps by force from \vars_of (subst t \') = {}\ and \vars_of (subst v \') = {}\ and rt have "vars_of (subst t' \') = {}" using ground_replacement [of t p v t' \'] unfolding ground_term_def by blast from \vars_of (subst t' \') = {}\ and \vars_of (subst s \') = {}\ have "vars_of_eq (subst_equation (Eq t' s) \') = {}" by auto from l' have "L' = (Pos (Eq t' s)) \ L' = (Neg (Eq t' s))" using mk_lit.elims by auto from this and \vars_of_eq (subst_equation (Eq t' s) \') = {}\ have "vars_of_lit (subst_lit L' \') = {}" by auto from \C \ S\ and \D \ S\ and \superposition C D R \' Ground ((Cl_C - { L1 }) \ ((Cl_D - { L2 }) \ { L' } ))\ have "derivable R { C,D } S \' Ground ((Cl_C - { L1 }) \ ((Cl_D - { L2 }) \ { L' } ))" unfolding derivable_def by auto have "ground_clause Cl_R" proof (rule ccontr) assume "\ground_clause Cl_R" then have "vars_of_cl Cl_R \ {}" by auto then obtain M where "M \ Cl_R" and "vars_of_lit M \ {}" by auto from \M \ Cl_R\ and clr obtain M' where "M = (subst_lit M' \')" and "M' \ ((Cl_C - { L1 }) \ ((Cl_D - { L2 }) \ { L' } )) " by auto show False proof (cases) assume "M' = L'" from this and \vars_of_lit (subst_lit L' \') = {}\ and \vars_of_lit M \ {}\ and \M = (subst_lit M' \')\ show False by auto next assume "M' \ L'" from this and l1 clc cld and \M' \(Cl_C - { L1 }) \ ((Cl_D - { L2 }) \ { L' } )\ have "M' \ (cl_ecl C) \ M' \ (cl_ecl D)" by auto then show False proof assume "M' \ (cl_ecl C)" from this and \ground_clause (subst_cl (cl_ecl C) \')\ have "vars_of_lit (subst_lit M' \') = {}" by auto from this and \M = (subst_lit M' \')\ and \vars_of_lit M \ {}\ show False by auto next assume "M' \ (cl_ecl D)" from this and \ground_clause (subst_cl (cl_ecl D) \')\ have "vars_of_lit (subst_lit M' \') = {}" by auto from this and \M = (subst_lit M' \')\ and \vars_of_lit M \ {}\ show False by auto qed qed qed from \ground_clause (subst_cl (cl_ecl C) \')\ and \ground_clause (subst_cl (cl_ecl D) \')\ have "grounding_set { C,D } \'" unfolding grounding_set_def by auto from \ground_clause Cl_R\ and \R = Ecl Cl_R nt_R\ have "ground_clause (cl_ecl R)" by auto from this and \derivable R { C,D } S \' Ground ((Cl_C - { L1 }) \ ((Cl_D - { L2 }) \ { L' } ))\ and \ground_inference_saturated S\ \grounding_set { C,D } \'\ have "(redundant_inference R S { C,D } \')" unfolding ground_inference_saturated_def by blast from this obtain S' where "S' \ (instances S)" and "(set_entails_clause (clset_instances S') (cl_ecl R))" and order: "(\x \ S'. (((fst x),(snd x)),(C,\')) \ ecl_ord \ (((fst x),(snd x)),(D,\')) \ ecl_ord)" and all_normalized_term_included: "(\x \ S'. (subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl R)))" unfolding redundant_inference_def by auto have all_smaller: "(\x \ S'. (((fst x),(snd x)),(C,\)) \ ecl_ord)" proof (rule ccontr) assume "\(\x \ S'. (((fst x),(snd x)),(C,\)) \ ecl_ord)" then obtain x where "x \ S'" and "(((fst x),(snd x)),(C,\)) \ ecl_ord" by auto from \x \ S'\ and order have "(((fst x),(snd x)),(C,\')) \ ecl_ord \ (((fst x),(snd x)),(D,\')) \ ecl_ord" by auto then have "(((fst x),(snd x)),(C,\')) \ ecl_ord" proof assume "(((fst x),(snd x)),(C,\')) \ ecl_ord" from this show ?thesis by auto next assume "(((fst x),(snd x)),(D,\')) \ ecl_ord" from this and \( (D,\'),(C,\')) \ ecl_ord\ show "(((fst x),(snd x)),(C,\')) \ ecl_ord" using ecl_ord_trans unfolding trans_def by metis qed from this have "((mset_ecl x), (mset_ecl (C,\'))) \ (mult (mult trm_ord))" unfolding ecl_ord_def by auto from \(coincide_on \ \' (vars_of_cl (cl_ecl C)))\ have "(mset_ecl (C,\')) = (mset_ecl (C,\))" using ecl_ord_coincide [of \ \' C] by auto from this and \((mset_ecl x), (mset_ecl (C,\'))) \ (mult (mult trm_ord))\ have "((mset_ecl x), (mset_ecl (C,\))) \ (mult (mult trm_ord))" by simp from this and \\(((fst x),(snd x)),(C,\)) \ ecl_ord\ show False unfolding ecl_ord_def by auto qed have "validate_clause_set ?I (clset_instances S')" proof (rule ccontr) assume "\ validate_clause_set ?I (clset_instances S')" then obtain x where "x \(clset_instances S')" and "\validate_clause ?I x" using validate_clause_set.simps by blast from \x \(clset_instances S')\ obtain pair' where "pair' \ S'" and "x = (subst_cl (cl_ecl (fst pair')) (snd pair'))" unfolding clset_instances_def by auto from all_smaller and \pair' \ S'\ have "(pair',(C,\)) \ ecl_ord" by auto from this and \C = fst pair\ and \\ = snd pair\ have "(pair',pair) \ ecl_ord" by auto from this and hyp_ind have "?P pair'" by blast from \pair' \ S'\ and all_normalized_term_included have "(subterms_inclusion (subst_set (trms_ecl (fst pair')) (snd pair')) (trms_ecl R))" by auto have i: "(all_trms_irreducible (subst_set (trms_ecl (fst pair')) (snd pair')) (\t. (trm_rep t S)))" proof (rule ccontr) assume "\(all_trms_irreducible (subst_set (trms_ecl (fst pair')) (snd pair')) (\t. (trm_rep t S)))" then obtain w w' where "w \ subst_set (trms_ecl (fst pair')) (snd pair')" and "occurs_in w' w" "trm_rep w' S \ w'" unfolding all_trms_irreducible_def by blast from \w \ subst_set (trms_ecl (fst pair')) (snd pair')\ and \(subterms_inclusion (subst_set (trms_ecl (fst pair')) (snd pair')) (trms_ecl R))\ obtain w'' where "w'' \ trms_ecl R" and "occurs_in w w''" unfolding subterms_inclusion_def by auto from \occurs_in w w''\ and \occurs_in w' w\ have "occurs_in w' w''" using occur_in_subterm by blast from ntr have "nt_R \ (subst_set ((trms_ecl C) \ (trms_ecl D) \ { r. \ q. (q,p) \ (pos_ord C t) \ (subterm t q r) }) \')" using dom_trms_subset [of Cl_R "(subst_set ((trms_ecl C) \ (trms_ecl D) \ { r. \ q. (q,p) \ (pos_ord C t) \ (subterm t q r) }) \')"] by blast from this and r have "trms_ecl R \ (subst_set ((trms_ecl C) \ (trms_ecl D) \ { r. \ q. (q,p) \ (pos_ord C t) \ (subterm t q r) }) \')" by auto from this and \w'' \ trms_ecl R\ have "w'' \ (subst_set ((trms_ecl C) \ (trms_ecl D) \ { r. \ q. (q,p) \ (pos_ord C t) \ (subterm t q r) }) \')" by blast from this obtain w''' where "w''' \ ((trms_ecl C) \ (trms_ecl D) \ { r. \ q. (q,p) \ (pos_ord C t) \ (subterm t q r) })" and "w'' = subst w''' \'" by auto from this and \occurs_in w' w''\ have "occurs_in w' (subst w''' \')" by auto have "\ (w''' \ trms_ecl C)" proof assume "w''' \ trms_ecl C" from this and \occurs_in w' w''\ and \w'' = subst w''' \'\ have "occurs_in w' (subst w''' \')" by auto from assms(3) and \C \ S\ and \w''' \ trms_ecl C\ have "vars_of w''' \ vars_of_cl (cl_ecl C)" using dom_trm_vars unfolding Ball_def well_constrained_def by blast from this and \coincide_on \ \' (vars_of_cl (cl_ecl C))\ have "coincide_on \ \' (vars_of w''')" unfolding coincide_on_def by auto from this have "subst w''' \ = subst w''' \'" using coincide_on_term by auto from this and \occurs_in w' (subst w''' \')\ have "occurs_in w' (subst w''' \)" by auto from this and \w''' \ trms_ecl C\ \(all_trms_irreducible (subst_set (trms_ecl C) \) (\t. (trm_rep t S)))\ have "trm_rep w' S = w'" unfolding all_trms_irreducible_def using subst_set.simps by blast from this and \trm_rep w' S \ w'\ show False by blast qed have "\ (w''' \ trms_ecl D)" proof assume "w''' \ trms_ecl D" from this and \occurs_in w' w''\ and \w'' = subst w''' \'\ have "occurs_in w' (subst w''' \')" by auto from this and \w''' \ trms_ecl D\ \(all_trms_irreducible (subst_set (trms_ecl D) \') (\t. (trm_rep t S)))\ have "trm_rep w' S = w'" unfolding all_trms_irreducible_def using subst_set.simps by blast from this and \trm_rep w' S \ w'\ show False by blast qed from this and \w''' \ ((trms_ecl C) \ (trms_ecl D) \ { r. \ q. (q,p) \ (pos_ord C t) \ (subterm t q r) })\ and \\ (w''' \ trms_ecl C)\ obtain q_w where "(subterm t q_w w''')" and "(q_w,p) \ (pos_ord C t)" by auto from \subterm t q_w w'''\ have "subterm (subst t \') q_w (subst w''' \')" using substs_preserve_subterms by auto from \occurs_in w' (subst w''' \')\ obtain q where "(subterm (subst w''' \') q w')" unfolding occurs_in_def by blast from this and \subterm (subst t \') q_w (subst w''' \')\ have "subterm (subst t \') (append q_w q) w'" using subterm_of_a_subterm_is_a_subterm by blast from this and \(subst t \) = (subst t \')\ have "subterm (subst t \) (append q_w q) w'" by auto from \(q_w,p) \ (pos_ord C t)\ have "((append q_w q),p) \ (pos_ord C t)" using pos_ord_prefix by blast from this and \minimal_redex p (subst t \) C S t\ and \subterm (subst t \) (append q_w q) w'\ have "trm_rep w' S = w'" unfolding minimal_redex_def by blast from this and \trm_rep w' S \ w'\ show False by blast qed from \S' \ (instances S)\ and \pair' \ S'\ have ii: "ground_clause (subst_cl (cl_ecl (fst pair')) (snd pair'))" unfolding instances_def [of S] by fastforce from \S' \ (instances S)\ and \pair' \ S'\ have iii: "(fst pair') \ S" unfolding instances_def [of S] by fastforce from \?P pair'\ and i ii iii have "validate_ground_clause ?I (subst_cl (cl_ecl (fst pair')) (snd pair'))" unfolding int_clset_def by blast from this and \x = (subst_cl (cl_ecl (fst pair')) (snd pair'))\ and \\validate_clause ?I x\ show False by (metis ii substs_preserve_ground_clause validate_clause.simps) qed from this and \fo_interpretation ?I\ and \(set_entails_clause (clset_instances S') (cl_ecl R))\ have "validate_clause ?I (cl_ecl R)" unfolding set_entails_clause_def by blast from this have "validate_ground_clause ?I (cl_ecl R)" by (metis \R = Ecl Cl_R nt_R\ \ground_clause Cl_R\ cl_ecl.simps substs_preserve_ground_clause validate_clause.simps) from this obtain L'' where "L'' \ (cl_ecl R)" and "validate_ground_lit ?I L''" using validate_ground_clause.simps by blast from \L'' \ (cl_ecl R)\ and \R = Ecl Cl_R nt_R\ and \Cl_R = (subst_cl ((Cl_C - { L1 }) \ ((Cl_D - { L2 }) \ { L' } )) \')\ obtain M where m: "M \ ((Cl_C - { L1 }) \ ((Cl_D - { L2 }) \ { L' } ))" and "L'' = subst_lit M \'" by auto have "M \ cl_ecl C" proof assume "M \ cl_ecl C" from this have "vars_of_lit M \ vars_of_cl (cl_ecl C)" by auto from this and \coincide_on \ \' (vars_of_cl (cl_ecl C))\ have "coincide_on \ \' (vars_of_lit M)" unfolding coincide_on_def by auto from this have "subst_lit M \ = subst_lit M \'" using coincide_on_lit by auto from this and \L'' = subst_lit M \'\have "L'' = subst_lit M \" by auto from \M \ cl_ecl C\ and \L'' = subst_lit M \\and \validate_ground_lit ?I L''\ have "validate_ground_clause ?I (subst_cl (cl_ecl C) \)" by (metis (mono_tags, lifting) subst_cl.simps mem_Collect_eq validate_ground_clause.simps) from this and cm show False by blast qed have "M \ Cl_D - { L2 }" proof assume "M \ Cl_D - { L2 }" from this and cld and \L'' = subst_lit M \'\ and \validate_ground_lit ?I L''\ have "validate_ground_clause ?I (subst_cl ( (cl_ecl D) - { L2 } ) \')" by (metis (mono_tags, lifting) subst_cl.simps mem_Collect_eq validate_ground_clause.simps) from this and \\validate_ground_clause ?I (subst_cl ( (cl_ecl D) - { L2 } ) \')\ show False by blast qed have "M \ L'" proof assume "M = L'" from \?I (subst u \') (subst v \')\ and e have "?I (subst u' \') (subst v \')" by metis from rt and st \fo_interpretation ?I\ \?I (subst u' \') (subst v \')\ have "?I (subst t \') (subst t' \')" using replacement_preserves_congruences [of ?I u' \' v t p t'] unfolding fo_interpretation_def by metis from l1 and cm have "\ (validate_ground_lit ?I (subst_lit L1 \'))" using \subst_cl (cl_ecl C) \' = subst_cl (cl_ecl C) \\ \subst_lit L1 \' \ subst_cl (cl_ecl C) \'\ validate_ground_clause.simps by blast from this and \?I (subst t \') (subst t' \')\ and \fo_interpretation ?I\ and l' \orient_lit_inst L1 t s polarity \'\ have "\validate_ground_lit ?I (subst_lit L' \')" using trm_rep_preserves_lit_semantics [of ?I t \' t' L1 s polarity \'] by metis from this and \M = L'\ and \validate_ground_lit ?I L''\ and \L'' = subst_lit M \'\ show False by blast qed from this and \M \ Cl_D - { L2 }\ \M \ (cl_ecl C)\ and m clc show False by auto qed text \Second, we show that the clause contains no contradictory literal (otherwise the reflexion rule applies).\ let ?no_cont = "\L t s. (L \ (cl_ecl C)) \ (eligible_literal L C \) \ (orient_lit_inst L t s neg \) \ (trm_rep (subst t \) S) = (subst t \) \ (subst t \) \ (subst s \)" have ?no_cont proof (rule ccontr) assume "\ ?no_cont" then obtain L t s where l: "L \ (cl_ecl C)" and e: "(eligible_literal L C \)" and o: "orient_lit_inst L t s neg \" and "(trm_rep (subst t \) S) = (subst t \)" and "(subst t \) = (subst s \)" by blast from \(subst t \) = (subst s \)\ have t: "ck_unifier t s \ Ground" unfolding ck_unifier_def Unifier_def by auto from l and e and o and t obtain R Cl_R nt_R where "R = Ecl Cl_R nt_R" and "Cl_R = (subst_cl ((cl_ecl C) - { L }) \)" and "reflexion C R \ Ground ((cl_ecl C) - { L })" and "trms_ecl R = (dom_trms (cl_ecl R) (subst_set ((trms_ecl C) \ { t }) \))" unfolding reflexion_def get_trms_def by fastforce from \C \ S\ and \reflexion C R \ Ground ((cl_ecl C) - { L })\ have "derivable R { C } S \ Ground ((cl_ecl C) - { L })" unfolding derivable_def by auto from \ground_clause (subst_cl (cl_ecl C) \)\ and \Cl_R = (subst_cl ((cl_ecl C) - { L }) \)\ have "ground_clause Cl_R" using ground_clause.simps by auto from this and \R = Ecl Cl_R nt_R\ have "ground_clause (cl_ecl R)" by auto from \ground_clause (subst_cl (cl_ecl C) \)\ have "grounding_set { C } \" unfolding grounding_set_def by auto from this and \ground_clause (cl_ecl R)\ \derivable R { C } S \ Ground ((cl_ecl C) - { L })\ and \ground_inference_saturated S\ have "(redundant_inference R S { C } \)" unfolding ground_inference_saturated_def by blast from this obtain S' where "S' \ (instances S)" and "(set_entails_clause (clset_instances S') (cl_ecl R))" and all_smaller: "(\x \ S'. (((fst x),(snd x)),(C,\)) \ ecl_ord)" and all_normalized_term_included: "(\x \ S'. (subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl R)))" unfolding redundant_inference_def by auto have "validate_clause_set ?I (clset_instances S')" proof (rule ccontr) assume "\ validate_clause_set ?I (clset_instances S')" then obtain x where "x \(clset_instances S')" and "\validate_clause ?I x" using validate_clause_set.simps by blast from \x \(clset_instances S')\ obtain pair' where "pair' \ S'" and "x = (subst_cl (cl_ecl (fst pair')) (snd pair'))" unfolding clset_instances_def by auto from all_smaller and \pair' \ S'\ have "(pair',(C,\)) \ ecl_ord" by auto from this and \C = fst pair\ and \\ = snd pair\ have "(pair',pair) \ ecl_ord" by auto from this and hyp_ind have "?P pair'" by blast from \trms_ecl R = (dom_trms (cl_ecl R) (subst_set ((trms_ecl C) \ { t }) \))\ have "trms_ecl R \ (subst_set ((trms_ecl C) \ { t }) \)" using dom_trms_subset by metis from \pair' \ S'\ and all_normalized_term_included have "(subterms_inclusion (subst_set (trms_ecl (fst pair')) (snd pair')) (trms_ecl R))" by auto have i: "(all_trms_irreducible (subst_set (trms_ecl (fst pair')) (snd pair')) (\t. (trm_rep t S)))" proof (rule ccontr) assume "\(all_trms_irreducible (subst_set (trms_ecl (fst pair')) (snd pair')) (\t. (trm_rep t S)))" then obtain t' t'' where "t' \ (subst_set (trms_ecl (fst pair')) (snd pair'))" "occurs_in t'' t'" and "trm_rep t'' S \ t''" unfolding all_trms_irreducible_def by blast from \t' \ (subst_set (trms_ecl (fst pair')) (snd pair'))\ and \(subterms_inclusion (subst_set (trms_ecl (fst pair')) (snd pair')) (trms_ecl R))\ obtain s' where "s' \ (trms_ecl R)" and "occurs_in t' s'" unfolding subterms_inclusion_def by auto from \s' \ (trms_ecl R)\ and \trms_ecl R \ (subst_set ((trms_ecl C) \ { t }) \)\ obtain s'' where "s' = subst s'' \" and "s'' \ ((trms_ecl C) \ { t })" by auto from \s'' \ ((trms_ecl C) \ { t })\ have "s'' \ trms_ecl C \ s'' = t" by auto thus False proof assume "s'' \ trms_ecl C" from this and \s' = subst s'' \\ have "s' \ (subst_set (trms_ecl C) \)" by auto from this and \(all_trms_irreducible (subst_set (trms_ecl C) \) (\t. (trm_rep t S)))\ and \occurs_in t' s'\ have "trm_rep t' S = t'" unfolding all_trms_irreducible_def by blast from this and \occurs_in t'' t'\ and \trm_rep t'' S \ t''\show False using occurs_in_def subts_of_irred_trms_are_irred by blast next assume "s'' = t" from this and \s' = subst s'' \\ have "s' = subst t \" by auto from this and \(trm_rep (subst t \) S) = (subst t \)\ have "trm_rep s' S = s'" by blast from \trm_rep s' S = s'\ \trm_rep t'' S \ t''\ \occurs_in t' s'\ \occurs_in t'' t'\ show False using occurs_in_def subts_of_irred_trms_are_irred by blast qed qed from \S' \ (instances S)\ and \pair' \ S'\ have ii: "ground_clause (subst_cl (cl_ecl (fst pair')) (snd pair'))" unfolding instances_def [of S] by fastforce from \S' \ (instances S)\ and \pair' \ S'\ have iii: "(fst pair') \ S" unfolding instances_def [of S] by fastforce from \?P pair'\ and i ii iii have "validate_ground_clause ?I (subst_cl (cl_ecl (fst pair')) (snd pair'))" unfolding int_clset_def by blast from this and \x = (subst_cl (cl_ecl (fst pair')) (snd pair'))\ and \\validate_clause ?I x\ show False by (metis ii substs_preserve_ground_clause validate_clause.simps) qed from this and \fo_interpretation ?I\ and \(set_entails_clause (clset_instances S') (cl_ecl R))\ have "validate_clause ?I (cl_ecl R)" unfolding set_entails_clause_def by blast from this have "validate_ground_clause ?I (cl_ecl R)" by (metis \R = Ecl Cl_R nt_R\ \ground_clause Cl_R\ cl_ecl.simps substs_preserve_ground_clause validate_clause.simps) from this obtain L' where "L' \ (cl_ecl R)" and "validate_ground_lit ?I L'" using validate_ground_clause.simps by blast from \L' \ (cl_ecl R)\ and \R = Ecl Cl_R nt_R\ and \Cl_R = (subst_cl ((cl_ecl C) - { L }) \)\ obtain L'' where "L'' \ cl_ecl C" and "L' = subst_lit L'' \" by auto from \L'' \ cl_ecl C\ and \L' = subst_lit L'' \\and \validate_ground_lit ?I L'\ have "validate_ground_clause ?I (subst_cl (cl_ecl C) \)" by (metis (mono_tags, lifting) subst_cl.simps mem_Collect_eq validate_ground_clause.simps) from this and cm show False unfolding int_clset_def by blast qed text \Third, we prove that the clause contains no pair of equations with the same left-hand side and equivalent right-hand sides (otherwise the factorization rule applies and a smaller false clause is derived).\ let ?no_fact = "\L1 L2 t s u v. (L1 \ (cl_ecl C)) \ (eligible_literal L1 C \) \ (L2 \ (cl_ecl C) - { L1 }) \ (orient_lit_inst L1 t s pos \) \ (orient_lit_inst L2 u v pos \) \ (subst t \) = (subst u \) \ (\ (proper_subterm_red t S \)) \ (trm_rep ((subst s) \) S) \ (trm_rep ((subst v) \) S)" have ?no_fact proof (rule ccontr) assume "\ ?no_fact" then obtain L1 L2 t s u v where l1: "L1 \ (cl_ecl C)" and l2: "L2 \ (cl_ecl C) - { L1 }" and e1: "(eligible_literal L1 C \)" and o1: "(orient_lit_inst L1 t s pos \)" and o2: "(orient_lit_inst L2 u v pos \)" and e: "(subst t \) = (subst u \)" and "(\ (proper_subterm_red t S \))" and i: "(trm_rep ((subst s) \) S) = (trm_rep ((subst v) \) S)" by blast from e have t: "ck_unifier t u \ Ground" unfolding ck_unifier_def Unifier_def using "inferences.distinct" by metis from \L1 \ (cl_ecl C)\ o1 \\(validate_ground_clause (int_clset S) (subst_cl (cl_ecl C) \))\ have "trm_rep (subst t \) S \ trm_rep (subst s \) S" using no_valid_literal by metis then have "subst t \ \ subst s \" by metis from \L2 \ (cl_ecl C) - { L1 }\ have "L2 \ (cl_ecl C)" by auto from this o2 \\(validate_ground_clause (int_clset S) (subst_cl (cl_ecl C) \))\ have "trm_rep (subst u \) S \ trm_rep (subst v \) S" using no_valid_literal by metis from this and e have "subst t \ \ subst v \" by metis obtain R Cl_R nt_R L' where ntr: "nt_R = (dom_trms Cl_R (subst_set ((trms_ecl C) \ (proper_subterms_of t)) \))" and r: "R = Ecl Cl_R nt_R" and clr: "Cl_R = (subst_cl ( ((cl_ecl C) - { L2 }) \ { L' } ) \)" and l': "L' = Neg (Eq s v)" by auto from ntr r l' clr l1 l2 o1 o2 e1 t \subst t \ \ subst s \\ \subst t \ \ subst v \\ have "factorization C R \ Ground (((cl_ecl C) - { L2 }) \ { L' } )" unfolding factorization_def get_trms_def using inferences.distinct by (metis cl_ecl.simps) from l2 have "(subst_lit L2 \) \ (subst_cl (cl_ecl C) \)" by auto from this and \ground_clause (subst_cl (cl_ecl C) \)\ have "vars_of_lit (subst_lit L2 \) = {}" by auto from this and o2 have "vars_of (subst v \) = {}" unfolding orient_lit_inst_def using vars_of_lit.simps vars_of_eq.simps by force from l1 have "(subst_lit L1 \) \ (subst_cl (cl_ecl C) \)" by auto from this and \ground_clause (subst_cl (cl_ecl C) \)\ have "vars_of_lit (subst_lit L1 \) = {}" by auto from this and o1 have "vars_of (subst s \) = {}" unfolding orient_lit_inst_def using vars_of_lit.simps vars_of_eq.simps by force from \vars_of (subst v \) = {}\ and \vars_of (subst s \) = {}\ and l' have "vars_of_lit (subst_lit L' \) = {}" by auto from \C \ S\ and \factorization C R \ Ground (((cl_ecl C) - { L2 }) \ { L' } )\ have "derivable R { C } S \ Ground (((cl_ecl C) - { L2 }) \ { L' } )" unfolding derivable_def by auto have "ground_clause Cl_R" proof (rule ccontr) assume "\ground_clause Cl_R" then have "vars_of_cl Cl_R \ {}" by auto then obtain M where "M \ Cl_R" and "vars_of_lit M \ {}" by auto from \M \ Cl_R\ and clr obtain M' where "M = (subst_lit M' \)" and "M' \((cl_ecl C) - { L2, L1 }) \ { L', L1 } " by auto show False proof (cases) assume "M' = L'" from this and \vars_of_lit (subst_lit L' \) = {}\ and \vars_of_lit M \ {}\ and \M = (subst_lit M' \)\ show False by auto next assume "M' \ L'" from this and l1 and \M' \((cl_ecl C) - { L2, L1 }) \ { L', L1 }\ have "M' \ (cl_ecl C)" by auto from this and \ground_clause (subst_cl (cl_ecl C) \)\ have "vars_of_lit (subst_lit M' \) = {}" by auto from this and \M = (subst_lit M' \)\ and \vars_of_lit M \ {}\ show False by auto qed qed from \ground_clause Cl_R\ and \R = Ecl Cl_R nt_R\ have "ground_clause (cl_ecl R)" by auto from \ground_clause (subst_cl (cl_ecl C) \)\ have "grounding_set { C } \" unfolding grounding_set_def by auto from this \ground_clause (cl_ecl R)\ and \derivable R { C } S \ Ground (((cl_ecl C) - { L2 }) \ { L' } )\ and \ground_inference_saturated S\ have "(redundant_inference R S { C } \)" unfolding ground_inference_saturated_def by blast from this obtain S' where "S' \ (instances S)" and "(set_entails_clause (clset_instances S') (cl_ecl R))" and all_smaller: "(\x \ S'. (((fst x),(snd x)),(C,\)) \ ecl_ord)" and all_normalized_term_included: "(\x \ S'. (subterms_inclusion (subst_set (trms_ecl (fst x)) (snd x)) (trms_ecl R)))" unfolding redundant_inference_def by auto have "validate_clause_set ?I (clset_instances S')" proof (rule ccontr) assume "\ validate_clause_set ?I (clset_instances S')" then obtain x where "x \(clset_instances S')" and "\validate_clause ?I x" using validate_clause_set.simps by blast from \x \(clset_instances S')\ obtain pair' where "pair' \ S'" and "x = (subst_cl (cl_ecl (fst pair')) (snd pair'))" unfolding clset_instances_def by auto from all_smaller and \pair' \ S'\ have "(pair',(C,\)) \ ecl_ord" by auto from this and \C = fst pair\ and \\ = snd pair\ have "(pair',pair) \ ecl_ord" by auto from this and hyp_ind have "?P pair'" by blast from r ntr have "trms_ecl R = (dom_trms (cl_ecl R) (subst_set ((trms_ecl C) \ (proper_subterms_of t)) \))" by auto from this have "trms_ecl R \ (subst_set ((trms_ecl C) \ (proper_subterms_of t)) \)" using dom_trms_subset by metis from \pair' \ S'\ and all_normalized_term_included have "(subterms_inclusion (subst_set (trms_ecl (fst pair')) (snd pair')) (trms_ecl R))" by auto have i: "(all_trms_irreducible (subst_set (trms_ecl (fst pair')) (snd pair')) (\t. (trm_rep t S)))" proof (rule ccontr) assume "\(all_trms_irreducible (subst_set (trms_ecl (fst pair')) (snd pair')) (\t. (trm_rep t S)))" then obtain t' t'' where "t' \ (subst_set (trms_ecl (fst pair')) (snd pair'))" "occurs_in t'' t'" and "trm_rep t'' S \ t''" unfolding all_trms_irreducible_def by blast from \t' \ (subst_set (trms_ecl (fst pair')) (snd pair'))\ and \(subterms_inclusion (subst_set (trms_ecl (fst pair')) (snd pair')) (trms_ecl R))\ obtain s' where "s' \ (trms_ecl R)" and "occurs_in t' s'" unfolding subterms_inclusion_def by auto from \s' \ (trms_ecl R)\ and \trms_ecl R \ (subst_set ((trms_ecl C) \ (proper_subterms_of t)) \)\ have "s' \(subst_set ((trms_ecl C) \ (proper_subterms_of t)) \)" by auto then obtain s'' where "s' = subst s'' \" and "s'' \ ((trms_ecl C) \ (proper_subterms_of t))" by auto from \s'' \ ((trms_ecl C) \ (proper_subterms_of t))\ have "s'' \ trms_ecl C \ s'' \ (proper_subterms_of t)" by auto thus False proof assume "s'' \ trms_ecl C" from this and \s' = subst s'' \\ have "s' \ (subst_set (trms_ecl C) \)" by auto from this and \(all_trms_irreducible (subst_set (trms_ecl C) \) (\t. (trm_rep t S)))\ and \occurs_in t' s'\ have "trm_rep t' S = t'" unfolding all_trms_irreducible_def by blast from this and \occurs_in t'' t'\ and \trm_rep t'' S \ t''\show False using occurs_in_def subts_of_irred_trms_are_irred by blast next assume " s'' \ (proper_subterms_of t)" from \occurs_in t' s'\ \occurs_in t'' t'\ \s' = s'' \ \\ \trm_rep t'' S \ t''\ have "trm_rep (subst s'' \) S \ (subst s'' \)" using occurs_in_def subts_of_irred_trms_are_irred by blast from this and \s'' \ (proper_subterms_of t)\ and \\ (proper_subterm_red t S \)\ show False using proper_subterm_red_def "proper_subterms_of.simps" by blast qed qed from \S' \ (instances S)\ and \pair' \ S'\ have ii: "ground_clause (subst_cl (cl_ecl (fst pair')) (snd pair'))" unfolding instances_def [of S] by fastforce from \S' \ (instances S)\ and \pair' \ S'\ have iii: "(fst pair') \ S" unfolding instances_def [of S] by fastforce from \?P pair'\ and i ii iii have "validate_ground_clause ?I (subst_cl (cl_ecl (fst pair')) (snd pair'))" unfolding int_clset_def by blast from this and \x = (subst_cl (cl_ecl (fst pair')) (snd pair'))\ and \\validate_clause ?I x\ show False by (metis ii substs_preserve_ground_clause validate_clause.simps) qed from this and \fo_interpretation ?I\ and \(set_entails_clause (clset_instances S') (cl_ecl R))\ have "validate_clause ?I (cl_ecl R)" unfolding set_entails_clause_def by blast from this have "validate_ground_clause ?I (cl_ecl R)" by (metis \R = Ecl Cl_R nt_R\ \ground_clause Cl_R\ cl_ecl.simps substs_preserve_ground_clause validate_clause.simps) from this obtain L'' where "L'' \ (cl_ecl R)" and "validate_ground_lit ?I L''" using validate_ground_clause.simps by blast from \L'' \ (cl_ecl R)\ and \R = Ecl Cl_R nt_R\ and \Cl_R = (subst_cl ( ((cl_ecl C) - { L2 }) \ { L' } ) \)\ obtain M where m: "M \ ( ((cl_ecl C) - { L2, L1 }) \ { L', L1 } )" and "L'' = subst_lit M \" by auto have "M \ cl_ecl C" proof (rule ccontr) assume "M \ cl_ecl C" from this and m and l1 have "M = L'" by auto from this and \L'' = subst_lit M \\ and \L' = (Neg (Eq s v))\ have "L'' = (Neg (Eq (subst s \) (subst v \)))" by auto from this and \validate_ground_lit ?I L''\ have "\(?I (subst s \) (subst v \))" using validate_ground_lit.simps(2) validate_ground_eq.simps by metis from this and i show False unfolding same_values_def int_clset_def by blast qed from \M \ cl_ecl C\ and \L'' = subst_lit M \\and \validate_ground_lit ?I L''\ have "validate_ground_clause ?I (subst_cl (cl_ecl C) \)" by (metis (mono_tags, lifting) subst_cl.simps mem_Collect_eq validate_ground_clause.simps) from this and cm show False by blast qed text \Now, it remains to prove that the considered clause yields a rule which can be used to reduce the left-hand side of the maximal equation, which (since no reduction is possible) entails that the left-hand side must be equivalent to the right-hand side (thus contradicting the fact that the clause is false).\ have "(finite (cl_ecl C))" by (simp add: \C \ S\ all_finite) have "(cl_ecl C) \ {}" by (simp add: \C \ S\ all_non_empty) from \finite (cl_ecl C)\ \(cl_ecl C) \ {}\ \ground_clause (subst_cl (cl_ecl C) \)\ obtain L where "L \ (cl_ecl C)" "eligible_literal L C \" using eligible_lit_exists by metis obtain t s p where "orient_lit_inst L t s p \" using literal.exhaust equation.exhaust using trm_ord_irrefl trm_ord_trans unfolding orient_lit_inst_def irrefl_def trans_def by metis text \We first show that the terms occurring inside variables are irreducible. To this aim, we need to consider the normal form of the substitution @{term "\"}, obtained by replacing the image of each variable by its normal form.\ have "\x y. ((x \ vars_of_cl (cl_ecl C)) \ occurs_in y (subst (Var x) \) \ trm_rep y S = y)" proof (rule ccontr) assume "\(\x y. (x \ vars_of_cl (cl_ecl C)) \ occurs_in y (subst (Var x) \) \ trm_rep y S = y)" then obtain x y where "(x \ vars_of_cl (cl_ecl C))" and "occurs_in y (subst (Var x) \)" and "trm_rep y S \ y" by blast from \occurs_in y (subst (Var x) \)\ obtain p where "subterm (subst (Var x) \) p y" unfolding occurs_in_def by auto from \subterm (subst (Var x) \) p y\ and \trm_rep y S \ y\ have "trm_rep (subst (Var x) \) S \ (subst (Var x) \)" using subts_of_irred_trms_are_irred by blast let ?\ = "map_subst (\x. (trm_rep x S)) \" have "equivalent_on \ ?\ (vars_of_cl (cl_ecl C)) ?I" proof (rule ccontr) assume "\equivalent_on \ ?\ (vars_of_cl (cl_ecl C)) ?I" then obtain z where "z \ vars_of_cl (cl_ecl C)" and "\ (?I (subst (Var z) \) (subst (Var z) ?\))" unfolding equivalent_on_def by blast from \\ (?I (subst (Var z) \) (subst (Var z) ?\))\ have "trm_rep (subst (Var z) \) S \ trm_rep (subst (Var z) ?\) S" unfolding same_values_def int_clset_def by blast from this have "trm_rep (trm_rep (subst (Var z) \) S) S \ trm_rep (subst (Var z) ?\) S" using trm_rep_involutive by metis from this have "(subst (Var z) \) = (subst (Var z) ?\)" using map_subst_lemma [of z \ "\x. (trm_rep x S)"] by metis from this and \\ (?I (subst (Var z) \) (subst (Var z) ?\))\ show False using \fo_interpretation ?I\ unfolding fo_interpretation_def congruence_def equivalence_relation_def reflexive_def by metis qed from this and \\ validate_ground_clause ?I (subst_cl (cl_ecl C) \)\ \fo_interpretation ?I\ have "\ validate_ground_clause ?I (subst_cl (cl_ecl C) ?\)" using equivalent_on_cl by metis have "lower_on ?\ \ (vars_of_cl (cl_ecl C))" proof (rule ccontr) assume "\lower_on ?\ \ (vars_of_cl (cl_ecl C))" then obtain z where "z \ vars_of_cl (cl_ecl C)" and "(subst (Var z) \) \ (subst (Var z) ?\)" and "((subst (Var z) ?\),(subst (Var z) \)) \ trm_ord" unfolding lower_on_def lower_or_eq_def by metis from \(subst (Var z) \) \ (subst (Var z) ?\)\ have "(trm_rep (subst (Var z) \) S) = (subst (Var z) ?\)" using map_subst_lemma [of z \ "\x. (trm_rep x S)"] by metis from this and \(subst (Var z) \) \ (subst (Var z) ?\)\ and \((subst (Var z) ?\),(subst (Var z) \)) \ trm_ord\ show False using trm_rep_is_lower by metis qed have "subst (Var x) \ \ (Var x)" proof assume "subst (Var x) \ = (Var x)" from this and \x \ vars_of_cl (cl_ecl C)\ have "\ (ground_on (vars_of_cl (cl_ecl C)) \)" unfolding ground_on_def ground_term_def by auto from this and \ground_clause (subst_cl (cl_ecl C) \)\ show False using ground_clauses_and_ground_substs by metis qed from \subst (Var x) \ \ (Var x)\ have "(trm_rep (subst (Var x) \) S) = (subst (Var x) ?\)" using map_subst_lemma [of x \ "\x. (trm_rep x S)"] by metis from this and \trm_rep (subst (Var x) \) S \ (subst (Var x) \)\ have "((subst (Var x) ?\),(subst (Var x) \)) \ trm_ord" using trm_rep_is_lower by metis from \lower_on ?\ \ (vars_of_cl (cl_ecl C))\ and \x \ vars_of_cl (cl_ecl C)\ \finite (cl_ecl C)\ and \((subst (Var x) ?\),(subst (Var x) \)) \ trm_ord\ have "((C,?\), (C, \)) \ ecl_ord" using lower_on_cl by blast from \C = fst pair\ \\ = snd pair\ have "pair = (C,\)" by auto from this and \((C,?\), (C, \)) \ ecl_ord\ have "((C,?\),pair) \ ecl_ord" by metis from this and hyp_ind have "?P (C,?\)" by blast from \(all_trms_irreducible (subst_set (trms_ecl C) \) (\t. (trm_rep t S)))\ \lower_on ?\ \ (vars_of_cl (cl_ecl C))\ \C \ S\ \fo_interpretation ?I\ \equivalent_on \ ?\ (vars_of_cl (cl_ecl C)) ?I\ assms(3) have "(all_trms_irreducible (subst_set (trms_ecl C) ?\) (\t. (trm_rep t S)))" using irred_terms_and_reduced_subst unfolding Ball_def well_constrained_def by metis have "ground_clause (subst_cl (cl_ecl C) ?\)" proof - from \ground_clause (subst_cl (cl_ecl C) \)\ have "ground_on (vars_of_cl (cl_ecl C)) \" using ground_clauses_and_ground_substs by auto from this and \lower_on ?\ \ (vars_of_cl (cl_ecl C))\ have "ground_on (vars_of_cl (cl_ecl C)) ?\" using lower_on_ground by meson from this show ?thesis using ground_substs_yield_ground_clause by metis qed from this \(all_trms_irreducible (subst_set (trms_ecl C) ?\) (\t. (trm_rep t S)))\ \?P (C,?\)\ \\ validate_ground_clause ?I (subst_cl (cl_ecl C) ?\)\ \C \ S\ show False unfolding int_clset_def by (metis fst_conv snd_conv) qed text \Next, we show that the eligible term @{term "t"} is in normal form. We first need to establish the result for proper subterms of @{term "t"} before considering the general case.\ have "\(proper_subterm_red t S \)" proof assume "(proper_subterm_red t S \)" from this have "trm_rep (subst t \) S \ subst t \" using proper_subterm_red_def substs_preserve_subterms subts_of_irred_trms_are_irred by blast from \(proper_subterm_red t S \)\ \\x y. ((x \ vars_of_cl (cl_ecl C)) \ occurs_in y (subst (Var x) \) \ trm_rep y S = y)\ \eligible_literal L C \\ \trm_rep (subst t \) S \ subst t \\ \L \ cl_ecl C\ \orient_lit_inst L t s p \\ \\x\S. finite (cl_ecl x)\ \ground_clause (subst_cl (cl_ecl C) \)\ \fo_interpretation (int_clset S)\ \Ball S well_constrained\ \C \ S\ \all_trms_irreducible (subst_set (trms_ecl C) \) (\t. trm_rep t S)\ \\ validate_ground_clause (int_clset S) (subst_cl (cl_ecl C) \)\ \closed_under_renaming S\ have " \\'' u u' pa v D L2. (reduction L C \'' t s p L2 u u' pa v D (same_values (\t. trm_rep t S)) S \ \ variable_disjoint C D)" using reduction_exists [of p t s C S \ L] unfolding int_clset_def by blast from this and \?nored\ show False unfolding int_clset_def by blast qed have "p = neg \ \ equivalent_eq_exists t s (cl_ecl C) (same_values (\x. trm_rep x S)) \ L" proof (rule ccontr) assume neg: "\ (p = neg \ \ equivalent_eq_exists t s (cl_ecl C) (same_values (\x. trm_rep x S)) \ L)" then have "p \ neg" by metis from neg have "equivalent_eq_exists t s (cl_ecl C) (same_values (\x. trm_rep x S)) \ L" by metis from \p \ neg\ have "p = pos" using sign.exhaust by auto from \equivalent_eq_exists t s (cl_ecl C) (same_values (\x. trm_rep x S)) \ L\ obtain L2 where "L2 \ (cl_ecl C) - { L }" and f:"\u v. orient_lit_inst L2 u v pos \ \ subst t \ = subst u \ \ trm_rep (subst s \) S = trm_rep (subst v \) S" unfolding equivalent_eq_exists_def unfolding same_values_def by metis from f obtain u v where f': "orient_lit_inst L2 u v pos \ \ subst t \ = subst u \ \ trm_rep (subst s \) S = trm_rep (subst v \) S" by blast from f' have "orient_lit_inst L2 u v pos \" by metis from f' have "subst t \ = subst u \" by metis from f' have "trm_rep (subst s \) S = trm_rep (subst v \) S" by metis from \orient_lit_inst L2 u v pos \\ \subst t \ = subst u \\ \trm_rep (subst s \) S = trm_rep (subst v \) S\ \orient_lit_inst L t s p \\ \p = pos\ \L \ (cl_ecl C)\ \L2 \ (cl_ecl C) - { L }\ \eligible_literal L C \\ \\(proper_subterm_red t S \)\ and \?no_fact\ show False by blast qed have "(trm_rep (subst t \) S) = (subst t \)" proof (rule ccontr) assume "(trm_rep (subst t \) S) \ (subst t \)" from \p = neg \ \ equivalent_eq_exists t s (cl_ecl C) (same_values (\x. trm_rep x S)) \ L\ \\x y. ((x \ vars_of_cl (cl_ecl C)) \ occurs_in y (subst (Var x) \) \ trm_rep y S = y)\ \eligible_literal L C \\ \trm_rep (subst t \) S \ subst t \\ \L \ cl_ecl C\ \orient_lit_inst L t s p \\ \\x\S. finite (cl_ecl x)\ \ground_clause (subst_cl (cl_ecl C) \)\ \fo_interpretation (int_clset S)\ \Ball S well_constrained\ \C \ S\ \all_trms_irreducible (subst_set (trms_ecl C) \) (\t. trm_rep t S)\ \\ validate_ground_clause (int_clset S) (subst_cl (cl_ecl C) \)\ \closed_under_renaming S\ have " \\'' u u' pa v D L2. (reduction L C \'' t s p L2 u u' pa v D (same_values (\t. trm_rep t S)) S \ \ variable_disjoint C D)" using reduction_exists [of p t s C S \ L] unfolding int_clset_def by blast from this and \?nored\ show False unfolding int_clset_def by blast qed from \orient_lit_inst L t s p \\ have "((subst t \),(subst s \)) \ trm_ord" unfolding orient_lit_inst_def by auto from \ground_clause (subst_cl (cl_ecl C) \)\ have "vars_of_cl (subst_cl (cl_ecl C) \) = {}" by auto from this and \L \ (cl_ecl C)\ have "vars_of_lit (subst_lit L \) = {}" by auto from \orient_lit_inst L t s p \\ have "orient_lit (subst_lit L \) (subst t \) (subst s \) p" using lift_orient_lit by auto from \orient_lit (subst_lit L \) (subst t \) (subst s \) p\ have "vars_of (subst t \) \ vars_of_lit (subst_lit L \)" using orient_lit_vars by auto from this and \vars_of_lit (subst_lit L \) = {}\ have "vars_of (subst t \) = {}" by auto from \orient_lit (subst_lit L \) (subst t \) (subst s \) p\ have "vars_of (subst s \) \ vars_of_lit (subst_lit L \)" using orient_lit_vars by auto from this and \vars_of_lit (subst_lit L \) = {}\ have "vars_of (subst s \) = {}" by auto from \((subst t \),(subst s \)) \ trm_ord\ \vars_of (subst t \) = {}\ \vars_of (subst s \) = {}\ have "(subst t \) = (subst s \) \ ((subst s \),(subst t \)) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by blast from \L \ (cl_ecl C)\ have "(subst_lit L \) \ (subst_cl (cl_ecl C) \)" by auto text \Using the fact that the eligible term is in normal form and that the eligible literal is false in the considered interpretation but is not a contradiction, we deduce that this literal must be positive.\ have "p = pos" proof (rule ccontr) assume "p \ pos" from this have "p = neg" using sign.exhaust by auto from \trm_rep (subst t \) S = (subst t \)\ \L \ (cl_ecl C)\ \eligible_literal L C \\ and \orient_lit_inst L t s p \\ \p = neg\ \?no_cont\ have "(subst t \) \ (subst s \)" by blast from this and \(subst t \) = (subst s \) \ ((subst s \),(subst t \)) \ trm_ord\ and \trm_rep (subst t \) S = subst t \\ have "((trm_rep (subst s \) S),(trm_rep (subst t \) S)) \ trm_ord" using trm_rep_is_lower [of " (subst s \)" S ] trm_ord_trans unfolding trans_def by metis from this have "(trm_rep (subst s \) S) \ (trm_rep (subst t \) S)" using trm_ord_irrefl irrefl_def by metis from this have "\validate_ground_eq ?I (Eq (subst t \) (subst s \))" unfolding same_values_def int_clset_def using validate_ground_eq.simps by (metis (mono_tags, lifting)) from \(trm_rep (subst s \) S) \ (trm_rep (subst t \) S)\ have "\validate_ground_eq ?I (Eq (subst s \) (subst t \))" unfolding same_values_def int_clset_def using validate_ground_eq.simps by (metis (mono_tags, lifting)) from \orient_lit_inst L t s p \\ and \p=neg\ have "L = (Neg (Eq t s)) \ L = (Neg (Eq s t))" unfolding orient_lit_inst_def by auto from this have "subst_lit L \ = (Neg (Eq (subst t \) (subst s \))) \ subst_lit L \ = (Neg (Eq (subst s \) (subst t \)))" by auto from this and \\validate_ground_eq ?I (Eq (subst s \) (subst t \))\ and \\validate_ground_eq ?I (Eq (subst t \) (subst s \))\ have "validate_ground_lit ?I (subst_lit L \)" using validate_ground_lit.simps(2) by metis from \(subst_lit L \) \ (subst_cl (cl_ecl C) \)\ and \validate_ground_lit ?I (subst_lit L \)\ have "validate_ground_clause ?I (subst_cl (cl_ecl C) \)" using validate_ground_clause.elims(3) by blast from this and \\ validate_ground_clause ?I (subst_cl (cl_ecl C) \)\ show False by blast qed text \This entails that the right-hand side of the eligible literal occurs in the set of possible values for the left-hand side @{term "t"}, which is impossible since this term is irreducible.\ from \L \ (cl_ecl C)\ \orient_lit_inst L t s p \\ \p = pos\ \\ validate_ground_clause ?I (subst_cl (cl_ecl C) \)\ have "trm_rep (subst t \) S \ trm_rep (subst s \) S" using no_valid_literal by metis from this have "(subst t \) \ (subst s \)" by metis from this and \(subst t \) = (subst s \) \ ((subst s \),(subst t \)) \ trm_ord\ have "((subst s \),(subst t \)) \ trm_ord" using trm_ord_ground_total unfolding ground_term_def by blast from \p=pos\ and \orient_lit_inst L t s p \\ have "\negative_literal L" unfolding orient_lit_inst_def by auto from this and \eligible_literal L C \\ have "sel(cl_ecl C) = {}" and "maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)" using sel_neg unfolding eligible_literal_def by auto from \\ validate_ground_clause ?I (subst_cl (cl_ecl C) \)\ have "smaller_lits_are_false (subst t \) (subst_cl (cl_ecl C) \) S" using smaller_lits_are_false_if_cl_not_valid [of S "(subst_cl (cl_ecl C) \)" ] by blast from \p = pos\ and \p = neg \ \ equivalent_eq_exists t s (cl_ecl C) (same_values (\x. trm_rep x S)) \ L\ have "\ equivalent_eq_exists t s (cl_ecl C) (int_clset S) \ L" unfolding int_clset_def using sign.distinct by metis from this \p=pos\ have "maximal_literal_is_unique (subst t \) (subst s \) (cl_ecl C) L S \" using maximal_literal_is_unique_lemma [of t s "(cl_ecl C)" S \ L] by blast from \all_trms_irreducible (subst_set (trms_ecl C) \) (\t. trm_rep t S)\ have "trms_irreducible C \ S (subst t \)" using trms_irreducible_lemma by blast have "(subst t \) \ subst_set (trms_ecl C) \" proof assume "(subst t \) \ subst_set (trms_ecl C) \" from this obtain t' where "t' \ trms_ecl C" and "(subst t' \) = (subst t \)" by auto from \t' \ trms_ecl C\ and assms(3) and \C \ S\ have "dom_trm t' (cl_ecl C)" unfolding Ball_def well_constrained_def by auto from this obtain M u v q where "M \ (cl_ecl C)" "decompose_literal M u v q" and "q = neg \ (u = t') \ ( (t',u) \ trm_ord)" unfolding dom_trm_def by blast obtain u' v' q' where "orient_lit_inst M u' v' q' \" using literal.exhaust equation.exhaust using trm_ord_irrefl trm_ord_trans unfolding orient_lit_inst_def irrefl_def trans_def by metis from \decompose_literal M u v q\ and \orient_lit_inst M u' v' q' \\ have "u = u' \ u = v'" unfolding decompose_literal_def orient_lit_inst_def by (metis atom.simps(2) decompose_equation_def equation.inject literal.distinct(1) literal.inject(1)) from \decompose_literal M u v q\ and \orient_lit_inst M u' v' q' \\ have "q = q'" unfolding decompose_literal_def orient_lit_inst_def by auto from \vars_of_cl (subst_cl (cl_ecl C) \) = {}\ and \M \ (cl_ecl C)\ have "vars_of_lit (subst_lit M \) = {}" by auto from \orient_lit_inst M u' v' q' \\ have "orient_lit (subst_lit M \) (subst u' \) (subst v' \) q'" using lift_orient_lit by auto from \orient_lit_inst L t s p \\ have "orient_lit (subst_lit L \) (subst t \) (subst s \) p" using lift_orient_lit by auto have "(t',u) \ trm_ord" proof assume "(t',u) \ trm_ord" then have "((subst t' \),(subst u \)) \ trm_ord" using trm_ord_subst by auto from this and \(subst t' \) = (subst t \)\ have "((subst t \),(subst u \)) \ trm_ord" by auto from \orient_lit (subst_lit M \) (subst u' \) (subst v' \) q'\ and \orient_lit (subst_lit L \) (subst t \) (subst s \) p\ and \((subst t \),(subst u \)) \ trm_ord\ and \vars_of_lit (subst_lit M \) = {}\ and \vars_of_lit (subst_lit L \) = {}\ and \u = u' \ u = v'\ have "((subst_lit L \),(subst_lit M \)) \ lit_ord" using lit_ord_dominating_term by metis from this and \maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)\ and \M \ (cl_ecl C)\ show False using maximal_literal_def by auto qed have "\ (q = neg \ (u = t'))" proof assume "q = neg \ (u = t')" then have "q = neg" and "u = t'" by auto from \orient_lit (subst_lit M \) (subst u' \) (subst v' \) q'\ and \orient_lit (subst_lit L \) (subst t \) (subst s \) p\ and \u = t'\ and \(subst t' \) = (subst t \)\ and \q = neg\ and \q = q'\ and \p = pos\ and \vars_of_lit (subst_lit M \) = {}\ and \vars_of_lit (subst_lit L \) = {}\ and \u = u' \ u = v'\ have "((subst_lit L \),(subst_lit M \)) \ lit_ord" using lit_ord_neg_lit_lhs lit_ord_neg_lit_rhs by metis from this and \maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)\ and \M \ (cl_ecl C)\ show False using maximal_literal_def by auto qed from this and \(t',u) \ trm_ord\ and \q = neg \ (u = t') \ ( (t',u) \ trm_ord)\ show False by auto qed from \C \ S\ \(subst s \, subst t \) \ trm_ord\ and \p=pos\ \orient_lit_inst L t s p \\ and \sel (cl_ecl C) = {}\ and \L \ cl_ecl C\ and \maximal_literal (subst_lit L \) (subst_cl (cl_ecl C) \)\ and \ground_clause (subst_cl (cl_ecl C) \)\ and \finite (cl_ecl C)\ and \smaller_lits_are_false (subst t \) (subst_cl (cl_ecl C) \) S\ and \maximal_literal_is_unique (subst t \) (subst s \) (cl_ecl C) L S \\ and \trms_irreducible C \ S (subst t \)\ and \(subst t \) \ subst_set (trms_ecl C) \\ have cv: "(candidate_values (trm_rep (subst s \) S) C (cl_ecl C) (subst_cl (cl_ecl C) \) (subst s \) (subst_lit L \) L \ t s (subst t \) S)" unfolding candidate_values_def by blast from cv have "(trm_rep (subst s \) S,(subst s \)) \ set_of_candidate_values S (subst t \)" unfolding set_of_candidate_values_def by blast from \trm_rep (subst t \) S = (subst t \)\ have "\(subterm_reduction_applicable S (subst t \))" using trm_rep_is_lower_subt_red trm_ord_irrefl irrefl_def by metis from \(trm_rep (subst s \) S, subst s \) \ set_of_candidate_values S (subst t \)\ have "set_of_candidate_values S (subst t \) \ {}" by blast from \(trm_rep (subst s \) S, subst s \) \ set_of_candidate_values S (subst t \)\ have "min_trms (set_of_candidate_values S (subst t \)) \ {}" using min_trms_not_empty by blast from \\(subterm_reduction_applicable S (subst t \))\ \min_trms (set_of_candidate_values S (subst t \)) \ {}\ have "(trm_rep (subst t \) S,(subst t \)) \ trm_ord" using trm_rep_is_lower_root_red [of S "subst t \"] by blast from this and \(trm_rep (subst t \) S) = (subst t \)\ show False using trm_ord_irrefl irrefl_def by metis qed qed text \As an immediate consequence of the previous lemma, we show that the set of clauses that are derivable from an unsatisfiable clause set must contain an empty clause (since this set is trivially saturated).\ lemma COMPLETENESS: assumes "\x. (x \ S \ (trms_ecl x = {}))" assumes "(\x\S. finite (cl_ecl x))" assumes "\ (satisfiable_clause_set (cl_ecl ` S))" shows "\x. (derivable_ecl x S) \ cl_ecl x = {}" proof (rule ccontr) assume "\ (\x. (derivable_ecl x S) \ cl_ecl x = {})" let ?S = "{ y. (derivable_ecl y S) }" let ?I = "same_values (\x. (trm_rep x ?S))" have "fo_interpretation ?I" using trm_rep_compatible_with_structure same_values_fo_int by metis have "\x \ ?S. (cl_ecl x) \ {}" proof (rule ccontr) assume "\ ?thesis" then obtain x where "x \ ?S" and "cl_ecl x = {}" by blast from \x \ ?S\ have "derivable_ecl x S" by (meson CollectD) from this \cl_ecl x = {}\ \\ (\x. (derivable_ecl x S) \ cl_ecl x = {})\ show False by metis qed have all_finite: "\x \ ?S. (finite (cl_ecl x))" proof (rule ccontr) assume "\ ?thesis" then obtain x where "x \ ?S" and "\ finite (cl_ecl x)" by blast from \x \ ?S\ have "derivable_ecl x S" by (meson CollectD) from this assms(2) \\ finite (cl_ecl x)\ show False using all_derived_clauses_are_finite by metis qed have "Ball S well_constrained" proof fix x assume "x \ S" from this assms(1) have "trms_ecl x = {}" by auto from this show "well_constrained x" unfolding well_constrained_def by blast qed have "Ball ?S well_constrained" proof fix x assume "x \ ?S" from this have "derivable_ecl x S" by (meson CollectD) from this assms(2) \Ball S well_constrained\ show "well_constrained x" using all_derived_clauses_are_wellconstrained by metis qed have "closed_under_renaming ?S" proof (rule ccontr) assume "\ ?thesis" then obtain C D where "C \ ?S" "renaming_cl C D" "D \ ?S" unfolding closed_under_renaming_def by metis from \C \ ?S\ have "derivable_ecl C S" by (meson CollectD) from \derivable_ecl C S\ \renaming_cl C D\ have "(derivable_ecl D S)" using derivable_ecl.intros(2) by metis from this and \D \ ?S\ show False by blast qed have "inference_closed ?S" proof (rule ccontr) assume "\ ?thesis" then obtain D P \ C' where "(derivable D P ?S \ FirstOrder C')" "D \ ?S" unfolding inference_closed_def by metis from \derivable D P ?S \ FirstOrder C'\ have "P \ ?S" using derivable_premisses by metis have "\x. x \ P \ derivable_ecl x S" proof ((rule allI),(rule impI)) fix x assume "x \ P" from this and \P \ ?S\ have "x \ ?S" by (meson rev_subsetD) from this show "derivable_ecl x S" by (meson CollectD) qed from this and \(derivable D P ?S \ FirstOrder C')\ have "derivable_ecl D S" using derivable_ecl.intros(3) [of P S D ?S \ C'] by meson from this and \D \ ?S\ show False by blast qed from this all_finite have "clause_saturated ?S" using inference_closed_sets_are_saturated by meson from this all_finite have "inference_saturated ?S" using clause_saturated_and_inference_saturated by meson from this have "ground_inference_saturated ?S" using lift_inference by metis have "validate_clause_set ?I (cl_ecl ` S)" proof (rule ccontr) assume "\ ?thesis" from this obtain Cl_C where clc: "Cl_C \ (cl_ecl ` S)" and "\ (validate_clause ?I Cl_C)" using validate_clause_set.simps by metis from clc obtain C where "C \ S" and "Cl_C = (cl_ecl C)" by blast from \C \ S\ have "derivable_ecl C S" using derivable_ecl.intros(1) by metis from this have "C \ ?S" by blast from \\ (validate_clause ?I Cl_C)\ obtain \ where "\ (validate_ground_clause ?I (subst_cl Cl_C \))" and "ground_clause (subst_cl Cl_C \)" using validate_clause.simps by metis let ?pair = "(C,\)" have "fst ?pair = C" by auto have "snd ?pair = \" by auto from \C \ S\ assms(1) have "trms_ecl C = {}" by auto then have "(subst_set (trms_ecl C) \) = {}" by auto then have n: "all_trms_irreducible (subst_set (trms_ecl C) \) (\t. trm_rep t {y. derivable_ecl y S})" unfolding all_trms_irreducible_def by blast from \ground_inference_saturated ?S\ all_finite \Ball ?S well_constrained\ \closed_under_renaming ?S\ \\x \ ?S. (cl_ecl x) \ {}\ have "\C \. fst ?pair = C \ \ = snd ?pair \ C \ {y. derivable_ecl y S} \ ground_clause (subst_cl (cl_ecl C) \) \ all_trms_irreducible (subst_set (trms_ecl C) \) (\t. trm_rep t {y. derivable_ecl y S}) \ validate_ground_clause ?I (subst_cl (cl_ecl C) \)" using int_clset_is_a_model [of ?S ?pair] by blast from this \fst ?pair = C\ \C \ ?S\ \snd ?pair = \\ \Cl_C = (cl_ecl C)\ \ground_clause (subst_cl Cl_C \)\ n have "validate_ground_clause ?I (subst_cl (cl_ecl C) \)" by metis from this and \\ (validate_ground_clause ?I (subst_cl Cl_C \))\ \Cl_C = (cl_ecl C)\ show False by metis qed from this and assms(3) \fo_interpretation ?I\ show False using satisfiable_clause_set_def by metis qed end end