Theory Resolution
section ‹More Terms and Literals›
theory Resolution imports TermsAndLiterals Tree begin
fun complement :: "'t literal ⇒ 't literal" (‹_⇧c› [300] 300) where
"(Pos P ts)⇧c = Neg P ts"
| "(Neg P ts)⇧c = Pos P ts"
lemma cancel_comp1: "(l⇧c)⇧c = l" by (cases l) auto
lemma cancel_comp2:
assumes asm: "l⇩1⇧c = l⇩2⇧c"
shows "l⇩1 = l⇩2"
proof -
from asm have "(l⇩1⇧c)⇧c = (l⇩2⇧c)⇧c" by auto
then have "l⇩1 = (l⇩2⇧c)⇧c" using cancel_comp1[of l⇩1] by auto
then show ?thesis using cancel_comp1[of l⇩2] by auto
qed
lemma comp_exi1: "∃l'. l' = l⇧c" by (cases l) auto
lemma comp_exi2: "∃l. l' = l⇧c"
proof
show "l' = (l'⇧c)⇧c" using cancel_comp1[of l'] by auto
qed
lemma comp_swap: "l⇩1⇧c = l⇩2 ⟷ l⇩1 = l⇩2⇧c"
proof -
have "l⇩1⇧c = l⇩2 ⟶ l⇩1 = l⇩2⇧c" using cancel_comp1[of l⇩1] by auto
moreover
have "l⇩1 = l⇩2⇧c ⟶ l⇩1⇧c = l⇩2" using cancel_comp1 by auto
ultimately
show ?thesis by auto
qed
lemma sign_comp: "sign l⇩1 ≠ sign l⇩2 ∧ get_pred l⇩1 = get_pred l⇩2 ∧ get_terms l⇩1 = get_terms l⇩2 ⟷ l⇩2 = l⇩1⇧c"
by (cases l⇩1; cases l⇩2) auto
lemma sign_comp_atom: "sign l⇩1 ≠ sign l⇩2 ∧ get_atom l⇩1 = get_atom l⇩2 ⟷ l⇩2 = l⇩1⇧c"
by (cases l⇩1; cases l⇩2) auto
section ‹Clauses›
type_synonym 't clause = "'t literal set"
abbreviation complementls :: "'t literal set ⇒ 't literal set" (‹_⇧C› [300] 300) where
"L⇧C ≡ complement ` L"
lemma cancel_compls1: "(L⇧C)⇧C = L"
apply (auto simp add: cancel_comp1)
apply (metis imageI cancel_comp1)
done
lemma cancel_compls2:
assumes asm: "L⇩1⇧C = L⇩2⇧C"
shows "L⇩1 = L⇩2"
proof -
from asm have "(L⇩1⇧C)⇧C = (L⇩2⇧C)⇧C" by auto
then show ?thesis using cancel_compls1[of L⇩1] cancel_compls1[of L⇩2] by simp
qed
fun vars⇩t :: "fterm ⇒ var_sym set" where
"vars⇩t (Var x) = {x}"
| "vars⇩t (Fun f ts) = (⋃t ∈ set ts. vars⇩t t)"
abbreviation vars⇩t⇩s :: "fterm list ⇒ var_sym set" where
"vars⇩t⇩s ts ≡ (⋃t ∈ set ts. vars⇩t t)"
definition vars⇩l :: "fterm literal ⇒ var_sym set" where
"vars⇩l l = vars⇩t⇩s (get_terms l)"
definition vars⇩l⇩s :: "fterm literal set ⇒ var_sym set" where
"vars⇩l⇩s L ≡ ⋃l∈L. vars⇩l l"
lemma ground_vars⇩t:
assumes "ground⇩t t"
shows "vars⇩t t = {}"
using assms by (induction t) auto
lemma ground⇩t⇩s_vars⇩t⇩s:
assumes "ground⇩t⇩s ts"
shows "vars⇩t⇩s ts = {}"
using assms ground_vars⇩t by auto
lemma ground⇩l_vars⇩l:
assumes "ground⇩l l"
shows "vars⇩l l = {}"
unfolding vars⇩l_def using assms ground_vars⇩t by auto
lemma ground⇩l⇩s_vars⇩l⇩s:
assumes "ground⇩l⇩s L"
shows "vars⇩l⇩s L = {}" unfolding vars⇩l⇩s_def using assms ground⇩l_vars⇩l by auto
lemma ground_comp: "ground⇩l (l⇧c) ⟷ ground⇩l l" by (cases l) auto
lemma ground_compls: "ground⇩l⇩s (L⇧C) ⟷ ground⇩l⇩s L" using ground_comp by auto
section ‹Semantics›
type_synonym 'u fun_denot = "fun_sym ⇒ 'u list ⇒ 'u"
type_synonym 'u pred_denot = "pred_sym ⇒ 'u list ⇒ bool"
type_synonym 'u var_denot = "var_sym ⇒ 'u"
fun eval⇩t :: "'u var_denot ⇒ 'u fun_denot ⇒ fterm ⇒ 'u" where
"eval⇩t E F (Var x) = E x"
| "eval⇩t E F (Fun f ts) = F f (map (eval⇩t E F) ts)"
abbreviation eval⇩t⇩s :: "'u var_denot ⇒ 'u fun_denot ⇒ fterm list ⇒ 'u list" where
"eval⇩t⇩s E F ts ≡ map (eval⇩t E F) ts"
fun eval⇩l :: "'u var_denot ⇒ 'u fun_denot ⇒ 'u pred_denot ⇒ fterm literal ⇒ bool" where
"eval⇩l E F G (Pos p ts) ⟷ G p (eval⇩t⇩s E F ts)"
| "eval⇩l E F G (Neg p ts) ⟷ ¬G p (eval⇩t⇩s E F ts)"
definition eval⇩c :: "'u fun_denot ⇒ 'u pred_denot ⇒ fterm clause ⇒ bool" where
"eval⇩c F G C ⟷ (∀E. ∃l ∈ C. eval⇩l E F G l)"
definition eval⇩c⇩s :: "'u fun_denot ⇒ 'u pred_denot ⇒ fterm clause set ⇒ bool" where
"eval⇩c⇩s F G Cs ⟷ (∀C ∈ Cs. eval⇩c F G C)"
subsection ‹Semantics of Ground Terms›
lemma ground_var_denott:
assumes "ground⇩t t"
shows "eval⇩t E F t = eval⇩t E' F t"
using assms proof (induction t)
case (Var x)
then have "False" by auto
then show ?case by auto
next
case (Fun f ts)
then have "∀t ∈ set ts. ground⇩t t" by auto
then have "∀t ∈ set ts. eval⇩t E F t = eval⇩t E' F t" using Fun by auto
then have "eval⇩t⇩s E F ts = eval⇩t⇩s E' F ts" by auto
then have "F f (map (eval⇩t E F) ts) = F f (map (eval⇩t E' F) ts)" by metis
then show ?case by simp
qed
lemma ground_var_denotts:
assumes "ground⇩t⇩s ts"
shows "eval⇩t⇩s E F ts = eval⇩t⇩s E' F ts"
using assms ground_var_denott by (metis map_eq_conv)
lemma ground_var_denot:
assumes "ground⇩l l"
shows "eval⇩l E F G l = eval⇩l E' F G l"
using assms proof (induction l)
case Pos then show ?case using ground_var_denotts by (metis eval⇩l.simps(1) literal.sel(3))
next
case Neg then show ?case using ground_var_denotts by (metis eval⇩l.simps(2) literal.sel(4))
qed
section ‹Substitutions›
type_synonym substitution = "var_sym ⇒ fterm"
fun sub :: "fterm ⇒ substitution ⇒ fterm" (infixl ‹⋅⇩t› 55) where
"(Var x) ⋅⇩t σ = σ x"
| "(Fun f ts) ⋅⇩t σ = Fun f (map (λt. t ⋅⇩t σ) ts)"
abbreviation subs :: "fterm list ⇒ substitution ⇒ fterm list" (infixl ‹⋅⇩t⇩s› 55) where
"ts ⋅⇩t⇩s σ ≡ (map (λt. t ⋅⇩t σ) ts)"
fun subl :: "fterm literal ⇒ substitution ⇒ fterm literal" (infixl ‹⋅⇩l› 55) where
"(Pos p ts) ⋅⇩l σ = Pos p (ts ⋅⇩t⇩s σ)"
| "(Neg p ts) ⋅⇩l σ = Neg p (ts ⋅⇩t⇩s σ)"
abbreviation subls :: "fterm literal set ⇒ substitution ⇒ fterm literal set" (infixl ‹⋅⇩l⇩s› 55) where
"L ⋅⇩l⇩s σ ≡ (λl. l ⋅⇩l σ) ` L"
lemma subls_def2: "L ⋅⇩l⇩s σ = {l ⋅⇩l σ|l. l ∈ L}" by auto
definition instance_of⇩t :: "fterm ⇒ fterm ⇒ bool" where
"instance_of⇩t t⇩1 t⇩2 ⟷ (∃σ. t⇩1 = t⇩2 ⋅⇩t σ)"
definition instance_of⇩t⇩s :: "fterm list ⇒ fterm list ⇒ bool" where
"instance_of⇩t⇩s ts⇩1 ts⇩2 ⟷ (∃σ. ts⇩1 = ts⇩2 ⋅⇩t⇩s σ)"
definition instance_of⇩l :: "fterm literal ⇒ fterm literal ⇒ bool" where
"instance_of⇩l l⇩1 l⇩2 ⟷ (∃σ. l⇩1 = l⇩2 ⋅⇩l σ)"
definition instance_of⇩l⇩s :: "fterm clause ⇒ fterm clause ⇒ bool" where
"instance_of⇩l⇩s C⇩1 C⇩2 ⟷ (∃σ. C⇩1 = C⇩2 ⋅⇩l⇩s σ)"
lemma comp_sub: "(l⇧c) ⋅⇩l σ=(l ⋅⇩l σ)⇧c"
by (cases l) auto
lemma compls_subls: "(L⇧C) ⋅⇩l⇩s σ=(L ⋅⇩l⇩s σ)⇧C"
using comp_sub apply auto
apply (metis image_eqI)
done
lemma subls_union: "(L⇩1 ∪ L⇩2) ⋅⇩l⇩s σ = (L⇩1 ⋅⇩l⇩s σ) ∪ (L⇩2 ⋅⇩l⇩s σ)" by auto
definition var_renaming_of :: "fterm clause ⇒ fterm clause ⇒ bool" where
"var_renaming_of C⇩1 C⇩2 ⟷ instance_of⇩l⇩s C⇩1 C⇩2 ∧ instance_of⇩l⇩s C⇩2 C⇩1"
subsection ‹The Empty Substitution›
abbreviation ε :: "substitution" where
"ε ≡ Var"
lemma empty_subt: "(t :: fterm) ⋅⇩t ε = t"
by (induction t) (auto simp add: map_idI)
lemma empty_subts: "ts ⋅⇩t⇩s ε = ts"
using empty_subt by auto
lemma empty_subl: "l ⋅⇩l ε = l"
using empty_subts by (cases l) auto
lemma empty_subls: "L ⋅⇩l⇩s ε = L"
using empty_subl by auto
lemma instance_of⇩t_self: "instance_of⇩t t t"
unfolding instance_of⇩t_def
proof
show "t = t ⋅⇩t ε" using empty_subt by auto
qed
lemma instance_of⇩t⇩s_self: "instance_of⇩t⇩s ts ts"
unfolding instance_of⇩t⇩s_def
proof
show "ts = ts ⋅⇩t⇩s ε" using empty_subts by auto
qed
lemma instance_of⇩l_self: "instance_of⇩l l l"
unfolding instance_of⇩l_def
proof
show "l = l ⋅⇩l ε" using empty_subl by auto
qed
lemma instance_of⇩l⇩s_self: "instance_of⇩l⇩s L L"
unfolding instance_of⇩l⇩s_def
proof
show "L = L ⋅⇩l⇩s ε" using empty_subls by auto
qed
subsection ‹Substitutions and Ground Terms›
lemma ground_sub:
assumes "ground⇩t t"
shows "t ⋅⇩t σ = t"
using assms by (induction t) (auto simp add: map_idI)
lemma ground_subs:
assumes "ground⇩t⇩s ts "
shows " ts ⋅⇩t⇩s σ = ts"
using assms ground_sub by (simp add: map_idI)
lemma ground⇩l_subs:
assumes "ground⇩l l "
shows " l ⋅⇩l σ = l"
using assms ground_subs by (cases l) auto
lemma ground⇩l⇩s_subls:
assumes ground: "ground⇩l⇩s L"
shows "L ⋅⇩l⇩s σ = L"
proof -
{
fix l
assume l_L: "l ∈ L"
then have "ground⇩l l" using ground by auto
then have "l = l ⋅⇩l σ" using ground⇩l_subs by auto
moreover
then have "l ⋅⇩l σ ∈ L ⋅⇩l⇩s σ" using l_L by auto
ultimately
have "l ∈ L ⋅⇩l⇩s σ" by auto
}
moreover
{
fix l
assume l_L: "l ∈ L ⋅⇩l⇩s σ"
then obtain l' where l'_p: "l' ∈ L ∧ l' ⋅⇩l σ = l" by auto
then have "l' = l" using ground ground⇩l_subs by auto
from l_L l'_p this have "l ∈ L" by auto
}
ultimately show ?thesis by auto
qed
subsection ‹Composition›
definition composition :: "substitution ⇒ substitution ⇒ substitution" (infixl ‹⋅› 55) where
"(σ⇩1 ⋅ σ⇩2) x = (σ⇩1 x) ⋅⇩t σ⇩2"
lemma composition_conseq2t: "(t ⋅⇩t σ⇩1) ⋅⇩t σ⇩2 = t ⋅⇩t (σ⇩1 ⋅ σ⇩2)"
proof (induction t)
case (Var x)
have "((Var x) ⋅⇩t σ⇩1) ⋅⇩t σ⇩2 = (σ⇩1 x) ⋅⇩t σ⇩2" by simp
also have " ... = (σ⇩1 ⋅ σ⇩2) x" unfolding composition_def by simp
finally show ?case by auto
next
case (Fun t ts)
then show ?case unfolding composition_def by auto
qed
lemma composition_conseq2ts: "(ts ⋅⇩t⇩s σ⇩1) ⋅⇩t⇩s σ⇩2 = ts ⋅⇩t⇩s (σ⇩1 ⋅ σ⇩2)"
using composition_conseq2t by auto
lemma composition_conseq2l: "(l ⋅⇩l σ⇩1) ⋅⇩l σ⇩2 = l ⋅⇩l (σ⇩1 ⋅ σ⇩2)"
using composition_conseq2t by (cases l) auto
lemma composition_conseq2ls: "(L ⋅⇩l⇩s σ⇩1) ⋅⇩l⇩s σ⇩2 = L ⋅⇩l⇩s (σ⇩1 ⋅ σ⇩2)"
using composition_conseq2l apply auto
apply (metis imageI)
done
lemma composition_assoc: "σ⇩1 ⋅ (σ⇩2 ⋅ σ⇩3) = (σ⇩1 ⋅ σ⇩2) ⋅ σ⇩3"
proof
fix x
show "(σ⇩1 ⋅ (σ⇩2 ⋅ σ⇩3)) x = ((σ⇩1 ⋅ σ⇩2) ⋅ σ⇩3) x"
by (simp only: composition_def composition_conseq2t)
qed
lemma empty_comp1: "(σ ⋅ ε) = σ"
proof
fix x
show "(σ ⋅ ε) x = σ x" unfolding composition_def using empty_subt by auto
qed
lemma empty_comp2: "(ε ⋅ σ) = σ"
proof
fix x
show "(ε ⋅ σ) x = σ x" unfolding composition_def by simp
qed
lemma instance_of⇩t_trans :
assumes t⇩1⇩2: "instance_of⇩t t⇩1 t⇩2"
assumes t⇩2⇩3: "instance_of⇩t t⇩2 t⇩3"
shows "instance_of⇩t t⇩1 t⇩3"
proof -
from t⇩1⇩2 obtain σ⇩1⇩2 where "t⇩1 = t⇩2 ⋅⇩t σ⇩1⇩2"
unfolding instance_of⇩t_def by auto
moreover
from t⇩2⇩3 obtain σ⇩2⇩3 where "t⇩2 = t⇩3 ⋅⇩t σ⇩2⇩3"
unfolding instance_of⇩t_def by auto
ultimately
have "t⇩1 = (t⇩3 ⋅⇩t σ⇩2⇩3) ⋅⇩t σ⇩1⇩2" by auto
then have "t⇩1 = t⇩3 ⋅⇩t (σ⇩2⇩3 ⋅ σ⇩1⇩2)" using composition_conseq2t by simp
then show ?thesis unfolding instance_of⇩t_def by auto
qed
lemma instance_of⇩t⇩s_trans :
assumes ts⇩1⇩2: "instance_of⇩t⇩s ts⇩1 ts⇩2"
assumes ts⇩2⇩3: "instance_of⇩t⇩s ts⇩2 ts⇩3"
shows "instance_of⇩t⇩s ts⇩1 ts⇩3"
proof -
from ts⇩1⇩2 obtain σ⇩1⇩2 where "ts⇩1 = ts⇩2 ⋅⇩t⇩s σ⇩1⇩2"
unfolding instance_of⇩t⇩s_def by auto
moreover
from ts⇩2⇩3 obtain σ⇩2⇩3 where "ts⇩2 = ts⇩3 ⋅⇩t⇩s σ⇩2⇩3"
unfolding instance_of⇩t⇩s_def by auto
ultimately
have "ts⇩1 = (ts⇩3 ⋅⇩t⇩s σ⇩2⇩3) ⋅⇩t⇩s σ⇩1⇩2" by auto
then have "ts⇩1 = ts⇩3 ⋅⇩t⇩s (σ⇩2⇩3 ⋅ σ⇩1⇩2)" using composition_conseq2ts by simp
then show ?thesis unfolding instance_of⇩t⇩s_def by auto
qed
lemma instance_of⇩l_trans :
assumes l⇩1⇩2: "instance_of⇩l l⇩1 l⇩2"
assumes l⇩2⇩3: "instance_of⇩l l⇩2 l⇩3"
shows "instance_of⇩l l⇩1 l⇩3"
proof -
from l⇩1⇩2 obtain σ⇩1⇩2 where "l⇩1 = l⇩2 ⋅⇩l σ⇩1⇩2"
unfolding instance_of⇩l_def by auto
moreover
from l⇩2⇩3 obtain σ⇩2⇩3 where "l⇩2 = l⇩3 ⋅⇩l σ⇩2⇩3"
unfolding instance_of⇩l_def by auto
ultimately
have "l⇩1 = (l⇩3 ⋅⇩l σ⇩2⇩3) ⋅⇩l σ⇩1⇩2" by auto
then have "l⇩1 = l⇩3 ⋅⇩l (σ⇩2⇩3 ⋅ σ⇩1⇩2)" using composition_conseq2l by simp
then show ?thesis unfolding instance_of⇩l_def by auto
qed
lemma instance_of⇩l⇩s_trans :
assumes L⇩1⇩2: "instance_of⇩l⇩s L⇩1 L⇩2"
assumes L⇩2⇩3: "instance_of⇩l⇩s L⇩2 L⇩3"
shows "instance_of⇩l⇩s L⇩1 L⇩3"
proof -
from L⇩1⇩2 obtain σ⇩1⇩2 where "L⇩1 = L⇩2 ⋅⇩l⇩s σ⇩1⇩2"
unfolding instance_of⇩l⇩s_def by auto
moreover
from L⇩2⇩3 obtain σ⇩2⇩3 where "L⇩2 = L⇩3 ⋅⇩l⇩s σ⇩2⇩3"
unfolding instance_of⇩l⇩s_def by auto
ultimately
have "L⇩1 = (L⇩3 ⋅⇩l⇩s σ⇩2⇩3) ⋅⇩l⇩s σ⇩1⇩2" by auto
then have "L⇩1 = L⇩3 ⋅⇩l⇩s (σ⇩2⇩3 ⋅ σ⇩1⇩2)" using composition_conseq2ls by simp
then show ?thesis unfolding instance_of⇩l⇩s_def by auto
qed
subsection ‹Merging substitutions›
lemma project_sub:
assumes inst_C:"C ⋅⇩l⇩s lmbd = C'"
assumes L'sub: "L' ⊆ C'"
shows "∃L ⊆ C. L ⋅⇩l⇩s lmbd = L' ∧ (C-L) ⋅⇩l⇩s lmbd = C' - L'"
proof -
let ?L = "{l ∈ C. ∃l' ∈ L'. l ⋅⇩l lmbd = l'}"
have "?L ⊆ C" by auto
moreover
have "?L ⋅⇩l⇩s lmbd = L'"
proof (rule Orderings.order_antisym; rule Set.subsetI)
fix l'
assume l'L: "l' ∈ L'"
from inst_C have "{l ⋅⇩l lmbd|l. l ∈ C} = C'" unfolding subls_def2 by -
then have "∃l. l' = l ⋅⇩l lmbd ∧ l ∈ C ∧ l ⋅⇩l lmbd ∈ L'" using L'sub l'L by auto
then have " l' ∈ {l ∈ C. l ⋅⇩l lmbd ∈ L'} ⋅⇩l⇩s lmbd" by auto
then show " l' ∈ {l ∈ C. ∃l'∈L'. l ⋅⇩l lmbd = l'} ⋅⇩l⇩s lmbd" by auto
qed auto
moreover
have "(C-?L) ⋅⇩l⇩s lmbd = C' - L'" using inst_C by auto
ultimately show ?thesis
by blast
qed
lemma relevant_vars_subt:
assumes "∀x ∈ vars⇩t t. σ⇩1 x = σ⇩2 x"
shows " t ⋅⇩t σ⇩1 = t ⋅⇩t σ⇩2"
using assms proof (induction t)
case (Fun f ts)
have f: "∀t. t ∈ set ts ⟶ vars⇩t t ⊆ vars⇩t⇩s ts" by (induction ts) auto
have "∀t∈set ts. t ⋅⇩t σ⇩1 = t ⋅⇩t σ⇩2"
proof
fix t
assume tints: "t ∈ set ts"
then have "∀x ∈ vars⇩t t. σ⇩1 x = σ⇩2 x" using f Fun(2) by auto
then show "t ⋅⇩t σ⇩1 = t ⋅⇩t σ⇩2" using Fun tints by auto
qed
then have "ts ⋅⇩t⇩s σ⇩1 = ts ⋅⇩t⇩s σ⇩2" by auto
then show ?case by auto
qed auto
lemma relevant_vars_subts:
assumes asm: "∀x ∈ vars⇩t⇩s ts. σ⇩1 x = σ⇩2 x"
shows "ts ⋅⇩t⇩s σ⇩1 = ts ⋅⇩t⇩s σ⇩2"
proof -
have f: "∀t. t ∈ set ts ⟶ vars⇩t t ⊆ vars⇩t⇩s ts" by (induction ts) auto
have "∀t∈set ts. t ⋅⇩t σ⇩1 = t ⋅⇩t σ⇩2"
proof
fix t
assume tints: "t ∈ set ts"
then have "∀x ∈ vars⇩t t. σ⇩1 x = σ⇩2 x" using f asm by auto
then show "t ⋅⇩t σ⇩1 = t ⋅⇩t σ⇩2" using relevant_vars_subt tints by auto
qed
then show ?thesis by auto
qed
lemma relevant_vars_subl:
assumes "∀x ∈ vars⇩l l. σ⇩1 x = σ⇩2 x "
shows "l ⋅⇩l σ⇩1 = l ⋅⇩l σ⇩2"
using assms proof (induction l)
case (Pos p ts)
then show ?case using relevant_vars_subts unfolding vars⇩l_def by auto
next
case (Neg p ts)
then show ?case using relevant_vars_subts unfolding vars⇩l_def by auto
qed
lemma relevant_vars_subls:
assumes asm: "∀x ∈ vars⇩l⇩s L. σ⇩1 x = σ⇩2 x"
shows "L ⋅⇩l⇩s σ⇩1 = L ⋅⇩l⇩s σ⇩2"
proof -
have f: "∀l. l ∈ L ⟶ vars⇩l l ⊆ vars⇩l⇩s L" unfolding vars⇩l⇩s_def by auto
have "∀l ∈ L. l ⋅⇩l σ⇩1 = l ⋅⇩l σ⇩2"
proof
fix l
assume linls: "l∈L"
then have "∀x∈vars⇩l l. σ⇩1 x = σ⇩2 x" using f asm by auto
then show "l ⋅⇩l σ⇩1 = l ⋅⇩l σ⇩2" using relevant_vars_subl linls by auto
qed
then show ?thesis by (meson image_cong)
qed
lemma merge_sub:
assumes dist: "vars⇩l⇩s C ∩ vars⇩l⇩s D = {}"
assumes CC': "C ⋅⇩l⇩s lmbd = C'"
assumes DD': "D ⋅⇩l⇩s μ = D'"
shows "∃η. C ⋅⇩l⇩s η = C' ∧ D ⋅⇩l⇩s η = D'"
proof -
let ?η = "λx. if x ∈ vars⇩l⇩s C then lmbd x else μ x"
have " ∀x∈vars⇩l⇩s C. ?η x = lmbd x" by auto
then have "C ⋅⇩l⇩s ?η = C ⋅⇩l⇩s lmbd" using relevant_vars_subls[of C ?η lmbd] by auto
then have "C ⋅⇩l⇩s ?η = C'" using CC' by auto
moreover
have " ∀x ∈ vars⇩l⇩s D. ?η x = μ x" using dist by auto
then have "D ⋅⇩l⇩s ?η = D ⋅⇩l⇩s μ" using relevant_vars_subls[of D ?η μ] by auto
then have "D ⋅⇩l⇩s ?η = D'" using DD' by auto
ultimately
show ?thesis by auto
qed
subsection ‹Standardizing apart›
abbreviation std⇩1 :: "fterm clause ⇒ fterm clause" where
"std⇩1 C ≡ C ⋅⇩l⇩s (λx. Var (''1'' @ x))"
abbreviation std⇩2 :: "fterm clause ⇒ fterm clause" where
"std⇩2 C ≡ C ⋅⇩l⇩s (λx. Var (''2'' @ x))"
lemma std_apart_apart'':
assumes "x ∈ vars⇩t (t ⋅⇩t (λx::char list. Var (y @ x)))"
shows "∃x'. x = y@x'"
using assms by (induction t) auto
lemma std_apart_apart':
assumes "x ∈ vars⇩l (l ⋅⇩l (λx. Var (y@x)))"
shows "∃x'. x = y@x'"
using assms unfolding vars⇩l_def using std_apart_apart'' by (cases l) auto
lemma std_apart_apart: "vars⇩l⇩s (std⇩1 C⇩1) ∩ vars⇩l⇩s (std⇩2 C⇩2) = {}"
proof -
{
fix x
assume xin: "x ∈ vars⇩l⇩s (std⇩1 C⇩1) ∩ vars⇩l⇩s (std⇩2 C⇩2)"
from xin have "x ∈ vars⇩l⇩s (std⇩1 C⇩1)" by auto
then have "∃x'. x=''1'' @ x'"
using std_apart_apart'[of x _ "''1''"] unfolding vars⇩l⇩s_def by auto
moreover
from xin have "x ∈ vars⇩l⇩s (std⇩2 C⇩2)" by auto
then have "∃x'. x= ''2'' @x' "
using std_apart_apart'[of x _ "''2''"] unfolding vars⇩l⇩s_def by auto
ultimately have "False" by auto
then have "x ∈ {}" by auto
}
then show ?thesis by auto
qed
lemma std_apart_instance_of⇩l⇩s1: "instance_of⇩l⇩s C⇩1 (std⇩1 C⇩1)"
proof -
have empty: "(λx. Var (''1''@x)) ⋅ (λx. Var (tl x)) = ε" using composition_def by auto
have "C⇩1 ⋅⇩l⇩s ε = C⇩1" using empty_subls by auto
then have "C⇩1 ⋅⇩l⇩s ((λx. Var (''1''@x)) ⋅ (λx. Var (tl x))) = C⇩1" using empty by auto
then have "(C⇩1 ⋅⇩l⇩s (λx. Var (''1''@x))) ⋅⇩l⇩s (λx. Var (tl x)) = C⇩1" using composition_conseq2ls by auto
then have "C⇩1 = (std⇩1 C⇩1) ⋅⇩l⇩s (λx. Var (tl x))" by auto
then show "instance_of⇩l⇩s C⇩1 (std⇩1 C⇩1)" unfolding instance_of⇩l⇩s_def by auto
qed
lemma std_apart_instance_of⇩l⇩s2: "instance_of⇩l⇩s C2 (std⇩2 C2)"
proof -
have empty: "(λx. Var (''2''@x)) ⋅ (λx. Var (tl x)) = ε" using composition_def by auto
have "C2 ⋅⇩l⇩s ε = C2" using empty_subls by auto
then have "C2 ⋅⇩l⇩s ((λx. Var (''2''@x)) ⋅ (λx. Var (tl x))) = C2" using empty by auto
then have "(C2 ⋅⇩l⇩s (λx. Var (''2''@x))) ⋅⇩l⇩s (λx. Var (tl x)) = C2" using composition_conseq2ls by auto
then have "C2 = (std⇩2 C2) ⋅⇩l⇩s (λx. Var (tl x))" by auto
then show "instance_of⇩l⇩s C2 (std⇩2 C2)" unfolding instance_of⇩l⇩s_def by auto
qed
section ‹Unifiers›
definition unifier⇩t⇩s :: "substitution ⇒ fterm set ⇒ bool" where
"unifier⇩t⇩s σ ts ⟷ (∃t'. ∀t ∈ ts. t ⋅⇩t σ = t')"
definition unifier⇩l⇩s :: "substitution ⇒ fterm literal set ⇒ bool" where
"unifier⇩l⇩s σ L ⟷ (∃l'. ∀l ∈ L. l ⋅⇩l σ = l')"
lemma unif_sub:
assumes unif: "unifier⇩l⇩s σ L"
assumes nonempty: "L ≠ {}"
shows "∃l. subls L σ = {subl l σ}"
proof -
from nonempty obtain l where "l ∈ L" by auto
from unif this have "L ⋅⇩l⇩s σ = {l ⋅⇩l σ}" unfolding unifier⇩l⇩s_def by auto
then show ?thesis by auto
qed
lemma unifiert_def2:
assumes L_elem: "ts ≠ {}"
shows "unifier⇩t⇩s σ ts ⟷ (∃l. (λt. sub t σ) ` ts ={l})"
proof
assume unif: "unifier⇩t⇩s σ ts"
from L_elem obtain t where "t ∈ ts" by auto
then have "(λt. sub t σ) ` ts = {t ⋅⇩t σ}" using unif unfolding unifier⇩t⇩s_def by auto
then show "∃l. (λt. sub t σ) ` ts = {l}" by auto
next
assume "∃l. (λt. sub t σ) ` ts ={l}"
then obtain l where "(λt. sub t σ) ` ts = {l}" by auto
then have "∀l' ∈ ts. l' ⋅⇩t σ = l" by auto
then show "unifier⇩t⇩s σ ts" unfolding unifier⇩t⇩s_def by auto
qed
lemma unifier⇩l⇩s_def2:
assumes L_elem: "L ≠ {}"
shows "unifier⇩l⇩s σ L ⟷ (∃l. L ⋅⇩l⇩s σ = {l})"
proof
assume unif: "unifier⇩l⇩s σ L"
from L_elem obtain l where "l ∈ L" by auto
then have "L ⋅⇩l⇩s σ = {l ⋅⇩l σ}" using unif unfolding unifier⇩l⇩s_def by auto
then show "∃l. L ⋅⇩l⇩s σ = {l}" by auto
next
assume "∃l. L ⋅⇩l⇩s σ ={l}"
then obtain l where "L ⋅⇩l⇩s σ = {l}" by auto
then have "∀l' ∈ L. l' ⋅⇩l σ = l" by auto
then show "unifier⇩l⇩s σ L" unfolding unifier⇩l⇩s_def by auto
qed
lemma ground⇩l⇩s_unif_singleton:
assumes ground⇩l⇩s: "ground⇩l⇩s L"
assumes unif: "unifier⇩l⇩s σ' L"
assumes empt: "L ≠ {}"
shows "∃l. L = {l}"
proof -
from unif empt have "∃l. L ⋅⇩l⇩s σ' = {l}" using unif_sub by auto
then show ?thesis using ground⇩l⇩s_subls ground⇩l⇩s by auto
qed
definition unifiablets :: "fterm set ⇒ bool" where
"unifiablets fs ⟷ (∃σ. unifier⇩t⇩s σ fs)"
definition unifiablels :: "fterm literal set ⇒ bool" where
"unifiablels L ⟷ (∃σ. unifier⇩l⇩s σ L)"
lemma unifier_comp[simp]: "unifier⇩l⇩s σ (L⇧C) ⟷ unifier⇩l⇩s σ L"
proof
assume "unifier⇩l⇩s σ (L⇧C)"
then obtain l'' where l''_p: "∀l ∈ L⇧C. l ⋅⇩l σ = l''"
unfolding unifier⇩l⇩s_def by auto
obtain l' where "(l')⇧c = l''" using comp_exi2[of l''] by auto
from this l''_p have l'_p:"∀l ∈ L⇧C. l ⋅⇩l σ = (l')⇧c" by auto
have "∀l ∈ L. l ⋅⇩l σ = l'"
proof
fix l
assume "l∈L"
then have "l⇧c ∈ L⇧C" by auto
then have "(l⇧c) ⋅⇩l σ = (l')⇧c" using l'_p by auto
then have "(l ⋅⇩l σ)⇧c = (l')⇧c" by (cases l) auto
then show "l ⋅⇩l σ = l'" using cancel_comp2 by blast
qed
then show "unifier⇩l⇩s σ L" unfolding unifier⇩l⇩s_def by auto
next
assume "unifier⇩l⇩s σ L"
then obtain l' where l'_p: "∀l ∈ L. l ⋅⇩l σ = l'" unfolding unifier⇩l⇩s_def by auto
have "∀l ∈ L⇧C. l ⋅⇩l σ = (l')⇧c"
proof
fix l
assume "l ∈ L⇧C"
then have "l⇧c ∈ L" using cancel_comp1 by (metis image_iff)
then show "l ⋅⇩l σ = (l')⇧c" using l'_p comp_sub cancel_comp1 by metis
qed
then show "unifier⇩l⇩s σ (L⇧C)" unfolding unifier⇩l⇩s_def by auto
qed
lemma unifier_sub1:
assumes "unifier⇩l⇩s σ L"
assumes "L' ⊆ L"
shows "unifier⇩l⇩s σ L' "
using assms unfolding unifier⇩l⇩s_def by auto
lemma unifier_sub2:
assumes asm: "unifier⇩l⇩s σ (L⇩1 ∪ L⇩2)"
shows "unifier⇩l⇩s σ L⇩1 ∧ unifier⇩l⇩s σ L⇩2 "
proof -
have "L⇩1 ⊆ (L⇩1 ∪ L⇩2) ∧ L⇩2 ⊆ (L⇩1 ∪ L⇩2)" by simp
from this asm show ?thesis using unifier_sub1 by auto
qed
subsection ‹Most General Unifiers›
definition mgu⇩t⇩s :: "substitution ⇒ fterm set ⇒ bool" where
"mgu⇩t⇩s σ ts ⟷ unifier⇩t⇩s σ ts ∧ (∀u. unifier⇩t⇩s u ts ⟶ (∃i. u = σ ⋅ i))"
definition mgu⇩l⇩s :: "substitution ⇒ fterm literal set ⇒ bool" where
"mgu⇩l⇩s σ L ⟷ unifier⇩l⇩s σ L ∧ (∀u. unifier⇩l⇩s u L ⟶ (∃i. u = σ ⋅ i))"
section ‹Resolution›
definition applicable :: " fterm clause ⇒ fterm clause
⇒ fterm literal set ⇒ fterm literal set
⇒ substitution ⇒ bool" where
"applicable C⇩1 C⇩2 L⇩1 L⇩2 σ ⟷
C⇩1 ≠ {} ∧ C⇩2 ≠ {} ∧ L⇩1 ≠ {} ∧ L⇩2 ≠ {}
∧ vars⇩l⇩s C⇩1 ∩ vars⇩l⇩s C⇩2 = {}
∧ L⇩1 ⊆ C⇩1 ∧ L⇩2 ⊆ C⇩2
∧ mgu⇩l⇩s σ (L⇩1 ∪ L⇩2⇧C)"
definition mresolution :: " fterm clause ⇒ fterm clause
⇒ fterm literal set ⇒ fterm literal set
⇒ substitution ⇒ fterm clause" where
"mresolution C⇩1 C⇩2 L⇩1 L⇩2 σ = ((C⇩1 ⋅⇩l⇩s σ)- (L⇩1 ⋅⇩l⇩s σ)) ∪ ((C⇩2 ⋅⇩l⇩s σ) - (L⇩2 ⋅⇩l⇩s σ))"
definition resolution :: " fterm clause ⇒ fterm clause
⇒ fterm literal set ⇒ fterm literal set
⇒ substitution ⇒ fterm clause" where
"resolution C⇩1 C⇩2 L⇩1 L⇩2 σ = ((C⇩1 - L⇩1) ∪ (C⇩2 - L⇩2)) ⋅⇩l⇩s σ"
inductive mresolution_step :: "fterm clause set ⇒ fterm clause set ⇒ bool" where
mresolution_rule:
"C⇩1 ∈ Cs ⟹ C⇩2 ∈ Cs ⟹ applicable C⇩1 C⇩2 L⇩1 L⇩2 σ ⟹
mresolution_step Cs (Cs ∪ {mresolution C⇩1 C⇩2 L⇩1 L⇩2 σ})"
| standardize_apart:
"C ∈ Cs ⟹ var_renaming_of C C' ⟹ mresolution_step Cs (Cs ∪ {C'})"
inductive resolution_step :: "fterm clause set ⇒ fterm clause set ⇒ bool" where
resolution_rule:
"C⇩1 ∈ Cs ⟹ C⇩2 ∈ Cs ⟹ applicable C⇩1 C⇩2 L⇩1 L⇩2 σ ⟹
resolution_step Cs (Cs ∪ {resolution C⇩1 C⇩2 L⇩1 L⇩2 σ})"
| standardize_apart:
"C ∈ Cs ⟹ var_renaming_of C C' ⟹ resolution_step Cs (Cs ∪ {C'})"
definition mresolution_deriv :: "fterm clause set ⇒ fterm clause set ⇒ bool" where
"mresolution_deriv = rtranclp mresolution_step"
definition resolution_deriv :: "fterm clause set ⇒ fterm clause set ⇒ bool" where
"resolution_deriv = rtranclp resolution_step"
section ‹Soundness›
definition evalsub :: "'u var_denot ⇒ 'u fun_denot ⇒ substitution ⇒ 'u var_denot" where
"evalsub E F σ = eval⇩t E F ∘ σ"
lemma substitutiont: "eval⇩t E F (t ⋅⇩t σ) = eval⇩t (evalsub E F σ) F t"
apply (induction t)
unfolding evalsub_def apply auto
apply (metis (mono_tags, lifting) comp_apply map_cong)
done
lemma substitutionts: "eval⇩t⇩s E F (ts ⋅⇩t⇩s σ) = eval⇩t⇩s (evalsub E F σ) F ts"
using substitutiont by auto
lemma substitution: "eval⇩l E F G (l ⋅⇩l σ) ⟷ eval⇩l (evalsub E F σ) F G l"
apply (induction l)
using substitutionts apply (metis eval⇩l.simps(1) subl.simps(1))
using substitutionts apply (metis eval⇩l.simps(2) subl.simps(2))
done
lemma subst_sound:
assumes asm: "eval⇩c F G C"
shows "eval⇩c F G (C ⋅⇩l⇩s σ)"
unfolding eval⇩c_def proof
fix E
from asm have "∀E'. ∃l ∈ C. eval⇩l E' F G l" using eval⇩c_def by blast
then have "∃l ∈ C. eval⇩l (evalsub E F σ) F G l" by auto
then show "∃l ∈ C ⋅⇩l⇩s σ. eval⇩l E F G l" using substitution by blast
qed
lemma simple_resolution_sound:
assumes C⇩1sat: "eval⇩c F G C⇩1"
assumes C⇩2sat: "eval⇩c F G C⇩2"
assumes l⇩1inc⇩1: "l⇩1 ∈ C⇩1"
assumes l⇩2inc⇩2: "l⇩2 ∈ C⇩2"
assumes comp: "l⇩1⇧c = l⇩2"
shows "eval⇩c F G ((C⇩1 - {l⇩1}) ∪ (C⇩2 - {l⇩2}))"
proof -
have "∀E. ∃l ∈ (((C⇩1 - {l⇩1}) ∪ (C⇩2 - {l⇩2}))). eval⇩l E F G l"
proof
fix E
have "eval⇩l E F G l⇩1 ∨ eval⇩l E F G l⇩2" using comp by (cases l⇩1) auto
then show "∃l ∈ (((C⇩1 - {l⇩1}) ∪ (C⇩2 - {l⇩2}))). eval⇩l E F G l"
proof
assume "eval⇩l E F G l⇩1"
then have "¬eval⇩l E F G l⇩2" using comp by (cases l⇩1) auto
then have "∃l⇩2'∈ C⇩2. l⇩2' ≠ l⇩2 ∧ eval⇩l E F G l⇩2'" using l⇩2inc⇩2 C⇩2sat unfolding eval⇩c_def by auto
then show "∃l∈(C⇩1 - {l⇩1}) ∪ (C⇩2 - {l⇩2}). eval⇩l E F G l" by auto
next
assume "eval⇩l E F G l⇩2"
then have "¬eval⇩l E F G l⇩1" using comp by (cases l⇩1) auto
then have "∃l⇩1'∈ C⇩1. l⇩1' ≠ l⇩1 ∧ eval⇩l E F G l⇩1'" using l⇩1inc⇩1 C⇩1sat unfolding eval⇩c_def by auto
then show "∃l∈(C⇩1 - {l⇩1}) ∪ (C⇩2 - {l⇩2}). eval⇩l E F G l" by auto
qed
qed
then show ?thesis unfolding eval⇩c_def by simp
qed
lemma mresolution_sound:
assumes sat⇩1: "eval⇩c F G C⇩1"
assumes sat⇩2: "eval⇩c F G C⇩2"
assumes appl: "applicable C⇩1 C⇩2 L⇩1 L⇩2 σ"
shows "eval⇩c F G (mresolution C⇩1 C⇩2 L⇩1 L⇩2 σ)"
proof -
from sat⇩1 have sat⇩1σ: "eval⇩c F G (C⇩1 ⋅⇩l⇩s σ)" using subst_sound by blast
from sat⇩2 have sat⇩2σ: "eval⇩c F G (C⇩2 ⋅⇩l⇩s σ)" using subst_sound by blast
from appl obtain l⇩1 where l⇩1_p: "l⇩1 ∈ L⇩1" unfolding applicable_def by auto
from l⇩1_p appl have "l⇩1 ∈ C⇩1" unfolding applicable_def by auto
then have inc⇩1σ: "l⇩1 ⋅⇩l σ ∈ C⇩1 ⋅⇩l⇩s σ" by auto
from l⇩1_p have unified⇩1: "l⇩1 ∈ (L⇩1 ∪ (L⇩2⇧C))" by auto
from l⇩1_p appl have l⇩1σisl⇩1σ: "{l⇩1 ⋅⇩l σ} = L⇩1 ⋅⇩l⇩s σ"
unfolding mgu⇩l⇩s_def unifier⇩l⇩s_def applicable_def by auto
from appl obtain l⇩2 where l⇩2_p: "l⇩2 ∈ L⇩2" unfolding applicable_def by auto
from l⇩2_p appl have "l⇩2 ∈ C⇩2" unfolding applicable_def by auto
then have inc⇩2σ: "l⇩2 ⋅⇩l σ ∈ C⇩2 ⋅⇩l⇩s σ" by auto
from l⇩2_p have unified⇩2: "l⇩2⇧c ∈ (L⇩1 ∪ (L⇩2⇧C))" by auto
from unified⇩1 unified⇩2 appl have "l⇩1 ⋅⇩l σ = (l⇩2⇧c) ⋅⇩l σ"
unfolding mgu⇩l⇩s_def unifier⇩l⇩s_def applicable_def by auto
then have comp: "(l⇩1 ⋅⇩l σ)⇧c = l⇩2 ⋅⇩l σ" using comp_sub comp_swap by auto
from appl have "unifier⇩l⇩s σ (L⇩2⇧C)"
using unifier_sub2 unfolding mgu⇩l⇩s_def applicable_def by blast
then have "unifier⇩l⇩s σ L⇩2" by auto
from this l⇩2_p have l⇩2σisl⇩2σ: "{l⇩2 ⋅⇩l σ} = L⇩2 ⋅⇩l⇩s σ" unfolding unifier⇩l⇩s_def by auto
from sat⇩1σ sat⇩2σ inc⇩1σ inc⇩2σ comp have "eval⇩c F G ((C⇩1 ⋅⇩l⇩s σ) - {l⇩1 ⋅⇩l σ} ∪ ((C⇩2 ⋅⇩l⇩s σ) - {l⇩2 ⋅⇩l σ}))" using simple_resolution_sound[of F G "C⇩1 ⋅⇩l⇩s σ" "C⇩2 ⋅⇩l⇩s σ" "l⇩1 ⋅⇩l σ" " l⇩2 ⋅⇩l σ"]
by auto
from this l⇩1σisl⇩1σ l⇩2σisl⇩2σ show ?thesis unfolding mresolution_def by auto
qed
lemma resolution_superset: "mresolution C⇩1 C⇩2 L⇩1 L⇩2 σ ⊆ resolution C⇩1 C⇩2 L⇩1 L⇩2 σ"
unfolding mresolution_def resolution_def by auto
lemma superset_sound:
assumes sup: "C ⊆ C'"
assumes sat: "eval⇩c F G C"
shows "eval⇩c F G C'"
proof -
have "∀E. ∃l ∈ C'. eval⇩l E F G l"
proof
fix E
from sat have "∀E. ∃l ∈ C. eval⇩l E F G l" unfolding eval⇩c_def by -
then have "∃l ∈ C . eval⇩l E F G l" by auto
then show "∃l ∈ C'. eval⇩l E F G l" using sup by auto
qed
then show "eval⇩c F G C'" unfolding eval⇩c_def by auto
qed
theorem resolution_sound:
assumes sat⇩1: "eval⇩c F G C⇩1"
assumes sat⇩2: "eval⇩c F G C⇩2"
assumes appl: "applicable C⇩1 C⇩2 L⇩1 L⇩2 σ"
shows "eval⇩c F G (resolution C⇩1 C⇩2 L⇩1 L⇩2 σ)"
proof -
from sat⇩1 sat⇩2 appl have "eval⇩c F G (mresolution C⇩1 C⇩2 L⇩1 L⇩2 σ)" using mresolution_sound by blast
then show ?thesis using superset_sound resolution_superset by metis
qed
lemma mstep_sound:
assumes "mresolution_step Cs Cs'"
assumes "eval⇩c⇩s F G Cs"
shows "eval⇩c⇩s F G Cs'"
using assms proof (induction rule: mresolution_step.induct)
case (mresolution_rule C⇩1 Cs C⇩2 l⇩1 l⇩2 σ)
then have "eval⇩c F G C⇩1 ∧ eval⇩c F G C⇩2" unfolding eval⇩c⇩s_def by auto
then have "eval⇩c F G (mresolution C⇩1 C⇩2 l⇩1 l⇩2 σ)"
using mresolution_sound mresolution_rule by auto
then show ?case using mresolution_rule unfolding eval⇩c⇩s_def by auto
next
case (standardize_apart C Cs C')
then have "eval⇩c F G C" unfolding eval⇩c⇩s_def by auto
then have "eval⇩c F G C'" using subst_sound standardize_apart unfolding var_renaming_of_def instance_of⇩l⇩s_def by metis
then show ?case using standardize_apart unfolding eval⇩c⇩s_def by auto
qed
theorem step_sound:
assumes "resolution_step Cs Cs'"
assumes "eval⇩c⇩s F G Cs"
shows "eval⇩c⇩s F G Cs'"
using assms proof (induction rule: resolution_step.induct)
case (resolution_rule C⇩1 Cs C⇩2 l⇩1 l⇩2 σ)
then have "eval⇩c F G C⇩1 ∧ eval⇩c F G C⇩2" unfolding eval⇩c⇩s_def by auto
then have "eval⇩c F G (resolution C⇩1 C⇩2 l⇩1 l⇩2 σ)"
using resolution_sound resolution_rule by auto
then show ?case using resolution_rule unfolding eval⇩c⇩s_def by auto
next
case (standardize_apart C Cs C')
then have "eval⇩c F G C" unfolding eval⇩c⇩s_def by auto
then have "eval⇩c F G C'" using subst_sound standardize_apart unfolding var_renaming_of_def instance_of⇩l⇩s_def by metis
then show ?case using standardize_apart unfolding eval⇩c⇩s_def by auto
qed
lemma mderivation_sound:
assumes "mresolution_deriv Cs Cs'"
assumes "eval⇩c⇩s F G Cs"
shows "eval⇩c⇩s F G Cs'"
using assms unfolding mresolution_deriv_def
proof (induction rule: rtranclp.induct)
case rtrancl_refl then show ?case by auto
next
case (rtrancl_into_rtrancl Cs⇩1 Cs⇩2 Cs⇩3) then show ?case using mstep_sound by auto
qed
theorem derivation_sound:
assumes "resolution_deriv Cs Cs'"
assumes "eval⇩c⇩s F G Cs"
shows"eval⇩c⇩s F G Cs'"
using assms unfolding resolution_deriv_def
proof (induction rule: rtranclp.induct)
case rtrancl_refl then show ?case by auto
next
case (rtrancl_into_rtrancl Cs⇩1 Cs⇩2 Cs⇩3) then show ?case using step_sound by auto
qed
theorem derivation_sound_refute:
assumes deriv: "resolution_deriv Cs Cs' ∧ {} ∈ Cs'"
shows "¬eval⇩c⇩s F G Cs"
proof -
from deriv have "eval⇩c⇩s F G Cs ⟶ eval⇩c⇩s F G Cs'" using derivation_sound by auto
moreover
from deriv have "eval⇩c⇩s F G Cs' ⟶ eval⇩c F G {}" unfolding eval⇩c⇩s_def by auto
moreover
then have "eval⇩c F G {} ⟶ False" unfolding eval⇩c_def by auto
ultimately show ?thesis by auto
qed
section ‹Herbrand Interpretations›
text ‹@{const HFun} is the Herbrand function denotation in which terms are mapped to themselves.›
term HFun
lemma eval_ground⇩t:
assumes "ground⇩t t"
shows "(eval⇩t E HFun t) = hterm_of_fterm t"
using assms by (induction t) auto
lemma eval_ground⇩t⇩s:
assumes "ground⇩t⇩s ts"
shows "(eval⇩t⇩s E HFun ts) = hterms_of_fterms ts"
unfolding hterms_of_fterms_def using assms eval_ground⇩t by (induction ts) auto
lemma eval⇩l_ground⇩t⇩s:
assumes asm: "ground⇩t⇩s ts"
shows "eval⇩l E HFun G (Pos P ts) ⟷ G P (hterms_of_fterms ts)"
proof -
have "eval⇩l E HFun G (Pos P ts) = G P (eval⇩t⇩s E HFun ts)" by auto
also have "... = G P (hterms_of_fterms ts)" using asm eval_ground⇩t⇩s by simp
finally show ?thesis by auto
qed
section ‹Partial Interpretations›
type_synonym partial_pred_denot = "bool list"
definition falsifies⇩l :: "partial_pred_denot ⇒ fterm literal ⇒ bool" where
"falsifies⇩l G l ⟷
ground⇩l l
∧ (let i = nat_of_fatom (get_atom l) in
i < length G ∧ G ! i = (¬sign l)
)"
text ‹A ground clause is falsified if it is actually ground and all its literals are falsified.›
abbreviation falsifies⇩g :: "partial_pred_denot ⇒ fterm clause ⇒ bool" where
"falsifies⇩g G C ≡ ground⇩l⇩s C ∧ (∀l ∈ C. falsifies⇩l G l)"
abbreviation falsifies⇩c :: "partial_pred_denot ⇒ fterm clause ⇒ bool" where
"falsifies⇩c G C ≡ (∃C'. instance_of⇩l⇩s C' C ∧ falsifies⇩g G C')"
abbreviation falsifies⇩c⇩s :: "partial_pred_denot ⇒ fterm clause set ⇒ bool" where
"falsifies⇩c⇩s G Cs ≡ (∃C ∈ Cs. falsifies⇩c G C)"
abbreviation extend :: "(nat ⇒ partial_pred_denot) ⇒ hterm pred_denot" where
"extend f P ts ≡ (
let n = nat_of_hatom (P, ts) in
f (Suc n) ! n
)"
fun sub_of_denot :: "hterm var_denot ⇒ substitution" where
"sub_of_denot E = fterm_of_hterm ∘ E"
lemma ground_sub_of_denott: "ground⇩t (t ⋅⇩t (sub_of_denot E))"
by (induction t) (auto simp add: ground_fterm_of_hterm)
lemma ground_sub_of_denotts: "ground⇩t⇩s (ts ⋅⇩t⇩s sub_of_denot E)"
using ground_sub_of_denott by simp
lemma ground_sub_of_denotl: "ground⇩l (l ⋅⇩l sub_of_denot E)"
proof -
have "ground⇩t⇩s (subs (get_terms l) (sub_of_denot E))"
using ground_sub_of_denotts by auto
then show ?thesis by (cases l) auto
qed
lemma sub_of_denot_equivx: "eval⇩t E HFun (sub_of_denot E x) = E x"
proof -
have "ground⇩t (sub_of_denot E x)" using ground_fterm_of_hterm by simp
then
have "eval⇩t E HFun (sub_of_denot E x) = hterm_of_fterm (sub_of_denot E x)"
using eval_ground⇩t(1) by auto
also have "... = hterm_of_fterm (fterm_of_hterm (E x))" by auto
also have "... = E x" by auto
finally show ?thesis by auto
qed
lemma sub_of_denot_equivt:
"eval⇩t E HFun (t ⋅⇩t (sub_of_denot E)) = eval⇩t E HFun t"
using sub_of_denot_equivx by (induction t) auto
lemma sub_of_denot_equivts: "eval⇩t⇩s E HFun (ts ⋅⇩t⇩s (sub_of_denot E)) = eval⇩t⇩s E HFun ts"
using sub_of_denot_equivt by simp
lemma sub_of_denot_equivl: "eval⇩l E HFun G (l ⋅⇩l sub_of_denot E) ⟷ eval⇩l E HFun G l"
proof (induction l)
case (Pos p ts)
have "eval⇩l E HFun G ((Pos p ts) ⋅⇩l sub_of_denot E) ⟷ G p (eval⇩t⇩s E HFun (ts ⋅⇩t⇩s (sub_of_denot E)))" by auto
also have " ... ⟷ G p (eval⇩t⇩s E HFun ts)" using sub_of_denot_equivts[of E ts] by metis
also have " ... ⟷ eval⇩l E HFun G (Pos p ts)" by simp
finally
show ?case by blast
next
case (Neg p ts)
have "eval⇩l E HFun G ((Neg p ts) ⋅⇩l sub_of_denot E) ⟷ ¬G p (eval⇩t⇩s E HFun (ts ⋅⇩t⇩s (sub_of_denot E)))" by auto
also have " ... ⟷ ¬G p (eval⇩t⇩s E HFun ts)" using sub_of_denot_equivts[of E ts] by metis
also have " ... = eval⇩l E HFun G (Neg p ts)" by simp
finally
show ?case by blast
qed
text ‹Under an Herbrand interpretation, an environment is equivalent to a substitution.›
lemma sub_of_denot_equiv_ground':
"eval⇩l E HFun G l = eval⇩l E HFun G (l ⋅⇩l sub_of_denot E) ∧ ground⇩l (l ⋅⇩l sub_of_denot E)"
using sub_of_denot_equivl ground_sub_of_denotl by auto
text ‹Under an Herbrand interpretation, an environment is similar to a substitution - also for partial interpretations.›
lemma partial_equiv_subst:
assumes "falsifies⇩c G (C ⋅⇩l⇩s τ)"
shows "falsifies⇩c G C"
proof -
from assms obtain C' where C'_p: "instance_of⇩l⇩s C' (C ⋅⇩l⇩s τ) ∧ falsifies⇩g G C'" by auto
then have "instance_of⇩l⇩s (C ⋅⇩l⇩s τ) C" unfolding instance_of⇩l⇩s_def by auto
then have "instance_of⇩l⇩s C' C" using C'_p instance_of⇩l⇩s_trans by auto
then show ?thesis using C'_p by auto
qed
text ‹Under an Herbrand interpretation, an environment is equivalent to a substitution.›
lemma sub_of_denot_equiv_ground:
"((∃l ∈ C. eval⇩l E HFun G l) ⟷ (∃l ∈ C ⋅⇩l⇩s sub_of_denot E. eval⇩l E HFun G l))
∧ ground⇩l⇩s (C ⋅⇩l⇩s sub_of_denot E)"
using sub_of_denot_equiv_ground' by auto
lemma std⇩1_falsifies: "falsifies⇩c G C⇩1 ⟷ falsifies⇩c G (std⇩1 C⇩1)"
proof
assume asm: "falsifies⇩c G C⇩1"
then obtain Cg where "instance_of⇩l⇩s Cg C⇩1 ∧ falsifies⇩g G Cg" by auto
moreover
then have "instance_of⇩l⇩s Cg (std⇩1 C⇩1)" using std_apart_instance_of⇩l⇩s1 instance_of⇩l⇩s_trans by blast
ultimately
show "falsifies⇩c G (std⇩1 C⇩1)" by auto
next
assume asm: "falsifies⇩c G (std⇩1 C⇩1)"
then have inst: "instance_of⇩l⇩s (std⇩1 C⇩1) C⇩1" unfolding instance_of⇩l⇩s_def by auto
from asm obtain Cg where "instance_of⇩l⇩s Cg (std⇩1 C⇩1) ∧ falsifies⇩g G Cg" by auto
moreover
then have "instance_of⇩l⇩s Cg C⇩1" using inst instance_of⇩l⇩s_trans by blast
ultimately
show "falsifies⇩c G C⇩1" by auto
qed
lemma std⇩2_falsifies: "falsifies⇩c G C⇩2 ⟷ falsifies⇩c G (std⇩2 C⇩2)"
proof
assume asm: "falsifies⇩c G C⇩2"
then obtain Cg where "instance_of⇩l⇩s Cg C⇩2 ∧ falsifies⇩g G Cg" by auto
moreover
then have "instance_of⇩l⇩s Cg (std⇩2 C⇩2)" using std_apart_instance_of⇩l⇩s2 instance_of⇩l⇩s_trans by blast
ultimately
show "falsifies⇩c G (std⇩2 C⇩2)" by auto
next
assume asm: "falsifies⇩c G (std⇩2 C⇩2)"
then have inst: "instance_of⇩l⇩s (std⇩2 C⇩2) C⇩2" unfolding instance_of⇩l⇩s_def by auto
from asm obtain Cg where "instance_of⇩l⇩s Cg (std⇩2 C⇩2) ∧ falsifies⇩g G Cg" by auto
moreover
then have "instance_of⇩l⇩s Cg C⇩2" using inst instance_of⇩l⇩s_trans by blast
ultimately
show "falsifies⇩c G C⇩2" by auto
qed
lemma std⇩1_renames: "var_renaming_of C⇩1 (std⇩1 C⇩1)"
proof -
have "instance_of⇩l⇩s C⇩1 (std⇩1 C⇩1)" using std_apart_instance_of⇩l⇩s1 by auto
moreover have "instance_of⇩l⇩s (std⇩1 C⇩1) C⇩1" unfolding instance_of⇩l⇩s_def by auto
ultimately show "var_renaming_of C⇩1 (std⇩1 C⇩1)" unfolding var_renaming_of_def by auto
qed
lemma std⇩2_renames: "var_renaming_of C⇩2 (std⇩2 C⇩2)"
proof -
have "instance_of⇩l⇩s C⇩2 (std⇩2 C⇩2)" using std_apart_instance_of⇩l⇩s2 by auto
moreover have "instance_of⇩l⇩s (std⇩2 C⇩2) C⇩2" unfolding instance_of⇩l⇩s_def by auto
ultimately show "var_renaming_of C⇩2 (std⇩2 C⇩2)" unfolding var_renaming_of_def by auto
qed
section ‹Semantic Trees›
abbreviation closed_branch :: "partial_pred_denot ⇒ tree ⇒ fterm clause set ⇒ bool" where
"closed_branch G T Cs ≡ branch G T ∧ falsifies⇩c⇩s G Cs"
abbreviation(input) open_branch :: "partial_pred_denot ⇒ tree ⇒ fterm clause set ⇒ bool" where
"open_branch G T Cs ≡ branch G T ∧ ¬falsifies⇩c⇩s G Cs"
definition closed_tree :: "tree ⇒ fterm clause set ⇒ bool" where
"closed_tree T Cs ⟷ anybranch T (λb. closed_branch b T Cs)
∧ anyinternal T (λp. ¬falsifies⇩c⇩s p Cs)"
section ‹Herbrand's Theorem›
lemma maximum:
assumes asm: "finite C"
shows "∃n :: nat. ∀l ∈ C. f l ≤ n"
proof
from asm show "∀l∈C. f l ≤ (Max (f ` C))" by auto
qed
lemma extend_preserves_model:
assumes f_infpath: "wf_infpath (f :: nat ⇒ partial_pred_denot)"
assumes C_ground: "ground⇩l⇩s C"
assumes C_sat: "¬falsifies⇩c (f (Suc n)) C"
assumes n_max: "∀l∈C. nat_of_fatom (get_atom l) ≤ n"
shows "eval⇩c HFun (extend f) C"
proof -
let ?F = "HFun"
let ?G = "extend f"
{
fix E
from C_sat have "∀C'. (¬instance_of⇩l⇩s C' C ∨ ¬falsifies⇩g (f (Suc n)) C')" by auto
then have "¬falsifies⇩g (f (Suc n)) C" using instance_of⇩l⇩s_self by auto
then obtain l where l_p: "l∈C ∧ ¬falsifies⇩l (f (Suc n)) l" using C_ground by blast
let ?i = "nat_of_fatom (get_atom l)"
from l_p have i_n: "?i ≤ n" using n_max by auto
then have j_n: "?i < length (f (Suc n))" using f_infpath infpath_length[of f] by auto
have "eval⇩l E HFun (extend f) l"
proof (cases l)
case (Pos P ts)
from Pos l_p C_ground have ts_ground: "ground⇩t⇩s ts" by auto
have "¬falsifies⇩l (f (Suc n)) l" using l_p by auto
then have "f (Suc n) ! ?i = True"
using j_n Pos ts_ground empty_subts[of ts] unfolding falsifies⇩l_def by auto
moreover have "f (Suc ?i) ! ?i = f (Suc n) ! ?i"
using f_infpath i_n j_n infpath_length[of f] ith_in_extension[of f] by simp
ultimately
have "f (Suc ?i) ! ?i = True" using Pos by auto
then have "?G P (hterms_of_fterms ts)" using Pos by (simp add: nat_of_fatom_def)
then show ?thesis using eval⇩l_ground⇩t⇩s[of ts _ ?G P] ts_ground Pos by auto
next
case (Neg P ts)
from Neg l_p C_ground have ts_ground: "ground⇩t⇩s ts" by auto
have "¬falsifies⇩l (f (Suc n)) l" using l_p by auto
then have "f (Suc n) ! ?i = False"
using j_n Neg ts_ground empty_subts[of ts] unfolding falsifies⇩l_def by auto
moreover have "f (Suc ?i) ! ?i = f (Suc n) ! ?i"
using f_infpath i_n j_n infpath_length[of f] ith_in_extension[of f] by simp
ultimately
have "f (Suc ?i) ! ?i = False" using Neg by auto
then have "¬?G P (hterms_of_fterms ts)" using Neg by (simp add: nat_of_fatom_def)
then show ?thesis using Neg eval⇩l_ground⇩t⇩s[of ts _ ?G P] ts_ground by auto
qed
then have "∃l ∈ C. eval⇩l E HFun (extend f) l" using l_p by auto
}
then have "eval⇩c HFun (extend f) C" unfolding eval⇩c_def by auto
then show ?thesis using instance_of⇩l⇩s_self by auto
qed
lemma extend_preserves_model2:
assumes f_infpath: "wf_infpath (f :: nat ⇒ partial_pred_denot)"
assumes C_ground: "ground⇩l⇩s C"
assumes fin_c: "finite C"
assumes model_C: "∀n. ¬falsifies⇩c (f n) C"
shows C_false: "eval⇩c HFun (extend f) C"
proof -
obtain n where largest: "∀l ∈ C. nat_of_fatom (get_atom l) ≤ n" using fin_c maximum[of C "λl. nat_of_fatom (get_atom l)"] by blast
moreover
then have "¬falsifies⇩c (f (Suc n)) C" using model_C by auto
ultimately show ?thesis using model_C f_infpath C_ground extend_preserves_model[of f C n ] by blast
qed
lemma extend_infpath:
assumes f_infpath: "wf_infpath (f :: nat ⇒ partial_pred_denot)"
assumes model_c: "∀n. ¬falsifies⇩c (f n) C"
assumes fin_c: "finite C"
shows "eval⇩c HFun (extend f) C"
unfolding eval⇩c_def proof
fix E
let ?G = "extend f"
let ?σ = "sub_of_denot E"
from fin_c have fin_cσ: "finite (C ⋅⇩l⇩s sub_of_denot E)" by auto
have groundcσ: "ground⇩l⇩s (C ⋅⇩l⇩s sub_of_denot E)" using sub_of_denot_equiv_ground by auto
from model_c have "∀n. ¬falsifies⇩c (f n) (C ⋅⇩l⇩s ?σ)" using partial_equiv_subst by blast
then have "eval⇩c HFun ?G (C ⋅⇩l⇩s ?σ)" using groundcσ f_infpath fin_cσ extend_preserves_model2[of f "C ⋅⇩l⇩s ?σ"] by blast
then have "∀E. ∃l ∈ (C ⋅⇩l⇩s ?σ). eval⇩l E HFun ?G l" unfolding eval⇩c_def by auto
then have "∃l ∈ (C ⋅⇩l⇩s ?σ). eval⇩l E HFun ?G l" by auto
then show "∃l ∈ C. eval⇩l E HFun ?G l" using sub_of_denot_equiv_ground[of C E "extend f"] by blast
qed
text ‹If we have a infpath of partial models, then we have a model.›
lemma infpath_model:
assumes f_infpath: "wf_infpath (f :: nat ⇒ partial_pred_denot)"
assumes model_cs: "∀n. ¬falsifies⇩c⇩s (f n) Cs"
assumes fin_cs: "finite Cs"
assumes fin_c: "∀C ∈ Cs. finite C"
shows "eval⇩c⇩s HFun (extend f) Cs"
proof -
let ?F = "HFun"
have "∀C ∈ Cs. eval⇩c ?F (extend f) C"
proof (rule ballI)
fix C
assume asm: "C ∈ Cs"
then have "∀n. ¬falsifies⇩c (f n) C" using model_cs by auto
then show "eval⇩c ?F (extend f) C" using fin_c asm f_infpath extend_infpath[of f C] by auto
qed
then show "eval⇩c⇩s ?F (extend f) Cs" unfolding eval⇩c⇩s_def by auto
qed
fun deeptree :: "nat ⇒ tree" where
"deeptree 0 = Leaf"
| "deeptree (Suc n) = Branching (deeptree n) (deeptree n)"
lemma branch_length:
assumes "branch b (deeptree n)"
shows "length b = n"
using assms proof (induction n arbitrary: b)
case 0 then show ?case using branch_inv_Leaf by auto
next
case (Suc n)
then have "branch b (Branching (deeptree n) (deeptree n))" by auto
then obtain a b' where p: "b = a#b'∧ branch b' (deeptree n)" using branch_inv_Branching[of b] by blast
then have "length b' = n" using Suc by auto
then show ?case using p by auto
qed
lemma infinity:
assumes inj: "∀n :: nat. undiago (diago n) = n"
assumes all_tree: "∀n :: nat. (diago n) ∈ tree"
shows "¬finite tree"
proof -
from inj all_tree have "∀n. n = undiago (diago n) ∧ (diago n) ∈ tree" by auto
then have "∀n. ∃ds. n = undiago ds ∧ ds ∈ tree" by auto
then have "undiago ` tree = (UNIV :: nat set)" by auto
then have "¬finite tree"by (metis finite_imageI infinite_UNIV_nat)
then show ?thesis by auto
qed
lemma longer_falsifies⇩l:
assumes "falsifies⇩l ds l"
shows "falsifies⇩l (ds@d) l"
proof -
let ?i = "nat_of_fatom (get_atom l)"
from assms have i_p: "ground⇩l l ∧ ?i < length ds ∧ ds ! ?i = (¬sign l)" unfolding falsifies⇩l_def by meson
moreover
from i_p have "?i < length (ds@d)" by auto
moreover
from i_p have "(ds@d) ! ?i = (¬sign l)" by (simp add: nth_append)
ultimately
show ?thesis unfolding falsifies⇩l_def by simp
qed
lemma longer_falsifies⇩g:
assumes "falsifies⇩g ds C"
shows "falsifies⇩g (ds @ d) C"
proof -
{
fix l
assume "l∈C"
then have "falsifies⇩l (ds @ d) l" using assms longer_falsifies⇩l by auto
} then show ?thesis using assms by auto
qed
lemma longer_falsifies⇩c:
assumes "falsifies⇩c ds C"
shows "falsifies⇩c (ds @ d) C"
proof -
from assms obtain C' where "instance_of⇩l⇩s C' C ∧ falsifies⇩g ds C'" by auto
moreover
then have "falsifies⇩g (ds @ d) C'" using longer_falsifies⇩g by auto
ultimately show ?thesis by auto
qed
text ‹We use this so that we can apply König's lemma.›
lemma longer_falsifies:
assumes "falsifies⇩c⇩s ds Cs"
shows "falsifies⇩c⇩s (ds @ d) Cs"
proof -
from assms obtain C where "C ∈ Cs ∧ falsifies⇩c ds C" by auto
moreover
then have "falsifies⇩c (ds @ d) C" using longer_falsifies⇩c[of C ds d] by blast
ultimately
show ?thesis by auto
qed
text ‹If all finite semantic trees have an open branch, then the set of clauses has a model.›
theorem herbrand':
assumes openb: "∀T. ∃G. open_branch G T Cs"
assumes finite_cs: "finite Cs" "∀C∈Cs. finite C"
shows "∃G. eval⇩c⇩s HFun G Cs"
proof -
let ?tree = "{G. ¬falsifies⇩c⇩s G Cs}"
let ?undiag = length
let ?diag = "(λl. SOME b. open_branch b (deeptree l) Cs) :: nat ⇒ partial_pred_denot"
from openb have diag_open: "∀l. open_branch (?diag l) (deeptree l) Cs"
using someI_ex[of "λb. open_branch b (deeptree _) Cs"] by auto
then have "∀n. ?undiag (?diag n) = n" using branch_length by auto
moreover
have "∀n. (?diag n) ∈ ?tree" using diag_open by auto
ultimately
have "¬finite ?tree" using infinity[of _ "λn. SOME b. open_branch b (_ n) Cs"] by simp
moreover
have "∀ds d. ¬falsifies⇩c⇩s (ds @ d) Cs ⟶ ¬falsifies⇩c⇩s ds Cs"
using longer_falsifies[of Cs] by blast
then have "(∀ds d. ds @ d ∈ ?tree ⟶ ds ∈ ?tree)" by auto
ultimately
have "∃c. wf_infpath c ∧ (∀n. c n ∈ ?tree)" using konig[of ?tree] by blast
then have "∃G. wf_infpath G ∧ (∀n. ¬ falsifies⇩c⇩s (G n) Cs)" by auto
then show "∃G. eval⇩c⇩s HFun G Cs" using infpath_model finite_cs by auto
qed
lemma shorter_falsifies⇩l:
assumes "falsifies⇩l (ds@d) l"
assumes "nat_of_fatom (get_atom l) < length ds"
shows "falsifies⇩l ds l"
proof -
let ?i = "nat_of_fatom (get_atom l)"
from assms have i_p: "ground⇩l l ∧ ?i < length (ds@d) ∧ (ds@d) ! ?i = (¬sign l)" unfolding falsifies⇩l_def by meson
moreover
then have "?i < length ds" using assms by auto
moreover
then have "ds ! ?i = (¬sign l)" using i_p nth_append[of ds d ?i] by auto
ultimately show ?thesis using assms unfolding falsifies⇩l_def by simp
qed
theorem herbrand'_contra:
assumes finite_cs: "finite Cs" "∀C∈Cs. finite C"
assumes unsat: "∀G. ¬eval⇩c⇩s HFun G Cs"
shows "∃T. ∀G. branch G T ⟶ closed_branch G T Cs"
proof -
from finite_cs unsat have "(∀T. ∃G. open_branch G T Cs) ⟶ (∃G. eval⇩c⇩s HFun G Cs)" using herbrand' by blast
then show ?thesis using unsat by blast
qed
theorem herbrand:
assumes unsat: "∀G. ¬eval⇩c⇩s HFun G Cs"
assumes finite_cs: "finite Cs" "∀C∈Cs. finite C"
shows "∃T. closed_tree T Cs"
proof -
from unsat finite_cs obtain T where "anybranch T (λb. closed_branch b T Cs)" using herbrand'_contra[of Cs] by blast
then have "∃T. anybranch T (λp. falsifies⇩c⇩s p Cs) ∧ anyinternal T (λp. ¬ falsifies⇩c⇩s p Cs)"
using cutoff_branch_internal[of T "λp. falsifies⇩c⇩s p Cs"] by blast
then show ?thesis unfolding closed_tree_def by auto
qed
end