Theory Simulation_SCLFOL_ORDRES
theory Simulation_SCLFOL_ORDRES
imports
Background
ORD_RES
ORD_RES_1
ORD_RES_2
ORD_RES_3
ORD_RES_4
ORD_RES_5
ORD_RES_6
ORD_RES_7
ORD_RES_8
ORD_RES_9
ORD_RES_10
ORD_RES_11
Clause_Could_Propagate
begin
section ‹ORD-RES-1 (deterministic)›
type_synonym 'f ord_res_1_state = "'f gclause fset"
context simulation_SCLFOL_ground_ordered_resolution begin
sublocale backward_simulation_with_measuring_function where
step1 = ord_res and
step2 = ord_res_1 and
final1 = ord_res_final and
final2 = ord_res_1_final and
order = "λ_ _. False" and
match = "(=)" and
measure = "λ_. ()"
proof unfold_locales
show "wfP (λ_ _. False)"
by simp
next
show "⋀N1 N2. N1 = N2 ⟹ ord_res_1_final N2 ⟹ ord_res_final N1"
unfolding ord_res_1_final_def by metis
next
fix N1 N2 N2' :: "'f ord_res_1_state"
assume match: "N1 = N2" and step2: "ord_res_1 N2 N2'"
show "(∃N1'. ord_res⇧+⇧+ N1 N1' ∧ N1' = N2') ∨ N1 = N2' ∧ False"
proof (intro disjI1 exI conjI)
have mempty_no_in: "{#} |∉| N2"
if C_least: "linorder_cls.is_least_in_fset {|C |∈| N2.
¬ ord_res.interp (fset N2) C ∪ ord_res.production (fset N2) C ⊫ C|} C" and
L_max: "linorder_lit.is_maximal_in_mset C L"
for C L
proof (rule notI)
assume "{#} |∈| N2"
moreover have "¬ ord_res.interp (fset N2) {#} ∪ ord_res.production (fset N2) {#} ⊫ {#}"
by simp
moreover have "⋀C. {#} ≼⇩c C"
using mempty_lesseq_cls by metis
ultimately have "C = {#}"
using C_least
by (metis (no_types, lifting) ffmember_filter linorder_cls.is_least_in_fset_iff
linorder_cls.less_le_not_le)
moreover have "L ∈# C"
using L_max by (simp add: linorder_lit.is_maximal_in_mset_iff)
ultimately show "False"
by simp
qed
have "ord_res N2 N2'"
using step2
proof (cases N2 N2' rule: ord_res_1.cases)
case hyps: (factoring C L C')
show ?thesis
proof (rule ord_res.factoring)
show "{#} |∉| N2"
using hyps mempty_no_in is_least_false_clause_def by simp
next
show "ex_false_clause (fset N2)"
unfolding ex_false_clause_def
using hyps is_least_false_clause_def
by (metis (no_types, lifting) linorder_cls.is_least_in_fset_ffilterD)
next
show "C |∈| N2"
using hyps is_least_false_clause_def linorder_cls.is_least_in_fset_ffilterD(1) by blast
next
show "ord_res.ground_factoring C C'"
using hyps by argo
next
show "N2' = finsert C' N2"
using hyps by argo
qed
next
case hyps: (resolution C L D CD)
show ?thesis
proof (rule ord_res.resolution)
show "{#} |∉| N2"
using hyps mempty_no_in is_least_false_clause_def by simp
next
show "ex_false_clause (fset N2)"
unfolding ex_false_clause_def
using hyps is_least_false_clause_def
by (metis (no_types, lifting) linorder_cls.is_least_in_fset_ffilterD)
next
show "C |∈| N2"
using hyps is_least_false_clause_def linorder_cls.is_least_in_fset_ffilterD(1) by blast
next
show "D |∈| N2"
using hyps by argo
next
show "ord_res.ground_resolution C D CD"
using hyps by argo
next
show "N2' = finsert CD N2"
using hyps by argo
qed
qed
thus "ord_res⇧+⇧+ N1 N2'"
unfolding match by simp
next
show "N2' = N2'" ..
qed
qed
end
section ‹ORD-RES-2 (full factorization)›
type_synonym 'f ord_res_2_state = "'f gclause fset × 'f gclause fset × 'f gclause fset"
context simulation_SCLFOL_ground_ordered_resolution begin
fun ord_res_1_matches_ord_res_2
:: "'f ord_res_1_state ⇒ _ ⇒ bool" where
"ord_res_1_matches_ord_res_2 S1 (N, (U⇩r, U⇩e⇩f)) ⟷ (∃U⇩f.
S1 = N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f ∧
(∀C⇩f |∈| U⇩f. ∃C |∈| N |∪| U⇩r |∪| U⇩e⇩f. ord_res.ground_factoring⇧+⇧+ C C⇩f ∧ C⇩f ≠ efac C⇩f ∧
(efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C)))"
lemma ord_res_1_matches_ord_res_2_simps':
"ord_res_1_matches_ord_res_2 S1 (N, (U⇩r, U⇩e⇩f)) ⟷
(∃U⇩f. S1 = N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f ∧
(∀C⇩f |∈| U⇩f. C⇩f ≠ efac C⇩f ∧ (∃C |∈| N |∪| U⇩r |∪| U⇩e⇩f. ord_res.ground_factoring⇧+⇧+ C C⇩f ∧
(efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C))))"
unfolding ord_res_1_matches_ord_res_2.simps by metis
lemma ord_res_1_matches_ord_res_2_simps'':
"ord_res_1_matches_ord_res_2 S1 (N, (U⇩r, U⇩e⇩f)) ⟷
(∃U⇩f. S1 = N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f ∧
(∀C⇩f |∈| U⇩f. C⇩f ≠ efac C⇩f ∧ (∃C |∈| N |∪| U⇩r |∪| U⇩e⇩f. ord_res.ground_factoring⇧+⇧+ C C⇩f ∧
(efac C |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C))))"
unfolding ord_res_1_matches_ord_res_2_simps'
by (metis ground_factorings_preserves_efac tranclp_into_rtranclp)
lemma ord_res_1_final_iff_ord_res_2_final:
assumes match: "ord_res_1_matches_ord_res_2 S⇩1 S⇩2"
shows "ord_res_1_final S⇩1 ⟷ ord_res_2_final S⇩2"
proof -
obtain N U⇩r U⇩e⇩f where "S⇩2 = (N, (U⇩r, U⇩e⇩f))"
by (metis prod.exhaust)
with match obtain U⇩f where
S⇩1_def: "S⇩1 = N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f" and
U⇩f_spec: "∀C⇩f |∈| U⇩f. ∃C |∈| N |∪| U⇩r |∪| U⇩e⇩f. ord_res.ground_factoring⇧+⇧+ C C⇩f ∧ C⇩f ≠ efac C⇩f ∧
(efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C)"
by auto
have U⇩f_unproductive: "∀C⇩f |∈| U⇩f. ord_res.production (fset (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f)) C⇩f = {}"
proof (intro ballI)
fix C⇩f
assume "C⇩f |∈| U⇩f"
hence "C⇩f ≠ efac C⇩f"
using U⇩f_spec by metis
hence "∄L. is_pos L ∧ ord_res.is_strictly_maximal_lit L C⇩f"
using nex_strictly_maximal_pos_lit_if_neq_efac by metis
thus "ord_res.production (fset (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f)) C⇩f = {}"
using unproductive_if_nex_strictly_maximal_pos_lit by metis
qed
have Interp_eq: "⋀C. ord_res_Interp (fset (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f)) C =
ord_res_Interp (fset (N |∪| U⇩r |∪| U⇩e⇩f)) C"
using Interp_union_unproductive[of "fset (N |∪| U⇩r |∪| U⇩e⇩f)" "fset U⇩f", folded union_fset,
OF finite_fset finite_fset U⇩f_unproductive] .
have "{#} |∈| N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f ⟷ {#} |∈| N |∪| U⇩r |∪| U⇩e⇩f"
proof (rule iffI)
assume "{#} |∈| N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f"
hence "{#} |∈| N |∪| U⇩r |∪| U⇩e⇩f ∨ {#} |∈| U⇩f"
by simp
thus "{#} |∈| N |∪| U⇩r |∪| U⇩e⇩f"
proof (elim disjE)
assume "{#} |∈| N |∪| U⇩r |∪| U⇩e⇩f"
thus "{#} |∈| N |∪| U⇩r |∪| U⇩e⇩f"
by assumption
next
assume "{#} |∈| U⇩f"
hence "{#} ≠ efac {#}"
using U⇩f_spec[rule_format, of "{#}"] by metis
hence False
by simp
thus "{#} |∈| N |∪| U⇩r |∪| U⇩e⇩f" ..
qed
next
assume "{#} |∈| N |∪| U⇩r |∪| U⇩e⇩f"
thus "{#} |∈| N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f"
by simp
qed
moreover have "¬ ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f)) ⟷
¬ ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f))"
proof (rule iffI; erule contrapos_nn)
assume "ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f))"
thus "ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f))"
unfolding ex_false_clause_def Interp_eq by auto
next
assume "ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f))"
then obtain C where
"C |∈| N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f" and
C_false: "¬ ord_res_Interp (fset (N |∪| U⇩r |∪| U⇩e⇩f)) C ⊫ C"
unfolding ex_false_clause_def Interp_eq by metis
hence "C |∈| N |∪| U⇩r |∪| U⇩e⇩f ∨ C |∈| U⇩f"
by simp
thus "ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f))"
proof (elim disjE)
assume "C |∈| N |∪| U⇩r |∪| U⇩e⇩f"
thus "ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f))"
unfolding ex_false_clause_def using C_false by metis
next
assume "C |∈| U⇩f"
then obtain C' where "C' |∈| N |∪| U⇩r |∪| U⇩e⇩f" and
"ord_res.ground_factoring⇧+⇧+ C' C" and
"C ≠ efac C" and
"efac C |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C'"
using U⇩f_spec[rule_format, of C] by metis
thus "ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f))"
proof (elim disjE exE conjE)
assume "efac C |∈| U⇩e⇩f"
show "ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f))"
proof (cases "ord_res_Interp (fset (N |∪| U⇩r |∪| U⇩e⇩f)) (efac C) ⊫ efac C")
case efac_C_true: True
have "efac C ⊆# C"
using efac_subset[of C] .
hence "efac C ≼⇩c C"
using subset_implies_reflclp_multp by metis
hence "ord_res_Interp (fset (N |∪| U⇩r |∪| U⇩e⇩f)) C ⊫ efac C"
using efac_C_true ord_res.entailed_clause_stays_entailed by fastforce
hence "ord_res_Interp (fset (N |∪| U⇩r |∪| U⇩e⇩f)) C ⊫ C"
using efac_C_true by (simp add: true_cls_def)
with C_false have False
by contradiction
thus ?thesis ..
next
case False
moreover have "efac C |∈| N |∪| U⇩r |∪| U⇩e⇩f"
using ‹efac C |∈| U⇩e⇩f› by simp
ultimately show "ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f))"
unfolding ex_false_clause_def by metis
qed
next
assume "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C'"
hence "C' |∈| N |∪| U⇩r |∪| U⇩e⇩f" and "¬ ord_res_Interp (fset (N |∪| U⇩r |∪| U⇩e⇩f)) C' ⊫ C'"
using linorder_cls.is_least_in_ffilter_iff is_least_false_clause_def by simp_all
thus "ex_false_clause (fset (N |∪| U⇩r |∪| U⇩e⇩f))"
unfolding ex_false_clause_def by metis
qed
qed
qed
ultimately show ?thesis
by (simp add: S⇩1_def ‹S⇩2 = (N, U⇩r, U⇩e⇩f)› ord_res_1_final_def ord_res_2_final.simps
ord_res_final_def)
qed
lemma safe_states_if_ord_res_1_matches_ord_res_2:
assumes match: "ord_res_1_matches_ord_res_2 S⇩1 S⇩2"
shows "safe_state ord_res_1 ord_res_1_final S⇩1 ∧ safe_state ord_res_2_step ord_res_2_final S⇩2"
proof -
have "safe_state ord_res_1 ord_res_1_final S⇩1"
using safe_state_if_all_states_safe ord_res_1_safe by metis
moreover have "safe_state ord_res_2_step ord_res_2_final S⇩2"
using safe_state_if_all_states_safe ord_res_2_step_safe by metis
ultimately show ?thesis
by argo
qed
definition ord_res_1_measure where
"ord_res_1_measure s1 =
(if ∃C. is_least_false_clause s1 C then
The (is_least_false_clause s1)
else
{#})"
lemma forward_simulation:
assumes match: "ord_res_1_matches_ord_res_2 s1 s2" and
step1: "ord_res_1 s1 s1'"
shows "(∃s2'. ord_res_2_step⇧+⇧+ s2 s2' ∧ ord_res_1_matches_ord_res_2 s1' s2') ∨
ord_res_1_matches_ord_res_2 s1' s2 ∧ ord_res_1_measure s1' ⊂# ord_res_1_measure s1"
proof -
let
?match = "ord_res_1_matches_ord_res_2" and
?measure = "ord_res_1_measure" and
?order = "(⊂#)"
obtain N U⇩r U⇩e⇩f :: "'f gterm clause fset" where
s2_def: "s2 = (N, (U⇩r, U⇩e⇩f))"
by (metis prod.exhaust)
from match obtain U⇩f where
s1_def: "s1 = N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f" and
U⇩f_spec: "∀C⇩f |∈| U⇩f. ∃C |∈| N |∪| U⇩r |∪| U⇩e⇩f. ord_res.ground_factoring⇧+⇧+ C C⇩f ∧ C⇩f ≠ efac C⇩f ∧
(efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C)"
unfolding s2_def ord_res_1_matches_ord_res_2.simps by metis
have U⇩f_unproductive: "∀C⇩f |∈| U⇩f. ord_res.production (fset (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f)) C⇩f = {}"
proof (intro ballI)
fix C⇩f
assume "C⇩f |∈| U⇩f"
hence "C⇩f ≠ efac C⇩f"
using U⇩f_spec by metis
hence "∄L. is_pos L ∧ ord_res.is_strictly_maximal_lit L C⇩f"
using nex_strictly_maximal_pos_lit_if_neq_efac by metis
thus "ord_res.production (fset (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f)) C⇩f = {}"
using unproductive_if_nex_strictly_maximal_pos_lit by metis
qed
have Interp_eq: "⋀C. ord_res_Interp (fset (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f)) C =
ord_res_Interp (fset (N |∪| U⇩r |∪| U⇩e⇩f)) C"
using Interp_union_unproductive[of "fset (N |∪| U⇩r |∪| U⇩e⇩f)" "fset U⇩f", folded union_fset,
OF finite_fset finite_fset U⇩f_unproductive] .
show "(∃s2'. ord_res_2_step⇧+⇧+ s2 s2' ∧ ?match s1' s2') ∨
?match s1' s2 ∧ ?order (?measure s1') (?measure s1)"
using step1
proof (cases s1 s1' rule: ord_res_1.cases)
case (factoring C L C')
have C_least_false: "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f) C"
using factoring
unfolding is_least_false_clause_def s1_def by argo
hence C_in: "C |∈| N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f"
unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff s1_def by argo
hence C_in_disj: "C |∈| N |∪| U⇩r |∪| U⇩e⇩f ∨ C |∈| U⇩f"
by simp
show ?thesis
proof (cases "C' = efac C'")
case True
let ?s2' = "(N, (U⇩r, finsert C' U⇩e⇩f))"
have "ord_res_2_step⇧+⇧+ s2 ?s2'"
proof (rule tranclp.r_into_trancl)
show "ord_res_2_step s2 (N, U⇩r, finsert C' U⇩e⇩f)"
using C_in_disj
proof (elim disjE)
assume "C |∈| N |∪| U⇩r |∪| U⇩e⇩f"
show ?thesis
unfolding s2_def
proof (intro ord_res_2_step.intros ord_res_2.factoring)
show "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C"
using is_least_false_clause_if_is_least_false_clause_in_union_unproductive[
OF U⇩f_unproductive ‹C |∈| N |∪| U⇩r |∪| U⇩e⇩f› C_least_false]
unfolding is_least_false_clause_def .
next
show "ord_res.is_maximal_lit L C"
using ‹ord_res.is_maximal_lit L C› .
next
show "is_pos L"
using ‹is_pos L› .
next
show "finsert C' U⇩e⇩f = finsert (efac C) U⇩e⇩f"
using True factoring ground_factoring_preserves_efac by metis
qed
next
assume "C |∈| U⇩f"
then obtain x where
"x |∈| N |∪| U⇩r |∪| U⇩e⇩f" and
"ord_res.ground_factoring⇧+⇧+ x C" and
"C ≠ efac C" and
"efac C |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x"
using U⇩f_spec by metis
show ?thesis
unfolding s2_def
proof (intro ord_res_2_step.intros ord_res_2.factoring)
have ‹efac C |∉| U⇩e⇩f›
proof (rule notI)
have "efac C ≼⇩c C"
using efac_subset[of C] subset_implies_reflclp_multp by metis
hence "efac C ≺⇩c C"
using ‹C ≠ efac C› by order
moreover assume "efac C |∈| U⇩e⇩f"
ultimately show False
using C_least_false[unfolded is_least_false_clause_def
linorder_cls.is_least_in_ffilter_iff]
by (metis ‹C ≠ efac C› funionCI linorder_cls.not_less_iff_gr_or_eq
ord_res.entailed_clause_stays_entailed set_mset_efac true_cls_def)
qed
thus "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x"
using ‹efac C |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x› by argo
next
show "ord_res.is_maximal_lit L x"
using ‹ord_res.ground_factoring⇧+⇧+ x C› ‹ord_res.is_maximal_lit L C›
using ord_res.ground_factorings_preserves_maximal_literal
by (metis tranclp_into_rtranclp)
next
show "is_pos L"
using ‹is_pos L› .
next
show "finsert C' U⇩e⇩f = finsert (efac x) U⇩e⇩f"
using ‹ord_res.ground_factoring⇧+⇧+ x C› ‹ord_res.ground_factoring C C'›
using True ground_factorings_preserves_efac ground_factoring_preserves_efac
by (metis tranclp_into_rtranclp)
qed
qed
qed
moreover have "?match s1' ?s2'"
proof -
have "s1' = N |∪| U⇩r |∪| finsert C' U⇩e⇩f |∪| U⇩f"
unfolding ‹s1' = finsert C' s1› s1_def by simp
moreover have "∃C |∈| N |∪| U⇩r |∪| finsert C' U⇩e⇩f.
ord_res.ground_factoring⇧+⇧+ C C⇩f ∧ C⇩f ≠ efac C⇩f ∧
(efac C⇩f |∈| finsert C' U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| finsert C' U⇩e⇩f) C)"
if "C⇩f |∈| U⇩f" for C⇩f
proof -
obtain x where
"x |∈| N |∪| U⇩r |∪| U⇩e⇩f" and
"ord_res.ground_factoring⇧+⇧+ x C⇩f" and
"C⇩f ≠ efac C⇩f" and
"efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x"
using ‹C⇩f |∈| U⇩f› U⇩f_spec by metis
show ?thesis
proof (intro bexI conjI)
show "x |∈| N |∪| U⇩r |∪| finsert C' U⇩e⇩f"
using ‹x |∈| N |∪| U⇩r |∪| U⇩e⇩f› by simp
next
show "ord_res.ground_factoring⇧+⇧+ x C⇩f"
using ‹ord_res.ground_factoring⇧+⇧+ x C⇩f› .
next
show "C⇩f ≠ efac C⇩f"
using ‹C⇩f ≠ efac C⇩f› .
next
show "efac C⇩f |∈| finsert C' U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| finsert C' U⇩e⇩f) x"
using ‹efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x›
proof (elim disjE)
assume "efac C⇩f |∈| U⇩e⇩f"
thus ?thesis
by simp
next
assume "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x"
show ?thesis
proof (cases "C' = efac x")
case True
moreover have "efac x = efac C⇩f"
using ‹ord_res.ground_factoring⇧+⇧+ x C⇩f› ground_factorings_preserves_efac
by (metis tranclp_into_rtranclp)
ultimately show ?thesis
by simp
next
case False
show ?thesis
using C_in_disj
proof (elim disjE)
assume "C |∈| N |∪| U⇩r |∪| U⇩e⇩f"
then show ?thesis
by (smt (verit) C_least_false True U⇩f_unproductive
‹is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x› ‹ord_res.ground_factoring⇧+⇧+ x C⇩f›
finsert_iff ground_factoring_preserves_efac ground_factorings_preserves_efac
linorder_cls.Uniq_is_least_in_fset local.factoring(4)
is_least_false_clause_def
is_least_false_clause_if_is_least_false_clause_in_union_unproductive
the1_equality' tranclp_into_rtranclp)
next
assume "C |∈| U⇩f"
then show ?thesis
using C_least_false
using is_least_false_clause_if_is_least_false_clause_in_union_unproductive[
OF U⇩f_unproductive]
by (smt (z3) True U⇩f_spec ‹is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x›
‹ord_res.ground_factoring⇧+⇧+ x C⇩f› finsert_absorb finsert_iff
ground_factoring_preserves_efac ground_factorings_preserves_efac
linorder_cls.Uniq_is_least_in_fset local.factoring(4)
is_least_false_clause_def the1_equality' tranclp_into_rtranclp)
qed
qed
qed
qed
qed
ultimately show ?thesis
by auto
qed
ultimately show ?thesis
by metis
next
case False
let ?U⇩f' = "finsert C' U⇩f"
have "?match s1' s2"
proof -
have "finsert C' s1 = N |∪| U⇩r |∪| U⇩e⇩f |∪| ?U⇩f'"
unfolding s1_def by simp
moreover have "∃C |∈| N |∪| U⇩r |∪| U⇩e⇩f.
ord_res.ground_factoring⇧+⇧+ C C⇩f ∧ C⇩f ≠ efac C⇩f ∧
(efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C)"
if "C⇩f |∈| ?U⇩f'" for C⇩f
proof -
from ‹C⇩f |∈| ?U⇩f'› have "C⇩f = C' ∨ C⇩f |∈| U⇩f"
by simp
thus ?thesis
proof (elim disjE)
assume "C⇩f = C'"
thus ?thesis
using C_in_disj
proof (elim disjE)
assume "C |∈| N |∪| U⇩r |∪| U⇩e⇩f"
show ?thesis
proof (intro bexI conjI)
show "C |∈| N |∪| U⇩r |∪| U⇩e⇩f"
using ‹C |∈| N |∪| U⇩r |∪| U⇩e⇩f› .
next
show "ord_res.ground_factoring⇧+⇧+ C C⇩f"
using ‹ord_res.ground_factoring C C'› ‹C⇩f = C'› by simp
next
show "C⇩f ≠ efac C⇩f"
using False ‹C⇩f = C'› by argo
next
have "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C"
using factoring
using Interp_eq ‹C |∈| N |∪| U⇩r |∪| U⇩e⇩f› linorder_cls.is_least_in_ffilter_iff
by (simp add: s1_def is_least_false_clause_def)
thus "efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C" ..
qed
next
assume "C |∈| U⇩f"
then obtain x where
"x |∈| N |∪| U⇩r |∪| U⇩e⇩f" and
"ord_res.ground_factoring⇧+⇧+ x C" and
"C ≠ efac C" and
"efac C |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x"
using U⇩f_spec by metis
show ?thesis
proof (intro bexI conjI)
show "x |∈| N |∪| U⇩r |∪| U⇩e⇩f"
using ‹x |∈| N |∪| U⇩r |∪| U⇩e⇩f› .
next
show "ord_res.ground_factoring⇧+⇧+ x C⇩f"
using ‹ord_res.ground_factoring⇧+⇧+ x C› ‹ord_res.ground_factoring C C'› ‹C⇩f = C'›
by simp
next
show "C⇩f ≠ efac C⇩f"
using False ‹C⇩f = C'› by argo
next
show "efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x"
using ‹efac C |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x›
proof (elim disjE)
assume "efac C |∈| U⇩e⇩f"
moreover have "efac C = efac C⇩f"
unfolding ‹C⇩f = C'›
using ‹ord_res.ground_factoring C C'› ground_factoring_preserves_efac by metis
ultimately show ?thesis
by argo
next
assume "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x"
thus ?thesis
by argo
qed
qed
qed
next
assume "C⇩f |∈| U⇩f"
thus ?thesis
using U⇩f_spec by metis
qed
qed
ultimately have "ord_res_1_matches_ord_res_2 (finsert C' s1) (N, (U⇩r, U⇩e⇩f))"
unfolding ord_res_1_matches_ord_res_2.simps by metis
thus ?thesis
unfolding s2_def ‹s1' = finsert C' s1› by simp
qed
moreover have "?order (?measure s1') (?measure s1)"
proof -
have "?measure s1 = C"
unfolding ord_res_1_measure_def
using C_least_false[folded s1_def]
by (metis (mono_tags, lifting) linorder_cls.Uniq_is_least_in_fset
is_least_false_clause_def the1_equality' the_equality)
moreover have "?measure s1' = C'"
proof -
have "C' ≺⇩c C"
using factoring ord_res.ground_factoring_smaller_conclusion by metis
have unproductive: "∀x∈{C'}. ord_res.production (fset s1 ∪ {C'}) x = {}"
using ‹C' ≠ efac C'›
by (simp add: nex_strictly_maximal_pos_lit_if_neq_efac
unproductive_if_nex_strictly_maximal_pos_lit)
have Interp_eq: "⋀D. ord_res_Interp (fset s1) D = ord_res_Interp (fset (finsert C' s1)) D"
using Interp_union_unproductive[of "fset s1" "{C'}", folded union_fset,
OF finite_fset _ unproductive]
by simp
have "is_least_false_clause (finsert C' s1) C'"
unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
proof (intro conjI ballI impI)
have "¬ ord_res_Interp (fset s1) C ⊫ C"
using C_least_false s1_def is_least_false_clause_def
linorder_cls.is_least_in_ffilter_iff by simp
thus "¬ ord_res_Interp (fset (finsert C' s1)) C' ⊫ C'"
by (metis Interp_eq ‹C' ≺⇩c C› local.factoring(4)
ord_res.entailed_clause_stays_entailed
ord_res.set_mset_eq_set_mset_if_ground_factoring subset_refl true_cls_mono)
next
fix y
assume "y |∈| finsert C' s1" and "y ≠ C'" and
y_false: "¬ ord_res_Interp (fset (finsert C' s1)) y ⊫ y"
hence "y |∈| s1"
by simp
moreover have "¬ ord_res_Interp (fset s1) y ⊫ y"
using y_false
unfolding Interp_eq .
ultimately have "C ≼⇩c y"
using C_least_false[folded s1_def, unfolded is_least_false_clause_def]
unfolding linorder_cls.is_least_in_ffilter_iff
by force
thus "C' ≺⇩c y"
using ‹C' ≺⇩c C› by order
qed simp
thus ?thesis
unfolding ord_res_1_measure_def ‹s1' = finsert C' s1›
by (metis (mono_tags, lifting) linorder_cls.Uniq_is_least_in_fset
is_least_false_clause_def the1_equality' the_equality)
qed
moreover have "C' ⊂# C"
using factoring ord_res.strict_subset_mset_if_ground_factoring by metis
ultimately show ?thesis
unfolding s1_def by simp
qed
ultimately show ?thesis
by argo
qed
next
case (resolution C L D CD)
have "is_least_false_clause s1 C"
using resolution unfolding is_least_false_clause_def by argo
hence
"C |∈| s1" and
"¬ ord_res_Interp (fset s1) C ⊫ C" and
"∀x |∈| s1. ¬ ord_res_Interp (fset s1) x ⊫ x ⟶ x ≠ C ⟶ C ≺⇩c x"
unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff by simp_all
have "C |∉| U⇩f"
proof (rule notI)
assume "C |∈| U⇩f"
then show False
by (metis U⇩f_spec Uniq_D is_pos_def linorder_lit.Uniq_is_maximal_in_mset local.resolution(2)
local.resolution(3) efac_spec)
qed
hence "C |∈| N |∪| U⇩r |∪| U⇩e⇩f"
using ‹C |∈| s1› by (simp add: s1_def)
have C_least_false: "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f) C"
using resolution s1_def by metis
hence C_least_false': "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C"
using is_least_false_clause_if_is_least_false_clause_in_union_unproductive[
OF U⇩f_unproductive ‹C |∈| N |∪| U⇩r |∪| U⇩e⇩f›] by argo
define s2' where
"s2' = (N, (finsert CD U⇩r, U⇩e⇩f))"
have "ord_res_2_step⇧+⇧+ s2 s2'"
proof -
have "D |∉| U⇩f"
proof (rule notI)
assume "D |∈| U⇩f"
thus False
using ‹ord_res.production (fset s1) D = {atm_of L}›
using U⇩f_unproductive s1_def by simp
qed
hence D_in: "D |∈| N |∪| U⇩r |∪| U⇩e⇩f"
using ‹D |∈| s1›[unfolded s1_def] by simp
have "ord_res_2 N (U⇩r, U⇩e⇩f) (finsert CD U⇩r, U⇩e⇩f)"
proof (rule ord_res_2.resolution)
show "is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C"
using C_least_false' .
next
show "ord_res.is_maximal_lit L C"
using resolution by argo
next
show "is_neg L"
using resolution by argo
next
show "D |∈| N |∪| U⇩r |∪| U⇩e⇩f"
using D_in .
next
show "D ≺⇩c C"
using resolution by argo
next
show "ord_res.production (fset (N |∪| U⇩r |∪| U⇩e⇩f)) D = {atm_of L}"
using resolution
unfolding s1_def
using production_union_unproductive[OF finite_fset finite_fset _ D_in] U⇩f_unproductive
by (metis (no_types, lifting) union_fset)
next
show "ord_res.ground_resolution C D CD"
using resolution by argo
qed simp_all
thus ?thesis
by (auto simp: s2_def s2'_def ord_res_2_step.simps)
qed
moreover have "?match s1' s2'"
proof -
have "finsert CD (N |∪| U⇩r |∪| U⇩e⇩f |∪| U⇩f) = N |∪| finsert CD U⇩r |∪| U⇩e⇩f |∪| U⇩f"
by simp
moreover have "∃C |∈| N |∪| finsert CD U⇩r |∪| U⇩e⇩f.
ord_res.ground_factoring⇧+⇧+ C C⇩f ∧ C⇩f ≠ efac C⇩f ∧
(efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| finsert CD U⇩r |∪| U⇩e⇩f) C)"
if "C⇩f |∈| U⇩f" for C⇩f
proof -
obtain x where
"x |∈| N |∪| U⇩r |∪| U⇩e⇩f" and
"ord_res.ground_factoring⇧+⇧+ x C⇩f" and
"C⇩f ≠ efac C⇩f" and
"efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x"
using ‹C⇩f |∈| U⇩f› U⇩f_spec by metis
show ?thesis
proof (intro bexI conjI)
show "x |∈| N |∪| finsert CD U⇩r |∪| U⇩e⇩f"
using ‹x |∈| N |∪| U⇩r |∪| U⇩e⇩f› by simp
next
show "ord_res.ground_factoring⇧+⇧+ x C⇩f"
using ‹ord_res.ground_factoring⇧+⇧+ x C⇩f› .
next
show "C⇩f ≠ efac C⇩f"
using ‹C⇩f ≠ efac C⇩f› .
next
show ‹efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| finsert CD U⇩r |∪| U⇩e⇩f) x›
using ‹efac C⇩f |∈| U⇩e⇩f ∨ is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) x› ‹x |∈| N |∪| U⇩r |∪| U⇩e⇩f›
by (metis (no_types, lifting) C_least_false' Uniq_D ‹ord_res.ground_factoring⇧+⇧+ x C⇩f›
is_least_false_clause_def is_pos_def linorder_cls.Uniq_is_least_in_fset
linorder_lit.Uniq_is_maximal_in_mset local.resolution(2) local.resolution(3)
ord_res.ground_factoring.cases tranclpD)
qed
qed
ultimately show ?thesis
unfolding s1_def resolution s2'_def by auto
qed
ultimately show ?thesis
by metis
qed
qed
theorem bisimulation_ord_res_1_ord_res_2:
defines "match ≡ λi s1 s2. i = ord_res_1_measure s1 ∧ ord_res_1_matches_ord_res_2 s1 s2"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_1_state ⇒ 'f ord_res_2_state ⇒ bool) ℛ.
bisimulation ord_res_1 ord_res_2_step ord_res_1_final ord_res_2_final ℛ MATCH"
proof (rule ex_bisimulation_from_forward_simulation)
show "right_unique ord_res_1"
using right_unique_ord_res_1 .
next
show "right_unique ord_res_2_step"
using right_unique_ord_res_2_step .
next
show "∀s1. ord_res_1_final s1 ⟶ (∄s1'. ord_res_1 s1 s1')"
using ord_res_1_semantics.final_finished
by (simp add: finished_def)
next
show "∀s2. ord_res_2_final s2 ⟶ (∄s2'. ord_res_2_step s2 s2')"
using ord_res_2_semantics.final_finished
by (simp add: finished_def)
next
show "∀i s1 s2. match i s1 s2 ⟶ ord_res_1_final s1 = ord_res_2_final s2"
using ord_res_1_final_iff_ord_res_2_final
by (simp add: match_def)
next
show "∀i s1 s2. match i s1 s2 ⟶
safe_state ord_res_1 ord_res_1_final s1 ∧
safe_state ord_res_2_step ord_res_2_final s2"
proof (intro allI impI)
fix i s1 S2
assume "match i s1 S2"
then obtain N s2 where
S2_def: "S2 = (N, s2)" and
"i = ord_res_1_measure s1" and
match: "ord_res_1_matches_ord_res_2 s1 S2"
unfolding match_def
by (metis prod.exhaust)
show "safe_state ord_res_1 ord_res_1_final s1 ∧ safe_state ord_res_2_step ord_res_2_final S2"
using safe_states_if_ord_res_1_matches_ord_res_2[OF match] .
qed
next
show "wfP (⊂#)"
using wfp_subset_mset .
next
show "∀i s1 s2 s1'. match i s1 s2 ⟶ ord_res_1 s1 s1' ⟶
(∃i' s2'. ord_res_2_step⇧+⇧+ s2 s2' ∧ match i' s1' s2') ∨ (∃i'. match i' s1' s2 ∧ i' ⊂# i)"
proof (intro allI impI)
fix i s1 S2 s1'
assume "match i s1 S2"
then obtain N s2 where
S2_def: "S2 = (N, s2)" and "i = ord_res_1_measure s1" and "ord_res_1_matches_ord_res_2 s1 S2"
unfolding match_def
by (metis prod.exhaust)
moreover assume "ord_res_1 s1 s1'"
ultimately have "(∃S2'. ord_res_2_step⇧+⇧+ S2 S2' ∧ ord_res_1_matches_ord_res_2 s1' S2') ∨
ord_res_1_matches_ord_res_2 s1' S2 ∧ ord_res_1_measure s1' ⊂# ord_res_1_measure s1"
using forward_simulation by metis
thus "(∃i' S2'. ord_res_2_step⇧+⇧+ S2 S2' ∧ match i' s1' S2') ∨ (∃i'. match i' s1' S2 ∧ i' ⊂# i)"
unfolding S2_def prod.case
using lift_tranclp_to_pairs_with_constant_fst[of ord_res_2 N s2]
by (metis (mono_tags, lifting) ‹i = ord_res_1_measure s1› match_def)
qed
qed
end
section ‹ORD-RES-3 (full resolve)›
type_synonym 'f ord_res_3_state = "'f gclause fset × 'f gclause fset × 'f gclause fset"
context simulation_SCLFOL_ground_ordered_resolution begin
inductive ord_res_2_matches_ord_res_3 :: "_ ⇒ 'f ord_res_3_state ⇒ bool" where
"(∀C |∈| U⇩p⇩r. ∃D1 |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f. ∃D2 |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f.
(ground_resolution D1)⇧+⇧+ D2 C ∧ C ≠ eres D1 D2 ∧ eres D1 D2 |∈| U⇩e⇩r) ⟹
ord_res_2_matches_ord_res_3 (N, (U⇩p⇩r |∪| U⇩e⇩r, U⇩e⇩f)) (N, (U⇩e⇩r, U⇩e⇩f))"
lemma ord_res_2_final_iff_ord_res_3_final:
assumes match: "ord_res_2_matches_ord_res_3 S⇩2 S⇩3"
shows "ord_res_2_final S⇩2 ⟷ ord_res_3_final S⇩3"
using match
proof (cases S⇩2 S⇩3 rule: ord_res_2_matches_ord_res_3.cases)
case match_hyps: (1 U⇩p⇩r N U⇩e⇩r U⇩e⇩f)
note invars = match_hyps(3-)
have U⇩p⇩r_spec: "∀C|∈|U⇩p⇩r. ∃D1|∈|N |∪| U⇩e⇩r |∪| U⇩e⇩f. ∃D2|∈|N |∪| U⇩e⇩r |∪| U⇩e⇩f.
(ground_resolution D1)⇧+⇧+ D2 C ∧ C ≠ eres D1 D2 ∧ eres D1 D2 |∈| U⇩e⇩r"
using invars by argo
have least_false_spec: "is_least_false_clause (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f) =
is_least_false_clause (N |∪| U⇩e⇩r |∪| U⇩e⇩f)"
using invars is_least_false_clause_conv_if_partial_resolution_invariant by metis
have U⇩p⇩r_unproductive: "∀C |∈| U⇩p⇩r. ord_res.production (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f |∪| U⇩p⇩r)) C = {}"
proof (intro ballI)
fix C
assume "C |∈| U⇩p⇩r"
hence "∄L. is_pos L ∧ ord_res.is_strictly_maximal_lit L C"
using U⇩p⇩r_spec
by (metis eres_eq_after_tranclp_ground_resolution nex_strictly_maximal_pos_lit_if_neq_eres)
thus "ord_res.production (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f |∪| U⇩p⇩r)) C = {}"
using unproductive_if_nex_strictly_maximal_pos_lit by metis
qed
hence Interp_N_U⇩r_U⇩e⇩f_eq_Interp_N_U⇩e⇩r_U⇩e⇩f: "⋀C.
ord_res_Interp (fset (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f)) C =
ord_res_Interp (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f)) C"
using Interp_union_unproductive[OF finite_fset finite_fset, folded union_fset,
of U⇩p⇩r "N |∪| U⇩e⇩r |∪| U⇩e⇩f"]
by (simp add: funion_left_commute sup_commute)
have "ex_false_clause (fset (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f)) ⟷
ex_false_clause (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f))"
proof (rule iffI)
assume "ex_false_clause (fset (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f))"
then obtain C where "is_least_false_clause (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f) C"
using obtains_least_false_clause_if_ex_false_clause by metis
thus "ex_false_clause (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f))"
using least_false_spec ex_false_clause_iff by metis
next
assume "ex_false_clause (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f))"
thus "ex_false_clause (fset (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f))"
unfolding ex_false_clause_def
unfolding Interp_N_U⇩r_U⇩e⇩f_eq_Interp_N_U⇩e⇩r_U⇩e⇩f
by auto
qed
moreover have "{#} |∈| N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f ⟷ {#} |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f"
proof (rule iffI)
assume "{#} |∈| N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f"
hence "{#} |∈| N |∪| U⇩e⇩f ∨ {#} |∈| U⇩p⇩r |∪| U⇩e⇩r"
by auto
thus "{#} |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f"
proof (elim disjE)
assume "{#} |∈| N |∪| U⇩e⇩f"
thus ?thesis
by auto
next
have "{#} |∉| U⇩p⇩r"
using U⇩p⇩r_spec[rule_format, of "{#}"]
by (metis eres_eq_after_tranclp_ground_resolution eres_mempty_right)
moreover assume "{#} |∈| U⇩p⇩r |∪| U⇩e⇩r"
ultimately show ?thesis
by simp
qed
next
assume "{#} |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f"
then show "{#} |∈| N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f"
by auto
qed
ultimately have "ord_res_final (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f) = ord_res_final (N |∪| U⇩e⇩r |∪| U⇩e⇩f)"
unfolding ord_res_final_def by argo
thus "ord_res_2_final S⇩2 ⟷ ord_res_3_final S⇩3"
unfolding match_hyps(1,2)
by (simp add: ord_res_2_final.simps ord_res_3_final.simps sup_assoc)
qed
definition ord_res_2_measure where
"ord_res_2_measure S1 =
(let (N, (U⇩r, U⇩e⇩f)) = S1 in
(if ∃C. is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f) C then
The (is_least_false_clause (N |∪| U⇩r |∪| U⇩e⇩f))
else
{#}))"
definition resolvent_at where
"resolvent_at C D i = (THE CD. (ground_resolution C ^^ i) D CD)"
lemma resolvent_at_0[simp]: "resolvent_at C D 0 = D"
by (simp add: resolvent_at_def)
lemma resolvent_at_less_cls_resolvent_at:
assumes reso_at: "(ground_resolution C ^^ n) D CD"
assumes "i < j" and "j ≤ n"
shows "resolvent_at C D j ≺⇩c resolvent_at C D i"
proof -
obtain j' where
"j = i + Suc j'"
using ‹i < j› by (metis less_iff_Suc_add nat_arith.suc1)
obtain n' where
"n = j + n'"
using ‹j ≤ n› by (metis le_add_diff_inverse)
obtain CD⇩i CD⇩j CD⇩n where
"(ground_resolution C ^^ i) D CD⇩i" and
"(ground_resolution C ^^ Suc j') CD⇩i CD⇩j"
"(ground_resolution C ^^ n') CD⇩j CD⇩n"
using reso_at ‹n = j + n'› ‹j = i + Suc j'› by (metis relpowp_plusD)
have *: "resolvent_at C D i = CD⇩i"
unfolding resolvent_at_def
using ‹(ground_resolution C ^^ i) D CD⇩i›
by (simp add: Uniq_ground_resolution Uniq_relpowp the1_equality')
have "(ground_resolution C ^^ j) D CD⇩j"
unfolding ‹j = i + Suc j'›
using ‹(ground_resolution C ^^ i) D CD⇩i› ‹(ground_resolution C ^^ Suc j') CD⇩i CD⇩j›
by (metis relpowp_trans)
hence **: "resolvent_at C D j = CD⇩j"
unfolding resolvent_at_def
by (simp add: Uniq_ground_resolution Uniq_relpowp the1_equality')
have "(ground_resolution C)⇧+⇧+ CD⇩i CD⇩j"
using ‹(ground_resolution C ^^ Suc j') CD⇩i CD⇩j›
by (metis Zero_not_Suc tranclp_if_relpowp)
hence "CD⇩j ≺⇩c CD⇩i"
using resolvent_lt_right_premise_if_tranclp_ground_resolution by metis
thus ?thesis
unfolding * ** .
qed
lemma
assumes reso_at: "(ground_resolution C ^^ n) D CD" and "i < n"
shows
left_premisse_lt_resolvent_at: "C ≺⇩c resolvent_at C D i" and
max_lit_resolvent_at:
"ord_res.is_maximal_lit L D ⟹ ord_res.is_maximal_lit L (resolvent_at C D i)" and
nex_pos_strictly_max_lit_in_resolvent_at:
"∄L. is_pos L ∧ ord_res.is_strictly_maximal_lit L (resolvent_at C D i)" and
ground_resolution_resolvent_at_resolvent_at_Suc:
"ground_resolution C (resolvent_at C D i) (resolvent_at C D (Suc i))" and
relpowp_to_resolvent_at: "(ground_resolution C ^^ i) D (resolvent_at C D i)"
proof -
obtain j where n_def: "n = i + Suc j"
using ‹i < n› less_natE by auto
obtain CD' where "(ground_resolution C ^^ i) D CD'" and "(ground_resolution C ^^ Suc j) CD' CD"
using reso_at n_def by (metis relpowp_plusD)
have "resolvent_at C D i = CD'"
unfolding resolvent_at_def
using ‹(ground_resolution C ^^ i) D CD'›
by (simp add: Uniq_ground_resolution Uniq_relpowp the1_equality')
have "C ≺⇩c CD'"
proof (rule left_premise_lt_right_premise_if_tranclp_ground_resolution)
show "(ground_resolution C)⇧+⇧+ CD' CD"
using ‹(ground_resolution C ^^ Suc j) CD' CD›
by (metis Zero_not_Suc tranclp_if_relpowp)
qed
thus "C ≺⇩c resolvent_at C D i"
unfolding ‹resolvent_at C D i = CD'› by argo
show "ord_res.is_maximal_lit L (resolvent_at C D i)" if "ord_res.is_maximal_lit L D"
unfolding ‹resolvent_at C D i = CD'›
using that
using ‹(ground_resolution C ^^ i) D CD'›
by (smt (verit, ccfv_SIG) Uniq_ground_resolution Uniq_relpowp Zero_not_Suc
‹⋀thesis. (⋀CD'. ⟦(ground_resolution C ^^ i) D CD'; (ground_resolution C ^^ Suc j) CD' CD⟧ ⟹ thesis) ⟹ thesis›
linorder_lit.Uniq_is_greatest_in_mset linorder_lit.Uniq_is_maximal_in_mset literal.sel(1)
n_def relpowp_ground_resolutionD reso_at the1_equality' zero_eq_add_iff_both_eq_0)
show "∄L. is_pos L ∧ ord_res.is_strictly_maximal_lit L (resolvent_at C D i)"
unfolding ‹resolvent_at C D i = CD'›
by (metis Zero_not_Suc ‹(ground_resolution C ^^ Suc j) CD' CD›
nex_strictly_maximal_pos_lit_if_resolvable tranclpD tranclp_if_relpowp)
show "ground_resolution C (resolvent_at C D i) (resolvent_at C D (Suc i))"
proof -
obtain CD'' where "ground_resolution C CD' CD''" and "(ground_resolution C ^^ j) CD'' CD"
using ‹(ground_resolution C ^^ Suc j) CD' CD› by (metis relpowp_Suc_D2)
hence "(ground_resolution C ^^ Suc i) D CD''"
using ‹(ground_resolution C ^^ i) D CD'› by auto
hence "resolvent_at C D (Suc i) = CD''"
unfolding resolvent_at_def
by (meson Uniq_ground_resolution Uniq_relpowp the1_equality')
show ?thesis
unfolding ‹resolvent_at C D i = CD'› ‹resolvent_at C D (Suc i) = CD''›
using ‹ground_resolution C CD' CD''› .
qed
show "(ground_resolution C ^^ i) D (resolvent_at C D i)"
using ‹(ground_resolution C ^^ i) D CD'› ‹resolvent_at C D i = CD'› by argo
qed
definition resolvents_upto where
"resolvents_upto C D n = resolvent_at C D |`| fset_upto (Suc 0) n"
lemma resolvents_upto_0[simp]:
"resolvents_upto C D 0 = {||}"
by (simp add: resolvents_upto_def)
lemma resolvents_upto_Suc[simp]:
"resolvents_upto C D (Suc n) = finsert (resolvent_at C D (Suc n)) (resolvents_upto C D n)"
by (simp add: resolvents_upto_def)
lemma resolvent_at_fmember_resolvents_upto:
assumes "k ≠ 0"
shows "resolvent_at C D k |∈| resolvents_upto C D k"
unfolding resolvents_upto_def
proof (rule fimageI)
show "k |∈| fset_upto (Suc 0) k"
using assms by simp
qed
lemma backward_simulation_2_to_3:
fixes match measure less
defines "match ≡ ord_res_2_matches_ord_res_3"
assumes
match: "match S2 S3" and
step2: "ord_res_3_step S3 S3'"
shows "(∃S2'. ord_res_2_step⇧+⇧+ S2 S2' ∧ match S2' S3')"
using match[unfolded match_def]
proof (cases S2 S3 rule: ord_res_2_matches_ord_res_3.cases)
case match_hyps: (1 U⇩p⇩r N U⇩e⇩r U⇩e⇩f)
note invars = match_hyps(3-)
have U⇩p⇩r_spec: "∀C|∈|U⇩p⇩r. ∃D1|∈|N |∪| U⇩e⇩r |∪| U⇩e⇩f. ∃D2|∈|N |∪| U⇩e⇩r |∪| U⇩e⇩f.
(ground_resolution D1)⇧+⇧+ D2 C ∧ C ≠ eres D1 D2 ∧ eres D1 D2 |∈| U⇩e⇩r"
using invars by argo
hence C_not_least_with_partial: "¬ is_least_false_clause (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f) C"
if C_in: "C |∈| U⇩p⇩r" for C
proof -
obtain D1 D2 where
"D1 |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f" and
"D2 |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f" and
"(ground_resolution D1)⇧+⇧+ D2 C" and
"C ≠ eres D1 D2" and
"eres D1 D2 |∈| U⇩e⇩r"
using U⇩p⇩r_spec C_in by metis
have "eres D1 C = eres D1 D2"
using ‹(ground_resolution D1)⇧+⇧+ D2 C› eres_eq_after_tranclp_ground_resolution by metis
hence "eres D1 C ≺⇩c C"
using eres_le[of D1 C] ‹C ≠ eres D1 D2› by order
show ?thesis
proof (cases "ord_res_Interp (fset (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f)) (eres D1 D2) ⊫ eres D1 D2")
case True
then show ?thesis
by (metis (no_types, lifting) ‹(ground_resolution D1)⇧+⇧+ D2 C› ‹eres D1 C = eres D1 D2›
clause_true_if_eres_true is_least_false_clause_def
linorder_cls.is_least_in_fset_ffilterD(2))
next
case False
then show ?thesis
by (metis (mono_tags, lifting) Un_iff ‹eres D1 C = eres D1 D2› ‹eres D1 C ≺⇩c C›
‹eres D1 D2 |∈| U⇩e⇩r› is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
linorder_cls.not_less_iff_gr_or_eq sup_fset.rep_eq)
qed
qed
have least_false_conv: "is_least_false_clause (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f) =
is_least_false_clause (N |∪| U⇩e⇩r |∪| U⇩e⇩f)"
using invars is_least_false_clause_conv_if_partial_resolution_invariant by metis
have U⇩p⇩r_unproductive: "⋀N. ∀C |∈| U⇩p⇩r. ord_res.production N C = {}"
proof (intro ballI)
fix C
assume "C |∈| U⇩p⇩r"
hence "∃D |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f. (∃C'. ground_resolution D C C')"
using U⇩p⇩r_spec by (metis eres_eq_after_tranclp_ground_resolution resolvable_if_neq_eres)
hence "∄L. is_pos L ∧ ord_res.is_strictly_maximal_lit L C"
using nex_strictly_maximal_pos_lit_if_resolvable by metis
thus "⋀N. ord_res.production N C = {}"
using unproductive_if_nex_strictly_maximal_pos_lit by metis
qed
hence Interp_N_U⇩r_U⇩e⇩f_eq_Interp_N_U⇩e⇩r_U⇩e⇩f:
"⋀C. ord_res_Interp (fset (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f)) C = ord_res_Interp (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f)) C"
using Interp_union_unproductive[OF finite_fset finite_fset, folded union_fset,
of U⇩p⇩r "N |∪| U⇩e⇩r |∪| U⇩e⇩f"]
by (simp add: funion_left_commute sup_commute)
have U⇩p⇩r_have_generalization: "∀Ca |∈| U⇩p⇩r. ∃D |∈| U⇩e⇩r. D ≺⇩c Ca ∧ {D} ⊫e {Ca}"
proof (intro ballI)
fix Ca
assume "Ca |∈| U⇩p⇩r"
then obtain D1 D2 where
"D1|∈|N |∪| U⇩e⇩r |∪| U⇩e⇩f" and
"D2|∈|N |∪| U⇩e⇩r |∪| U⇩e⇩f" and
"(ground_resolution D1)⇧+⇧+ D2 Ca" and
"Ca ≠ eres D1 D2" and
"eres D1 D2 |∈| U⇩e⇩r"
using U⇩p⇩r_spec by metis
have "eres D1 D2 = eres D1 Ca"
using ‹(ground_resolution D1)⇧+⇧+ D2 Ca› eres_eq_after_tranclp_ground_resolution by metis
show "∃D |∈| U⇩e⇩r. D ≺⇩c Ca ∧ {D} ⊫e {Ca}"
proof (intro bexI conjI)
have "eres D1 Ca ≼⇩c Ca"
using eres_le .
thus "eres D1 D2 ≺⇩c Ca"
using ‹Ca ≠ eres D1 D2› ‹eres D1 D2 = eres D1 Ca› by order
next
show "{eres D1 D2} ⊫e {Ca}"
using ‹(ground_resolution D1)⇧+⇧+ D2 Ca› eres_entails_resolvent by metis
next
show "eres D1 D2 |∈| U⇩e⇩r"
using ‹eres D1 D2 |∈| U⇩e⇩r› by simp
qed
qed
from step2 obtain s3' where S3'_def: "S3' = (N, s3')" and "ord_res_3 N (U⇩e⇩r, U⇩e⇩f) s3'"
by (auto simp: match_hyps(1,2) elim: ord_res_3_step.cases)
show ?thesis
using ‹ord_res_3 N (U⇩e⇩r, U⇩e⇩f) s3'›
proof (cases N "(U⇩e⇩r, U⇩e⇩f)" s3' rule: ord_res_3.cases)
case (factoring C L U⇩e⇩f')
define S2' where
"S2' = (N, (U⇩p⇩r |∪| U⇩e⇩r, finsert (efac C) U⇩e⇩f))"
have "ord_res_2_step⇧+⇧+ S2 S2'"
unfolding match_hyps(1,2) S2'_def
proof (intro tranclp.r_into_trancl ord_res_2_step.intros ord_res_2.factoring)
show "is_least_false_clause (N |∪| (U⇩p⇩r |∪| U⇩e⇩r) |∪| U⇩e⇩f) C"
using ‹is_least_false_clause (N |∪| U⇩e⇩r |∪| U⇩e⇩f) C›
using least_false_conv
by (metis sup_assoc)
next
show "ord_res.is_maximal_lit L C"
using factoring by metis
next
show "is_pos L"
using factoring by metis
next
show "finsert (efac C) U⇩e⇩f = finsert (efac C) U⇩e⇩f"
by argo
qed
moreover have "match S2' S3'"
unfolding S2'_def S3'_def
unfolding factoring
unfolding match_def
proof (rule ord_res_2_matches_ord_res_3.intros)
show "∀Ca|∈|U⇩p⇩r.
∃D1|∈|N |∪| U⇩e⇩r |∪| finsert (efac C) U⇩e⇩f. ∃D2|∈|N |∪| U⇩e⇩r |∪| finsert (efac C) U⇩e⇩f.
(ground_resolution D1)⇧+⇧+ D2 Ca ∧ Ca ≠ eres D1 D2 ∧ eres D1 D2 |∈| U⇩e⇩r"
using U⇩p⇩r_spec by auto
qed
ultimately show ?thesis
by metis
next
case (resolution C L D U⇩r⇩r')
have "(ground_resolution D)⇧*⇧* C (eres D C)" "∄x. ground_resolution D (eres D C) x"
unfolding atomize_conj
by (metis ex1_eres_eq_full_run_ground_resolution full_run_def)
moreover have "∃x. ground_resolution D C x"
unfolding ground_resolution_def
using resolution
by (metis Neg_atm_of_iff ex_ground_resolutionI ord_res.mem_productionE singletonI)
ultimately have "(ground_resolution D)⇧+⇧+ C (eres D C)"
by (metis rtranclpD)
then obtain n where "(ground_resolution D ^^ Suc n) C (eres D C)"
by (metis not0_implies_Suc not_gr_zero tranclp_power)
hence "resolvent_at D C (Suc n) = eres D C"
by (metis Uniq_ground_resolution Uniq_relpowp resolvent_at_def the1_equality')
have steps: "k ≤ Suc n ⟹ (ord_res_2_step ^^ k)
(N, U⇩p⇩r |∪| U⇩e⇩r, U⇩e⇩f) (N, U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C k, U⇩e⇩f)" for k
proof (induction k)
case 0
show ?case
by simp
next
case (Suc k)
have "k < Suc n"
using ‹Suc k ≤ Suc n› by presburger
hence "k ≤ Suc n"
by presburger
hence "(ord_res_2_step ^^ k) (N, U⇩p⇩r |∪| U⇩e⇩r, U⇩e⇩f)
(N, U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C k, U⇩e⇩f)"
using Suc.IH by metis
moreover have "ord_res_2_step
(N, U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C k, U⇩e⇩f)
(N, U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc k), U⇩e⇩f)"
unfolding resolvents_upto_Suc
proof (intro ord_res_2_step.intros ord_res_2.resolution)
show "is_least_false_clause (N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C k) |∪| U⇩e⇩f)
(resolvent_at D C k)"
using ‹k < Suc n›
proof (induction k)
case 0
have "is_least_false_clause (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f) C"
using ‹is_least_false_clause (N |∪| U⇩e⇩r |∪| U⇩e⇩f) C›
unfolding least_false_conv .
thus ?case
unfolding funion_fempty_right funion_assoc[symmetric]
by simp
next
case (Suc k')
have "⋀x. ord_res_Interp (fset (N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc k')) |∪| U⇩e⇩f)) x =
ord_res_Interp (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f) ∪ fset (U⇩p⇩r |∪| resolvents_upto D C (Suc k'))) x"
by (simp add: funion_left_commute sup_assoc sup_commute)
also have "⋀x. ord_res_Interp (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f) ∪ fset (U⇩p⇩r |∪| resolvents_upto D C (Suc k'))) x =
ord_res_Interp (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f)) x"
proof (intro Interp_union_unproductive ballI)
fix x y assume "y |∈| U⇩p⇩r |∪| resolvents_upto D C (Suc k')"
hence "y |∈| U⇩p⇩r ∨ y |∈| resolvents_upto D C (Suc k')"
by blast
thus "ord_res.production (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f) ∪ fset (U⇩p⇩r |∪| resolvents_upto D C (Suc k'))) y = {}"
proof (elim disjE)
assume "y |∈| U⇩p⇩r"
thus ?thesis
using U⇩p⇩r_unproductive by metis
next
assume "y |∈| resolvents_upto D C (Suc k')"
then obtain i where "i |∈| fset_upto (Suc 0) (Suc k')" and "y = resolvent_at D C i"
unfolding resolvents_upto_def by blast
have "∄L. is_pos L ∧ ord_res.is_strictly_maximal_lit L (resolvent_at D C i)"
proof (rule nex_pos_strictly_max_lit_in_resolvent_at)
show "(ground_resolution D ^^ Suc n) C (eres D C)"
using ‹(ground_resolution D ^^ Suc n) C (eres D C)› .
next
have "i ≤ Suc k'"
using ‹i |∈| fset_upto (Suc 0) (Suc k')› by auto
thus "i < Suc n"
using ‹Suc k' < Suc n› by presburger
qed
then show ?thesis
using ‹y = resolvent_at D C i› unproductive_if_nex_strictly_maximal_pos_lit
by metis
qed
qed simp_all
finally have Interp_simp: "⋀x.
ord_res_Interp (fset (N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc k')) |∪| U⇩e⇩f)) x =
ord_res_Interp (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f)) x" .
show ?case
unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
proof (intro conjI ballI impI)
have "resolvent_at D C (Suc k') |∈| resolvents_upto D C (Suc k')"
using resolvent_at_fmember_resolvents_upto by simp
thus "resolvent_at D C (Suc k') |∈| N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc k')) |∪| U⇩e⇩f"
by simp
next
show "¬ ord_res_Interp (fset (N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc k')) |∪| U⇩e⇩f))
(resolvent_at D C (Suc k')) ⊫ resolvent_at D C (Suc k')"
unfolding Interp_simp
by (metis (no_types, lifting) Suc.prems Zero_not_Suc
‹(ground_resolution D ^^ Suc n) C (eres D C)› clause_true_if_resolved_true
insert_not_empty is_least_false_clause_def
linorder_cls.is_least_in_fset_ffilterD(2) local.resolution(2) local.resolution(7)
relpowp_to_resolvent_at tranclp_if_relpowp)
next
fix y
assume "y ≠ resolvent_at D C (Suc k')"
assume "¬ ord_res_Interp (fset (N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc k')) |∪| U⇩e⇩f)) y ⊫ y"
hence "¬ ord_res_Interp (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f)) y ⊫ y"
unfolding Interp_simp .
hence "¬ ord_res_Interp (fset (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f)) y ⊫ y"
using Interp_N_U⇩r_U⇩e⇩f_eq_Interp_N_U⇩e⇩r_U⇩e⇩f by metis
assume "y |∈| N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc k')) |∪| U⇩e⇩f"
hence "y |∈| N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f ∨ y |∈| resolvents_upto D C (Suc k')"
by auto
thus "resolvent_at D C (Suc k') ≺⇩c y"
proof (elim disjE)
assume "y |∈| N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f"
have "C ≼⇩c y"
proof (cases "y = C")
case True
thus ?thesis
by order
next
case False
thus ?thesis
using ‹y |∈| N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f›
using ‹¬ ord_res_Interp (fset (N |∪| U⇩p⇩r |∪| U⇩e⇩r |∪| U⇩e⇩f)) y ⊫ y›
using ‹is_least_false_clause (N |∪| U⇩e⇩r |∪| U⇩e⇩f) C›
unfolding least_false_conv[symmetric]
unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
by simp
qed
moreover have "resolvent_at D C (Suc k') ≺⇩c C"
by (metis Suc.prems ‹(ground_resolution D ^^ Suc n) C (eres D C)› less_or_eq_imp_le
resolvent_at_less_cls_resolvent_at resolvent_at_0 zero_less_Suc)
ultimately show "resolvent_at D C (Suc k') ≺⇩c y"
by order
next
assume "y |∈| resolvents_upto D C (Suc k')"
then obtain i where
i_in: "i |∈| fset_upto (Suc 0) (Suc k')" and y_def: "y = resolvent_at D C i"
unfolding resolvents_upto_def by blast
hence "i < Suc k'"
using ‹y ≠ resolvent_at D C (Suc k')›
by auto
show "resolvent_at D C (Suc k') ≺⇩c y"
unfolding y_def
proof (rule resolvent_at_less_cls_resolvent_at)
show "(ground_resolution D ^^ Suc n) C (eres D C)"
using ‹(ground_resolution D ^^ Suc n) C (eres D C)› .
next
show "i < Suc k'"
using ‹y ≠ resolvent_at D C (Suc k')› i_in y_def by auto
next
show "Suc k' ≤ Suc n"
using ‹Suc k' < Suc n› by presburger
qed
qed
qed
qed
next
show "ord_res.is_maximal_lit L (resolvent_at D C k)"
proof (rule max_lit_resolvent_at)
show "(ground_resolution D ^^ Suc n) C (eres D C)"
using ‹(ground_resolution D ^^ Suc n) C (eres D C)› .
next
show "k < Suc n"
using ‹k < Suc n› .
next
show "ord_res.is_maximal_lit L C"
using ‹ord_res.is_maximal_lit L C› .
qed
next
show "is_neg L"
using ‹is_neg L› .
next
show "D |∈| N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C k) |∪| U⇩e⇩f"
using ‹D |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f› by auto
next
show "D ≺⇩c resolvent_at D C k"
proof (rule left_premisse_lt_resolvent_at)
show "(ground_resolution D ^^ Suc n) C (eres D C)"
using ‹(ground_resolution D ^^ Suc n) C (eres D C)› .
next
show "k < Suc n"
using ‹k < Suc n› .
qed
next
have "ord_res.production (fset (N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C k) |∪| U⇩e⇩f)) D =
ord_res.production (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f) ∪ fset (U⇩p⇩r |∪| resolvents_upto D C k)) D"
by (simp add: funion_left_commute sup_assoc sup_commute)
also have "… = ord_res.production (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f)) D"
proof (intro production_union_unproductive ballI)
fix x
assume "x |∈| U⇩p⇩r |∪| resolvents_upto D C k"
hence "∄L. is_pos L ∧ ord_res.is_strictly_maximal_lit L x"
unfolding funion_iff
proof (elim disjE)
assume "x |∈| U⇩p⇩r"
thus ?thesis
using U⇩p⇩r_spec
by (metis eres_eq_after_tranclp_ground_resolution nex_strictly_maximal_pos_lit_if_neq_eres)
next
assume "x |∈| resolvents_upto D C k"
then obtain i where "i |∈| fset_upto (Suc 0) k" and x_def: "x = resolvent_at D C i"
unfolding resolvents_upto_def by auto
have "0 < i" and "i ≤ k"
using ‹i |∈| fset_upto (Suc 0) k› by simp_all
show ?thesis
unfolding x_def
proof (rule nex_pos_strictly_max_lit_in_resolvent_at)
show "(ground_resolution D ^^ Suc n) C (eres D C)"
using ‹(ground_resolution D ^^ Suc n) C (eres D C)› .
next
show "i < Suc n"
using ‹i ≤ k› ‹k < Suc n› by presburger
qed
qed
thus "ord_res.production (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f) ∪
fset (U⇩p⇩r |∪| resolvents_upto D C k)) x = {}"
using unproductive_if_nex_strictly_maximal_pos_lit by metis
next
show "D |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f"
using ‹D |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f› .
qed simp_all
finally show "ord_res.production (fset (N |∪| (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C k) |∪| U⇩e⇩f)) D =
{atm_of L}"
using ‹ord_res.production (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f)) D = {atm_of L}› by argo
next
show "ord_res.ground_resolution (resolvent_at D C k) D (resolvent_at D C (Suc k))"
unfolding ground_resolution_def[symmetric]
proof (rule ground_resolution_resolvent_at_resolvent_at_Suc)
show "(ground_resolution D ^^ Suc n) C (eres D C)"
using ‹(ground_resolution D ^^ Suc n) C (eres D C)› .
next
show "k < Suc n"
using ‹k < Suc n› .
qed
next
show "U⇩p⇩r |∪| U⇩e⇩r |∪| finsert (resolvent_at D C (Suc k)) (resolvents_upto D C k) =
finsert (resolvent_at D C (Suc k)) (U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C k)"
by simp
qed
ultimately show ?case
by (meson relpowp_Suc_I)
qed
hence "(ord_res_2_step ^^ Suc n) S2 (N, U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc n), U⇩e⇩f)"
unfolding match_hyps(1,2) by blast
moreover have "match (N, U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc n), U⇩e⇩f) S3'"
proof -
have 1: "S3' = (N, finsert (eres D C) U⇩e⇩r, U⇩e⇩f)"
unfolding S3'_def ‹s3' = (U⇩r⇩r', U⇩e⇩f)› ‹U⇩r⇩r' = finsert (eres D C) U⇩e⇩r› ..
have 2: "U⇩p⇩r |∪| U⇩e⇩r |∪| resolvents_upto D C (Suc n) =
U⇩p⇩r |∪| resolvents_upto D C n |∪| finsert (eres D C) U⇩e⇩r"
by (auto simp: ‹resolvent_at D C (Suc n) = eres D C›)
show ?thesis
unfolding match_def 1 2
proof (rule ord_res_2_matches_ord_res_3.intros)
show "∀E|∈|U⇩p⇩r |∪| resolvents_upto D C n.
∃D1|∈|N |∪| finsert (eres D C) U⇩e⇩r |∪| U⇩e⇩f. ∃D2|∈|N |∪| finsert (eres D C) U⇩e⇩r |∪| U⇩e⇩f.
(ground_resolution D1)⇧+⇧+ D2 E ∧ E ≠ eres D1 D2 ∧ eres D1 D2 |∈| finsert (eres D C) U⇩e⇩r"
proof (intro ballI)
fix Ca
assume "Ca |∈| U⇩p⇩r |∪| resolvents_upto D C n"
hence "Ca |∈| U⇩p⇩r ∨ Ca |∈| resolvents_upto D C n"
by simp
thus "∃D1|∈|N |∪| finsert (eres D C) U⇩e⇩r |∪| U⇩e⇩f. ∃D2|∈|N |∪| finsert (eres D C) U⇩e⇩r |∪| U⇩e⇩f.
(ground_resolution D1)⇧+⇧+ D2 Ca ∧ Ca ≠ eres D1 D2 ∧ eres D1 D2 |∈| finsert (eres D C) U⇩e⇩r"
proof (elim disjE)
show "Ca |∈| U⇩p⇩r ⟹ ?thesis"
using U⇩p⇩r_spec by auto
next
assume "Ca |∈| resolvents_upto D C n"
then obtain i where i_in: "i |∈| fset_upto (Suc 0) n" and Ca_def:"Ca = resolvent_at D C i"
unfolding resolvents_upto_def by auto
from i_in have "0 < i" "i ≤ n"
by simp_all
show "?thesis"
proof (intro bexI conjI)
have "(ground_resolution D ^^ i) C Ca"
unfolding ‹Ca = resolvent_at D C i›
proof (rule relpowp_to_resolvent_at)
show "(ground_resolution D ^^ Suc n) C (eres D C)"
using ‹(ground_resolution D ^^ Suc n) C (eres D C)› .
next
show "i < Suc n"
using ‹i ≤ n› by presburger
qed
thus "(ground_resolution D)⇧+⇧+ C Ca"
using ‹0 < i› by (simp add: tranclp_if_relpowp)
next
show "Ca ≠ eres D C"
by (metis Ca_def ‹(ground_resolution D ^^ Suc n) C (eres D C)›
‹∄x. ground_resolution D (eres D C) x› ‹i ≤ n›
ground_resolution_resolvent_at_resolvent_at_Suc less_Suc_eq_le)
next
show "eres D C |∈| finsert (eres D C) U⇩e⇩r"
by simp
next
show "D |∈| N |∪| finsert (eres D C) U⇩e⇩r |∪| U⇩e⇩f"
using resolution by simp
next
have "C |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f"
using resolution
by (simp add: is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff)
thus "C |∈| N |∪| finsert (eres D C) U⇩e⇩r |∪| U⇩e⇩f"
by simp
qed
qed
qed
qed
qed
ultimately have "∃S2'. (ord_res_2_step ^^ Suc n) S2 S2' ∧ match S2' S3'"
by metis
thus "∃S2'. ord_res_2_step⇧+⇧+ S2 S2' ∧ match S2' S3'"
by (metis Zero_neq_Suc tranclp_if_relpowp)
qed
qed
lemma safe_states_if_ord_res_2_matches_ord_res_3:
assumes match: "ord_res_2_matches_ord_res_3 S⇩2 S⇩3"
shows
"safe_state ord_res_2_step ord_res_2_final S⇩2"
"safe_state ord_res_3_step ord_res_3_final S⇩3"
proof -
show "safe_state ord_res_2_step ord_res_2_final S⇩2"
using safe_state_if_all_states_safe ord_res_2_step_safe by metis
show "safe_state ord_res_3_step ord_res_3_final S⇩3"
using safe_state_if_all_states_safe ord_res_3_step_safe by metis
qed
theorem bisimulation_ord_res_2_ord_res_3:
defines "match ≡ λ_ S2 S3. ord_res_2_matches_ord_res_3 S2 S3"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_2_state ⇒ 'f ord_res_3_state ⇒ bool) ℛ.
bisimulation ord_res_2_step ord_res_3_step ord_res_2_final ord_res_3_final ℛ MATCH"
proof (rule ex_bisimulation_from_backward_simulation)
show "right_unique ord_res_2_step"
using right_unique_ord_res_2_step .
next
show "right_unique ord_res_3_step"
using right_unique_ord_res_3_step .
next
show "∀s1. ord_res_2_final s1 ⟶ (∄s1'. ord_res_2_step s1 s1')"
by (metis finished_def ord_res_2_semantics.final_finished)
next
show "∀s2. ord_res_3_final s2 ⟶ (∄s2'. ord_res_3_step s2 s2')"
by (metis finished_def ord_res_3_semantics.final_finished)
next
show "∀i s1 s2. match i s1 s2 ⟶ ord_res_2_final s1 = ord_res_3_final s2"
unfolding match_def
using ord_res_2_final_iff_ord_res_3_final by metis
next
show "∀i s1 s2. match i s1 s2 ⟶
safe_state ord_res_2_step ord_res_2_final s1 ∧ safe_state ord_res_3_step ord_res_3_final s2"
unfolding match_def
using safe_states_if_ord_res_2_matches_ord_res_3 by metis
next
show "wfP (λ_ _. False)"
by simp
next
show "∀i s1 s2 s2'.
match i s1 s2 ⟶
ord_res_3_step s2 s2' ⟶
(∃i' s1'. ord_res_2_step⇧+⇧+ s1 s1' ∧ match i' s1' s2') ∨ (∃i'. match i' s1 s2' ∧ False)"
unfolding match_def
using backward_simulation_2_to_3 by metis
qed
end
section ‹ORD-RES-4 (implicit factorization)›
type_synonym 'f ord_res_4_state = "'f gclause fset × 'f gclause fset × 'f gclause fset"
context simulation_SCLFOL_ground_ordered_resolution begin
inductive ord_res_3_matches_ord_res_4 :: "'f ord_res_3_state ⇒ 'f ord_res_4_state ⇒ bool" where
"ℱ |⊆| N |∪| U⇩e⇩r ⟹ U⇩e⇩f = iefac ℱ |`| {|C |∈| N |∪| U⇩e⇩r. iefac ℱ C ≠ C|} ⟹
ord_res_3_matches_ord_res_4 (N, (U⇩e⇩r, U⇩e⇩f)) (N, U⇩e⇩r, ℱ)"
lemma ord_res_3_final_iff_ord_res_4_final:
assumes match: "ord_res_3_matches_ord_res_4 S3 S4"
shows "ord_res_3_final S3 ⟷ ord_res_4_final S4"
using match
proof (cases S3 S4 rule: ord_res_3_matches_ord_res_4.cases)
case match_hyps: (1 ℱ N U⇩e⇩r U⇩e⇩f)
note invars = match_hyps(3-)
have "{#} |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f ⟷ {#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using invars by (auto simp: iefac_def)
moreover have "ex_false_clause (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f)) ⟷
ex_false_clause (fset (iefac ℱ |`| (N |∪| U⇩e⇩r)))"
unfolding ex_false_clause_iff
unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF invars(2)]
unfolding is_least_false_clause_with_iefac_conv ..
ultimately have "ord_res_final (N |∪| U⇩e⇩r |∪| U⇩e⇩f) ⟷ ord_res_final (iefac ℱ |`| (N |∪| U⇩e⇩r))"
unfolding ord_res_final_def by argo
thus ?thesis
unfolding match_hyps(1,2)
by (simp add: ord_res_3_final.simps ord_res_4_final.simps)
qed
lemma forward_simulation_between_3_and_4:
assumes
match: "ord_res_3_matches_ord_res_4 S3 S4" and
step: "ord_res_3_step S3 S3'"
shows "(∃S4'. ord_res_4_step⇧+⇧+ S4 S4' ∧ ord_res_3_matches_ord_res_4 S3' S4')"
using match
proof (cases S3 S4 rule: ord_res_3_matches_ord_res_4.cases)
case match_hyps: (1 ℱ N U⇩e⇩r U⇩e⇩f)
note match_invars = match_hyps(3-)
from step obtain s3' where step': "ord_res_3 N (U⇩e⇩r, U⇩e⇩f) s3'" and "S3' = (N, s3')"
unfolding match_hyps(1,2)
by (auto elim: ord_res_3_step.cases)
from step' show ?thesis
proof (cases N "(U⇩e⇩r, U⇩e⇩f)" s3' rule: ord_res_3.cases)
case (factoring C L U⇩e⇩f')
have "¬ ord_res.is_strictly_maximal_lit L C"
using ‹is_least_false_clause (N |∪| U⇩e⇩r |∪| U⇩e⇩f) C› ‹ord_res.is_maximal_lit L C› ‹is_pos L›
by (metis (no_types, lifting) is_least_false_clause_def is_pos_def
pos_lit_not_greatest_if_maximal_in_least_false_clause)
have "C |∈| N |∪| U⇩e⇩r"
proof -
have "C |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f"
using ‹is_least_false_clause (N |∪| U⇩e⇩r |∪| U⇩e⇩f) C›
by (simp add: is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff)
moreover have "C |∉| U⇩e⇩f"
proof (rule notI)
assume "C |∈| U⇩e⇩f"
then obtain C⇩0 where "C = iefac ℱ C⇩0" and "C⇩0 |∈| N |∪| U⇩e⇩r" and "iefac ℱ C⇩0 ≠ C⇩0"
using match_invars(2) by force
then show False
by (metis Uniq_D ‹¬ ord_res.is_strictly_maximal_lit L C› iefac_def
linorder_lit.Uniq_is_maximal_in_mset
linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset local.factoring(3)
obtains_positive_greatest_lit_if_efac_not_ident)
qed
ultimately show ?thesis
by simp
qed
show ?thesis
proof (intro exI conjI)
show "ord_res_4_step⇧+⇧+ S4 (N, U⇩e⇩r, finsert C ℱ)"
unfolding match_hyps(1,2)
proof (intro tranclp.r_into_trancl ord_res_4_step.intros ord_res_4.factoring)
have "is_least_false_clause (N |∪| U⇩e⇩r |∪| U⇩e⇩f) C"
using factoring by argo
hence "is_least_false_clause (N |∪| U⇩e⇩r |∪| iefac ℱ |`| (N |∪| U⇩e⇩r)) C"
unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF match_invars(2)] .
thus "is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r)) C"
unfolding is_least_false_clause_with_iefac_conv .
next
show "ord_res.is_maximal_lit L C"
using ‹ord_res.is_maximal_lit L C› .
next
show "is_pos L"
using ‹is_pos L› .
qed (rule refl)+
next
show "ord_res_3_matches_ord_res_4 S3' (N, U⇩e⇩r, finsert C ℱ)"
unfolding ‹S3' = (N, s3')› ‹s3' = (U⇩e⇩r, U⇩e⇩f')› ‹U⇩e⇩f' = finsert (efac C) U⇩e⇩f›
proof (rule ord_res_3_matches_ord_res_4.intros)
show "finsert C ℱ |⊆| N |∪| U⇩e⇩r"
using match_invars ‹C |∈| N |∪| U⇩e⇩r› by simp
next
have "∃C'. ord_res.ground_factoring C C'"
using ‹ord_res.is_maximal_lit L C› ‹is_pos L›
by (metis ‹¬ ord_res.is_strictly_maximal_lit L C› ex_ground_factoringI is_pos_def)
hence "efac C ≠ C"
by (metis ex1_efac_eq_factoring_chain)
hence "iefac (finsert C ℱ) C ≠ C"
by (simp add: iefac_def)
have "{|Ca |∈| N |∪| U⇩e⇩r. iefac (finsert C ℱ) Ca ≠ Ca|} =
finsert C {|Ca |∈| N |∪| U⇩e⇩r. iefac ℱ Ca ≠ Ca|}"
proof (intro fsubset_antisym fsubsetI)
fix x
assume "x |∈| {|Ca |∈| N |∪| U⇩e⇩r. iefac (finsert C ℱ) Ca ≠ Ca|}"
hence "x |∈| N |∪| U⇩e⇩r" and "iefac (finsert C ℱ) x ≠ x"
by simp_all
then show "x |∈| finsert C {|Ca |∈| N |∪| U⇩e⇩r. iefac ℱ Ca ≠ Ca|}"
by (smt (verit, best) ffmember_filter finsert_iff iefac_def)
next
fix x
assume "x |∈| finsert C {|Ca |∈| N |∪| U⇩e⇩r. iefac ℱ Ca ≠ Ca|}"
hence "x = C ∨ x |∈| N |∪| U⇩e⇩r ∧ iefac ℱ x ≠ x"
by auto
thus "x |∈| {|Ca |∈| N |∪| U⇩e⇩r. iefac (finsert C ℱ) Ca ≠ Ca|}"
proof (elim disjE conjE)
assume "x = C"
thus ?thesis
using ‹C |∈| N |∪| U⇩e⇩r› ‹iefac (finsert C ℱ) C ≠ C› by auto
next
assume "x |∈| N |∪| U⇩e⇩r" and "iefac ℱ x ≠ x"
thus ?thesis
by (smt (verit, best) ffmember_filter finsertCI iefac_def)
qed
qed
thus "finsert (efac C) U⇩e⇩f =
iefac (finsert C ℱ) |`| {|Ca |∈| N |∪| U⇩e⇩r. iefac (finsert C ℱ) Ca ≠ Ca|}"
using iefac_def match_invars(2) by auto
qed
qed
next
case (resolution C L D U⇩r⇩r')
have "D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
proof -
have "D |∈| N |∪| U⇩e⇩r |∪| U⇩e⇩f"
using resolution by argo
hence "D |∈| N |∪| U⇩e⇩r |∪| iefac ℱ |`| (N |∪| U⇩e⇩r)"
unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF match_invars(2)] .
moreover have "D |∉| N |∪| U⇩e⇩r - iefac ℱ |`| (N |∪| U⇩e⇩r)"
by (metis clauses_for_iefac_are_unproductive insert_not_empty local.resolution(7))
ultimately show ?thesis
by blast
qed
show ?thesis
proof (intro exI conjI)
show "ord_res_4_step⇧+⇧+ S4 (N, finsert (eres D C) U⇩e⇩r, ℱ)"
unfolding match_hyps(1,2)
proof (intro tranclp.r_into_trancl ord_res_4_step.intros ord_res_4.resolution)
have "is_least_false_clause (N |∪| U⇩e⇩r |∪| U⇩e⇩f) C"
using resolution by argo
hence "is_least_false_clause (N |∪| U⇩e⇩r |∪| iefac ℱ |`| (N |∪| U⇩e⇩r)) C"
unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF match_invars(2)] .
thus "is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r)) C"
unfolding is_least_false_clause_with_iefac_conv .
next
show "ord_res.is_maximal_lit L C"
using resolution by argo
next
show "is_neg L"
using resolution by argo
next
show "D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using ‹D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)› .
next
show "D ≺⇩c C"
using resolution by argo
next
have "ord_res.production (fset (N |∪| U⇩e⇩r |∪| U⇩e⇩f)) D =
ord_res.production (fset (N |∪| U⇩e⇩r |∪| iefac ℱ |`| (N |∪| U⇩e⇩r))) D"
unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF match_invars(2)] ..
also have "… = ord_res.production (fset (iefac ℱ |`| (N |∪| U⇩e⇩r)) ∪ fset (N |∪| U⇩e⇩r)) D"
by (simp add: sup.commute)
also have "… = ord_res.production (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) D"
proof (rule production_union_unproductive_strong)
show "∀x ∈ fset (N |∪| U⇩e⇩r) - fset (iefac ℱ |`| (N |∪| U⇩e⇩r)).
ord_res.production (fset (iefac ℱ |`| (N |∪| U⇩e⇩r)) ∪ fset (N |∪| U⇩e⇩r)) x = {}"
using clauses_for_iefac_are_unproductive[of "N |∪| U⇩e⇩r" ℱ] by simp
next
show "D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using ‹D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)› .
qed (rule finite_fset)+
finally show "ord_res.production (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) D = {atm_of L}"
using resolution by argo
qed (rule refl)+
next
show "ord_res_3_matches_ord_res_4 S3' (N, finsert (eres D C) U⇩e⇩r, ℱ)"
unfolding ‹S3' = (N, s3')› ‹s3' = (U⇩r⇩r', U⇩e⇩f)› ‹U⇩r⇩r' = finsert (eres D C) U⇩e⇩r›
proof (rule ord_res_3_matches_ord_res_4.intros)
show "ℱ |⊆| N |∪| finsert (eres D C) U⇩e⇩r"
using match_invars by auto
next
show "U⇩e⇩f = iefac ℱ |`| {|C |∈| N |∪| finsert (eres D C) U⇩e⇩r. iefac ℱ C ≠ C|} "
proof (cases "eres D C |∈| ℱ")
case True
then show ?thesis
using ‹ℱ |⊆| N |∪| U⇩e⇩r›
using match_invars by force
next
case False
hence "iefac ℱ (eres D C) = eres D C"
by (simp add: iefac_def)
hence "{|C |∈| N |∪| finsert (eres D C) U⇩e⇩r. iefac ℱ C ≠ C|} = {|C |∈| N |∪| U⇩e⇩r. iefac ℱ C ≠ C|}"
using ffilter_eq_ffilter_minus_singleton by auto
thus ?thesis
using match_invars by argo
qed
qed
qed
qed
qed
theorem bisimulation_ord_res_3_ord_res_4:
defines "match ≡ λ_ S3 S4. ord_res_3_matches_ord_res_4 S3 S4"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_3_state ⇒ 'f ord_res_4_state ⇒ bool) ℛ.
bisimulation ord_res_3_step ord_res_4_step ord_res_3_final ord_res_4_final ℛ MATCH"
proof (rule ex_bisimulation_from_forward_simulation)
show "right_unique ord_res_3_step"
using right_unique_ord_res_3_step .
next
show "right_unique ord_res_4_step"
using right_unique_ord_res_4_step .
next
show "∀s1. ord_res_3_final s1 ⟶ (∄s1'. ord_res_3_step s1 s1')"
by (metis finished_def ord_res_3_semantics.final_finished)
next
show "∀s2. ord_res_4_final s2 ⟶ (∄s2'. ord_res_4_step s2 s2')"
by (metis finished_def ord_res_4_semantics.final_finished)
next
show "∀i s1 s2. match i s1 s2 ⟶ ord_res_3_final s1 ⟷ ord_res_4_final s2"
unfolding match_def
using ord_res_3_final_iff_ord_res_4_final by metis
next
show "∀i s1 s2. match i s1 s2 ⟶
safe_state ord_res_3_step ord_res_3_final s1 ∧ safe_state ord_res_4_step ord_res_4_final s2"
using ord_res_3_step_safe ord_res_4_step_safe
by (simp add: safe_state_if_all_states_safe)
next
show "wfP (λi' i. False)"
by simp
next
show "∀i s1 s2 s1'. match i s1 s2 ⟶ ord_res_3_step s1 s1' ⟶
(∃i' s2'. ord_res_4_step⇧+⇧+ s2 s2' ∧ match i' s1' s2') ∨ (∃i'. match i' s1' s2 ∧ False)"
unfolding match_def
using forward_simulation_between_3_and_4 by metis
qed
end
section ‹ORD-RES-5 (explicit model construction)›
type_synonym 'f ord_res_5_state = "'f gclause fset × 'f gclause fset × 'f gclause fset ×
('f gterm ⇒ 'f gclause option) × 'f gclause option"
context simulation_SCLFOL_ground_ordered_resolution begin
inductive ord_res_4_matches_ord_res_5 :: "'f ord_res_4_state ⇒ 'f ord_res_5_state ⇒ bool" where
"ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞) ⟹
(∀C. 𝒞 = Some C ⟷ is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r)) C) ⟹
ord_res_4_matches_ord_res_5 (N, U⇩e⇩r, ℱ) (N, U⇩e⇩r, ℱ, ℳ, 𝒞)"
lemma ord_res_4_final_iff_ord_res_5_final:
assumes match: "ord_res_4_matches_ord_res_5 S4 S5"
shows "ord_res_4_final S4 ⟷ ord_res_5_final S5"
using match
proof (cases S4 S5 rule: ord_res_4_matches_ord_res_5.cases)
case match_hyps: (1 N U⇩e⇩r ℱ ℳ 𝒞)
show ?thesis
unfolding match_hyps(1,2,3)
proof (intro iffI ord_res_5_final.intros)
assume "ord_res_4_final (N, U⇩e⇩r, ℱ)"
hence "{#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r) ∨ ¬ ex_false_clause (fset (iefac ℱ |`| (N |∪| U⇩e⇩r)))"
by (simp add: ord_res_4_final.simps ord_res_final_def)
thus "ord_res_5_final (N, U⇩e⇩r, ℱ, ℳ, 𝒞)"
proof (elim disjE)
assume "{#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
hence "is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r)) {#}"
using is_least_false_clause_mempty by metis
hence "𝒞 = Some {#}"
by (smt (verit) all_smaller_clauses_true_wrt_respective_Interp_def is_least_false_clause_def
linorder_cls.is_least_in_ffilter_iff linorder_cls.le_imp_less_or_eq match_hyps(3)
mempty_lesseq_cls ord_res_5_invars_def)
thus ?thesis
using ord_res_5_final.contradiction_found by metis
next
assume "¬ ex_false_clause (fset (iefac ℱ |`| (N |∪| U⇩e⇩r)))"
hence "𝒞 = None"
using match_hyps(2-)
by (metis ex_false_clause_if_least_false_clause option.exhaust)
thus ?thesis
using ord_res_5_final.model_found by metis
qed
next
assume "ord_res_5_final (N, U⇩e⇩r, ℱ, ℳ, 𝒞)"
thus "ord_res_4_final (N, U⇩e⇩r, ℱ)"
proof (cases "(N, U⇩e⇩r, ℱ, ℳ, 𝒞)" rule: ord_res_5_final.cases)
case model_found
have "all_smaller_clauses_true_wrt_respective_Interp N (U⇩e⇩r, ℱ, ℳ, 𝒞)"
using ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
unfolding ord_res_5_invars_def by metis
hence "∀C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). ord_res_Interp (iefac ℱ ` (fset N ∪ fset U⇩e⇩r)) C ⊫ C"
by (simp add: model_found all_smaller_clauses_true_wrt_respective_Interp_def)
hence "¬ ex_false_clause (fset (iefac ℱ |`| (N |∪| U⇩e⇩r)))"
by (simp add: ex_false_clause_def)
then show ?thesis
by (metis ord_res_4_final.intros ord_res_final_def)
next
case contradiction_found
hence "{#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
by (metis next_clause_in_factorized_clause_def ord_res_5_invars_def)
then show ?thesis
by (metis ord_res_4_final.intros ord_res_final_def)
qed
qed
qed
lemma forward_simulation_between_4_and_5:
fixes S4 S4' S5
assumes match: "ord_res_4_matches_ord_res_5 S4 S5" and step: "ord_res_4_step S4 S4'"
shows "∃S5'. ord_res_5_step⇧+⇧+ S5 S5' ∧ ord_res_4_matches_ord_res_5 S4' S5'"
using match
proof (cases S4 S5 rule: ord_res_4_matches_ord_res_5.cases)
case match_hyps: (1 N U⇩e⇩r ℱ ℳ 𝒞)
hence
S4_def: "S4 = (N, U⇩e⇩r, ℱ)" and
S5_def: "S5 = (N, U⇩e⇩r, ℱ, ℳ, 𝒞)"
unfolding atomize_conj by metis
have dom_ℳ_eq: "⋀C. 𝒞 = Some C ⟹ dom ℳ = ord_res.interp (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) C"
using match_hyps unfolding ord_res_5_invars_def model_eq_interp_upto_next_clause_def by simp
obtain s4' where S4'_def: "S4' = (N, s4')" and step': "ord_res_4 N (U⇩e⇩r, ℱ) s4'"
using step unfolding S4_def by (auto simp: ord_res_4_step.simps)
show ?thesis
using step'
proof (cases N "(U⇩e⇩r, ℱ)" s4' rule: ord_res_4.cases)
case step_hyps: (factoring NN C L ℱ')
have "𝒞 = Some C"
using match_hyps(3-) step_hyps by metis
define ℳ' :: "'f gterm ⇒ 'f gterm literal multiset option" where
"ℳ' = (λ_. None)"
define 𝒞' :: "'f gclause option" where
"𝒞' = The_optional (linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| U⇩e⇩r)))"
have ord_res_5_step: "ord_res_5 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ', ℳ', 𝒞')"
proof (rule ord_res_5.factoring)
have "is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r)) C"
using step_hyps by argo
then show "¬ dom ℳ ⊫ C"
using dom_ℳ_eq[OF ‹𝒞 = Some C›]
by (metis (mono_tags, lifting) is_least_false_clause_def
linorder_cls.is_least_in_ffilter_iff ord_res_Interp_entails_if_greatest_lit_is_pos
unproductive_if_nex_strictly_maximal_pos_lit sup_bot.right_neutral)
next
show "ord_res.is_maximal_lit L C"
using step_hyps by metis
next
show "is_pos L"
using step_hyps by metis
next
show "¬ ord_res.is_strictly_maximal_lit L C"
using step_hyps
by (metis (no_types, lifting) is_least_false_clause_def literal.collapse(1)
pos_lit_not_greatest_if_maximal_in_least_false_clause)
next
show "ℱ' = finsert C ℱ"
using step_hyps by metis
qed (simp_all add: ℳ'_def 𝒞'_def)
moreover have "∃ℳ'' 𝒞''.
(ord_res_5 N)⇧*⇧* (U⇩e⇩r, ℱ', ℳ', 𝒞') (U⇩e⇩r, ℱ', ℳ'', 𝒞'') ∧
(∀C. (𝒞'' = Some C) ⟷ is_least_false_clause (iefac ℱ' |`| (N |∪| U⇩e⇩r)) C)"
proof (rule ord_res_5_construct_model_upto_least_false_clause)
show "ord_res_5_invars N (U⇩e⇩r, ℱ', ℳ', 𝒞')"
using ord_res_5_step ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)› ‹𝒞 = Some C›
by (metis ord_res_5_preserves_invars)
qed
ultimately obtain ℳ'' 𝒞'' where
s5_steps: "(ord_res_5 N)⇧+⇧+ (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ', ℳ'', 𝒞'')" and
next_clause_least_false:
"(∀C. (𝒞'' = Some C) ⟷ is_least_false_clause (iefac ℱ' |`| (N |∪| U⇩e⇩r)) C)"
by (meson rtranclp_into_tranclp2)
have "ord_res_5_step⇧+⇧+ S5 (N, U⇩e⇩r, ℱ', ℳ'', 𝒞'')"
unfolding S5_def ‹𝒞 = Some C›
using s5_steps by (metis tranclp_ord_res_5_step_if_tranclp_ord_res_5)
moreover have "ord_res_4_matches_ord_res_5 S4' (N, U⇩e⇩r, ℱ', ℳ'', 𝒞'')"
unfolding S4'_def ‹s4' = (U⇩e⇩r, ℱ')›
proof (intro ord_res_4_matches_ord_res_5.intros)
show "ord_res_5_invars N (U⇩e⇩r, ℱ', ℳ'', 𝒞'')"
using s5_steps ‹𝒞 = Some C› ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
by (smt (verit, best) ord_res_5_preserves_invars tranclp_induct)
next
show "∀C. (𝒞'' = Some C) = is_least_false_clause (iefac ℱ' |`| (N |∪| U⇩e⇩r)) C"
using next_clause_least_false .
qed
ultimately show ?thesis
by metis
next
case step_hyps: (resolution NN C L D U⇩e⇩r')
have "𝒞 = Some C"
using match_hyps(3-) step_hyps by metis
define ℳ' :: "'f gterm ⇒ 'f gterm literal multiset option" where
"ℳ' = (λ_. None)"
define 𝒞' :: "'f gclause option" where
"𝒞' = The_optional (linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')))"
have ord_res_5_step: "ord_res_5 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r', ℱ, ℳ', 𝒞')"
proof (rule ord_res_5.resolution)
have "is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r)) C"
using step_hyps by argo
then show "¬ dom ℳ ⊫ C"
using dom_ℳ_eq[OF ‹𝒞 = Some C›]
by (metis (mono_tags, lifting) is_least_false_clause_def
linorder_cls.is_least_in_ffilter_iff ord_res_Interp_entails_if_greatest_lit_is_pos
unproductive_if_nex_strictly_maximal_pos_lit sup_bot.right_neutral)
next
show "ord_res.is_maximal_lit L C"
using step_hyps by metis
next
show "is_neg L"
using step_hyps by metis
next
show "ℳ (atm_of L) = Some D"
using step_hyps
by (smt (verit) ‹𝒞 = Some C› all_produced_atoms_in_model_def insertI1 match_hyps(3)
ord_res_5_invars_def)
next
show "U⇩e⇩r' = finsert (eres D C) U⇩e⇩r"
using step_hyps by metis
qed (simp_all add: ℳ'_def 𝒞'_def)
moreover have "∃ℳ'' 𝒞''.
(ord_res_5 N)⇧*⇧* (U⇩e⇩r', ℱ, ℳ', 𝒞') (U⇩e⇩r', ℱ, ℳ'', 𝒞'') ∧
(∀C. (𝒞'' = Some C) ⟷ is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r')) C)"
proof (rule ord_res_5_construct_model_upto_least_false_clause)
show "ord_res_5_invars N (U⇩e⇩r', ℱ, ℳ', 𝒞')"
using ord_res_5_step ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)› ‹𝒞 = Some C›
by (metis ord_res_5_preserves_invars)
qed
ultimately obtain ℳ'' 𝒞'' where
s5_steps: "(ord_res_5 N)⇧+⇧+ (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r', ℱ, ℳ'', 𝒞'')" and
next_clause_least_false:
"(∀C. (𝒞'' = Some C) ⟷ is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r')) C)"
by (meson rtranclp_into_tranclp2)
have "ord_res_5_step⇧+⇧+ S5 (N, U⇩e⇩r', ℱ, ℳ'', 𝒞'')"
unfolding S5_def ‹𝒞 = Some C›
using s5_steps by (metis tranclp_ord_res_5_step_if_tranclp_ord_res_5)
moreover have "ord_res_4_matches_ord_res_5 S4' (N, U⇩e⇩r', ℱ, ℳ'', 𝒞'')"
unfolding S4'_def ‹s4' = (U⇩e⇩r', ℱ)›
proof (intro ord_res_4_matches_ord_res_5.intros)
show "ord_res_5_invars N (U⇩e⇩r', ℱ, ℳ'', 𝒞'')"
using s5_steps ‹𝒞 = Some C› ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
by (smt (verit, best) ord_res_5_preserves_invars tranclp_induct)
next
show "∀C. (𝒞'' = Some C) = is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r')) C"
using next_clause_least_false .
qed
ultimately show ?thesis
by metis
qed
qed
theorem bisimulation_ord_res_4_ord_res_5:
defines "match ≡ λ_. ord_res_4_matches_ord_res_5"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_4_state ⇒ 'f ord_res_5_state ⇒ bool) ℛ.
bisimulation ord_res_4_step ord_res_5_step ord_res_4_final ord_res_5_final ℛ MATCH"
proof (rule ex_bisimulation_from_forward_simulation)
show "right_unique ord_res_4_step"
using right_unique_ord_res_4_step .
next
show "right_unique ord_res_5_step"
using right_unique_ord_res_5_step .
next
show "∀s. ord_res_4_final s ⟶ (∄s'. ord_res_4_step s s')"
by (metis finished_def ord_res_4_semantics.final_finished)
next
show "∀s. ord_res_5_final s ⟶ (∄s'. ord_res_5_step s s')"
by (metis finished_def ord_res_5_semantics.final_finished)
next
show "∀i s4 s5. match i s4 s5 ⟶ ord_res_4_final s4 ⟷ ord_res_5_final s5"
unfolding match_def
using ord_res_4_final_iff_ord_res_5_final by metis
next
show "∀i S4 S5. match i S4 S5 ⟶
safe_state ord_res_4_step ord_res_4_final S4 ∧ safe_state ord_res_5_step ord_res_5_final S5"
proof (intro allI impI conjI)
fix i S4 S5
show "safe_state ord_res_4_step ord_res_4_final S4"
using ord_res_4_step_safe safe_state_if_all_states_safe by metis
assume "match i S4 S5"
thus "safe_state ord_res_5_step ord_res_5_final S5"
using ‹match i S4 S5›
using ord_res_5_safe_state_if_invars
using match_def ord_res_4_matches_ord_res_5.cases by metis
qed
next
show "wfp (λ_ _. False)"
by simp
next
show "∀i s1 s2 s1'.
match i s1 s2 ⟶
ord_res_4_step s1 s1' ⟶
(∃i' s2'. ord_res_5_step⇧+⇧+ s2 s2' ∧ match i' s1' s2') ∨ (∃i'. match i' s1' s2 ∧ False)"
unfolding match_def
using forward_simulation_between_4_and_5 by metis
qed
end
section ‹ORD-RES-6 (model backjump)›
type_synonym 'f ord_res_6_state = "'f gclause fset × 'f gclause fset × 'f gclause fset ×
('f gterm ⇒ 'f gclause option) × 'f gclause option"
context simulation_SCLFOL_ground_ordered_resolution begin
inductive ord_res_5_matches_ord_res_6 :: "'f ord_res_5_state ⇒ 'f ord_res_6_state ⇒ bool" where
"ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞) ⟹
ord_res_5_matches_ord_res_6 (N, U⇩e⇩r, ℱ, ℳ, 𝒞) (N, U⇩e⇩r, ℱ, ℳ, 𝒞)"
lemma ord_res_5_final_iff_ord_res_6_final:
fixes i S5 S6
assumes match: "ord_res_5_matches_ord_res_6 S5 S6"
shows "ord_res_5_final S5 ⟷ ord_res_6_final S6"
using match
proof (cases S5 S6 rule: ord_res_5_matches_ord_res_6.cases)
case (1 N U⇩e⇩r ℱ ℳ 𝒞)
thus ?thesis
by (metis (no_types, opaque_lifting) ord_res_5_final.simps ord_res_6_final.cases
ord_res_6_final.contradiction_found ord_res_6_final.model_found)
qed
lemma backward_simulation_between_5_and_6:
fixes S5 S6 S6'
assumes match: "ord_res_5_matches_ord_res_6 S5 S6" and step: "ord_res_6_step S6 S6'"
shows "∃S5'. ord_res_5_step⇧+⇧+ S5 S5' ∧ ord_res_5_matches_ord_res_6 S5' S6'"
using match
proof (cases S5 S6 rule: ord_res_5_matches_ord_res_6.cases)
case match_hyps: (1 N U⇩e⇩r ℱ ℳ 𝒞)
hence S5_def: "S5 = (N, U⇩e⇩r, ℱ, ℳ, 𝒞)" and S6_def: "S6 = (N, U⇩e⇩r, ℱ, ℳ, 𝒞)"
by metis+
obtain s6' where S6'_def: "S6' = (N, s6')" and step': "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, 𝒞) s6'"
using step unfolding S6_def
using ord_res_6_step.simps by auto
show ?thesis
using step'
proof (cases N "(U⇩e⇩r, ℱ, ℳ, 𝒞)" s6' rule: ord_res_6.cases)
case step_hyps: (skip C 𝒞')
define S5' where
"S5' = (N, U⇩e⇩r, ℱ, ℳ, 𝒞')"
show ?thesis
proof (intro exI conjI)
have step5: "ord_res_5 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ, ℳ, 𝒞')"
using ord_res_5.skip step_hyps by metis
hence "ord_res_5_step S5 S5'"
unfolding S5_def S5'_def
by (metis ord_res_5_step.simps step_hyps(1))
thus "ord_res_5_step⇧+⇧+ S5 S5'"
by simp
have "ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞')"
using step5 match_hyps(3) ord_res_5_preserves_invars step_hyps(1) by metis
thus "ord_res_5_matches_ord_res_6 S5' S6'"
unfolding S5'_def S6'_def ‹s6' = (U⇩e⇩r, ℱ, ℳ, 𝒞')›
using ord_res_5_matches_ord_res_6.intros by metis
qed
next
case step_hyps: (production C L ℳ' 𝒞')
define S5' where
"S5' = (N, U⇩e⇩r, ℱ, ℳ', 𝒞')"
show ?thesis
proof (intro exI conjI)
have step5: "ord_res_5 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ, ℳ', 𝒞')"
using ord_res_5.production step_hyps by metis
hence "ord_res_5_step S5 S5'"
unfolding S5_def S5'_def
by (metis ord_res_5_step.simps step_hyps(1))
thus "ord_res_5_step⇧+⇧+ S5 S5'"
by simp
have "ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ', 𝒞')"
using step5 match_hyps(3) ord_res_5_preserves_invars step_hyps(1) by metis
thus "ord_res_5_matches_ord_res_6 S5' S6'"
unfolding S5'_def S6'_def ‹s6' = (U⇩e⇩r, ℱ, ℳ', 𝒞')›
using ord_res_5_matches_ord_res_6.intros by metis
qed
next
case step_hyps: (factoring D K ℱ')
define S5' where
"S5' = (N, U⇩e⇩r, ℱ', ℳ, Some (efac D))"
have "D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
by (metis match_hyps(3) next_clause_in_factorized_clause_def ord_res_5_invars_def step_hyps(1))
hence "iefac ℱ' |`| (N |∪| U⇩e⇩r) ≠ {||}"
by blast
then obtain C where C_least: "linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| U⇩e⇩r)) C"
by (metis linorder_cls.ex1_least_in_fset)
have "efac D ≠ D"
by (metis ex1_efac_eq_factoring_chain is_pos_def ex_ground_factoringI step_hyps(4,5,6))
show ?thesis
proof (intro exI conjI)
have "The_optional (linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| U⇩e⇩r))) = Some C"
proof (rule The_optional_eq_SomeI)
show "∃⇩≤⇩1 x. linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| U⇩e⇩r)) x"
by blast
next
show "linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| U⇩e⇩r)) C"
using C_least .
qed
hence step5: "ord_res_5 N (U⇩e⇩r, ℱ, ℳ, Some D) (U⇩e⇩r, ℱ', Map.empty, Some C)"
using ord_res_5.factoring step_hyps by metis
moreover have "(ord_res_5 N)⇧*⇧* … (U⇩e⇩r, ℱ', ℳ, Some (efac D))"
proof (rule full_rtranclp_ord_res_5_run_upto)
show "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some D) (U⇩e⇩r, ℱ', ℳ, Some (efac D))"
using step' S6_def S6'_def ‹s6' = (U⇩e⇩r, ℱ', ℳ, Some (efac D))› ‹𝒞 = Some D› by argo
next
show "ord_res_5_invars N (U⇩e⇩r, ℱ', ℳ, Some (efac D))"
using match_hyps(3) ord_res_6_preserves_invars step' step_hyps(2) by blast
next
have "iefac ℱ D = D" and "D |∈| N |∪| U⇩e⇩r"
unfolding atomize_conj
using ‹efac D ≠ D› ‹D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)›[unfolded fimage_iff]
unfolding iefac_def
by (metis ex1_efac_eq_factoring_chain factorizable_if_neq_efac)
have iefac_ℱ'_eq: "iefac ℱ' = (iefac ℱ)(D := efac D)"
unfolding ‹ℱ' = finsert D ℱ› iefac_def by auto
have fimage_iefac_ℱ'_eq:
"iefac ℱ' |`| (N |∪| U⇩e⇩r) = finsert (efac D) (iefac ℱ |`| (N |∪| U⇩e⇩r - {|D|}))"
unfolding iefac_ℱ'_eq
unfolding fun_upd_fimage[of "iefac ℱ" D "efac D"] ‹D |∈| N |∪| U⇩e⇩r›
using ‹D |∈| N |∪| U⇩e⇩r› by argo
have "{|C |∈| iefac ℱ' |`| (N |∪| U⇩e⇩r). C ≺⇩c efac D|} =
{|C |∈| finsert (efac D) (iefac ℱ |`| (N |∪| U⇩e⇩r - {|D|})). C ≺⇩c efac D|}"
unfolding fimage_iefac_ℱ'_eq ..
also have "… = {|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r - {|D|}). C ≺⇩c efac D|}"
by auto
also have "… = {|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). C ≺⇩c efac D|}"
by (smt (verit, ccfv_SIG) ‹iefac ℱ D = D› efac_properties_if_not_ident(1)
ffilter_eq_ffilter_minus_singleton fimage_finsert finsertI1 finsert_fminus1
finsert_fminus_single linorder_cls.less_imp_not_less)
finally have "{|C |∈| iefac ℱ' |`| (N |∪| U⇩e⇩r). C ≺⇩c efac D|} =
{|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). C ≺⇩c efac D|}" .
next
have dom_ℳ_eq: "dom ℳ = ord_res.interp (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) D"
using ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)› ‹𝒞 = Some D›
unfolding ord_res_5_invars_def model_eq_interp_upto_next_clause_def
by metis
have "atm_of K ∉ dom ℳ"
by (metis linorder_lit.is_maximal_in_mset_iff literal.collapse(1)
pos_literal_in_imp_true_cls step_hyps(3) step_hyps(4) step_hyps(5))
have "A ≺⇩t atm_of K" if "A ∈ dom ℳ" for A
proof -
obtain C where
"C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
"C ≺⇩c D" and
"A ∈ ord_res.production (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) C"
using ‹A ∈ dom ℳ› unfolding dom_ℳ_eq
unfolding ord_res.interp_def UN_iff
by blast
hence "ord_res.is_strictly_maximal_lit (Pos A) C"
using ord_res.mem_productionE by metis
hence "Pos A ≼⇩l K"
using ‹ord_res.is_maximal_lit K D› ‹C ≺⇩c D›
by (metis ord_res.asymp_less_lit ord_res.transp_less_lit linorder_cls.less_asym
linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset linorder_lit.leI
linorder_lit.multp⇩H⇩O_if_maximal_less_that_maximal multp_eq_multp⇩H⇩O)
hence "A ≼⇩t atm_of K"
by (metis literal.collapse(1) literal.sel(1) ord_res.less_lit_simps(1) reflclp_iff
step_hyps(5))
moreover have "A ≠ atm_of K"
using ‹atm_of K ∉ dom ℳ› ‹A ∈ dom ℳ› by metis
ultimately show ?thesis
by order
qed
hence "dom ℳ ⊆ {A. ∃K. ord_res.is_maximal_lit K (efac D) ∧ A ≺⇩t atm_of K}"
using linorder_lit.is_maximal_in_mset_iff step_hyps(4) by auto
thus "ℳ = restrict_map ℳ {A. ∃K. ord_res.is_maximal_lit K (efac D) ∧ A ≺⇩t atm_of K}"
using restrict_map_ident_if_dom_subset by fastforce
next
show "linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| U⇩e⇩r)) C"
using C_least .
qed
ultimately have steps5: "(ord_res_5 N)⇧+⇧+ (U⇩e⇩r, ℱ, ℳ, Some D) (U⇩e⇩r, ℱ', ℳ, Some (efac D))"
by simp
thus "ord_res_5_step⇧+⇧+ S5 S5'"
using S5'_def S5_def step_hyps(1) tranclp_ord_res_5_step_if_tranclp_ord_res_5 by metis
have "ord_res_5_invars N (U⇩e⇩r, ℱ', ℳ, Some (efac D))"
using steps5 match_hyps(3) tranclp_ord_res_5_preserves_invars step_hyps(1) by metis
thus "ord_res_5_matches_ord_res_6 S5' S6'"
unfolding S5'_def S6'_def ‹s6' = (U⇩e⇩r, ℱ', ℳ, Some (efac D))›
using ord_res_5_matches_ord_res_6.intros by metis
qed
next
case step_hyps: (resolution_bot C L D U⇩e⇩r' ℳ')
define S5' :: "_ × _ × _ × ('f gterm ⇒ 'f gclause option) × 'f gclause option" where
"S5' = (N, U⇩e⇩r', ℱ, ℳ', Some {#})"
show ?thesis
proof (intro exI conjI)
have "{#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r')"
using ‹U⇩e⇩r' = finsert (eres D C) U⇩e⇩r› ‹eres D C = {#}›
using iefac_def by simp
hence "linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')) {#}"
by (metis linorder_cls.is_minimal_in_fset_eq_is_least_in_fset
linorder_cls.is_minimal_in_fset_iff linorder_cls.leD mempty_lesseq_cls)
hence "The_optional (linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r'))) = Some {#}"
by (metis linorder_cls.Uniq_is_least_in_fset The_optional_eq_SomeI)
hence step5: "ord_res_5 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r', ℱ, ℳ', Some {#})"
using ord_res_5.resolution step_hyps by metis
thus "ord_res_5_step⇧+⇧+ S5 S5'"
unfolding S5_def ‹𝒞 = Some C› S5'_def
by (simp only: ord_res_5_step.intros tranclp.r_into_trancl)
show "ord_res_5_matches_ord_res_6 S5' S6'"
using step5
by (metis S5'_def S6'_def match_hyps(3) ord_res_5_matches_ord_res_6.intros
ord_res_5_preserves_invars step_hyps(1) step_hyps(2))
qed
next
case step_hyps: (resolution_pos E L D U⇩e⇩r' ℳ' K)
define S5' :: "_ × _ × _ × ('f gterm ⇒ 'f gclause option) × 'f gclause option" where
"S5' = (N, U⇩e⇩r', ℱ, ℳ', Some (eres D E))"
hence "iefac ℱ |`| (N |∪| U⇩e⇩r') ≠ {||}"
using ‹U⇩e⇩r' = finsert (eres D E) U⇩e⇩r› by simp
then obtain C where C_least: "linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')) C"
by (metis linorder_cls.ex1_least_in_fset)
show ?thesis
proof (intro exI conjI)
have "The_optional (linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r'))) = Some C"
proof (rule The_optional_eq_SomeI)
show "∃⇩≤⇩1 x. linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')) x"
by blast
next
show "linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')) C"
using C_least .
qed
hence step5: "ord_res_5 N (U⇩e⇩r, ℱ, ℳ, Some E) (U⇩e⇩r', ℱ, Map.empty, Some C)"
using ord_res_5.resolution step_hyps by metis
moreover have "(ord_res_5 N)⇧*⇧* … (U⇩e⇩r', ℱ, ℳ', Some (eres D E))"
proof (rule full_rtranclp_ord_res_5_run_upto)
show "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some E) (U⇩e⇩r', ℱ, ℳ', Some (eres D E))"
using step' ‹𝒞 = Some E› ‹s6' = (U⇩e⇩r', ℱ, ℳ', Some (eres D E))› by argo
next
show "ord_res_5_invars N (U⇩e⇩r', ℱ, ℳ', Some (eres D E))"
using match_hyps(3) ord_res_6_preserves_invars step' step_hyps(2) by blast
next
have "eres D E ≠ E"
using step_hyps by (metis linorder_lit.Uniq_is_maximal_in_mset the1_equality')
moreover have "eres D E ≼⇩c E"
using eres_le .
ultimately have "eres D E ≺⇩c E"
by order
have "∀F. is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r)) F ⟶ E ≼⇩c F"
using ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
unfolding ord_res_5_invars_def ‹𝒞 = Some E›
using next_clause_lt_least_false_clause[of N "(U⇩e⇩r, ℱ, ℳ, Some E)"]
by simp
have E_least_false: "is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r)) E"
unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
proof (intro conjI ballI impI)
show "E |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
unfolding ord_res_5_invars_def ‹𝒞 = Some E›
by (metis next_clause_in_factorized_clause_def)
next
have "¬ ord_res.interp (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) E ⊫ E"
using ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
unfolding ord_res_5_invars_def ‹𝒞 = Some E›
using ‹¬ dom ℳ ⊫ E› by (metis model_eq_interp_upto_next_clause_def)
moreover have "ord_res.production (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) E = {}"
proof -
have "∄L. is_pos L ∧ ord_res.is_strictly_maximal_lit L E"
using ‹ord_res.is_maximal_lit L E› ‹is_neg L›
by (metis Uniq_D linorder_lit.Uniq_is_maximal_in_mset
linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset)
thus ?thesis
using unproductive_if_nex_strictly_maximal_pos_lit by metis
qed
ultimately show "¬ ord_res_Interp (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) E ⊫ E"
by simp
next
fix F
assume F_in: "F |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and "F ≠ E" and
F_false: "¬ ord_res_Interp (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) F ⊫ F"
have "¬ F ≺⇩c E"
using ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
unfolding ord_res_5_invars_def ‹𝒞 = Some E›
unfolding all_smaller_clauses_true_wrt_respective_Interp_def
using F_in F_false
by (metis option.inject)
thus "E ≺⇩c F"
using ‹F ≠ E› by order
qed
have D_prod: "ord_res.production (fset (iefac ℱ |`| (N |∪| U⇩e⇩r))) D ≠ {}"
using ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
unfolding ord_res_5_invars_def ‹𝒞 = Some E›
by (metis atoms_in_model_were_produced_def empty_iff step_hyps(6))
have "iefac ℱ (eres D E) = eres D E"
using E_least_false D_prod
by (smt (verit, ccfv_threshold)
‹∀F. is_least_false_clause (iefac ℱ |`| (N |∪| U⇩e⇩r)) F ⟶ (≺⇩c)⇧=⇧= E F›
‹eres D E ≺⇩c E› clause_true_if_resolved_true ex1_eres_eq_full_run_ground_resolution
fimage_finsert finsert_absorb finsert_iff full_run_def funion_finsert_right
is_least_false_clause_def is_least_false_clause_finsert_smaller_false_clause
linorder_cls.is_least_in_fset_ffilterD(2) linorder_cls.leD match_hyps(3)
next_clause_in_factorized_clause_def ord_res_5_invars_def ord_res_6_preserves_invars
rtranclpD step' step_hyps(2) step_hyps(7))
hence "{|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r'). C ≺⇩c eres D E|} =
{|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). C ≺⇩c eres D E|}"
unfolding ‹U⇩e⇩r' = finsert (eres D E) U⇩e⇩r› by auto
next
show "ℳ' = restrict_map ℳ {A. ∃K. ord_res.is_maximal_lit K (eres D E) ∧ A ≺⇩t atm_of K}"
using ‹ℳ' = restrict_map ℳ {A. A ≺⇩t atm_of K}›
by (smt (verit, ccfv_SIG) Collect_cong linorder_lit.Uniq_is_maximal_in_mset step_hyps(10)
the1_equality')
next
show "linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')) C"
using C_least .
qed
ultimately have steps5: "(ord_res_5 N)⇧+⇧+ (U⇩e⇩r, ℱ, ℳ, Some E) (U⇩e⇩r', ℱ, ℳ', Some (eres D E))"
by simp
thus "ord_res_5_step⇧+⇧+ S5 S5'"
unfolding S5_def ‹𝒞 = Some E› S5'_def
by (metis tranclp_ord_res_5_step_if_tranclp_ord_res_5)
show "ord_res_5_matches_ord_res_6 S5' S6'"
unfolding S5'_def S6'_def ‹s6' = (U⇩e⇩r', ℱ, ℳ', Some (eres D E))›
using steps5
using match_hyps(3) ord_res_5_matches_ord_res_6.intros ord_res_6_preserves_invars step'
step_hyps(2) by metis
qed
next
case step_hyps: (resolution_neg E L D U⇩e⇩r' ℳ' K C)
define S5' :: "_ × _ × _ × ('f gterm ⇒ 'f gclause option) × 'f gclause option" where
"S5' = (N, U⇩e⇩r', ℱ, ℳ', Some C)"
hence "iefac ℱ |`| (N |∪| U⇩e⇩r') ≠ {||}"
using ‹U⇩e⇩r' = finsert (eres D E) U⇩e⇩r› by simp
then obtain B where B_least: "linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')) B"
by (metis linorder_cls.ex1_least_in_fset)
show ?thesis
proof (intro exI conjI)
have "The_optional (linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r'))) = Some B"
proof (rule The_optional_eq_SomeI)
show "∃⇩≤⇩1 x. linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')) x"
by blast
next
show "linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')) B"
using B_least .
qed
hence step5: "ord_res_5 N (U⇩e⇩r, ℱ, ℳ, Some E) (U⇩e⇩r', ℱ, Map.empty, Some B)"
using ord_res_5.resolution step_hyps by metis
moreover have "(ord_res_5 N)⇧*⇧* … (U⇩e⇩r', ℱ, ℳ', Some C)"
proof (rule full_rtranclp_ord_res_5_run_upto)
show "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some E) (U⇩e⇩r', ℱ, ℳ', Some C)"
using step' ‹𝒞 = Some E› ‹s6' = (U⇩e⇩r', ℱ, ℳ', Some C)› by argo
next
show "ord_res_5_invars N (U⇩e⇩r', ℱ, ℳ', Some C)"
using match_hyps(3) ord_res_6_preserves_invars step' step_hyps(2) by blast
next
have "ord_res.is_strictly_maximal_lit (Pos (atm_of K)) C"
using ‹ℳ (atm_of K) = Some C›
‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›[unfolded ord_res_5_invars_def
atoms_in_model_were_produced_def, simplified]
using ord_res.mem_productionE by blast
moreover have "Pos (atm_of K) ≺⇩l K"
using ‹is_neg K› by (cases K) simp_all
ultimately have "C ≺⇩c eres D E"
using ‹ord_res.is_maximal_lit K (eres D E)›
by (metis ord_res.asymp_less_lit ord_res.transp_less_lit
linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset
linorder_lit.multp⇩H⇩O_if_maximal_less_that_maximal multp_eq_multp⇩H⇩O)
hence "C ≺⇩c E"
using eres_le[of D E] by order
have "C ≺⇩c efac (eres D E)"
by (metis Uniq_D ‹C ≺⇩c eres D E› efac_spec is_pos_def linorder_lit.Uniq_is_maximal_in_mset
step_hyps(10) step_hyps(11))
moreover have "efac (eres D E) ≼⇩c eres D E"
by (metis efac_subset subset_implies_reflclp_multp)
ultimately have "C ≺⇩c iefac ℱ (eres D E)"
unfolding iefac_def by auto
hence "{|Ca |∈| iefac ℱ |`| (N |∪| U⇩e⇩r'). Ca ≺⇩c C|} =
{|Ca |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). Ca ≺⇩c C|}"
unfolding ‹U⇩e⇩r' = finsert (eres D E) U⇩e⇩r› by auto
have "(∃K. ord_res.is_maximal_lit K C ∧ A ≺⇩t atm_of K) ⟷ A ≺⇩t atm_of K" for A
using ‹ord_res.is_strictly_maximal_lit (Pos (atm_of K)) C›
by (metis Uniq_def linorder_lit.Uniq_is_maximal_in_mset
linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset literal.sel(1))
hence "{A. ∃K. ord_res.is_maximal_lit K C ∧ A ≺⇩t atm_of K} = {A. A ≺⇩t atm_of K}"
by metis
thus "ℳ' = restrict_map ℳ {A. ∃K. ord_res.is_maximal_lit K C ∧ A ≺⇩t atm_of K}"
using ‹ℳ' = restrict_map ℳ {A. A ≺⇩t atm_of K}› by argo
next
show "linorder_cls.is_least_in_fset (iefac ℱ |`| (N |∪| U⇩e⇩r')) B"
using B_least .
qed
ultimately have steps5: "(ord_res_5 N)⇧+⇧+ (U⇩e⇩r, ℱ, ℳ, Some E) (U⇩e⇩r', ℱ, ℳ', Some C)"
by simp
thus "ord_res_5_step⇧+⇧+ S5 S5'"
unfolding S5_def ‹𝒞 = Some E› S5'_def
by (metis tranclp_ord_res_5_step_if_tranclp_ord_res_5)
show "ord_res_5_matches_ord_res_6 S5' S6'"
unfolding S5'_def S6'_def ‹s6' = (U⇩e⇩r', ℱ, ℳ', Some C)›
using steps5
using match_hyps(3) ord_res_5_matches_ord_res_6.intros ord_res_6_preserves_invars step'
step_hyps(2) by metis
qed
qed
qed
theorem bisimulation_ord_res_5_ord_res_6:
defines "match ≡ λ_. ord_res_5_matches_ord_res_6"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_5_state ⇒ 'f ord_res_6_state ⇒ bool) ℛ.
bisimulation ord_res_5_step ord_res_6_step ord_res_5_final ord_res_6_final ℛ MATCH"
proof (rule ex_bisimulation_from_backward_simulation)
show "right_unique ord_res_5_step"
using right_unique_ord_res_5_step .
next
show "right_unique ord_res_6_step"
using right_unique_ord_res_6_step .
next
show "∀s. ord_res_5_final s ⟶ (∄s'. ord_res_5_step s s')"
by (metis finished_def ord_res_5_semantics.final_finished)
next
show "∀s. ord_res_6_final s ⟶ (∄s'. ord_res_6_step s s')"
by (metis finished_def ord_res_6_semantics.final_finished)
next
show "∀i S5 S6. match i S5 S6 ⟶ ord_res_5_final S5 ⟷ ord_res_6_final S6"
unfolding match_def
using ord_res_5_final_iff_ord_res_6_final by metis
next
show "∀i S5 S6.
match i S5 S6 ⟶
safe_state ord_res_5_step ord_res_5_final S5 ∧ safe_state ord_res_6_step ord_res_6_final S6"
proof (intro allI impI conjI)
fix i S5 S6
assume "match i S5 S6"
show "safe_state ord_res_5_step ord_res_5_final S5"
using ‹match i S5 S6›
using ord_res_5_safe_state_if_invars
using match_def ord_res_5_matches_ord_res_6.cases by metis
show "safe_state ord_res_6_step ord_res_6_final S6"
using ‹match i S5 S6›
using ord_res_6_safe_state_if_invars
using match_def ord_res_5_matches_ord_res_6.cases by metis
qed
next
show "wfp (λ_ _. False)"
by simp
next
show "∀i S5 S6 S6'.
match i S5 S6 ⟶
ord_res_6_step S6 S6' ⟶
(∃i' S5'. ord_res_5_step⇧+⇧+ S5 S5' ∧ match i' S5' S6') ∨ (∃i'. match i' S5 S6' ∧ False)"
unfolding match_def
using backward_simulation_between_5_and_6 by metis
qed
end
section ‹ORD-RES-7 (clause-guided literal trail construction)›
type_synonym 'f ord_res_7_state =
"'f gclause fset ×'f gclause fset × 'f gclause fset × ('f gliteral × 'f gclause option) list ×
'f gclause option"
context simulation_SCLFOL_ground_ordered_resolution begin
inductive ord_res_6_matches_ord_res_7 ::
"'f gterm fset ⇒ 'f ord_res_6_state ⇒ 'f ord_res_7_state ⇒ bool" where
"ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞) ⟹
ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞) ⟹
(∀A C. ℳ A = Some C ⟷ map_of Γ (Pos A) = Some (Some C)) ⟹
(∀A. ℳ A = None ⟷ map_of Γ (Neg A) ≠ None ∨ A |∉| trail_atms Γ) ⟹
i = atms_of_clss (N |∪| U⇩e⇩r) - trail_atms Γ ⟹
ord_res_6_matches_ord_res_7 i (N, U⇩e⇩r, ℱ, ℳ, 𝒞) (N, U⇩e⇩r, ℱ, Γ, 𝒞)"
lemma ord_res_6_final_iff_ord_res_7_final:
fixes i S6 S7
assumes match: "ord_res_6_matches_ord_res_7 i S6 S7"
shows "ord_res_6_final S6 ⟷ ord_res_7_final S7"
using match
proof (cases i S6 S7 rule: ord_res_6_matches_ord_res_7.cases)
case match_hyps: (1 N U⇩e⇩r ℱ ℳ 𝒞 Γ)
show "ord_res_6_final S6 ⟷ ord_res_7_final S7"
proof (rule iffI)
assume "ord_res_6_final S6"
thus "ord_res_7_final S7"
unfolding ‹S6 = (N, U⇩e⇩r, ℱ, ℳ, 𝒞)›
proof (cases "(N, U⇩e⇩r, ℱ, ℳ, 𝒞)" rule: ord_res_6_final.cases)
case model_found
thus "ord_res_7_final S7"
unfolding ‹S7 = (N, U⇩e⇩r, ℱ, Γ, 𝒞)›
using ord_res_7_final.model_found
by metis
next
case contradiction_found
thus "ord_res_7_final S7"
unfolding ‹S7 = (N, U⇩e⇩r, ℱ, Γ, 𝒞)›
using ord_res_7_final.contradiction_found
by metis
qed
next
assume "ord_res_7_final S7"
thus "ord_res_6_final S6"
unfolding ‹S7 = (N, U⇩e⇩r, ℱ, Γ, 𝒞)›
proof (cases "(N, U⇩e⇩r, ℱ, Γ, 𝒞)" rule: ord_res_7_final.cases)
case model_found
thus "ord_res_6_final S6"
unfolding ‹S6 = (N, U⇩e⇩r, ℱ, ℳ, 𝒞)›
using ord_res_6_final.model_found
by metis
next
case contradiction_found
thus "ord_res_6_final S6"
unfolding ‹S6 = (N, U⇩e⇩r, ℱ, ℳ, 𝒞)›
using ord_res_6_final.contradiction_found
by metis
qed
qed
qed
lemma backward_simulation_between_6_and_7:
fixes i S6 S7 S7'
assumes match: "ord_res_6_matches_ord_res_7 i S6 S7" and step: "constant_context ord_res_7 S7 S7'"
shows "
(∃i' S6'. ord_res_6_step⇧+⇧+ S6 S6' ∧ ord_res_6_matches_ord_res_7 i' S6' S7') ∨
(∃i'. ord_res_6_matches_ord_res_7 i' S6 S7' ∧ i' |⊂| i)"
using match
proof (cases i S6 S7 rule: ord_res_6_matches_ord_res_7.cases)
case match_hyps: (1 N U⇩e⇩r ℱ ℳ 𝒞 Γ)
note S6_def = ‹S6 = (N, U⇩e⇩r, ℱ, ℳ, 𝒞)›
note invars_6 = ‹ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞)›
note invars_7 = ‹ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞)›[
unfolded ord_res_7_invars_def, rule_format, OF refl]
have Γ_sorted: "sorted_wrt (λx y. atm_of (fst y) ≺⇩t atm_of (fst x)) Γ"
using invars_7 by argo
have Γ_consistent: "trail_consistent Γ"
using invars_7 by (metis trail_consistent_if_sorted_wrt_atoms)
hence Γ_distinct_atoms: "distinct (map fst Γ)"
using distinct_lits_if_trail_consistent by iprover
have clause_true_wrt_model_if_true_wrt_Γ: "dom ℳ ⊫ D"
if D_true: "trail_true_cls Γ D" for D
proof -
obtain L where "L ∈# D" and L_true: "trail_true_lit Γ L"
using D_true unfolding trail_true_cls_def by auto
have "∃𝒞. (L, 𝒞) ∈ set Γ"
using L_true unfolding trail_true_lit_def by auto
show ?thesis
proof (cases L)
case (Pos A)
then obtain C where "(Pos A, Some C) ∈ set Γ"
using invars_7 ‹∃𝒞. (L, 𝒞) ∈ set Γ›
by (metis fst_conv literal.disc(1) not_None_eq snd_conv)
hence "map_of Γ (Pos A) = Some (Some C)"
using Γ_distinct_atoms by (metis map_of_is_SomeI)
hence "ℳ A = Some C"
using ‹∀A C. (ℳ A = Some C) = (map_of Γ (Pos A) = Some (Some C))› by metis
hence "A ∈ dom ℳ"
by blast
then show ?thesis
using ‹L ∈# D› ‹L = Pos A› by blast
next
case (Neg A)
hence "(Neg A, None) ∈ set Γ"
using invars_7 ‹∃𝒞. (L, 𝒞) ∈ set Γ›
by (metis fst_conv literal.disc(2) snd_conv)
hence "map_of Γ (Neg A) ≠ None"
by (simp add: weak_map_of_SomeI)
hence "ℳ A = None"
using ‹∀A. (ℳ A = None) = (map_of Γ (Neg A) ≠ None ∨ A |∉| trail_atms Γ)› by metis
hence "A ∉ dom ℳ"
by blast
then show ?thesis
using ‹L ∈# D› ‹L = Neg A› by blast
qed
qed
have clause_false_wrt_model_if_false_wrt_Γ: "¬ dom ℳ ⊫ D"
if D_false: "trail_false_cls Γ D" for D
unfolding true_cls_def
proof (intro notI , elim bexE)
fix L :: "'f gterm literal"
assume "L ∈# D" and "dom ℳ ⊫l L"
have "trail_false_lit Γ L"
using ‹L ∈# D› D_false unfolding trail_false_cls_def by metis
hence "¬ trail_true_lit Γ L" and "trail_defined_lit Γ L"
unfolding atomize_conj
using Γ_consistent ‹L ∈# D› not_trail_true_cls_and_trail_false_cls that
trail_defined_lit_iff_true_or_false trail_true_cls_def by blast
show False
proof (cases L)
case (Pos A)
hence "ℳ A ≠ None"
using ‹dom ℳ ⊫l L› by blast
hence "map_of Γ (Pos A) ≠ None"
using ‹∀A C. (ℳ A = Some C) = (map_of Γ (Pos A) = Some (Some C))› by blast
hence "Pos A ∈ fst ` set Γ"
by (simp add: map_of_eq_None_iff)
hence "trail_true_lit Γ (Pos A)"
unfolding trail_true_lit_def .
moreover have "¬ trail_true_lit Γ (Pos A)"
using ‹¬ trail_true_lit Γ L› ‹L = Pos A› by argo
ultimately show False
by contradiction
next
case (Neg A)
hence "ℳ A = None"
using ‹dom ℳ ⊫l L› by blast
hence "map_of Γ (Neg A) ≠ None ∨ A |∉| trail_atms Γ"
using ‹∀A. (ℳ A = None) = (map_of Γ (Neg A) ≠ None ∨ A |∉| trail_atms Γ)› by blast
hence "trail_true_lit Γ (Neg A) ∨ ¬ trail_defined_lit Γ (Neg A)"
unfolding map_of_eq_None_iff not_not
unfolding trail_true_lit_def trail_defined_lit_iff_trail_defined_atm literal.sel
.
then show ?thesis
using ‹¬ trail_true_lit Γ L› ‹trail_defined_lit Γ L› ‹L = Neg A› by argo
qed
qed
obtain s7' where
"S7' = (N, s7')" and
step': "ord_res_7 N (U⇩e⇩r, ℱ, Γ, 𝒞) s7'"
using step unfolding ‹S7 = (N, U⇩e⇩r, ℱ, Γ, 𝒞)›
by (auto elim: constant_context.cases)
have invars_s7': "ord_res_7_invars N s7'"
using ord_res_7_preserves_invars[OF step' ‹ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞)›] .
show ?thesis
using step'
proof (cases N "(U⇩e⇩r, ℱ, Γ, 𝒞)" s7' rule: ord_res_7.cases)
case step_hyps: (decide_neg C L A Γ')
define i' where
"i' = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ'"
have "A |∈| atms_of_clss (N |∪| U⇩e⇩r)" and "A ≺⇩t atm_of L" and "A |∉| trail_atms Γ"
using step_hyps unfolding atomize_conj linorder_trm.is_least_in_ffilter_iff by argo
have "ord_res_6_matches_ord_res_7 i' S6 S7'"
unfolding S6_def ‹𝒞 = Some C› ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r, ℱ, Γ', Some C)›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, Some C)"
using invars_6 unfolding ‹𝒞 = Some C› .
next
show "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', Some C)"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r, ℱ, Γ', Some C)› .
next
show "∀A C. (ℳ A = Some C) = (map_of Γ' (Pos A) = Some (Some C))"
using match_hyps unfolding ‹Γ' = (Neg A, None) # Γ› by simp
next
show "∀A. (ℳ A = None) = (map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ')"
unfolding ‹Γ' = (Neg A, None) # Γ›
using match_hyps ‹A |∉| trail_atms Γ› by force
next
show "i' = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ'"
unfolding i'_def ..
qed
moreover have "i' |⊂| i"
proof -
have "i = finsert A i'"
unfolding match_hyps i'_def
using ‹A |∈| atms_of_clss (N |∪| U⇩e⇩r)› ‹A |∉| trail_atms Γ› step_hyps(6) by force
moreover have "A |∉| i'"
unfolding i'_def
using step_hyps(6) by fastforce
ultimately show ?thesis
by auto
qed
ultimately show ?thesis
by metis
next
case step_hyps: (skip_defined C L 𝒞')
define S6' where
"S6' = (N, U⇩e⇩r, ℱ, ℳ, 𝒞')"
have C_almost_defined: "trail_defined_cls Γ {#x ∈# C. x ≠ L#}"
using step_hyps by (metis clause_almost_definedI invars_7)
hence C_defined: "trail_defined_cls Γ C"
using step_hyps unfolding trail_defined_cls_def by auto
hence C_true: "trail_true_cls Γ C"
using step_hyps by (metis trail_true_or_false_cls_if_defined)
have step6: "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ, ℳ, 𝒞')"
proof (rule ord_res_6.skip)
show "dom ℳ ⊫ C"
using clause_true_wrt_model_if_true_wrt_Γ[OF C_true] .
next
show "𝒞' = The_optional (linorder_cls.is_least_in_fset
(ffilter ((≺⇩c) C) (iefac ℱ |`| (N |∪| U⇩e⇩r))))"
using step_hyps by argo
qed
hence "ord_res_6_step⇧+⇧+ S6 S6'"
using S6_def ‹𝒞 = Some C› S6'_def ord_res_6_step.intros by blast
moreover have "ord_res_6_matches_ord_res_7 i S6' S7'"
unfolding S6'_def ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r, ℱ, Γ, 𝒞')›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞')"
using invars_6 unfolding ‹𝒞 = Some C›
using ord_res_6_preserves_invars[OF step6] by argo
next
show "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞')"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r, ℱ, Γ, 𝒞')› .
next
show "∀A C. (ℳ A = Some C) = (map_of Γ (Pos A) = Some (Some C))"
using match_hyps by argo
next
show "∀A. (ℳ A = None) = (map_of Γ (Neg A) ≠ None ∨ A |∉| trail_atms Γ)"
using match_hyps by argo
next
show "i = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ"
using match_hyps by argo
qed
ultimately show ?thesis
by metis
next
case step_hyps: (skip_undefined_neg C L Γ' 𝒞')
define S6' where
"S6' = (N, U⇩e⇩r, ℱ, ℳ, 𝒞')"
define i' where
"i' = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ'"
have "trail_true_lit Γ' L"
unfolding ‹Γ' = (L, None) # Γ› by (simp add: trail_true_lit_def)
hence C_true: "trail_true_cls Γ' C"
using step_hyps unfolding linorder_lit.is_maximal_in_mset_iff trail_true_cls_def by metis
have step6: "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ, ℳ, 𝒞')"
proof (rule ord_res_6.skip)
show "dom ℳ ⊫ C"
using C_true
by (metis domIff linorder_lit.is_maximal_in_mset_iff literal.collapse(2) match_hyps(6)
step_hyps(4) step_hyps(6) step_hyps(7) trail_defined_lit_iff_trail_defined_atm
true_cls_def true_lit_simps(2))
next
show "𝒞' = The_optional (linorder_cls.is_least_in_fset
(ffilter ((≺⇩c) C) (iefac ℱ |`| (N |∪| U⇩e⇩r))))"
using step_hyps by argo
qed
hence "ord_res_6_step⇧+⇧+ S6 S6'"
using S6_def ‹𝒞 = Some C› S6'_def ord_res_6_step.intros by blast
moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
unfolding S6'_def ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r, ℱ, Γ', 𝒞')›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, 𝒞')"
using invars_6 unfolding ‹𝒞 = Some C›
using ord_res_6_preserves_invars[OF step6] by argo
next
show "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞')"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r, ℱ, Γ', 𝒞')› .
next
show "∀A C. (ℳ A = Some C) = (map_of Γ' (Pos A) = Some (Some C))"
using match_hyps
unfolding ‹Γ' = (L, None) # Γ›
by (metis literal.disc(1) map_of_Cons_code(2) step_hyps(7))
next
show "∀A. (ℳ A = None) = (map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ')"
using match_hyps
unfolding ‹Γ' = (L, None) # Γ›
by (metis finsert_iff literal.collapse(2) literal.sel(2) map_of_Cons_code(2) option.discI
prod.sel(1) step_hyps(6) step_hyps(7) trail_atms.simps(2)
trail_defined_lit_iff_trail_defined_atm)
next
show "i' = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ'"
using i'_def .
qed
ultimately show ?thesis
by metis
next
case step_hyps: (skip_undefined_pos C L D)
define S6' where
"S6' = (N, U⇩e⇩r, ℱ, ℳ, Some D)"
have "trail_defined_cls Γ {#x ∈# C. x ≠ L ∧ x ≠ - L#}"
proof (rule clause_almost_almost_definedI)
show "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using invars_7 step_hyps by metis
next
show "ord_res.is_maximal_lit L C"
using step_hyps by argo
next
show "¬ (∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ)"
using step_hyps by argo
qed
moreover have "- L ∉# C"
by (metis atm_of_uminus is_pos_def linorder_lit.is_maximal_in_mset_iff linorder_lit.neqE
linorder_trm.less_irrefl literal.collapse(2) literal.sel(1) ord_res.less_lit_simps(4)
step_hyps(4) step_hyps(7) uminus_not_id')
ultimately have "trail_defined_cls Γ {#x ∈# C. x ≠ L#}"
unfolding trail_defined_cls_def by auto
hence "trail_true_cls Γ {#x ∈# C. x ≠ L#}"
using ‹¬ trail_false_cls Γ {#K ∈# C. K ≠ L#}› by (metis trail_true_or_false_cls_if_defined)
hence C_true: "trail_true_cls Γ C"
by (auto simp: trail_true_cls_def)
have step6: "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ, ℳ, Some D)"
proof (rule ord_res_6.skip)
show "dom ℳ ⊫ C"
using clause_true_wrt_model_if_true_wrt_Γ[OF C_true] .
next
show "Some D = The_optional (linorder_cls.is_least_in_fset
(ffilter ((≺⇩c) C) (iefac ℱ |`| (N |∪| U⇩e⇩r))))"
using linorder_cls.Uniq_is_least_in_fset step_hyps(9) The_optional_eq_SomeI
by fastforce
qed
hence "ord_res_6_step⇧+⇧+ S6 S6'"
using S6_def ‹𝒞 = Some C› S6'_def ord_res_6_step.intros by blast
moreover have "ord_res_6_matches_ord_res_7 i S6' S7'"
unfolding S6'_def ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r, ℱ, Γ, Some D)›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, Some D)"
using invars_6 unfolding ‹𝒞 = Some C›
using ord_res_6_preserves_invars[OF step6] by argo
next
show "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, Some D)"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r, ℱ, Γ, Some D)› .
next
show "∀A C. (ℳ A = Some C) = (map_of Γ (Pos A) = Some (Some C))"
using match_hyps by argo
next
show "∀A. (ℳ A = None) = (map_of Γ (Neg A) ≠ None ∨ A |∉| trail_atms Γ)"
using match_hyps by argo
next
show "i = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ"
using match_hyps by argo
qed
ultimately show ?thesis
by metis
next
case step_hyps: (skip_undefined_pos_ultimate C L Γ')
define S6' where
"S6' = (N, U⇩e⇩r, ℱ, ℳ, None :: 'f gclause option)"
define i' where
"i' = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ'"
have "trail_defined_cls Γ {#x ∈# C. x ≠ L ∧ x ≠ - L#}"
proof (rule clause_almost_almost_definedI)
show "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using invars_7 step_hyps by metis
next
show "ord_res.is_maximal_lit L C"
using step_hyps by argo
next
show "¬ (∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ)"
using step_hyps by argo
qed
moreover have "- L ∉# C"
by (metis atm_of_uminus is_pos_def linorder_lit.is_maximal_in_mset_iff linorder_lit.neqE
linorder_trm.less_irrefl literal.collapse(2) literal.sel(1) ord_res.less_lit_simps(4)
step_hyps(4) step_hyps(7) uminus_not_id')
ultimately have "trail_defined_cls Γ {#x ∈# C. x ≠ L#}"
unfolding trail_defined_cls_def by auto
hence "trail_true_cls Γ {#x ∈# C. x ≠ L#}"
using ‹¬ trail_false_cls Γ {#K ∈# C. K ≠ L#}› by (metis trail_true_or_false_cls_if_defined)
hence C_true: "trail_true_cls Γ C"
by (auto simp: trail_true_cls_def)
have step6: "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ, ℳ, None)"
proof (rule ord_res_6.skip)
show "dom ℳ ⊫ C"
using clause_true_wrt_model_if_true_wrt_Γ[OF C_true] .
next
have "¬ (∃D. linorder_cls.is_least_in_fset (ffilter ((≺⇩c) C) (iefac ℱ |`| (N |∪| U⇩e⇩r))) D)"
using ‹¬ fBex (iefac ℱ |`| (N |∪| U⇩e⇩r)) ((≺⇩c) C)›
by (meson linorder_cls.is_least_in_ffilter_iff)
thus "None = The_optional (linorder_cls.is_least_in_fset
(ffilter ((≺⇩c) C) (iefac ℱ |`| (N |∪| U⇩e⇩r))))"
unfolding The_optional_def by metis
qed
hence "ord_res_6_step⇧+⇧+ S6 S6'"
using S6_def ‹𝒞 = Some C› S6'_def ord_res_6_step.intros by blast
moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
unfolding S6'_def ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r, ℱ, Γ', None)›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ, None)"
using invars_6 unfolding ‹𝒞 = Some C›
using ord_res_6_preserves_invars[OF step6] by argo
next
show "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', None)"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r, ℱ, Γ', None)› .
next
show "∀A C. (ℳ A = Some C) = (map_of Γ' (Pos A) = Some (Some C))"
using match_hyps(3-)
unfolding ‹Γ' = (- L, None) # Γ›
by (metis is_pos_neg_not_is_pos literal.disc(1) map_of_Cons_code(2) step_hyps(7))
next
show "∀A. (ℳ A = None) = (map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ')"
using match_hyps(3-)
unfolding ‹Γ' = (- L, None) # Γ›
by (metis (no_types, opaque_lifting) atm_of_eq_atm_of eq_fst_iff fset_simps(2) insertCI
insertE literal.discI(2) literal.sel(2) map_of_Cons_code(2) option.distinct(1)
trail_defined_lit_iff_trail_defined_atm step_hyps(6) step_hyps(7) trail_atms.simps(2))
next
show "i' = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ'"
using i'_def .
qed
ultimately show ?thesis
by metis
next
case step_hyps: (production C L Γ' 𝒞')
define S6' where
"S6' = (N, U⇩e⇩r, ℱ, ℳ(atm_of L ↦ C), 𝒞')"
define i' where
"i' = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ'"
have "L ∈# C"
using step_hyps unfolding linorder_lit.is_maximal_in_mset_iff by argo
have step6: "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ, ℳ(atm_of L ↦ C), 𝒞')"
proof (rule ord_res_6.production)
have "atm_of L |∉| trail_atms Γ"
using ‹¬ trail_defined_lit Γ L›
unfolding trail_defined_lit_iff_trail_defined_atm .
hence "ℳ (atm_of L) = None"
using match_hyps(3-) by metis
hence "atm_of L ∉ dom ℳ"
unfolding dom_def by simp
hence "¬ dom ℳ ⊫l L"
using ‹is_pos L› unfolding true_lit_def by metis
moreover have "¬ dom ℳ ⊫ {#K ∈# C. K ≠ L#}"
using clause_false_wrt_model_if_false_wrt_Γ[OF ‹trail_false_cls Γ {#K ∈# C. K ≠ L#}›] .
ultimately show "¬ dom ℳ ⊫ C"
using ‹L ∈# C›
unfolding true_cls_def by auto
next
show "ord_res.is_maximal_lit L C"
using step_hyps by argo
next
show "is_pos L"
using step_hyps by argo
next
show "ord_res.is_strictly_maximal_lit L C"
using step_hyps by argo
next
show "ℳ(atm_of L ↦ C) = ℳ(atm_of L ↦ C)" ..
next
show "𝒞' = The_optional (linorder_cls.is_least_in_fset
(ffilter ((≺⇩c) C) (iefac ℱ |`| (N |∪| U⇩e⇩r))))"
using step_hyps by argo
qed
hence "ord_res_6_step⇧+⇧+ S6 S6'"
using S6_def ‹𝒞 = Some C› S6'_def ord_res_6_step.intros by blast
moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
unfolding S6'_def ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r, ℱ, Γ', 𝒞')›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r, ℱ, ℳ(atm_of L ↦ C), 𝒞')"
using invars_6 unfolding ‹𝒞 = Some C›
using ord_res_6_preserves_invars[OF step6] by argo
next
show "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞')"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r, ℱ, Γ', 𝒞')› .
next
show "∀A D. ((ℳ(atm_of L ↦ C)) A = Some D) = (map_of Γ' (Pos A) = Some (Some D))"
using match_hyps(3-)
unfolding ‹Γ' = (L, Some C) # Γ›
using step_hyps(7) by auto
next
show "∀A. ((ℳ(atm_of L ↦ C)) A = None) = (map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ')"
using match_hyps(3-)
unfolding ‹Γ' = (L, Some C) # Γ›
by (metis (no_types, opaque_lifting) domI domIff finsert_iff fun_upd_apply
literal.collapse(1) literal.discI(2) map_of_Cons_code(2) map_of_eq_None_iff prod.sel(1)
step_hyps(6) step_hyps(7) trail_atms.simps(2) trail_defined_lit_def uminus_Pos)
next
show "i' = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ'"
using i'_def .
qed
ultimately show ?thesis
by metis
next
case step_hyps: (factoring C L ℱ')
define S6' where
"S6' = (N, U⇩e⇩r, ℱ', ℳ, Some (efac C))"
have "L ∈# C"
using step_hyps unfolding linorder_lit.is_maximal_in_mset_iff by argo
have step6: "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some C) (U⇩e⇩r, ℱ', ℳ, Some (efac C))"
proof (rule ord_res_6.factoring)
have "atm_of L |∉| trail_atms Γ"
using ‹¬ trail_defined_lit Γ L›
unfolding trail_defined_lit_iff_trail_defined_atm .
hence "ℳ (atm_of L) = None"
using match_hyps(3-) by metis
hence "atm_of L ∉ dom ℳ"
unfolding dom_def by simp
hence "¬ dom ℳ ⊫l L"
using ‹is_pos L› unfolding true_lit_def by metis
moreover have "¬ dom ℳ ⊫ {#K ∈# C. K ≠ L#}"
using clause_false_wrt_model_if_false_wrt_Γ[OF ‹trail_false_cls Γ {#K ∈# C. K ≠ L#}›] .
ultimately show "¬ dom ℳ ⊫ C"
using ‹L ∈# C›
unfolding true_cls_def by auto
next
show "ord_res.is_maximal_lit L C"
using step_hyps by argo
next
show "is_pos L"
using step_hyps by argo
next
show "¬ ord_res.is_strictly_maximal_lit L C"
using step_hyps by argo
next
show "ℱ' = finsert C ℱ"
using step_hyps by argo
qed
hence "ord_res_6_step⇧+⇧+ S6 S6'"
using S6_def ‹𝒞 = Some C› S6'_def ord_res_6_step.intros by blast
moreover have "ord_res_6_matches_ord_res_7 i S6' S7'"
unfolding S6'_def ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r, ℱ', Γ, Some (efac C))›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r, ℱ', ℳ, Some (efac C))"
using invars_6 unfolding ‹𝒞 = Some C›
using ord_res_6_preserves_invars[OF step6] by argo
next
show "ord_res_7_invars N (U⇩e⇩r, ℱ', Γ, Some (efac C))"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r, ℱ', Γ, Some (efac C))› .
next
show "∀A C. (ℳ A = Some C) = (map_of Γ (Pos A) = Some (Some C))"
using match_hyps(3-) by argo
next
show "∀A. (ℳ A = None) = (map_of Γ (Neg A) ≠ None ∨ A |∉| trail_atms Γ)"
using match_hyps(3-) by argo
next
show "i = atms_of_clss (N |∪| U⇩e⇩r) |-| trail_atms Γ"
using match_hyps(3-) by argo
qed
ultimately show ?thesis
by metis
next
case step_hyps: (resolution_bot E L D U⇩e⇩r' Γ')
define S6' where
"S6' = (N, U⇩e⇩r', ℱ, (λ_. None) :: 'f gterm ⇒ 'f gclause option, Some ({#} :: 'f gclause))"
define i' where
"i' = atms_of_clss (N |∪| U⇩e⇩r') |-| trail_atms Γ'"
have step6: "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some E) (U⇩e⇩r', ℱ, λ_. None, Some {#})"
proof (rule ord_res_6.resolution_bot)
show "¬ dom ℳ ⊫ E"
using clause_false_wrt_model_if_false_wrt_Γ[OF ‹trail_false_cls Γ E›] .
next
show "ord_res.is_maximal_lit L E"
using step_hyps by argo
next
show "is_neg L"
using step_hyps by argo
next
show "ℳ (atm_of L) = Some D"
by (metis literal.collapse(2) match_hyps(5) step_hyps(5) step_hyps(6) uminus_Neg)
next
show "U⇩e⇩r' = finsert (eres D E) U⇩e⇩r"
using step_hyps by argo
next
show "eres D E = {#}"
using step_hyps by argo
next
show "((λ_. None)) = (λ_. None)" ..
qed
hence "ord_res_6_step⇧+⇧+ S6 S6'"
using S6_def ‹𝒞 = Some E› S6'_def ord_res_6_step.intros by blast
moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
unfolding S6'_def ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r', ℱ, Γ', Some {#})›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r', ℱ, λ_. None, Some {#})"
using invars_6 unfolding ‹𝒞 = Some E›
using ord_res_6_preserves_invars[OF step6] by argo
next
show "ord_res_7_invars N (U⇩e⇩r', ℱ, Γ', Some {#})"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r', ℱ, Γ', Some {#})› .
next
show "∀A C. (None = Some C) = (map_of Γ' (Pos A) = Some (Some C))"
unfolding ‹Γ' = []› by simp
next
show "∀A. (None = None) = (map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ')"
unfolding ‹Γ' = []› by simp
next
show "i' = atms_of_clss (N |∪| U⇩e⇩r') |-| trail_atms Γ'"
using i'_def .
qed
ultimately show ?thesis
by metis
next
case step_hyps: (resolution_pos E L D U⇩e⇩r' Γ' K)
define S6' where
"S6' = (N, U⇩e⇩r', ℱ, restrict_map ℳ {A. A ≺⇩t atm_of K}, Some (eres D E))"
define i' where
"i' = atms_of_clss (N |∪| U⇩e⇩r') |-| trail_atms Γ'"
have mem_set_Γ'_iff: "⋀x. (x ∈ set Γ') = (atm_of (fst x) ≺⇩t atm_of K ∧ x ∈ set Γ)"
unfolding ‹Γ' = dropWhile (λLn. atm_of K ≼⇩t atm_of (fst Ln)) Γ›
unfolding mem_set_dropWhile_conv_if_list_sorted_and_pred_monotone[OF Γ_sorted mono_atms_lt]
by auto
have step6: "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some E) (U⇩e⇩r', ℱ, restrict_map ℳ {A. A ≺⇩t atm_of K}, Some (eres D E))"
proof (rule ord_res_6.resolution_pos)
show "¬ dom ℳ ⊫ E"
using clause_false_wrt_model_if_false_wrt_Γ[OF ‹trail_false_cls Γ E›] .
next
show "ord_res.is_maximal_lit L E"
using step_hyps by argo
next
show "is_neg L"
using step_hyps by argo
next
show "ℳ (atm_of L) = Some D"
by (metis literal.collapse(2) match_hyps(5) step_hyps(5) step_hyps(6) uminus_Neg)
next
show "U⇩e⇩r' = finsert (eres D E) U⇩e⇩r"
using step_hyps by argo
next
show "eres D E ≠ {#}"
using step_hyps by argo
next
show "restrict_map ℳ {A. A ≺⇩t atm_of K} = restrict_map ℳ {A. A ≺⇩t atm_of K}"
..
next
show "ord_res.is_maximal_lit K (eres D E)"
using step_hyps by argo
next
show "is_pos K"
using step_hyps by argo
qed
hence "ord_res_6_step⇧+⇧+ S6 S6'"
using S6_def ‹𝒞 = Some E› S6'_def ord_res_6_step.intros by blast
moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
unfolding S6'_def ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r', ℱ, Γ', Some (eres D E))›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r', ℱ, restrict_map ℳ {A. A ≺⇩t atm_of K}, Some (eres D E))"
using invars_6 unfolding ‹𝒞 = Some E›
using ord_res_6_preserves_invars[OF step6] by argo
next
show "ord_res_7_invars N (U⇩e⇩r', ℱ, Γ', Some (eres D E))"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r', ℱ, Γ', Some (eres D E))› .
next
show "∀A C. (restrict_map ℳ {A. A ≺⇩t atm_of K} A = Some C) =
(map_of Γ' (Pos A) = Some (Some C))"
proof (intro allI)
fix A :: "'f gterm" and C :: "'f gclause"
show "restrict_map ℳ {A. A ≺⇩t atm_of K} A = Some C ⟷ map_of Γ' (Pos A) = Some (Some C)"
proof (cases "A ∈ dom ℳ ∧ A ≺⇩t atm_of K")
case True
have "restrict_map ℳ {A. A ≺⇩t atm_of K} A = Some C ⟷ ℳ A = Some C"
using True by simp
also have "… ⟷ map_of Γ (Pos A) = Some (Some C)"
using match_hyps(3-) by metis
also have "… ⟷ map_of Γ' (Pos A) = Some (Some C)"
proof -
have "Pos A ∈ fst ` set Γ"
using True
by (metis domIff map_of_eq_None_iff match_hyps(5) not_None_eq)
hence "∃𝒞. (Pos A, 𝒞) ∈ set Γ"
by fastforce
hence "∃𝒞. (Pos A, 𝒞) ∈ set Γ ∧ (Pos A, 𝒞) ∈ set Γ'"
using True unfolding mem_set_Γ'_iff prod.sel literal.sel by metis
moreover have "distinct (map fst Γ')"
using Γ_distinct_atoms
proof (rule distinct_suffix)
show "suffix (map fst Γ') (map fst Γ)"
using map_mono_suffix step_hyps(9) suffix_dropWhile by blast
qed
ultimately have "map_of Γ (Pos A) = map_of Γ' (Pos A)"
using Γ_distinct_atoms by (auto dest: map_of_is_SomeI)
thus ?thesis
by argo
qed
finally show ?thesis .
next
case False
have "restrict_map ℳ {A. A ≺⇩t atm_of K} A = None"
using False unfolding restrict_map_def by auto
moreover have "map_of Γ' (Pos A) ≠ Some (Some C)"
using False unfolding de_Morgan_conj
proof (elim disjE)
assume "A ∉ dom ℳ"
hence "⋀𝒞. (Pos A, 𝒞) ∉ set Γ"
using match_hyps(5)
by (metis (no_types, opaque_lifting) domIff fst_eqD invars_7 is_pos_def map_of_SomeD
not_None_eq snd_conv weak_map_of_SomeI)
hence "⋀𝒞. (Pos A, 𝒞) ∉ set Γ'"
unfolding mem_set_Γ'_iff by simp
then show "map_of Γ' (Pos A) ≠ Some (Some C)"
by (meson map_of_SomeD)
next
assume "¬ A ≺⇩t atm_of K"
hence "⋀𝒞. (Pos A, 𝒞) ∉ set Γ'"
unfolding mem_set_Γ'_iff by simp
then show "map_of Γ' (Pos A) ≠ Some (Some C)"
by (meson map_of_SomeD)
qed
ultimately show ?thesis
by simp
qed
qed
next
show "∀A. (restrict_map ℳ {A. A ≺⇩t atm_of K} A = None) =
(map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ')"
proof (intro allI)
fix A :: "'f gterm"
show "(restrict_map ℳ {A. A ≺⇩t atm_of K} A = None) =
(map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ')"
proof (cases "A ≺⇩t atm_of K")
case True
have "restrict_map ℳ {A. A ≺⇩t atm_of K} A = None ⟷ ℳ A = None"
using True by simp
also have "… ⟷ map_of Γ (Neg A) ≠ None ∨ A |∉| trail_atms Γ"
using match_hyps(6) by metis
also have "… ⟷ map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ"
using True mem_set_Γ'_iff
by (metis eq_fst_iff literal.sel(2) map_of_SomeD not_None_eq weak_map_of_SomeI)
also have "… ⟷ map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ'"
using True mem_set_Γ'_iff
by (smt (verit, best) fset_trail_atms image_iff)
finally show ?thesis .
next
case False
have "restrict_map ℳ {A. A ≺⇩t atm_of K} A = None"
using False by simp
moreover have "A |∉| trail_atms Γ'"
using False mem_set_Γ'_iff
by (smt (verit, ccfv_threshold) fset_trail_atms image_iff)
ultimately show ?thesis
by metis
qed
qed
next
show "i' = atms_of_clss (N |∪| U⇩e⇩r') |-| trail_atms Γ'"
using i'_def .
qed
ultimately show ?thesis
by metis
next
case step_hyps: (resolution_neg E L D U⇩e⇩r' Γ' K C)
define S6' where
"S6' = (N, U⇩e⇩r', ℱ, restrict_map ℳ {A. A ≺⇩t atm_of K}, Some C)"
define i' where
"i' = atms_of_clss (N |∪| U⇩e⇩r') |-| trail_atms Γ'"
have mem_set_Γ'_iff: "⋀x. (x ∈ set Γ') = (atm_of (fst x) ≺⇩t atm_of K ∧ x ∈ set Γ)"
unfolding ‹Γ' = dropWhile (λLn. atm_of K ≼⇩t atm_of (fst Ln)) Γ›
unfolding mem_set_dropWhile_conv_if_list_sorted_and_pred_monotone[OF Γ_sorted mono_atms_lt]
by auto
have step6: "ord_res_6 N (U⇩e⇩r, ℱ, ℳ, Some E) (U⇩e⇩r', ℱ, restrict_map ℳ {A. A ≺⇩t atm_of K}, Some C)"
proof (rule ord_res_6.resolution_neg)
show "¬ dom ℳ ⊫ E"
using clause_false_wrt_model_if_false_wrt_Γ[OF ‹trail_false_cls Γ E›] .
next
show "ord_res.is_maximal_lit L E"
using step_hyps by argo
next
show "is_neg L"
using step_hyps by argo
next
show "ℳ (atm_of L) = Some D"
by (metis literal.collapse(2) match_hyps(5) step_hyps(5) step_hyps(6) uminus_Neg)
next
show "U⇩e⇩r' = finsert (eres D E) U⇩e⇩r"
using step_hyps by argo
next
show "eres D E ≠ {#}"
using step_hyps by argo
next
show "restrict_map ℳ {A. A ≺⇩t atm_of K} = restrict_map ℳ {A. A ≺⇩t atm_of K}"
..
next
show "ord_res.is_maximal_lit K (eres D E)"
using step_hyps by argo
next
show "is_neg K"
using step_hyps by argo
next
show "ℳ (atm_of K) = Some C"
using match_hyps(3-)
by (metis (mono_tags, lifting) step_hyps(11) step_hyps(12) uminus_literal_def)
qed
hence "ord_res_6_step⇧+⇧+ S6 S6'"
using S6_def ‹𝒞 = Some E› S6'_def ord_res_6_step.intros by blast
moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
unfolding S6'_def ‹S7' = (N, s7')› ‹s7' = (U⇩e⇩r', ℱ, Γ', Some C)›
proof (rule ord_res_6_matches_ord_res_7.intros)
show "ord_res_5_invars N (U⇩e⇩r', ℱ, restrict_map ℳ {A. A ≺⇩t atm_of K}, Some C)"
using invars_6 unfolding ‹𝒞 = Some E›
using ord_res_6_preserves_invars[OF step6] by argo
next
show "ord_res_7_invars N (U⇩e⇩r', ℱ, Γ', Some C)"
using invars_s7' unfolding ‹s7' = (U⇩e⇩r', ℱ, Γ', Some C)› .
next
show "∀A C. (restrict_map ℳ {A. A ≺⇩t atm_of K} A = Some C) =
(map_of Γ' (Pos A) = Some (Some C))"
proof (intro allI)
fix A :: "'f gterm" and C :: "'f gclause"
show "restrict_map ℳ {A. A ≺⇩t atm_of K} A = Some C ⟷ map_of Γ' (Pos A) = Some (Some C)"
proof (cases "A ∈ dom ℳ ∧ A ≺⇩t atm_of K")
case True
have "restrict_map ℳ {A. A ≺⇩t atm_of K} A = Some C ⟷ ℳ A = Some C"
using True by simp
also have "… ⟷ map_of Γ (Pos A) = Some (Some C)"
using match_hyps(3-) by metis
also have "… ⟷ map_of Γ' (Pos A) = Some (Some C)"
proof -
have "Pos A ∈ fst ` set Γ"
using True
by (metis domIff map_of_eq_None_iff match_hyps(5) not_None_eq)
hence "∃𝒞. (Pos A, 𝒞) ∈ set Γ"
by fastforce
hence "∃𝒞. (Pos A, 𝒞) ∈ set Γ ∧ (Pos A, 𝒞) ∈ set Γ'"
using True unfolding mem_set_Γ'_iff prod.sel literal.sel by metis
moreover have "distinct (map fst Γ')"
using Γ_distinct_atoms
proof (rule distinct_suffix)
show "suffix (map fst Γ') (map fst Γ)"
using map_mono_suffix step_hyps(9) suffix_dropWhile by blast
qed
ultimately have "map_of Γ (Pos A) = map_of Γ' (Pos A)"
using Γ_distinct_atoms by (auto dest: map_of_is_SomeI)
thus ?thesis
by argo
qed
finally show ?thesis .
next
case False
have "restrict_map ℳ {A. A ≺⇩t atm_of K} A = None"
using False unfolding restrict_map_def by auto
moreover have "map_of Γ' (Pos A) ≠ Some (Some C)"
using False unfolding de_Morgan_conj
proof (elim disjE)
assume "A ∉ dom ℳ"
hence "⋀𝒞. (Pos A, 𝒞) ∉ set Γ"
using match_hyps(5)
by (metis (no_types, opaque_lifting) domIff fst_eqD invars_7 is_pos_def map_of_SomeD
not_None_eq snd_conv weak_map_of_SomeI)
hence "⋀𝒞. (Pos A, 𝒞) ∉ set Γ'"
unfolding mem_set_Γ'_iff by simp
then show "map_of Γ' (Pos A) ≠ Some (Some C)"
by (meson map_of_SomeD)
next
assume "¬ A ≺⇩t atm_of K"
hence "⋀𝒞. (Pos A, 𝒞) ∉ set Γ'"
unfolding mem_set_Γ'_iff by simp
then show "map_of Γ' (Pos A) ≠ Some (Some C)"
by (meson map_of_SomeD)
qed
ultimately show ?thesis
by simp
qed
qed
next
show "∀A. (restrict_map ℳ {A. A ≺⇩t atm_of K} A = None) =
(map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ')"
proof (intro allI)
fix A :: "'f gterm"
show "(restrict_map ℳ {A. A ≺⇩t atm_of K} A = None) =
(map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ')"
proof (cases "A ≺⇩t atm_of K")
case True
have "restrict_map ℳ {A. A ≺⇩t atm_of K} A = None ⟷ ℳ A = None"
using True by simp
also have "… ⟷ map_of Γ (Neg A) ≠ None ∨ A |∉| trail_atms Γ"
using match_hyps(6) by metis
also have "… ⟷ map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ"
using True mem_set_Γ'_iff
by (metis eq_fst_iff literal.sel(2) map_of_SomeD not_None_eq weak_map_of_SomeI)
also have "… ⟷ map_of Γ' (Neg A) ≠ None ∨ A |∉| trail_atms Γ'"
using True mem_set_Γ'_iff
by (smt (verit, best) fset_trail_atms image_iff)
finally show ?thesis .
next
case False
have "restrict_map ℳ {A. A ≺⇩t atm_of K} A = None"
using False by simp
moreover have "A |∉| trail_atms Γ'"
using False mem_set_Γ'_iff
by (smt (verit, ccfv_threshold) fset_trail_atms image_iff)
ultimately show ?thesis
by metis
qed
qed
next
show "i' = atms_of_clss (N |∪| U⇩e⇩r') |-| trail_atms Γ'"
using i'_def .
qed
ultimately show ?thesis
by metis
qed
qed
theorem bisimulation_ord_res_6_ord_res_7:
defines "match ≡ ord_res_6_matches_ord_res_7"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_6_state ⇒ 'f ord_res_7_state ⇒ bool) ℛ.
bisimulation ord_res_6_step (constant_context ord_res_7) ord_res_6_final ord_res_7_final
ℛ MATCH"
proof (rule ex_bisimulation_from_backward_simulation)
show "right_unique ord_res_6_step"
using right_unique_ord_res_6_step .
next
show "right_unique (constant_context ord_res_7)"
using right_unique_constant_context right_unique_ord_res_7 by metis
next
show "∀S. ord_res_6_final S ⟶ (∄S'. ord_res_6_step S S')"
by (metis finished_def ord_res_6_semantics.final_finished)
next
show "∀S. ord_res_7_final S ⟶ (∄S'. constant_context ord_res_7 S S')"
by (metis finished_def ord_res_7_semantics.final_finished)
next
show "∀i S6 S7. match i S6 S7 ⟶ ord_res_6_final S6 ⟷ ord_res_7_final S7"
unfolding match_def
using ord_res_6_final_iff_ord_res_7_final by metis
next
show "∀i S6 S7. match i S6 S7 ⟶
safe_state ord_res_6_step ord_res_6_final S6 ∧
safe_state (constant_context ord_res_7) ord_res_7_final S7"
proof (intro allI impI conjI)
fix i S6 S7
assume "match i S6 S7"
show "safe_state ord_res_6_step ord_res_6_final S6"
using ‹match i S6 S7›[unfolded match_def]
using ord_res_6_safe_state_if_invars
using ord_res_6_matches_ord_res_7.simps by auto
show "safe_state (constant_context ord_res_7) ord_res_7_final S7"
using ‹match i S6 S7›[unfolded match_def]
using ord_res_7_safe_state_if_invars
using ord_res_6_matches_ord_res_7.simps by auto
qed
next
show "wfp (|⊂|)"
using wfP_pfsubset .
next
show "∀i S6 S7 S7'. match i S6 S7 ⟶ constant_context ord_res_7 S7 S7' ⟶
(∃i' S6'. ord_res_6_step⇧+⇧+ S6 S6' ∧ match i' S6' S7') ∨
(∃i'. match i' S6 S7' ∧ i' |⊂| i)"
unfolding match_def
using backward_simulation_between_6_and_7 by metis
qed
end
section ‹ORD-RES-8 (atom-guided literal trail construction)›
context simulation_SCLFOL_ground_ordered_resolution begin
inductive ord_res_8_can_decide_neg where
"¬ trail_false_cls Γ C ⟹
linorder_lit.is_maximal_in_mset C L ⟹
linorder_trm.is_least_in_fset {|A |∈| atms_of_clss (N |∪| U⇩e⇩r).
A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ|} A ⟹
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ C"
inductive ord_res_8_can_skip_undefined_neg where
"¬ trail_false_cls Γ C ⟹
linorder_lit.is_maximal_in_mset C L ⟹
¬(∃A |∈| atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ) ⟹
¬ trail_defined_lit Γ L ⟹
is_neg L ⟹
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ C"
inductive ord_res_8_can_skip_undefined_pos_ultimate where
"¬ trail_false_cls Γ C ⟹
linorder_lit.is_maximal_in_mset C L ⟹
¬(∃A |∈| atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ) ⟹
¬ trail_defined_lit Γ L ⟹
is_pos L ⟹
¬ trail_false_cls Γ {#K ∈# C. K ≠ L#} ⟹
¬(∃D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). C ≺⇩c D) ⟹
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ C"
inductive ord_res_8_can_produce where
"¬ trail_false_cls Γ C ⟹
linorder_lit.is_maximal_in_mset C L ⟹
¬ (∃A |∈| atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ) ⟹
¬ trail_defined_lit Γ L ⟹
is_pos L ⟹
trail_false_cls Γ {#K ∈# C. K ≠ L#} ⟹
linorder_lit.is_greatest_in_mset C L ⟹
ord_res_8_can_produce N U⇩e⇩r ℱ Γ C"
inductive ord_res_8_can_factorize where
"¬ trail_false_cls Γ C ⟹
linorder_lit.is_maximal_in_mset C L ⟹
¬ (∃A |∈| atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ) ⟹
¬ trail_defined_lit Γ L ⟹
is_pos L ⟹
trail_false_cls Γ {#K ∈# C. K ≠ L#} ⟹
¬ linorder_lit.is_greatest_in_mset C L ⟹
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ C"
definition is_least_nonskipped_clause where
"is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C ⟷
linorder_cls.is_least_in_fset {|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r).
trail_false_cls Γ C ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ C|} C"
lemma is_least_nonskipped_clause_mempty:
assumes bot_in: "{#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
shows "is_least_nonskipped_clause N U⇩e⇩r ℱ Γ {#}"
unfolding is_least_nonskipped_clause_def linorder_cls.is_least_in_ffilter_iff
proof (intro conjI ballI impI)
show "{#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using bot_in .
next
show "trail_false_cls Γ {#} ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ {#} ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ {#} ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ {#} ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ {#} ∨ ord_res_8_can_factorize N U⇩e⇩r ℱ Γ {#}"
by simp
next
fix C :: "'f gclause"
assume "C ≠ {#}"
thus "{#} ≺⇩c C"
using mempty_lesseq_cls by blast
qed
lemma nex_is_least_nonskipped_clause_if:
assumes
no_undef_atom: "¬ (∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A |∉| trail_atms Γ)" and
no_false_clause: "¬ fBex (iefac ℱ |`| (N |∪| U⇩e⇩r)) (trail_false_cls Γ)"
shows "∄C. is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C"
unfolding not_ex
proof (intro allI notI)
fix C :: "'f gclause"
assume "is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C"
hence C_in: "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
"ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ C"
unfolding atomize_conj
unfolding is_least_nonskipped_clause_def
unfolding linorder_cls.is_least_in_ffilter_iff
using no_false_clause by metis
hence "∃L. linorder_lit.is_maximal_in_mset C L ∧ ¬ trail_defined_lit Γ L"
proof (elim disjE)
assume "ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ C"
then show ?thesis
by (metis (mono_tags, lifting) assms(1) linorder_trm.is_least_in_ffilter_iff
ord_res_8_can_decide_neg.cases)
qed (auto elim:
ord_res_8_can_skip_undefined_neg.cases
ord_res_8_can_skip_undefined_pos_ultimate.cases
ord_res_8_can_produce.cases
ord_res_8_can_factorize.cases)
hence "∃L. atm_of L |∈| atms_of_clss (N |∪| U⇩e⇩r) ∧ atm_of L |∉| trail_atms Γ"
using C_in
unfolding linorder_lit.is_maximal_in_mset_iff trail_defined_lit_iff_trail_defined_atm
by (metis atm_of_in_atms_of_clssI)
thus False
using no_undef_atom by metis
qed
lemma MAGIC5:
assumes invars: "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞)" and
no_more_steps: "∄𝒞'. ord_res_7 N (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ, 𝒞')"
shows "(∀C. 𝒞 = Some C ⟷ is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C)"
proof (cases 𝒞)
case None
have "trail_atms Γ = atms_of_clss (N |∪| U⇩e⇩r)"
using None invars[unfolded ord_res_7_invars_def] by simp
have "∀C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). trail_true_cls Γ C"
using None invars[unfolded ord_res_7_invars_def] by simp
hence no_false: "∀C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). ¬ trail_false_cls Γ C"
using invars[unfolded ord_res_7_invars_def]
by (meson invars not_trail_true_cls_and_trail_false_cls
ord_res_7_invars_implies_trail_consistent)
have "∄C. is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C"
proof (rule notI, elim exE)
fix C
assume "is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C"
hence
C_in: "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
C_spec_disj: "
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ C"
unfolding atomize_conj
unfolding is_least_nonskipped_clause_def
unfolding linorder_cls.is_least_in_ffilter_iff
using no_false by metis
thus False
proof (elim disjE)
assume "ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ C"
thus ?thesis
using ‹trail_atms Γ = atms_of_clss (N |∪| U⇩e⇩r)›
using ord_res_8_can_decide_neg.cases
by (metis (no_types, lifting) linorder_trm.is_least_in_ffilter_iff)
next
assume "ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ C"
thus ?thesis
using ‹trail_atms Γ = atms_of_clss (N |∪| U⇩e⇩r)›
using ord_res_8_can_skip_undefined_neg.cases
by (metis C_in atm_of_in_atms_of_clssI linorder_lit.is_maximal_in_mset_iff
trail_defined_lit_iff_trail_defined_atm)
next
assume "ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ C"
thus ?thesis
using ‹trail_atms Γ = atms_of_clss (N |∪| U⇩e⇩r)›
using ord_res_8_can_skip_undefined_pos_ultimate.cases
by (metis C_in atm_of_in_atms_of_clssI linorder_lit.is_maximal_in_mset_iff
trail_defined_lit_iff_trail_defined_atm)
next
assume "ord_res_8_can_produce N U⇩e⇩r ℱ Γ C"
thus False
using ‹trail_atms Γ = atms_of_clss (N |∪| U⇩e⇩r)›
using ord_res_8_can_produce.cases
by (metis C_in atm_of_in_atms_of_clssI linorder_lit.is_maximal_in_mset_iff
trail_defined_lit_iff_trail_defined_atm)
next
assume "ord_res_8_can_factorize N U⇩e⇩r ℱ Γ C"
thus False
using ‹trail_atms Γ = atms_of_clss (N |∪| U⇩e⇩r)›
using ord_res_8_can_factorize.cases
by (metis C_in atm_of_in_atms_of_clssI linorder_lit.is_maximal_in_mset_iff
trail_defined_lit_iff_trail_defined_atm)
qed
qed
thus ?thesis
using None by simp
next
case (Some D)
have D_in: "D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using Some invars[unfolded ord_res_7_invars_def] by simp
have "is_least_nonskipped_clause N U⇩e⇩r ℱ Γ D"
proof (cases "D = {#}")
case True
thus ?thesis
using D_in is_least_nonskipped_clause_mempty by metis
next
case False
then obtain L⇩D where D_max_lit: "linorder_lit.is_maximal_in_mset D L⇩D"
using linorder_lit.ex_maximal_in_mset by presburger
show ?thesis
unfolding is_least_nonskipped_clause_def
unfolding linorder_cls.is_least_in_ffilter_iff
proof (intro conjI ballI impI)
show "D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using D_in .
next
show "trail_false_cls Γ D ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ D ∨ ord_res_8_can_factorize N U⇩e⇩r ℱ Γ D"
proof (cases "trail_false_cls Γ D")
case True
then show ?thesis
by metis
next
case D_not_false: False
then obtain L where D_max_lit: "linorder_lit.is_maximal_in_mset D L"
by (metis linorder_lit.ex_maximal_in_mset trail_false_cls_mempty)
show ?thesis
proof (cases "∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ")
case True
hence "ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D"
using ord_res_8_can_decide_neg.intros[OF D_not_false D_max_lit]
by (metis (no_types, lifting) equalsffemptyD ffmember_filter linorder_trm.ex_least_in_fset)
thus ?thesis
by argo
next
case no_undef_atm: False
show ?thesis
proof (cases "trail_defined_lit Γ L")
case L_defined: True
hence "∃𝒞'. ord_res_7 N (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ, 𝒞')"
unfolding ‹𝒞 = Some D›
using ord_res_7.skip_defined[OF D_not_false D_max_lit no_undef_atm]
by metis
hence False
using no_more_steps by contradiction
thus ?thesis ..
next
case L_undef: False
show ?thesis
proof (cases L)
case (Pos A)
hence L_pos: "is_pos L"
by simp
show ?thesis
proof (cases "trail_false_cls Γ {#K ∈# D. K ≠ L#}")
case D_almost_false: True
thus ?thesis
using ord_res_8_can_factorize.intros[
OF D_not_false D_max_lit no_undef_atm L_undef L_pos]
using ord_res_8_can_produce.intros[
OF D_not_false D_max_lit no_undef_atm L_undef L_pos]
by metis
next
case D_not_flagrantly_false: False
show ?thesis
proof (cases "∃E |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). D ≺⇩c E")
case True
hence "∃𝒞'. ord_res_7 N (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ, 𝒞')"
unfolding ‹𝒞 = Some D›
using ord_res_7.skip_undefined_pos[
OF D_not_false D_max_lit no_undef_atm L_undef L_pos D_not_flagrantly_false]
by (metis femptyE ffmember_filter linorder_cls.ex_least_in_fset)
hence False
using no_more_steps by contradiction
thus ?thesis ..
next
case False
hence "ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D"
using ord_res_8_can_skip_undefined_pos_ultimate.intros[
OF D_not_false D_max_lit no_undef_atm L_undef L_pos D_not_flagrantly_false]
by metis
thus ?thesis
by argo
qed
qed
next
case (Neg A)
hence L_neg: "is_neg L"
by simp
hence "ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D"
unfolding ‹𝒞 = Some D›
using ord_res_8_can_skip_undefined_neg.intros[
OF D_not_false D_max_lit no_undef_atm L_undef]
by metis
thus ?thesis
by argo
qed
qed
qed
qed
fix E :: "'f gterm literal multiset"
assume
E_in: "E |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
E_neq: "E ≠ D" and
E_spec: "trail_false_cls Γ E ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ E ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ E ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ E ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ E ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ E"
have true_cls_if_lt_D:
"∀C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). C ≺⇩c D ⟶ trail_true_cls Γ C"
using invars[unfolded ord_res_7_invars_def] Some by simp
have Γ_lower_set: "linorder_trm.is_lower_fset (trail_atms Γ) (atms_of_clss (N |∪| U⇩e⇩r))"
using invars[unfolded ord_res_7_invars_def] by simp
have FOO:
"∀C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r). C ≺⇩c D ⟶
(∀K. ord_res.is_maximal_lit K C ⟶
¬ (∃A |∈| atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of K ∧ A |∉| trail_atms Γ))"
using invars[unfolded ord_res_7_invars_def] Some E_in by simp
hence BAR:
"∀C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). C ≺⇩c D ⟶
(∀K. linorder_lit.is_maximal_in_mset C K ⟶
¬ trail_defined_lit Γ K ⟶ (trail_true_cls Γ {#x ∈# C. x ≠ K#} ∧ is_pos K))"
using invars[unfolded ord_res_7_invars_def] Some by simp
show "D ≺⇩c E"
using E_spec
proof (elim disjE)
assume "trail_false_cls Γ E"
hence "¬ trail_true_cls Γ E"
using invars not_trail_true_cls_and_trail_false_cls
ord_res_7_invars_implies_trail_consistent by blast
thus "D ≺⇩c E"
using E_in E_neq true_cls_if_lt_D by force
next
assume "ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ E"
thus "D ≺⇩c E"
proof (cases N U⇩e⇩r ℱ Γ E rule: ord_res_8_can_decide_neg.cases)
case (1 L⇩E A)
hence "∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L⇩E ∧ A |∉| trail_atms Γ"
unfolding linorder_trm.is_least_in_ffilter_iff by metis
thus ?thesis
using FOO[rule_format, OF E_in _ ‹ord_res.is_maximal_lit L⇩E E›] E_in E_neq
by force
qed
next
assume "ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ E"
thus "D ≺⇩c E"
proof (cases N U⇩e⇩r ℱ Γ E rule: ord_res_8_can_skip_undefined_neg.cases)
case hyps: (1 L⇩E)
thus ?thesis
using BAR[rule_format, OF E_in _ ‹ord_res.is_maximal_lit L⇩E E›]
using invars[unfolded ord_res_7_invars_def Some, rule_format, OF refl] Some E_in
using E_neq by fastforce
qed
next
assume "ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ E"
thus "D ≺⇩c E"
proof (cases N U⇩e⇩r ℱ Γ E rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
case (1 L)
then show ?thesis using E_neq D_in by force
qed
next
assume "ord_res_8_can_produce N U⇩e⇩r ℱ Γ E"
hence "¬ trail_true_cls Γ E"
proof (cases N U⇩e⇩r ℱ Γ E rule: ord_res_8_can_produce.cases)
case (1 L)
then show ?thesis
using invars[THEN ord_res_7_invars_implies_trail_consistent]
by (smt (verit, ccfv_SIG) mem_Collect_eq not_trail_true_cls_and_trail_false_cls
set_mset_filter trail_defined_lit_iff_true_or_false trail_true_cls_def)
qed
thus "D ≺⇩c E"
using E_in E_neq true_cls_if_lt_D by force
next
assume "ord_res_8_can_factorize N U⇩e⇩r ℱ Γ E"
hence "¬ trail_true_cls Γ E"
proof (cases N U⇩e⇩r ℱ Γ E rule: ord_res_8_can_factorize.cases)
case (1 L)
then show ?thesis
using invars[THEN ord_res_7_invars_implies_trail_consistent]
by (smt (verit, ccfv_SIG) mem_Collect_eq not_trail_true_cls_and_trail_false_cls
set_mset_filter trail_defined_lit_iff_true_or_false trail_true_cls_def)
qed
thus "D ≺⇩c E"
using E_in E_neq true_cls_if_lt_D by force
qed
qed
qed
moreover have "Uniq (is_least_nonskipped_clause N U⇩e⇩r ℱ Γ)"
unfolding is_least_nonskipped_clause_def
using linorder_cls.Uniq_is_least_in_fset
by simp
ultimately show ?thesis
using Some
by (metis (no_types) The_optional_eq_SomeD The_optional_eq_SomeI)
qed
lemma MAGIC6:
assumes invars: "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞)"
shows "∃𝒞'. (ord_res_7 N)⇧*⇧* (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ, 𝒞') ∧
(∄𝒞''. ord_res_7 N (U⇩e⇩r, ℱ, Γ, 𝒞') (U⇩e⇩r, ℱ, Γ, 𝒞''))"
proof -
define R where
"R = (λ𝒞 𝒞'. ord_res_7 N (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ, 𝒞'))"
define f :: "'f gclause option ⇒ 'f gclause fset" where
"f = (λ𝒞. case 𝒞 of None ⇒ {||} | Some C ⇒ {|D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). C ≼⇩c D|})"
let ?less_f = "(|⊂|)"
have "Wellfounded.wfp_on {x'. R⇧*⇧* 𝒞 x'} R¯¯"
proof (rule wfp_on_if_convertible_to_wfp_on)
have "wfp (|⊂|)"
by auto
thus "Wellfounded.wfp_on (f ` {x'. R⇧*⇧* 𝒞 x'}) ?less_f"
using Wellfounded.wfp_on_subset subset_UNIV by metis
next
fix 𝒞⇩x 𝒞⇩y :: "'f gclause option"
have rtranclp_with_constsD: "(λy y'. R (x, y) (x, y'))⇧*⇧* y y' ⟹
R⇧*⇧* (x, y) (x, y')" for R x y y'
proof (induction y arbitrary: rule: converse_rtranclp_induct)
case base
show ?case
by simp
next
case (step w)
thus ?case
by force
qed
assume "𝒞⇩x ∈ {x'. R⇧*⇧* 𝒞 x'}" and "𝒞⇩y ∈ {x'. R⇧*⇧* 𝒞 x'}"
hence
"(ord_res_7 N)⇧*⇧* (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ, 𝒞⇩x)" and
"(ord_res_7 N)⇧*⇧* (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ, 𝒞⇩y)"
unfolding atomize_conj mem_Collect_eq R_def
by (auto intro: rtranclp_with_constsD)
hence
x_invars: "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞⇩x)" and
y_invars: "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞⇩y)"
using ord_res_7_preserves_invars
using invars by (metis rtranclp_ord_res_7_preserves_ord_res_7_invars)+
have Γ_consistent: "trail_consistent Γ"
using x_invars by (metis ord_res_7_invars_implies_trail_consistent)
have less_f_if: "?less_f (f 𝒞⇩y) (f 𝒞⇩x)"
if "𝒞⇩x = Some C" and
𝒞⇩y_disj: "𝒞⇩y = None ∨ 𝒞⇩y = Some D ∧ C ≺⇩c D" and
C_in: "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
for C D
proof -
have f_x: "f 𝒞⇩x = {|D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). C ≼⇩c D|}"
by (auto simp add: ‹𝒞⇩x = Some C› f_def)
moreover have "C |∈| f 𝒞⇩x"
using C_in f_x by simp
moreover have "C |∉| f 𝒞⇩y ∧ f 𝒞⇩y |⊆| f 𝒞⇩x"
using 𝒞⇩y_disj
proof (elim disjE conjE; intro conjI)
assume "𝒞⇩y = None"
thus "C |∉| f 𝒞⇩y" and "f 𝒞⇩y |⊆| f 𝒞⇩x"
unfolding f_x
by (simp_all add: f_def)
next
assume "𝒞⇩y = Some D" and "C ≺⇩c D"
have "⋀x. D ≼⇩c x ⟹ C ≼⇩c x"
using ‹C ≺⇩c D› by auto
moreover have fst_f_y: "f 𝒞⇩y = {|E |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). D ≼⇩c E|}"
by (auto simp add: ‹𝒞⇩y = Some D› f_def)
ultimately show "f 𝒞⇩y |⊆| f 𝒞⇩x"
using f_x by auto
show "C |∉| f 𝒞⇩y"
using ‹C ≺⇩c D› fst_f_y by auto
qed
ultimately have "f 𝒞⇩y |⊂| f 𝒞⇩x"
by blast
thus ?thesis
by (simp add: lex_prodp_def)
qed
have eres_not_in_if: "eres D E |∉| U⇩e⇩r"
if "𝒞⇩x = Some E" and E_false: "trail_false_cls Γ E" and
E_max_lit: "ord_res.is_maximal_lit L E" and L_neg: "is_neg L"
"map_of Γ (- L) = Some (Some D)"
for D E L
proof -
have
clauses_lt_E_true:
"∀C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). C ≺⇩c E ⟶ trail_true_cls Γ C" and
Γ_prop_greatest:
"∀Ln ∈ set Γ. ∀C. snd Ln = Some C ⟶ linorder_lit.is_greatest_in_mset C (fst Ln)"
using x_invars unfolding ‹𝒞⇩x = Some E› ord_res_7_invars_def by simp_all
have "(- L, Some D) ∈ set Γ"
using ‹map_of Γ (- L) = Some (Some D)› by (metis map_of_SomeD)
hence D_greatest_lit: "linorder_lit.is_greatest_in_mset D (- L)"
using Γ_prop_greatest by fastforce
have "eres D E ≺⇩c E"
using eres_lt_if
using E_max_lit L_neg D_greatest_lit
by metis
hence "eres D E ≠ E"
by order
have "L ∈# E"
using E_max_lit unfolding linorder_lit.is_maximal_in_mset_iff by metis
hence "- L ∉# E"
using not_both_lit_and_comp_lit_in_false_clause_if_trail_consistent
using Γ_consistent E_false by metis
hence "∀K ∈# eres D E. atm_of K ≺⇩t atm_of L"
using lit_in_eres_lt_greatest_lit_in_grestest_resolvant[OF ‹eres D E ≠ E› E_max_lit]
by metis
hence "∀K ∈# eres D E. K ≠ L ∧ K ≠ - L"
by fastforce
moreover have "∀L ∈# eres D E. L ∈# D ∨ L ∈# E"
using lit_in_one_of_resolvents_if_in_eres by metis
moreover have D_almost_false: "trail_false_cls Γ {#K ∈# D. K ≠ - L#}"
using ord_res_7_invars_implies_propagated_clause_almost_false
using ‹(- L, Some D) ∈ set Γ› x_invars
by metis
ultimately have "trail_false_cls Γ (eres D E)"
using E_false unfolding trail_false_cls_def by fastforce
have "eres D E |∉| N |∪| U⇩e⇩r"
using eres_not_in_known_clauses_if_trail_false_cls
using Γ_consistent clauses_lt_E_true ‹eres D E ≺⇩c E› ‹trail_false_cls Γ (eres D E)›
by metis
thus "eres D E |∉| U⇩e⇩r"
by blast
qed
assume "R¯¯ 𝒞⇩y 𝒞⇩x"
hence "ord_res_7 N (U⇩e⇩r, ℱ, Γ, 𝒞⇩x) (U⇩e⇩r, ℱ, Γ, 𝒞⇩y)"
unfolding conversep_iff R_def .
thus "?less_f (f 𝒞⇩y) (f 𝒞⇩x)"
proof (cases N "(U⇩e⇩r, ℱ, Γ, 𝒞⇩x)" "(U⇩e⇩r, ℱ, Γ, 𝒞⇩y)" rule: ord_res_7.cases)
case step_hyps: (decide_neg C L A)
hence False by simp
thus ?thesis ..
next
case step_hyps: (skip_defined C L)
have "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using ord_res_7_invars_def step_hyps(1) x_invars by presburger
moreover have "∃D. 𝒞⇩y = None ∨ 𝒞⇩y = Some D ∧ C ≺⇩c D"
proof (cases 𝒞⇩y)
case None
thus ?thesis
by simp
next
case (Some D)
thus ?thesis
using step_hyps
by (metis linorder_cls.is_least_in_ffilter_iff Some_eq_The_optionalD)
qed
ultimately show ?thesis
using less_f_if step_hyps by metis
next
case step_hyps: (skip_undefined_neg C L)
hence False by simp
thus ?thesis ..
next
case step_hyps: (skip_undefined_pos C L D)
have "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using ord_res_7_invars_def step_hyps(1) x_invars by presburger
moreover have "∃D. 𝒞⇩y = None ∨ 𝒞⇩y = Some D ∧ C ≺⇩c D"
using step_hyps by (metis linorder_cls.is_least_in_ffilter_iff)
ultimately show ?thesis
using less_f_if step_hyps by metis
next
case step_hyps: (skip_undefined_pos_ultimate C L)
hence False by simp
thus ?thesis ..
next
case step_hyps: (production C L)
hence False by simp
thus ?thesis ..
next
case step_hyps: (factoring C L)
have "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using ord_res_7_invars_def step_hyps(1) x_invars by presburger
moreover have "efac C ≠ C"
using step_hyps by (metis greatest_literal_in_efacI)
ultimately have "C |∉| ℱ"
by (smt (verit, ccfv_threshold) fimage_iff iefac_def ex1_efac_eq_factoring_chain
factorizable_if_neq_efac)
hence False
using ‹ℱ = finsert C ℱ› by blast
thus ?thesis ..
next
case step_hyps: (resolution_bot E L D)
hence "eres D E |∉| U⇩e⇩r"
using eres_not_in_if by metis
hence False
using ‹U⇩e⇩r = finsert (eres D E) U⇩e⇩r› by blast
thus ?thesis ..
next
case (resolution_pos E L D K)
hence "eres D E |∉| U⇩e⇩r"
using eres_not_in_if by metis
hence False
using ‹U⇩e⇩r = finsert (eres D E) U⇩e⇩r› by blast
thus ?thesis ..
next
case (resolution_neg E L D K C)
hence "eres D E |∉| U⇩e⇩r"
using eres_not_in_if by metis
hence False
using ‹U⇩e⇩r = finsert (eres D E) U⇩e⇩r› by blast
thus ?thesis ..
qed
qed
then obtain 𝒞' where "R⇧*⇧* 𝒞 𝒞'" and "∄z. R 𝒞' z"
using ex_terminating_rtranclp_strong by metis
show ?thesis
proof (intro exI conjI)
show "(ord_res_7 N)⇧*⇧* (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ, 𝒞')"
using ‹R⇧*⇧* 𝒞 𝒞'›
by (induction 𝒞 rule: converse_rtranclp_induct) (auto simp: R_def)
next
show "∄𝒞''. ord_res_7 N (U⇩e⇩r, ℱ, Γ, 𝒞') (U⇩e⇩r, ℱ, Γ, 𝒞'')"
using ‹∄z. R 𝒞' z›
by (simp add: R_def)
qed
qed
inductive ord_res_7_matches_ord_res_8 :: "'f ord_res_7_state ⇒ 'f ord_res_8_state ⇒ bool" where
"ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞) ⟹
ord_res_8_invars N (U⇩e⇩r, ℱ, Γ) ⟹
(∀C. 𝒞 = Some C ⟷ is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C) ⟹
ord_res_7_matches_ord_res_8 (N, U⇩e⇩r, ℱ, Γ, 𝒞) (N, U⇩e⇩r, ℱ, Γ)"
lemma ord_res_7_final_iff_ord_res_8_final:
fixes S7 S8
assumes match: "ord_res_7_matches_ord_res_8 S7 S8"
shows "ord_res_7_final S7 ⟷ ord_res_8_final S8"
using match
proof (cases S7 S8 rule: ord_res_7_matches_ord_res_8.cases)
case match_hyps: (1 N U⇩e⇩r ℱ Γ 𝒞)
note invars7 = ‹ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞)›[unfolded ord_res_7_invars_def,
rule_format, OF refl]
have Γ_consistent: "trail_consistent Γ"
using invars7 by (metis trail_consistent_if_sorted_wrt_atoms)
show "ord_res_7_final S7 ⟷ ord_res_8_final S8"
proof (rule iffI)
assume "ord_res_7_final S7"
thus "ord_res_8_final S8"
unfolding ‹S7 = (N, U⇩e⇩r, ℱ, Γ, 𝒞)›
proof (cases "(N, U⇩e⇩r, ℱ, Γ, 𝒞)" rule: ord_res_7_final.cases)
case model_found
show "ord_res_8_final S8"
unfolding ‹S8 = (N, U⇩e⇩r, ℱ, Γ)›
proof (rule ord_res_8_final.model_found)
have "𝒞 = None ⟶ trail_atms Γ = atms_of_clss (N |∪| U⇩e⇩r)"
using invars7 by argo
hence "trail_atms Γ = atms_of_clss (N |∪| U⇩e⇩r)"
using model_found by argo
thus "¬ (∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A |∉| trail_atms Γ)"
by metis
next
have "∀C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). trail_true_cls Γ C"
using invars7 model_found by simp
moreover have "¬ (trail_true_cls Γ C ∧ trail_false_cls Γ C)" for C
using not_trail_true_cls_and_trail_false_cls[OF Γ_consistent] .
ultimately show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r). trail_false_cls Γ C)"
by metis
qed
next
case contradiction_found
show "ord_res_8_final S8"
unfolding ‹S8 = (N, U⇩e⇩r, ℱ, Γ)›
proof (rule ord_res_8_final.contradiction_found)
show "{#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using invars7 ‹𝒞 = Some {#}› by metis
qed
qed
next
assume "ord_res_8_final S8"
thus "ord_res_7_final S7"
unfolding ‹S8 = (N, U⇩e⇩r, ℱ, Γ)›
proof (cases "(N, U⇩e⇩r, ℱ, Γ)" rule: ord_res_8_final.cases)
case model_found
hence "∄C. is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C"
using nex_is_least_nonskipped_clause_if by metis
hence "𝒞 = None"
using match_hyps by simp
thus "ord_res_7_final S7"
unfolding ‹S7 = (N, U⇩e⇩r, ℱ, Γ, 𝒞)›
using ord_res_7_final.model_found by metis
next
case contradiction_found
hence "is_least_nonskipped_clause N U⇩e⇩r ℱ Γ {#}"
using is_least_nonskipped_clause_mempty by metis
hence "𝒞 = Some {#}"
using match_hyps by presburger
thus "ord_res_7_final S7"
unfolding ‹S7 = (N, U⇩e⇩r, ℱ, Γ, 𝒞)›
using ord_res_7_final.contradiction_found by metis
qed
qed
qed
lemma backward_simulation_between_7_and_8:
fixes i S7 S8 S8'
assumes match: "ord_res_7_matches_ord_res_8 S7 S8" and step: "constant_context ord_res_8 S8 S8'"
shows "∃S7'. (constant_context ord_res_7)⇧+⇧+ S7 S7' ∧ ord_res_7_matches_ord_res_8 S7' S8'"
using match
proof (cases S7 S8 rule: ord_res_7_matches_ord_res_8.cases)
case match_hyps: (1 N U⇩e⇩r ℱ Γ 𝒞)
note S7_def = ‹S7 = (N, U⇩e⇩r, ℱ, Γ, 𝒞)›
note invars7 = ‹ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞)›[unfolded ord_res_7_invars_def,
rule_format, OF refl]
have Γ_sorted: "sorted_wrt (λx y. atm_of (fst y) ≺⇩t atm_of (fst x)) Γ"
using invars7 by argo
have Γ_consistent: "trail_consistent Γ"
using trail_consistent_if_sorted_wrt_atoms[OF Γ_sorted] .
have Γ_lower_set: "linorder_trm.is_lower_fset (trail_atms Γ) (atms_of_clss (N |∪| U⇩e⇩r))"
using invars7 by argo
have 𝒞_eq_None_implies_all_atoms_defined: "𝒞 = None ⟶ trail_atms Γ = atms_of_clss (N |∪| U⇩e⇩r)"
using invars7 by argo
obtain s8' where
S8'_def: "S8' = (N, s8')" and
step': "ord_res_8 N (U⇩e⇩r, ℱ, Γ) s8'"
using step unfolding ‹S8 = (N, U⇩e⇩r, ℱ, Γ)›
by (auto elim: constant_context.cases)
have invars_s8': "ord_res_8_invars N s8'"
using ord_res_8_preserves_invars[OF step' ‹ord_res_8_invars N (U⇩e⇩r, ℱ, Γ)›] .
show ?thesis
using step'
proof (cases N "(U⇩e⇩r, ℱ, Γ)" s8' rule: ord_res_8.cases)
case step_hyps: (decide_neg A Γ')
have
A_in: "A |∈| atms_of_clss (N |∪| U⇩e⇩r)" and
A_undef: "A |∉| trail_atms Γ" and
A_least: "∀y|∈|atms_of_clss (N |∪| U⇩e⇩r). y ≠ A ⟶ (∀A⇩1|∈|trail_atms Γ. A⇩1 ≺⇩t y) ⟶ A ≺⇩t y"
using step_hyps(3) unfolding linorder_trm.is_least_in_fset_iff by auto
have "𝒞 ≠ None"
using 𝒞_eq_None_implies_all_atoms_defined A_in A_undef by metis
then obtain D :: "'f gclause" where "𝒞 = Some D"
by blast
hence D_in: "D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
by (metis ‹𝒞 = Some D› invars7)
have "is_least_nonskipped_clause N U⇩e⇩r ℱ Γ D"
using match_hyps ‹𝒞 = Some D› by metis
moreover have D_not_false: "¬ trail_false_cls Γ D"
using D_in step_hyps by metis
moreover have "¬ ord_res_8_can_produce N U⇩e⇩r ℱ Γ D"
proof (rule notI)
assume "ord_res_8_can_produce N U⇩e⇩r ℱ Γ D"
thus False
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_produce.cases)
case (1 L')
thus ?thesis
by (metis A_in A_least A_undef D_in atm_of_in_atms_of_clssI
atoms_of_trail_lt_atom_of_propagatable_literal clause_could_propagate_def invars7
linorder_lit.is_maximal_in_mset_iff literal.collapse(1) step_hyps(4))
qed
qed
moreover have "¬ ord_res_8_can_factorize N U⇩e⇩r ℱ Γ D"
proof (rule notI)
assume "ord_res_8_can_factorize N U⇩e⇩r ℱ Γ D"
thus False
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_factorize.cases)
case (1 L')
thus False
by (metis A_in A_least A_undef D_in atm_of_in_atms_of_clssI
atoms_of_trail_lt_atom_of_propagatable_literal clause_could_propagate_def invars7
linorder_lit.is_maximal_in_mset_iff literal.collapse(1) step_hyps(4))
qed
qed
ultimately have "ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D"
unfolding is_least_nonskipped_clause_def
unfolding linorder_cls.is_least_in_ffilter_iff
by argo
then obtain 𝒞' where first_step7: "ord_res_7 N (U⇩e⇩r, ℱ, Γ, Some D) (U⇩e⇩r, ℱ, Γ', 𝒞')"
proof (elim disjE; atomize_elim)
assume "ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D"
thus "∃𝒞'. ord_res_7 N (U⇩e⇩r, ℱ, Γ, Some D) (U⇩e⇩r, ℱ, Γ', 𝒞')"
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_decide_neg.cases)
case hyps: (1 L A')
hence "A' = A"
by (smt (verit, del_insts) Γ_lower_set linorder_trm.is_least_in_ffilter_iff
linorder_trm.neq_iff linorder_trm.not_in_lower_setI linorder_trm.order.strict_trans
step_hyps(3))
thus ?thesis
using hyps ‹Γ' = (Neg A, None) # Γ›
using ord_res_7.decide_neg[of Γ D _ N U⇩e⇩r A Γ' ℱ] by blast
qed
next
assume "ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D"
thus "∃𝒞'. ord_res_7 N (U⇩e⇩r, ℱ, Γ, Some D) (U⇩e⇩r, ℱ, Γ', 𝒞')"
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_skip_undefined_neg.cases)
case hyps: (1 L)
hence "L = Neg A"
by (smt (verit) A_in A_least A_undef D_in Γ_lower_set atm_of_in_atms_of_clssI
linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
linorder_trm.not_in_lower_setI literal.disc(2) literal.expand literal.sel(2)
trail_defined_lit_iff_trail_defined_atm)
thus ?thesis
using hyps ‹Γ' = (Neg A, None) # Γ›
using ord_res_7.skip_undefined_neg by blast
qed
next
assume "ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D"
thus "∃𝒞'. ord_res_7 N (U⇩e⇩r, ℱ, Γ, Some D) (U⇩e⇩r, ℱ, Γ', 𝒞')"
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
case hyps: (1 L)
hence "L = Pos A"
by (smt (verit, best) A_in A_least A_undef D_in atm_of_in_atms_of_clssI invars7
linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
linorder_trm.not_in_lower_setI literal.disc(1) literal.expand literal.sel(1)
trail_defined_lit_iff_trail_defined_atm)
thus ?thesis
using hyps ‹Γ' = (Neg A, None) # Γ›
using ord_res_7.skip_undefined_pos_ultimate by fastforce
qed
qed
moreover have "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞')"
using ‹𝒞 = Some D› first_step7 match_hyps(3) ord_res_7_preserves_invars by blast
ultimately obtain 𝒞'' where
following_steps7: "(ord_res_7 N)⇧*⇧* (U⇩e⇩r, ℱ, Γ', 𝒞') (U⇩e⇩r, ℱ, Γ', 𝒞'')" and
no_more_step7: "(∄𝒞'''. ord_res_7 N (U⇩e⇩r, ℱ, Γ', 𝒞'') (U⇩e⇩r, ℱ, Γ', 𝒞'''))"
using MAGIC6 by metis
show ?thesis
proof (intro exI conjI)
have "(ord_res_7 N)⇧+⇧+ (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ', 𝒞'')"
unfolding ‹𝒞 = Some D›
using first_step7 following_steps7 by simp
thus "(constant_context ord_res_7)⇧+⇧+ S7 (N, U⇩e⇩r, ℱ, Γ', 𝒞'')"
unfolding S7_def by (simp add: tranclp_constant_context)
show "ord_res_7_matches_ord_res_8 (N, U⇩e⇩r, ℱ, Γ', 𝒞'') S8'"
unfolding S8'_def ‹s8' = (U⇩e⇩r, ℱ, Γ')›
proof (intro ord_res_7_matches_ord_res_8.intros allI)
show "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞'')"
using ‹ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞')› following_steps7
rtranclp_ord_res_7_preserves_ord_res_7_invars by blast
show "ord_res_8_invars N (U⇩e⇩r, ℱ, Γ')"
using invars_s8' step_hyps(1) by blast
fix C :: "'f gclause"
show "𝒞'' = Some C ⟷ is_least_nonskipped_clause N U⇩e⇩r ℱ Γ' C"
using MAGIC5 ‹ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞'')› no_more_step7 by metis
qed
qed
next
case step_hyps: (propagate A C Γ')
have
A_in: "A |∈| atms_of_clss (N |∪| U⇩e⇩r)" and
A_undef: "A |∉| trail_atms Γ" and
A_least: "∀y|∈|atms_of_clss (N |∪| U⇩e⇩r). y ≠ A ⟶ (∀A⇩1|∈|trail_atms Γ. A⇩1 ≺⇩t y) ⟶ A ≺⇩t y"
using step_hyps(3) unfolding linorder_trm.is_least_in_fset_iff by auto
have
C_in: "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
C_can_prop:"clause_could_propagate Γ C (Pos A)" and
C_least: "∀D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r).
D ≠ C ⟶ clause_could_propagate Γ D (Pos A) ⟶ C ≺⇩c D"
using step_hyps unfolding atomize_conj linorder_cls.is_least_in_ffilter_iff by argo
hence
Pos_A_undef: "¬ trail_defined_lit Γ (Pos A)" and
C_max_lit: "linorder_lit.is_maximal_in_mset C (Pos A)" and
C_almost_false: "trail_false_cls Γ {#K ∈# C. K ≠ Pos A#}"
unfolding atomize_conj clause_could_propagate_def by argo
have "is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C"
unfolding is_least_nonskipped_clause_def
unfolding linorder_cls.is_least_in_ffilter_iff
proof (intro conjI ballI impI)
show "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using C_in .
next
have "ord_res_8_can_produce N U⇩e⇩r ℱ Γ C"
proof (rule ord_res_8_can_produce.intros)
show "¬ trail_false_cls Γ C"
using step_hyps C_in by metis
next
show "ord_res.is_maximal_lit (Pos A) C"
using step_hyps by blast
next
show "¬ (∃Aa|∈|atms_of_clss (N |∪| U⇩e⇩r). Aa ≺⇩t atm_of (Pos A) ∧ Aa |∉| trail_atms Γ)"
unfolding literal.sel
using step_hyps
by (smt (verit, ccfv_threshold) Γ_lower_set linorder_trm.dual_order.asym
linorder_trm.is_least_in_ffilter_iff linorder_trm.is_lower_set_iff
linorder_trm.neq_iff)
next
show "¬ trail_defined_lit Γ (Pos A)"
using A_undef unfolding trail_defined_lit_iff_trail_defined_atm literal.sel .
next
show "is_pos (Pos A)"
by simp
next
show "trail_false_cls Γ {#K ∈# C. K ≠ Pos A#}"
using ‹clause_could_propagate Γ C (Pos A)›
unfolding clause_could_propagate_def by argo
next
show "ord_res.is_strictly_maximal_lit (Pos A) C"
using step_hyps by argo
qed
thus "trail_false_cls Γ C ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ C ∨ ord_res_8_can_factorize N U⇩e⇩r ℱ Γ C"
by argo
next
fix D :: "'f gclause"
assume
D_in: "D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
D_neq: "D ≠ C" and
D_spec_disj: "trail_false_cls Γ D ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ D"
hence "
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ D"
using step_hyps D_in by metis
thus "C ≺⇩c D"
proof (elim disjE)
assume "ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D"
thus "C ≺⇩c D"
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_decide_neg.cases)
case hyps: (1 L A')
hence "A = A'"
using step_hyps
by (smt (verit, del_insts) Γ_lower_set linorder_trm.antisym_conv3
linorder_trm.dual_order.strict_implies_not_eq linorder_trm.dual_order.strict_trans
linorder_trm.is_least_in_ffilter_iff linorder_trm.not_in_lower_setI)
hence "A ≺⇩t atm_of L"
using hyps
unfolding linorder_trm.is_least_in_ffilter_iff
by argo
hence "Pos A ≺⇩l L"
by (cases L) simp_all
thus ?thesis
using C_max_lit ‹linorder_lit.is_maximal_in_mset D L›
by (metis linorder_lit.multp_if_maximal_less_that_maximal)
qed
next
assume "ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D"
thus "C ≺⇩c D"
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_skip_undefined_neg.cases)
case hyps: (1 L)
hence "atm_of L = A"
using step_hyps
by (smt (verit, best)
A_in A_least A_undef D_in Γ_lower_set atm_of_in_atms_of_clssI
linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
linorder_trm.not_in_lower_setI trail_defined_lit_iff_trail_defined_atm)
hence "Pos A ≺⇩l L"
using ‹is_neg L› by (cases L) simp_all
thus ?thesis
using C_max_lit ‹linorder_lit.is_maximal_in_mset D L›
by (metis linorder_lit.multp_if_maximal_less_that_maximal)
qed
next
assume "ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D"
thus "C ≺⇩c D"
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
case hyps: (1 L)
thus ?thesis
by (meson C_in D_neq linorder_cls.linorder_cases)
qed
next
assume "ord_res_8_can_produce N U⇩e⇩r ℱ Γ D"
then obtain L where
D_max_lit: "ord_res.is_maximal_lit L D" and
"¬ (∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ)" and
L_undef: "¬ trail_defined_lit Γ L" and
D_almost_false: "trail_false_cls Γ {#K ∈# D. K ≠ L#}" and
"is_pos L"
by (auto elim: ord_res_8_can_factorize.cases ord_res_8_can_produce.cases)
hence "atm_of L = A"
using step_hyps
by (smt (verit, ccfv_SIG) A_in A_least A_undef D_in Γ_lower_set
linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
linorder_trm.not_in_lower_setI atm_of_in_atms_of_clssI
trail_defined_lit_iff_trail_defined_atm)
hence "L = Pos A"
using ‹is_pos L› by (cases L) simp_all
hence "clause_could_propagate Γ D (Pos A)"
unfolding clause_could_propagate_def
using D_almost_false D_max_lit L_undef by metis
thus "C ≺⇩c D"
using D_in D_neq C_least by metis
next
assume "ord_res_8_can_factorize N U⇩e⇩r ℱ Γ D"
then obtain L where
D_max_lit: "ord_res.is_maximal_lit L D" and
"¬ (∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ)" and
L_undef: "¬ trail_defined_lit Γ L" and
D_almost_false: "trail_false_cls Γ {#K ∈# D. K ≠ L#}" and
"is_pos L"
by (auto elim: ord_res_8_can_factorize.cases ord_res_8_can_produce.cases)
hence "atm_of L = A"
using step_hyps
by (smt (verit, ccfv_SIG) A_in A_least A_undef D_in Γ_lower_set
linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
linorder_trm.not_in_lower_setI atm_of_in_atms_of_clssI
trail_defined_lit_iff_trail_defined_atm)
hence "L = Pos A"
using ‹is_pos L› by (cases L) simp_all
hence "clause_could_propagate Γ D (Pos A)"
unfolding clause_could_propagate_def
using D_almost_false D_max_lit L_undef by metis
thus "C ≺⇩c D"
using D_in D_neq C_least by metis
qed
qed
hence "𝒞 = Some C"
using match_hyps by metis
define 𝒞' where
"𝒞' = The_optional (linorder_cls.is_least_in_fset
(ffilter ((≺⇩c) C) (iefac ℱ |`| (N |∪| U⇩e⇩r))))"
have first_step7: "ord_res_7 N (U⇩e⇩r, ℱ, Γ, Some C) (U⇩e⇩r, ℱ, Γ', 𝒞')"
proof (rule ord_res_7.production)
show "¬ trail_false_cls Γ C"
using C_in step_hyps(2) by blast
next
show "ord_res.is_maximal_lit (Pos A) C"
using C_max_lit by force
next
show "¬ (∃Aa|∈|atms_of_clss (N |∪| U⇩e⇩r). Aa ≺⇩t atm_of (Pos A) ∧ Aa |∉| trail_atms Γ)"
by (metis A_least Γ_lower_set linorder_trm.dual_order.asym linorder_trm.neq_iff
linorder_trm.not_in_lower_setI literal.sel(1))
next
show "¬ trail_defined_lit Γ (Pos A)"
using Pos_A_undef .
next
show "is_pos (Pos A)"
by simp
next
show "trail_false_cls Γ {#K ∈# C. K ≠ Pos A#}"
using C_almost_false .
next
show "ord_res.is_strictly_maximal_lit (Pos A) C"
using step_hyps by argo
next
show "Γ' = (Pos A, Some C) # Γ"
using step_hyps by argo
next
show "𝒞' = The_optional (linorder_cls.is_least_in_fset (ffilter ((≺⇩c) C) (iefac ℱ |`| (N |∪| U⇩e⇩r))))"
using 𝒞'_def .
qed
moreover have "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞')"
using ‹𝒞 = Some C› first_step7 match_hyps(3) ord_res_7_preserves_invars by blast
ultimately obtain 𝒞'' where
following_steps7: "(ord_res_7 N)⇧*⇧* (U⇩e⇩r, ℱ, Γ', 𝒞') (U⇩e⇩r, ℱ, Γ', 𝒞'')" and
no_more_step7: "(∄𝒞'''. ord_res_7 N (U⇩e⇩r, ℱ, Γ', 𝒞'') (U⇩e⇩r, ℱ, Γ', 𝒞'''))"
using MAGIC6 by metis
show ?thesis
proof (intro exI conjI)
have "(ord_res_7 N)⇧+⇧+ (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ, Γ', 𝒞'')"
unfolding ‹𝒞 = Some C›
using first_step7 following_steps7 by simp
thus "(constant_context ord_res_7)⇧+⇧+ S7 (N, U⇩e⇩r, ℱ, Γ', 𝒞'')"
unfolding S7_def by (simp add: tranclp_constant_context)
show "ord_res_7_matches_ord_res_8 (N, U⇩e⇩r, ℱ, Γ', 𝒞'') S8'"
unfolding S8'_def ‹s8' = (U⇩e⇩r, ℱ, Γ')›
proof (intro ord_res_7_matches_ord_res_8.intros allI)
show "ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞'')"
using ‹ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞')› following_steps7
rtranclp_ord_res_7_preserves_ord_res_7_invars by blast
show "ord_res_8_invars N (U⇩e⇩r, ℱ, Γ')"
using invars_s8' step_hyps(1) by blast
fix C :: "'f gclause"
show "𝒞'' = Some C ⟷ is_least_nonskipped_clause N U⇩e⇩r ℱ Γ' C"
using MAGIC5 ‹ord_res_7_invars N (U⇩e⇩r, ℱ, Γ', 𝒞'')› no_more_step7 by metis
qed
qed
next
case step_hyps: (factorize A C ℱ')
have
A_in: "A |∈| atms_of_clss (N |∪| U⇩e⇩r)" and
A_undef: "A |∉| trail_atms Γ" and
A_least: "∀y|∈|atms_of_clss (N |∪| U⇩e⇩r). y ≠ A ⟶ (∀A⇩1|∈|trail_atms Γ. A⇩1 ≺⇩t y) ⟶ A ≺⇩t y"
using step_hyps(3) unfolding linorder_trm.is_least_in_fset_iff by auto
have
C_in: "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
C_can_prop:"clause_could_propagate Γ C (Pos A)" and
C_least: "∀D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r).
D ≠ C ⟶ clause_could_propagate Γ D (Pos A) ⟶ C ≺⇩c D"
using step_hyps unfolding atomize_conj linorder_cls.is_least_in_ffilter_iff by argo
hence
Pos_A_undef: "¬ trail_defined_lit Γ (Pos A)" and
C_max_lit: "linorder_lit.is_maximal_in_mset C (Pos A)" and
C_almost_false: "trail_false_cls Γ {#K ∈# C. K ≠ Pos A#}"
unfolding atomize_conj clause_could_propagate_def by argo
have C_not_false: "¬ trail_false_cls Γ C"
using C_in step_hyps by metis
have no_undef_atm_lt_A:
"¬ (∃Aa|∈|atms_of_clss (N |∪| U⇩e⇩r). Aa ≺⇩t A ∧ Aa |∉| trail_atms Γ)"
by (metis A_least Γ_lower_set linorder_trm.dual_order.asym linorder_trm.neq_iff
linorder_trm.not_in_lower_setI)
have "is_least_nonskipped_clause N U⇩e⇩r ℱ Γ C"
unfolding is_least_nonskipped_clause_def
unfolding linorder_cls.is_least_in_ffilter_iff
proof (intro conjI ballI impI)
show "C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using C_in .
next
have "ord_res_8_can_factorize N U⇩e⇩r ℱ Γ C"
proof (intro ord_res_8_can_factorize.intros)
show "¬ trail_false_cls Γ C"
using C_not_false .
next
show "ord_res.is_maximal_lit (Pos A) C"
using C_max_lit .
next
show "¬ (∃Aa|∈|atms_of_clss (N |∪| U⇩e⇩r). Aa ≺⇩t atm_of (Pos A) ∧ Aa |∉| trail_atms Γ)"
using no_undef_atm_lt_A by simp
next
show "¬ trail_defined_lit Γ (Pos A)"
using Pos_A_undef .
next
show "is_pos (Pos A)"
by simp
next
show "trail_false_cls Γ {#K ∈# C. K ≠ Pos A#}"
using C_almost_false .
next
show "¬ ord_res.is_strictly_maximal_lit (Pos A) C"
using step_hyps by argo
qed
thus "trail_false_cls Γ C ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ C ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ C"
using ‹ord_res_8_can_factorize N U⇩e⇩r ℱ Γ C› by argo
next
fix D :: "'f gclause"
assume
D_in: "D |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
D_neq: "D ≠ C" and
D_spec_disj: "trail_false_cls Γ D ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ D"
hence "
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ D ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ D"
using step_hyps D_in by metis
thus "C ≺⇩c D"
proof (elim disjE)
assume "ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ D"
thus "C ≺⇩c D"
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_decide_neg.cases)
case hyps: (1 L A')
hence "A = A'"
using step_hyps
by (smt (verit, del_insts) Γ_lower_set linorder_trm.antisym_conv3
linorder_trm.dual_order.strict_implies_not_eq linorder_trm.dual_order.strict_trans
linorder_trm.is_least_in_ffilter_iff linorder_trm.not_in_lower_setI)
hence "A ≺⇩t atm_of L"
using hyps
unfolding linorder_trm.is_least_in_ffilter_iff
by argo
hence "Pos A ≺⇩l L"
by (cases L) simp_all
thus ?thesis
using C_max_lit ‹linorder_lit.is_maximal_in_mset D L›
by (metis linorder_lit.multp_if_maximal_less_that_maximal)
qed
next
assume "ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ D"
thus "C ≺⇩c D"
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_skip_undefined_neg.cases)
case hyps: (1 L)
hence "atm_of L = A"
using step_hyps
by (smt (verit, best)
A_in A_least A_undef D_in Γ_lower_set atm_of_in_atms_of_clssI
linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
linorder_trm.not_in_lower_setI trail_defined_lit_iff_trail_defined_atm)
hence "Pos A ≺⇩l L"
using ‹is_neg L› by (cases L) simp_all
thus ?thesis
using C_max_lit ‹linorder_lit.is_maximal_in_mset D L›
by (metis linorder_lit.multp_if_maximal_less_that_maximal)
qed
next
assume "ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ D"
thus "C ≺⇩c D"
proof (cases N U⇩e⇩r ℱ Γ D rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
case hyps: (1 L)
thus ?thesis
by (meson C_in D_neq linorder_cls.linorder_cases)
qed
next
assume "ord_res_8_can_produce N U⇩e⇩r ℱ Γ D"
then obtain L where
D_max_lit: "ord_res.is_maximal_lit L D" and
"¬ (∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ)" and
L_undef: "¬ trail_defined_lit Γ L" and
D_almost_false: "trail_false_cls Γ {#K ∈# D. K ≠ L#}" and
"is_pos L"
by (auto elim: ord_res_8_can_factorize.cases ord_res_8_can_produce.cases)
hence "atm_of L = A"
using step_hyps
by (smt (verit, ccfv_SIG) A_in A_least A_undef D_in Γ_lower_set
linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
linorder_trm.not_in_lower_setI atm_of_in_atms_of_clssI
trail_defined_lit_iff_trail_defined_atm)
hence "L = Pos A"
using ‹is_pos L› by (cases L) simp_all
hence "clause_could_propagate Γ D (Pos A)"
unfolding clause_could_propagate_def
using D_almost_false D_max_lit L_undef by metis
thus "C ≺⇩c D"
using D_in D_neq C_least by metis
next
assume "ord_res_8_can_factorize N U⇩e⇩r ℱ Γ D"
then obtain L where
D_max_lit: "ord_res.is_maximal_lit L D" and
"¬ (∃A|∈|atms_of_clss (N |∪| U⇩e⇩r). A ≺⇩t atm_of L ∧ A |∉| trail_atms Γ)" and
L_undef: "¬ trail_defined_lit Γ L" and
D_almost_false: "trail_false_cls Γ {#K ∈# D. K ≠ L#}" and
"is_pos L"
by (auto elim: ord_res_8_can_factorize.cases ord_res_8_can_produce.cases)
hence "atm_of L = A"
using step_hyps
by (smt (verit, ccfv_SIG) A_in A_least A_undef D_in Γ_lower_set
linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
linorder_trm.not_in_lower_setI atm_of_in_atms_of_clssI
trail_defined_lit_iff_trail_defined_atm)
hence "L = Pos A"
using ‹is_pos L› by (cases L) simp_all
hence "clause_could_propagate Γ D (Pos A)"
unfolding clause_could_propagate_def
using D_almost_false D_max_lit L_undef by metis
thus "C ≺⇩c D"
using D_in D_neq C_least by metis
qed
qed
hence "𝒞 = Some C"
using match_hyps by metis
define 𝒞' where
"𝒞' = Some (efac C)"
have first_step7: "ord_res_7 N (U⇩e⇩r, ℱ, Γ, Some C) (U⇩e⇩r, ℱ', Γ, 𝒞')"
unfolding 𝒞'_def
proof (rule ord_res_7.factoring)
show "¬ trail_false_cls Γ C"
using C_not_false .
next
show "ord_res.is_maximal_lit (Pos A) C"
using C_max_lit .
next
show "¬ (∃Aa|∈|atms_of_clss (N |∪| U⇩e⇩r). Aa ≺⇩t atm_of (Pos A) ∧ Aa |∉| trail_atms Γ)"
using no_undef_atm_lt_A by simp
next
show "¬ trail_defined_lit Γ (Pos A)"
using Pos_A_undef .
next
show "is_pos (Pos A)"
by simp
next
show "trail_false_cls Γ {#K ∈# C. K ≠ Pos A#}"
using C_almost_false .
next
show "¬ ord_res.is_strictly_maximal_lit (Pos A) C"
using step_hyps by argo
next
show "ℱ' = finsert C ℱ"
using step_hyps by argo
qed
moreover have "ord_res_7_invars N (U⇩e⇩r, ℱ', Γ, 𝒞')"
using ‹𝒞 = Some C› first_step7 match_hyps(3) ord_res_7_preserves_invars by blast
ultimately obtain 𝒞'' where
following_steps7: "(ord_res_7 N)⇧*⇧* (U⇩e⇩r, ℱ', Γ, 𝒞') (U⇩e⇩r, ℱ', Γ, 𝒞'')" and
no_more_step7: "(∄𝒞'''. ord_res_7 N (U⇩e⇩r, ℱ', Γ, 𝒞'') (U⇩e⇩r, ℱ', Γ, 𝒞'''))"
using MAGIC6 by metis
show ?thesis
proof (intro exI conjI)
have "(ord_res_7 N)⇧+⇧+ (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r, ℱ', Γ, 𝒞'')"
unfolding ‹𝒞 = Some C›
using first_step7 following_steps7 by simp
thus "(constant_context ord_res_7)⇧+⇧+ S7 (N, U⇩e⇩r, ℱ', Γ, 𝒞'')"
unfolding S7_def by (simp add: tranclp_constant_context)
show "ord_res_7_matches_ord_res_8 (N, U⇩e⇩r, ℱ', Γ, 𝒞'') S8'"
unfolding S8'_def ‹s8' = (U⇩e⇩r, ℱ', Γ)›
proof (intro ord_res_7_matches_ord_res_8.intros allI)
show "ord_res_7_invars N (U⇩e⇩r, ℱ', Γ, 𝒞'')"
using ‹ord_res_7_invars N (U⇩e⇩r, ℱ', Γ, 𝒞')› following_steps7
rtranclp_ord_res_7_preserves_ord_res_7_invars by blast
show "ord_res_8_invars N (U⇩e⇩r, ℱ', Γ)"
using invars_s8' step_hyps(1) by blast
fix C :: "'f gclause"
show "𝒞'' = Some C ⟷ is_least_nonskipped_clause N U⇩e⇩r ℱ' Γ C"
using MAGIC5 ‹ord_res_7_invars N (U⇩e⇩r, ℱ', Γ, 𝒞'')› no_more_step7 by metis
qed
qed
next
case step_hyps: (resolution E A D U⇩e⇩r' Γ')
note E_max_lit = ‹ord_res.is_maximal_lit (Neg A) E›
have
E_in: "E |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
E_false: "trail_false_cls Γ E" and
E_least: "∀F |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). F ≠ E ⟶ trail_false_cls Γ F ⟶ E ≺⇩c F"
using step_hyps
unfolding atomize_conj
unfolding linorder_cls.is_least_in_ffilter_iff
by argo
have "is_least_nonskipped_clause N U⇩e⇩r ℱ Γ E"
unfolding is_least_nonskipped_clause_def
unfolding linorder_cls.is_least_in_ffilter_iff
proof (intro conjI ballI impI)
show "E |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)"
using E_in .
next
show "trail_false_cls Γ E ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ E ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ E ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ E ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ E ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ E"
using E_false by argo
next
fix F :: "'f gclause"
assume
F_in: "F |∈| iefac ℱ |`| (N |∪| U⇩e⇩r)" and
F_neq: "F ≠ E" and
D_spec_disj: "trail_false_cls Γ F ∨
ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ F ∨
ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ F ∨
ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ F ∨
ord_res_8_can_produce N U⇩e⇩r ℱ Γ F ∨
ord_res_8_can_factorize N U⇩e⇩r ℱ Γ F"
show "E ≺⇩c F"
using D_spec_disj
proof (elim disjE)
assume "trail_false_cls Γ F"
thus "E ≺⇩c F"
using E_least F_in F_neq by metis
next
assume "ord_res_8_can_decide_neg N U⇩e⇩r ℱ Γ F"
thus "E ≺⇩c F"
proof (cases N U⇩e⇩r ℱ Γ F rule: ord_res_8_can_decide_neg.cases)
case hyps: (1 L' A')
thus ?thesis
using no_undef_atom_le_max_lit_if_lt_false_clause[
OF Γ_lower_set E_in E_false E_max_lit F_in ‹ord_res.is_maximal_lit L' F›]
by (metis (no_types, lifting) F_neq linorder_cls.neq_iff
linorder_trm.is_least_in_ffilter_iff reflclp_iff)
qed
next
assume "ord_res_8_can_skip_undefined_neg N U⇩e⇩r ℱ Γ F"
thus "E ≺⇩c F"
proof (cases N U⇩e⇩r ℱ Γ F rule: ord_res_8_can_skip_undefined_neg.cases)
case (1 L')
thus ?thesis
using no_undef_atom_le_max_lit_if_lt_false_clause[
OF Γ_lower_set E_in E_false E_max_lit F_in ‹ord_res.is_maximal_lit L' F›]
by (metis F_in F_neq atm_of_in_atms_of_clssI linorder_cls.not_less_iff_gr_or_eq
linorder_lit.is_maximal_in_mset_iff reflclp_iff
trail_defined_lit_iff_trail_defined_atm)
qed
next
assume "ord_res_8_can_skip_undefined_pos_ultimate N U⇩e⇩r ℱ Γ F"
thus "E ≺⇩c F"
proof (cases N U⇩e⇩r ℱ Γ F rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
case (1 L')
thus ?thesis
using no_undef_atom_le_max_lit_if_lt_false_clause[
OF Γ_lower_set E_in E_false E_max_lit F_in ‹ord_res.is_maximal_lit L' F›]
by (metis F_in F_neq atm_of_in_atms_of_clssI linorder_cls.not_less_iff_gr_or_eq
linorder_lit.is_maximal_in_mset_iff reflclp_iff
trail_defined_lit_iff_trail_defined_atm)
qed
next
assume "ord_res_8_can_produce N U⇩e⇩r ℱ Γ F"
thus "E ≺⇩c F"
proof (cases N U⇩e⇩r ℱ Γ F rule: ord_res_8_can_produce.cases)
case (1 L')
thus ?thesis
using no_undef_atom_le_max_lit_if_lt_false_clause[
OF Γ_lower_set E_in E_false E_max_lit F_in ‹ord_res.is_maximal_lit L' F›]
by (metis F_in F_neq atm_of_in_atms_of_clssI linorder_cls.not_less_iff_gr_or_eq
linorder_lit.is_maximal_in_mset_iff reflclp_iff
trail_defined_lit_iff_trail_defined_atm)
qed
next
assume "ord_res_8_can_factorize N U⇩e⇩r ℱ Γ F"
thus "E ≺⇩c F"
proof (cases N U⇩e⇩r ℱ Γ F rule: ord_res_8_can_factorize.cases)
case (1 L')
thus ?thesis
using no_undef_atom_le_max_lit_if_lt_false_clause[
OF Γ_lower_set E_in E_false E_max_lit F_in ‹ord_res.is_maximal_lit L' F›]
by (metis F_in F_neq atm_of_in_atms_of_clssI linorder_cls.not_less_iff_gr_or_eq
linorder_lit.is_maximal_in_mset_iff reflclp_iff
trail_defined_lit_iff_trail_defined_atm)
qed
qed
qed
hence "𝒞 = Some E"
using match_hyps by metis
obtain 𝒞' where first_step7: "ord_res_7 N (U⇩e⇩r, ℱ, Γ, Some E) (U⇩e⇩r', ℱ, Γ', 𝒞')"
proof atomize_elim
have "(Pos A, Some D) ∈ set Γ"
using ‹map_of Γ (Pos A) = Some (Some D)› by (metis map_of_SomeD)
hence D_almost_false: "trail_false_cls Γ {#K ∈# D. K ≠ Pos A#}"
using ord_res_7_invars_implies_propagated_clause_almost_false
‹ord_res_7_invars N (U⇩e⇩r, ℱ, Γ, 𝒞)› by metis
have eres_false: "trail_false_cls Γ (eres D E)"
unfolding trail_false_cls_def
proof (intro ballI)
fix K :: "'f gliteral"
assume "K ∈# eres D E"
hence "K ∈# D ∧ K ≠ Pos A ∨ K ∈# E"
using strong_lit_in_one_of_resolvents_if_in_eres[OF E_max_lit] by simp
thus "trail_false_lit Γ K"
proof (elim disjE conjE)
assume "K ∈# D" and "K ≠ Pos A"
thus "trail_false_lit Γ K"
using D_almost_false unfolding trail_false_cls_def by simp
next
assume "K ∈# E"
thus "trail_false_lit Γ K"
using E_false unfolding trail_false_cls_def by simp
qed
qed
show "∃𝒞'. ord_res_7 N (U⇩e⇩r, ℱ, Γ, Some E) (U⇩e⇩r', ℱ, Γ', 𝒞')"
proof (cases "eres D E = {#}")
case True
hence "⋀Ln. ∀K. ord_res.is_maximal_lit K (eres D E) ⟶ atm_of K ≼⇩t atm_of (fst Ln)"
unfolding linorder_lit.is_maximal_in_mset_iff
by simp
hence "Γ' = dropWhile (λLn. True) Γ"
using step_hyps by meson
hence "Γ' = []"
by simp
show ?thesis
proof (intro exI)
show "ord_res_7 N (U⇩e⇩r, ℱ, Γ, Some E) (U⇩e⇩r', ℱ, Γ', Some {#})"
using ord_res_7.resolution_bot[OF E_false E_max_lit]
‹map_of Γ (Pos A) = Some (Some D)› ‹U⇩e⇩r' = finsert (eres D E) U⇩e⇩r› True ‹Γ' = []›
by simp
qed
next
case False
then obtain K where eres_max_lit: "ord_res.is_maximal_lit K (eres D E)"
using linorder_lit.ex_maximal_in_mset by presburger
hence "⋀Ln. (∀K. ord_res.is_maximal_lit K (eres D E) ⟶ atm_of K ≼⇩t atm_of (fst Ln)) ⟷
atm_of K ≼⇩t atm_of (fst Ln)"
by (metis linorder_lit.Uniq_is_maximal_in_mset the1_equality')
hence Γ'_eq: "Γ' = dropWhile (λLn. atm_of K ≼⇩t atm_of (fst Ln)) Γ"
using step_hyps by meson
show ?thesis
proof (cases K)
case (Pos A⇩K)
hence K_pos: "is_pos K"
by simp
then show ?thesis
using ord_res_7.resolution_pos[OF E_false E_max_lit _ _ _ False Γ'_eq
eres_max_lit K_pos]
using step_hyps by fastforce
next
case (Neg A⇩K)
hence K_neg: "is_neg K"
by simp
have "trail_false_lit Γ K"
using eres_false eres_max_lit
unfolding linorder_lit.is_maximal_in_mset_iff trail_false_cls_def by metis
hence "∃opt. (- K, opt) ∈ set Γ"
unfolding trail_false_lit_def by auto
moreover have "∀Ln∈set Γ. is_neg (fst Ln) = (snd Ln = None)"
using invars7 by argo
ultimately obtain C where "(- K, Some C) ∈ set Γ"
unfolding Neg uminus_Neg by fastforce
hence "map_of Γ (- K) = Some (Some C)"
proof (rule map_of_is_SomeI[rotated])
show "distinct (map fst Γ)"
using Γ_consistent
by (metis distinct_lits_if_trail_consistent)
qed
then show ?thesis
using ord_res_7.resolution_neg[OF E_false E_max_lit _ _ _ False Γ'_eq
eres_max_lit K_neg]
using step_hyps by fastforce
qed
qed
qed
moreover have "ord_res_7_invars N (U⇩e⇩r', ℱ, Γ', 𝒞')"
using ‹𝒞 = Some E› first_step7 match_hyps(3) ord_res_7_preserves_invars by blast
ultimately obtain 𝒞'' where
following_steps7: "(ord_res_7 N)⇧*⇧* (U⇩e⇩r', ℱ, Γ', 𝒞') (U⇩e⇩r', ℱ, Γ', 𝒞'')" and
no_more_step7: "(∄𝒞'''. ord_res_7 N (U⇩e⇩r', ℱ, Γ', 𝒞'') (U⇩e⇩r', ℱ, Γ', 𝒞'''))"
using MAGIC6 by metis
show ?thesis
proof (intro exI conjI)
have "(ord_res_7 N)⇧+⇧+ (U⇩e⇩r, ℱ, Γ, 𝒞) (U⇩e⇩r', ℱ, Γ', 𝒞'')"
unfolding ‹𝒞 = Some E›
using first_step7 following_steps7 by simp
thus "(constant_context ord_res_7)⇧+⇧+ S7 (N, U⇩e⇩r', ℱ, Γ', 𝒞'')"
unfolding S7_def by (simp add: tranclp_constant_context)
show "ord_res_7_matches_ord_res_8 (N, U⇩e⇩r', ℱ, Γ', 𝒞'') S8'"
unfolding S8'_def ‹s8' = (U⇩e⇩r', ℱ, Γ')›
proof (intro ord_res_7_matches_ord_res_8.intros allI)
show "ord_res_7_invars N (U⇩e⇩r', ℱ, Γ', 𝒞'')"
using ‹ord_res_7_invars N (U⇩e⇩r', ℱ, Γ', 𝒞')› following_steps7
rtranclp_ord_res_7_preserves_ord_res_7_invars by blast
show "ord_res_8_invars N (U⇩e⇩r', ℱ, Γ')"
using invars_s8' step_hyps(1) by blast
fix C :: "'f gclause"
show "𝒞'' = Some C ⟷ is_least_nonskipped_clause N U⇩e⇩r' ℱ Γ' C"
using MAGIC5 ‹ord_res_7_invars N (U⇩e⇩r', ℱ, Γ', 𝒞'')› no_more_step7 by metis
qed
qed
qed
qed
theorem bisimulation_ord_res_7_ord_res_8:
defines "match ≡ λ_. ord_res_7_matches_ord_res_8"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_7_state ⇒ 'f ord_res_8_state ⇒ bool) ℛ.
bisimulation
(constant_context ord_res_7) (constant_context ord_res_8)
ord_res_7_final ord_res_8_final
ℛ MATCH"
proof (rule ex_bisimulation_from_backward_simulation)
show "right_unique (constant_context ord_res_7)"
using right_unique_constant_context right_unique_ord_res_7 by metis
next
show "right_unique (constant_context ord_res_8)"
using right_unique_constant_context right_unique_ord_res_8 by metis
next
show "∀S. ord_res_7_final S ⟶ (∄S'. constant_context ord_res_7 S S')"
by (metis finished_def ord_res_7_semantics.final_finished)
next
show "∀S. ord_res_8_final S ⟶ (∄S'. constant_context ord_res_8 S S')"
by (metis finished_def ord_res_8_semantics.final_finished)
next
show "∀i S7 S8. match i S7 S8 ⟶ ord_res_7_final S7 ⟷ ord_res_8_final S8"
unfolding match_def
using ord_res_7_final_iff_ord_res_8_final by metis
next
show "∀i S7 S8. match i S7 S8 ⟶
safe_state (constant_context ord_res_7) ord_res_7_final S7 ∧
safe_state (constant_context ord_res_8) ord_res_8_final S8"
proof (intro allI impI conjI)
fix i S7 S8
assume match: "match i S7 S8"
show "safe_state (constant_context ord_res_7) ord_res_7_final S7"
using match[unfolded match_def]
using ord_res_7_safe_state_if_invars
using ord_res_7_matches_ord_res_8.simps by auto
show "safe_state (constant_context ord_res_8) ord_res_8_final S8"
using match[unfolded match_def]
using ord_res_8_safe_state_if_invars
using ord_res_7_matches_ord_res_8.simps by auto
qed
next
show "wfp (λ_ _. False)"
by simp
next
show "∀i S7 S8 S8'. match i S7 S8 ⟶ constant_context ord_res_8 S8 S8' ⟶
(∃i' S7'. (constant_context ord_res_7)⇧+⇧+ S7 S7' ∧ match i' S7' S8') ∨
(∃i'. match i' S7 S8' ∧ False)"
unfolding match_def
using backward_simulation_between_7_and_8 by metis
qed
end
section ‹ORD-RES-9 (factorize when propagating)›
type_synonym 'f ord_res_9_state =
"'f gclause fset ×'f gclause fset × 'f gclause fset × ('f gliteral × 'f gclause option) list"
context simulation_SCLFOL_ground_ordered_resolution begin
inductive ord_res_8_matches_ord_res_9 :: "'f ord_res_8_state ⇒ 'f ord_res_9_state ⇒ bool" where
"ord_res_8_invars N (U⇩e⇩r, ℱ, Γ) ⟹
ord_res_8_matches_ord_res_9 (N, U⇩e⇩r, ℱ, Γ) (N, U⇩e⇩r, ℱ, Γ)"
lemma ord_res_8_final_iff_ord_res_9_final:
fixes S8 S9
assumes match: "ord_res_8_matches_ord_res_9 S8 S9"
shows "ord_res_8_final S8 ⟷ ord_res_8_final S9"
using match
proof (cases S8 S9 rule: ord_res_8_matches_ord_res_9.cases)
case (1 N U⇩e⇩r ℱ Γ)
then show ?thesis
by argo
qed
lemma backward_simulation_between_8_and_9:
fixes S8 S9 S9'
assumes match: "ord_res_8_matches_ord_res_9 S8 S9" and step: "constant_context ord_res_9 S9 S9'"
shows "∃S8'. (constant_context ord_res_8)⇧+⇧+ S8 S8' ∧ ord_res_8_matches_ord_res_9 S8' S9'"
using match
proof (cases S8 S9 rule: ord_res_8_matches_ord_res_9.cases)
case match_hyps: (1 N U⇩e⇩r ℱ Γ)
note S8_def = ‹S8 = (N, U⇩e⇩r, ℱ, Γ)›
note S9_def = ‹S9 = (N, U⇩e⇩r, ℱ, Γ)›
note invars = ‹ord_res_8_invars N (U⇩e⇩r, ℱ, Γ)›
obtain s9' where S9'_def: "S9' = (N, s9')" and step': "ord_res_9 N (U⇩e⇩r, ℱ, Γ) s9'"
using step unfolding S9_def
using constant_context.cases by blast
have "ord_res_8 N (U⇩e⇩r, ℱ, Γ) s9' ∨ (ord_res_8 N OO ord_res_8 N) (U⇩e⇩r, ℱ, Γ) s9'"
using step' ord_res_9_is_one_or_two_ord_res_9_steps by metis
hence steps8: "(ord_res_8 N)⇧+⇧+ (U⇩e⇩r, ℱ, Γ) s9'"
by auto
show ?thesis
proof (intro exI conjI)
show "(constant_context ord_res_8)⇧+⇧+ S8 (N, s9')"
unfolding S8_def
using steps8 by (simp add: tranclp_constant_context)
next
have "ord_res_8_invars N s9'"
using invars steps8 tranclp_ord_res_8_preserves_invars by metis
thus "ord_res_8_matches_ord_res_9 (N, s9') S9'"
unfolding S9'_def
by (metis ord_res_8_matches_ord_res_9.intros prod_cases3)
qed
qed
theorem bisimulation_ord_res_8_ord_res_9:
defines "match ≡ λ_. ord_res_8_matches_ord_res_9"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_8_state ⇒ 'f ord_res_9_state ⇒ bool) ℛ.
bisimulation
(constant_context ord_res_8) (constant_context ord_res_9)
ord_res_8_final ord_res_8_final
ℛ MATCH"
proof (rule ex_bisimulation_from_backward_simulation)
show "right_unique (constant_context ord_res_8)"
using right_unique_constant_context right_unique_ord_res_8 by metis
next
show "right_unique (constant_context ord_res_9)"
using right_unique_constant_context right_unique_ord_res_9 by metis
next
show "∀S. ord_res_8_final S ⟶ (∄S'. constant_context ord_res_8 S S')"
by (metis finished_def ord_res_8_semantics.final_finished)
next
show "∀S. ord_res_8_final S ⟶ (∄S'. constant_context ord_res_9 S S')"
by (metis finished_def ord_res_9_semantics.final_finished)
next
show "∀i S8 S9. match i S8 S9 ⟶ ord_res_8_final S8 ⟷ ord_res_8_final S9"
unfolding match_def
using ord_res_8_final_iff_ord_res_9_final by metis
next
show "∀i S8 S9. match i S8 S9 ⟶
safe_state (constant_context ord_res_8) ord_res_8_final S8 ∧
safe_state (constant_context ord_res_9) ord_res_8_final S9"
proof (intro allI impI conjI)
fix i S8 S9
assume match: "match i S8 S9"
show "safe_state (constant_context ord_res_8) ord_res_8_final S8"
using match[unfolded match_def]
using ord_res_8_safe_state_if_invars
using ord_res_8_matches_ord_res_9.simps by auto
show "safe_state (constant_context ord_res_9) ord_res_8_final S9"
using match[unfolded match_def]
using ord_res_9_safe_state_if_invars
using ord_res_8_matches_ord_res_9.simps by auto
qed
next
show "wfp (λ_ _. False)"
by simp
next
show "∀i S8 S9 S9'. match i S8 S9 ⟶ constant_context ord_res_9 S9 S9' ⟶
(∃i' S8'. (constant_context ord_res_8)⇧+⇧+ S8 S8' ∧ match i' S8' S9') ∨
(∃i'. match i' S8 S9' ∧ False)"
unfolding match_def
using backward_simulation_between_8_and_9 by metis
qed
end
section ‹ORD-RES-10 (propagate iff a conflict is produced)›
context simulation_SCLFOL_ground_ordered_resolution begin
inductive ord_res_9_matches_ord_res_10 :: "'f ord_res_9_state ⇒ 'f ord_res_10_state ⇒ bool" where
"ord_res_8_invars N (U⇩e⇩r, ℱ, Γ⇩9) ⟹
ord_res_10_invars N (U⇩e⇩r, ℱ, Γ⇩1⇩0) ⟹
list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0 ⟹
list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9 Γ⇩1⇩0 ⟹
ord_res_9_matches_ord_res_10 (N, U⇩e⇩r, ℱ, Γ⇩9) (N, U⇩e⇩r, ℱ, Γ⇩1⇩0)"
lemma ord_res_9_final_iff_ord_res_10_final:
fixes S9 S10
assumes match: "ord_res_9_matches_ord_res_10 S9 S10"
shows "ord_res_8_final S9 ⟷ ord_res_8_final S10"
using match
proof (cases S9 S10 rule: ord_res_9_matches_ord_res_10.cases)
case match_hyps: (1 N U⇩e⇩r ℱ Γ⇩9 Γ⇩1⇩0)
then show ?thesis
using trail_atms_eq_trail_atms_if_same_lits[OF ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0›]
using trail_false_cls_eq_trail_false_cls_if_same_lits[OF ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0›]
unfolding ord_res_8_final.simps
by simp
qed
lemma backward_simulation_between_9_and_10:
fixes S9 S10 S10'
assumes
match: "ord_res_9_matches_ord_res_10 S9 S10" and
step: "constant_context ord_res_10 S10 S10'"
shows "∃S9'. (constant_context ord_res_9)⇧+⇧+ S9 S9' ∧ ord_res_9_matches_ord_res_10 S9' S10'"
using match
proof (cases S9 S10 rule: ord_res_9_matches_ord_res_10.cases)
case match_hyps: (1 N U⇩e⇩r ℱ Γ⇩9 Γ⇩1⇩0)
note S9_def = ‹S9 = (N, U⇩e⇩r, ℱ, Γ⇩9)›
note S10_def = ‹S10 = (N, U⇩e⇩r, ℱ, Γ⇩1⇩0)›
note invars9 = ‹ord_res_8_invars N (U⇩e⇩r, ℱ, Γ⇩9)›
note invars10 = ‹ord_res_10_invars N (U⇩e⇩r, ℱ, Γ⇩1⇩0)›
have "trail_atms Γ⇩9 = trail_atms Γ⇩1⇩0"
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0› trail_atms_eq_trail_atms_if_same_lits
by metis
have "trail_false_lit Γ⇩9 = trail_false_lit Γ⇩1⇩0"
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0› trail_false_lit_eq_trail_false_lit_if_same_lits
by metis
have "trail_false_cls Γ⇩9 = trail_false_cls Γ⇩1⇩0"
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0› trail_false_cls_eq_trail_false_cls_if_same_lits
by metis
have "trail_defined_lit Γ⇩9 = trail_defined_lit Γ⇩1⇩0"
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0›
trail_defined_lit_eq_trail_defined_lit_if_same_lits by metis
have "trail_defined_cls Γ⇩9 = trail_defined_cls Γ⇩1⇩0"
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0›
trail_defined_cls_eq_trail_defined_cls_if_same_lits by metis
have "clause_could_propagate Γ⇩9 = clause_could_propagate Γ⇩1⇩0"
unfolding clause_could_propagate_def
unfolding ‹trail_defined_lit Γ⇩9 = trail_defined_lit Γ⇩1⇩0›
unfolding ‹trail_false_cls Γ⇩9 = trail_false_cls Γ⇩1⇩0› ..
have Γ⇩9_sorted: "sorted_wrt (λx y. atm_of (fst y) ≺⇩t atm_of (fst x)) Γ⇩9"
using invars9[unfolded ord_res_8_invars_def trail_is_sorted_def, simplified] by argo
obtain s10' where "S10' = (N, s10')" and step10: "ord_res_10 N (U⇩e⇩r, ℱ, Γ⇩1⇩0) s10'"
using step unfolding S10_def by (auto elim: constant_context.cases)
show ?thesis
using step10
proof (cases N "(U⇩e⇩r, ℱ, Γ⇩1⇩0)" s10' rule: ord_res_10.cases)
case step_hyps: (decide_neg A Γ⇩1⇩0')
define Γ⇩9' where
"Γ⇩9' = (Neg A, None) # Γ⇩9"
show ?thesis
proof (intro exI conjI)
have step9: "ord_res_9 N (U⇩e⇩r, ℱ, Γ⇩9) (U⇩e⇩r, ℱ, Γ⇩9')"
proof (rule ord_res_9.decide_neg)
show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r). trail_false_cls Γ⇩9 C)"
using step_hyps ‹trail_false_cls Γ⇩9 = trail_false_cls Γ⇩1⇩0› by argo
next
show "linorder_trm.is_least_in_fset
{|A⇩2 |∈| atms_of_clss (N |∪| U⇩e⇩r). ∀A⇩1 |∈| trail_atms Γ⇩9. A⇩1 ≺⇩t A⇩2|} A"
using step_hyps ‹trail_atms Γ⇩9 = trail_atms Γ⇩1⇩0› by metis
next
show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r). clause_could_propagate Γ⇩9 C (Pos A))"
using step_hyps ‹clause_could_propagate Γ⇩9 = clause_could_propagate Γ⇩1⇩0› by metis
next
show "Γ⇩9' = (Neg A, None) # Γ⇩9"
using Γ⇩9'_def .
qed
thus "(constant_context ord_res_9)⇧+⇧+ S9 (N, U⇩e⇩r, ℱ, Γ⇩9')"
unfolding S9_def by (auto intro: constant_context.intros)
show "ord_res_9_matches_ord_res_10 (N, U⇩e⇩r, ℱ, Γ⇩9') S10'"
unfolding ‹S10' = (N, s10')› ‹s10' = (U⇩e⇩r, ℱ, Γ⇩1⇩0')›
proof (rule ord_res_9_matches_ord_res_10.intros)
show "ord_res_8_invars N (U⇩e⇩r, ℱ, Γ⇩9')"
using invars9 step9 ord_res_9_preserves_invars by metis
next
show "ord_res_10_invars N (U⇩e⇩r, ℱ, Γ⇩1⇩0')"
using invars10 step10 ord_res_10_preserves_invars ‹s10' = (U⇩e⇩r, ℱ, Γ⇩1⇩0')› by metis
next
show "list_all2 (λx y. fst x = fst y) Γ⇩9' Γ⇩1⇩0'"
unfolding ‹Γ⇩9' = (Neg A, None) # Γ⇩9› ‹Γ⇩1⇩0' = (Neg A, None) # Γ⇩1⇩0›
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0› by simp
next
show "list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9' Γ⇩1⇩0'"
unfolding ‹Γ⇩9' = (Neg A, None) # Γ⇩9› ‹Γ⇩1⇩0' = (Neg A, None) # Γ⇩1⇩0›
using ‹list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9 Γ⇩1⇩0›
by simp
qed
qed
next
case step_hyps: (decide_pos A C Γ⇩1⇩0' ℱ')
define Γ⇩9' where
"Γ⇩9' = (Pos A, Some (efac C)) # Γ⇩9"
show ?thesis
proof (intro exI conjI)
have step9: "ord_res_9 N (U⇩e⇩r, ℱ, Γ⇩9) (U⇩e⇩r, ℱ', Γ⇩9')"
proof (rule ord_res_9.propagate)
show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r). trail_false_cls Γ⇩9 C)"
using step_hyps ‹trail_false_cls Γ⇩9 = trail_false_cls Γ⇩1⇩0› by argo
next
show "linorder_trm.is_least_in_fset
{|A⇩2 |∈| atms_of_clss (N |∪| U⇩e⇩r). ∀A⇩1|∈|trail_atms Γ⇩9. A⇩1 ≺⇩t A⇩2|} A"
using step_hyps ‹trail_atms Γ⇩9 = trail_atms Γ⇩1⇩0› by metis
next
show "linorder_cls.is_least_in_fset
{|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). clause_could_propagate Γ⇩9 C (Pos A)|} C"
using step_hyps ‹clause_could_propagate Γ⇩9 = clause_could_propagate Γ⇩1⇩0› by metis
next
show "Γ⇩9' = (Pos A, Some (efac C)) # Γ⇩9"
using Γ⇩9'_def .
next
show "ℱ' = (if ord_res.is_strictly_maximal_lit (Pos A) C then ℱ else finsert C ℱ)"
using step_hyps by argo
qed
thus "(constant_context ord_res_9)⇧+⇧+ S9 (N, U⇩e⇩r, ℱ', Γ⇩9')"
unfolding S9_def by (auto intro: constant_context.intros)
show "ord_res_9_matches_ord_res_10 (N, U⇩e⇩r, ℱ', Γ⇩9') S10'"
unfolding ‹S10' = (N, s10')› ‹s10' = (U⇩e⇩r, ℱ', Γ⇩1⇩0')›
proof (rule ord_res_9_matches_ord_res_10.intros)
show "ord_res_8_invars N (U⇩e⇩r, ℱ', Γ⇩9')"
using invars9 step9 ord_res_9_preserves_invars by metis
next
show "ord_res_10_invars N (U⇩e⇩r, ℱ', Γ⇩1⇩0')"
using invars10 step10 ord_res_10_preserves_invars ‹s10' = (U⇩e⇩r, ℱ', Γ⇩1⇩0')› by metis
next
show "list_all2 (λx y. fst x = fst y) Γ⇩9' Γ⇩1⇩0'"
unfolding ‹Γ⇩9' = (Pos A, Some (efac C)) # Γ⇩9› ‹Γ⇩1⇩0' = (Pos A, None) # Γ⇩1⇩0›
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0› by simp
next
show "list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9' Γ⇩1⇩0'"
unfolding ‹Γ⇩9' = (Pos A, Some (efac C)) # Γ⇩9› ‹Γ⇩1⇩0' = (Pos A, None) # Γ⇩1⇩0›
using ‹list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9 Γ⇩1⇩0›
by simp
qed
qed
next
case step_hyps: (propagate A C Γ⇩1⇩0' ℱ')
define Γ⇩9' where
"Γ⇩9' = (Pos A, Some (efac C)) # Γ⇩9"
show ?thesis
proof (intro exI conjI)
have step9: "ord_res_9 N (U⇩e⇩r, ℱ, Γ⇩9) (U⇩e⇩r, ℱ', Γ⇩9')"
proof (rule ord_res_9.propagate)
show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r). trail_false_cls Γ⇩9 C)"
using step_hyps ‹trail_false_cls Γ⇩9 = trail_false_cls Γ⇩1⇩0› by argo
next
show "linorder_trm.is_least_in_fset
{|A⇩2 |∈| atms_of_clss (N |∪| U⇩e⇩r). ∀A⇩1|∈|trail_atms Γ⇩9. A⇩1 ≺⇩t A⇩2|} A"
using step_hyps ‹trail_atms Γ⇩9 = trail_atms Γ⇩1⇩0› by metis
next
show "linorder_cls.is_least_in_fset
{|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r). clause_could_propagate Γ⇩9 C (Pos A)|} C"
using step_hyps ‹clause_could_propagate Γ⇩9 = clause_could_propagate Γ⇩1⇩0› by metis
next
show "Γ⇩9' = (Pos A, Some (efac C)) # Γ⇩9"
using Γ⇩9'_def .
next
show "ℱ' = (if ord_res.is_strictly_maximal_lit (Pos A) C then ℱ else finsert C ℱ)"
using step_hyps by argo
qed
thus "(constant_context ord_res_9)⇧+⇧+ S9 (N, U⇩e⇩r, ℱ', Γ⇩9')"
unfolding S9_def by (auto intro: constant_context.intros)
show "ord_res_9_matches_ord_res_10 (N, U⇩e⇩r, ℱ', Γ⇩9') S10'"
unfolding ‹S10' = (N, s10')› ‹s10' = (U⇩e⇩r, ℱ', Γ⇩1⇩0')›
proof (rule ord_res_9_matches_ord_res_10.intros)
show "ord_res_8_invars N (U⇩e⇩r, ℱ', Γ⇩9')"
using invars9 step9 ord_res_9_preserves_invars by metis
next
show "ord_res_10_invars N (U⇩e⇩r, ℱ', Γ⇩1⇩0')"
using invars10 step10 ord_res_10_preserves_invars ‹s10' = (U⇩e⇩r, ℱ', Γ⇩1⇩0')› by metis
next
show "list_all2 (λx y. fst x = fst y) Γ⇩9' Γ⇩1⇩0'"
unfolding ‹Γ⇩9' = (Pos A, Some (efac C)) # Γ⇩9› ‹Γ⇩1⇩0' = (Pos A, Some (efac C)) # Γ⇩1⇩0›
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0› by simp
next
show "list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9' Γ⇩1⇩0'"
unfolding ‹Γ⇩9' = (Pos A, Some (efac C)) # Γ⇩9› ‹Γ⇩1⇩0' = (Pos A, Some (efac C)) # Γ⇩1⇩0›
using ‹list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9 Γ⇩1⇩0›
by simp
qed
qed
next
case step_hyps: (resolution D A C U⇩e⇩r' Γ⇩1⇩0')
have "∀Ln Γ'. Γ⇩1⇩0 = Ln # Γ' ⟶
(snd Ln ≠ None) = fBex (iefac ℱ |`| (N |∪| U⇩e⇩r)) (trail_false_cls Γ⇩1⇩0) ∧
(∀x∈set Γ'. snd x = None)"
using invars10 by (simp add: ord_res_10_invars.simps)
then obtain Γ⇩1⇩0'' where "Γ⇩1⇩0 = (Pos A, Some C) # Γ⇩1⇩0''"
using ‹map_of Γ⇩1⇩0 (Pos A) = Some (Some C)›
by (metis list.set_cases map_of_SomeD not_Some_eq snd_conv)
then obtain Γ⇩9'' where "Γ⇩9 = (Pos A, Some C) # Γ⇩9''"
using ‹list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9 Γ⇩1⇩0›
by (smt (verit, best) list_all2_Cons2 option.discI snd_conv)
define Γ⇩9' where
"Γ⇩9' = dropWhile (λLn. ∀K. ord_res.is_maximal_lit K (eres C D) ⟶
atm_of K ≼⇩t atm_of (fst Ln)) Γ⇩9"
show ?thesis
proof (intro exI conjI)
have step9: "ord_res_9 N (U⇩e⇩r, ℱ, Γ⇩9) (U⇩e⇩r', ℱ, Γ⇩9')"
proof (rule ord_res_9.resolution)
show "linorder_cls.is_least_in_fset (ffilter (trail_false_cls Γ⇩9) (iefac ℱ |`| (N |∪| U⇩e⇩r))) D"
using step_hyps ‹trail_false_cls Γ⇩9 = trail_false_cls Γ⇩1⇩0› by argo
next
show "ord_res.is_maximal_lit (Neg A) D"
using step_hyps by argo
next
show "map_of Γ⇩9 (Pos A) = Some (Some C)"
unfolding ‹Γ⇩9 = (Pos A, Some C) # Γ⇩9''› by simp
next
show "U⇩e⇩r' = finsert (eres C D) U⇩e⇩r"
using step_hyps by argo
next
show "Γ⇩9' = dropWhile (λLn. ∀K. ord_res.is_maximal_lit K (eres C D) ⟶
atm_of K ≼⇩t atm_of (fst Ln)) Γ⇩9"
using Γ⇩9'_def .
qed
thus "(constant_context ord_res_9)⇧+⇧+ S9 (N, U⇩e⇩r', ℱ, Γ⇩9')"
unfolding S9_def by (auto intro: constant_context.intros)
show "ord_res_9_matches_ord_res_10 (N, U⇩e⇩r', ℱ, Γ⇩9') S10'"
unfolding ‹S10' = (N, s10')› ‹s10' = (U⇩e⇩r', ℱ, Γ⇩1⇩0')›
proof (rule ord_res_9_matches_ord_res_10.intros)
show "ord_res_8_invars N (U⇩e⇩r', ℱ, Γ⇩9')"
using invars9 step9 ord_res_9_preserves_invars by metis
next
show "ord_res_10_invars N (U⇩e⇩r', ℱ, Γ⇩1⇩0')"
using invars10 step10 ord_res_10_preserves_invars ‹s10' = (U⇩e⇩r', ℱ, Γ⇩1⇩0')› by metis
next
define P :: "'f gterm literal × 'f gterm literal multiset option ⇒ bool" where
"P ≡ λLn. ∀K. ord_res.is_maximal_lit K (eres C D) ⟶ atm_of K ≼⇩t atm_of (fst Ln)"
have "length (takeWhile P Γ⇩9) = length (takeWhile P Γ⇩1⇩0)"
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0›
proof (induction Γ⇩9 Γ⇩1⇩0 rule: list.rel_induct)
case Nil
show ?case
by simp
next
case (Cons x xs y ys)
then show ?case
by (simp add: P_def)
qed
moreover have "Γ⇩9 = takeWhile P Γ⇩9 @ Γ⇩9'"
unfolding takeWhile_dropWhile_id
unfolding P_def ‹Γ⇩9' = dropWhile _ Γ⇩9› by simp
moreover have "Γ⇩1⇩0 = takeWhile P Γ⇩1⇩0 @ Γ⇩1⇩0'"
unfolding takeWhile_dropWhile_id
unfolding P_def ‹Γ⇩1⇩0' = dropWhile _ Γ⇩1⇩0› by simp
ultimately have "⋀Q. list_all2 Q Γ⇩9 Γ⇩1⇩0 ⟷
(list_all2 Q (takeWhile P Γ⇩9) (takeWhile P Γ⇩1⇩0) ∧ list_all2 Q Γ⇩9' Γ⇩1⇩0')"
using list_all2_append by metis
thus
"list_all2 (λx y. fst x = fst y) Γ⇩9' Γ⇩1⇩0'"
"list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9' Γ⇩1⇩0'"
unfolding atomize_conj
using ‹list_all2 (λx y. fst x = fst y) Γ⇩9 Γ⇩1⇩0›
using ‹list_all2 (λx y. snd y ≠ None ⟶ x = y) Γ⇩9 Γ⇩1⇩0›
by (simp only:)
qed
qed
qed
qed
theorem bisimulation_ord_res_9_ord_res_10:
defines "match ≡ λ_. ord_res_9_matches_ord_res_10"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_8_state ⇒ 'f ord_res_9_state ⇒ bool) ℛ.
bisimulation
(constant_context ord_res_9) (constant_context ord_res_10)
ord_res_8_final ord_res_8_final
ℛ MATCH"
proof (rule ex_bisimulation_from_backward_simulation)
show "right_unique (constant_context ord_res_9)"
using right_unique_constant_context right_unique_ord_res_9 by metis
next
show "right_unique (constant_context ord_res_10)"
using right_unique_constant_context right_unique_ord_res_10 by metis
next
show "∀S. ord_res_8_final S ⟶ (∄S'. constant_context ord_res_9 S S')"
by (metis finished_def ord_res_9_semantics.final_finished)
next
show "∀S. ord_res_8_final S ⟶ (∄S'. constant_context ord_res_10 S S')"
by (metis finished_def ord_res_10_semantics.final_finished)
next
show "∀i S9 S10. match i S9 S10 ⟶ ord_res_8_final S9 ⟷ ord_res_8_final S10"
unfolding match_def
using ord_res_9_final_iff_ord_res_10_final by metis
next
show "∀i S9 S10. match i S9 S10 ⟶
safe_state (constant_context ord_res_9) ord_res_8_final S9 ∧
safe_state (constant_context ord_res_10) ord_res_8_final S10"
proof (intro allI impI conjI)
fix i S9 S10
assume match: "match i S9 S10"
show "safe_state (constant_context ord_res_9) ord_res_8_final S9"
using match[unfolded match_def]
using ord_res_9_safe_state_if_invars
using ord_res_9_matches_ord_res_10.simps by auto
show "safe_state (constant_context ord_res_10) ord_res_8_final S10"
using match[unfolded match_def]
using ord_res_10_safe_state_if_invars
using ord_res_9_matches_ord_res_10.simps by auto
qed
next
show "wfp (λ_ _. False)"
by simp
next
show "∀i S9 S10 S10'. match i S9 S10 ⟶ constant_context ord_res_10 S10 S10' ⟶
(∃i' S9'. (constant_context ord_res_9)⇧+⇧+ S9 S9' ∧ match i' S9' S10') ∨
(∃i'. match i' S9 S10' ∧ False)"
unfolding match_def
using backward_simulation_between_9_and_10 by metis
qed
end
section ‹ORD-RES-11 (SCL strategy)›
context simulation_SCLFOL_ground_ordered_resolution begin
inductive ord_res_10_matches_ord_res_11 :: "'f ord_res_10_state ⇒ 'f ord_res_11_state ⇒ bool" where
"ord_res_10_invars N (U⇩e⇩r⇩1⇩0, ℱ, Γ) ⟹
ord_res_11_invars N (U⇩e⇩r⇩1⇩1, ℱ, Γ, 𝒞) ⟹
U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0 - {|{#}|} ⟹
if {#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0) then Γ = [] ∧ 𝒞 = Some {#} else 𝒞 = None ⟹
ord_res_10_matches_ord_res_11 (N, U⇩e⇩r⇩1⇩0, ℱ, Γ) (N, U⇩e⇩r⇩1⇩1, ℱ, Γ, 𝒞)"
lemma ord_res_10_final_iff_ord_res_11_final:
fixes S10 S11
assumes match: "ord_res_10_matches_ord_res_11 S10 S11"
shows "ord_res_8_final S10 ⟷ ord_res_11_final S11"
using match
proof (cases S10 S11 rule: ord_res_10_matches_ord_res_11.cases)
case match_hyps: (1 N U⇩e⇩r⇩1⇩0 ℱ Γ U⇩e⇩r⇩1⇩1 𝒞)
show ?thesis
proof (rule iffI)
assume "ord_res_8_final S10"
thus "ord_res_11_final S11"
unfolding ‹S10 = _›
proof (cases "(N, U⇩e⇩r⇩1⇩0, ℱ, Γ)" rule: ord_res_8_final.cases)
case model_found
hence "{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)"
using trail_false_cls_mempty by blast
hence "𝒞 = None"
using match_hyps by argo
moreover have "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0"
using match_hyps
by (metis ‹{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)› fimage_eqI finsert_fminus1 finsert_iff
fminus_finsert_absorb funionI2 iefac_mempty)
ultimately show ?thesis
unfolding ‹S11 = _›
using model_found
using ord_res_11_final.model_found
by metis
next
case contradiction_found
hence "Γ = [] ∧ 𝒞 = Some {#}"
using match_hyps by argo
thus ?thesis
unfolding ‹S11 = _›
using ord_res_11_final.contradiction_found by metis
qed
next
assume "ord_res_11_final S11"
thus "ord_res_8_final S10"
unfolding ‹S11 = _›
proof (cases "(N, U⇩e⇩r⇩1⇩1, ℱ, Γ, 𝒞)" rule: ord_res_11_final.cases)
case model_found
show ?thesis
unfolding ‹S10 = _›
proof (rule ord_res_8_final.model_found)
show "¬ (∃A |∈| atms_of_clss (N |∪| U⇩e⇩r⇩1⇩0). A |∉| trail_atms Γ)"
by (metis (no_types, lifting) fimage_iff fminus_finsert_absorb fminus_idemp funionCI
iefac_mempty local.model_found(1,2) match_hyps(5,6) option.simps(3))
next
show "¬ (∃C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0). trail_false_cls Γ C)"
by (metis finsertCI finsert_fminus1 fminus_finsert_absorb funionI2 iefac_mempty
local.model_found(1,3) match_hyps(5,6) option.simps(3) rev_fimage_eqI)
qed
next
case contradiction_found
show ?thesis
unfolding ‹S10 = _›
proof (rule ord_res_8_final.contradiction_found)
show "{#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)"
using match_hyps contradiction_found
by auto
qed
qed
qed
qed
lemma forward_simulation_between_10_and_11:
fixes S10 S11 S10'
assumes
match: "ord_res_10_matches_ord_res_11 S10 S11" and
step: "constant_context ord_res_10 S10 S10'"
shows "∃S11'. (constant_context ord_res_11)⇧+⇧+ S11 S11' ∧ ord_res_10_matches_ord_res_11 S10' S11'"
using match
proof (cases S10 S11 rule: ord_res_10_matches_ord_res_11.cases)
case match_hyps: (1 N U⇩e⇩r⇩1⇩0 ℱ Γ U⇩e⇩r⇩1⇩1 𝒞)
note S10_def = ‹S10 = (N, U⇩e⇩r⇩1⇩0, ℱ, Γ)›
note S11_def = ‹S11 = (N, U⇩e⇩r⇩1⇩1, ℱ, Γ, 𝒞)›
note invars10 = ‹ord_res_10_invars N (U⇩e⇩r⇩1⇩0, ℱ, Γ)›
note invars11 = ‹ord_res_11_invars N (U⇩e⇩r⇩1⇩1, ℱ, Γ, 𝒞)›
have mempty_not_in_if_no_false_cls: "{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)"
if "¬ fBex (iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)) (trail_false_cls Γ)"
using that by force
have 𝒞_eq_None_if_no_false_cls: "𝒞 = None"
if "¬ fBex (iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)) (trail_false_cls Γ)"
using match_hyps mempty_not_in_if_no_false_cls[OF that] by argo
have mempty_not_in_if: "{#} |∉| N |∪| U⇩e⇩r⇩1⇩0"
if "¬ fBex (iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)) (trail_false_cls Γ)"
using that
by (metis (no_types, opaque_lifting) fimageI iefac_mempty trail_false_cls_mempty)
have U⇩e⇩r⇩1⇩1_eq_U⇩e⇩r⇩1⇩0_if: "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0"
if "¬ fBex (iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)) (trail_false_cls Γ)"
using mempty_not_in_if[OF that] ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0 |-| {|{#}|}›
by (metis (no_types, opaque_lifting) finsertI1 finsert_ident fminusD2 funionCI
funion_fempty_right funion_finsert_right funion_fminus_cancel2)
obtain s10' where "S10' = (N, s10')" and step10: "ord_res_10 N (U⇩e⇩r⇩1⇩0, ℱ, Γ) s10'"
using step unfolding S10_def by (auto elim: constant_context.cases)
show ?thesis
using step10
proof (cases N "(U⇩e⇩r⇩1⇩0, ℱ, Γ)" s10' rule: ord_res_10.cases)
case step_hyps: (decide_neg A Γ')
have "𝒞 = None"
using step_hyps 𝒞_eq_None_if_no_false_cls by argo
have "{#} |∉| N |∪| U⇩e⇩r⇩1⇩0"
using step_hyps mempty_not_in_if by argo
have "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0"
using step_hyps U⇩e⇩r⇩1⇩1_eq_U⇩e⇩r⇩1⇩0_if by argo
show ?thesis
proof (intro exI conjI)
have step11: "ord_res_11 N (U⇩e⇩r⇩1⇩1, ℱ, Γ, None) (U⇩e⇩r⇩1⇩1, ℱ, Γ', None)"
proof (rule ord_res_11.decide_neg)
show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩1). trail_false_cls Γ C)"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "linorder_trm.is_least_in_fset
{|A⇩2 |∈| atms_of_clss (N |∪| U⇩e⇩r⇩1⇩1). ∀A⇩1|∈|trail_atms Γ. A⇩1 ≺⇩t A⇩2|} A"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩1). clause_could_propagate Γ C (Pos A))"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "Γ' = (Neg A, None) # Γ"
using step_hyps by argo
qed
thus "(constant_context ord_res_11)⇧+⇧+ S11 (N, U⇩e⇩r⇩1⇩1, ℱ, Γ', None)"
unfolding S11_def ‹𝒞 = None› by (auto intro: constant_context.intros)
show "ord_res_10_matches_ord_res_11 S10' (N, U⇩e⇩r⇩1⇩1, ℱ, Γ', None)"
unfolding ‹S10' = (N, s10')› ‹s10' = _›
proof (rule ord_res_10_matches_ord_res_11.intros)
show "ord_res_10_invars N (U⇩e⇩r⇩1⇩0, ℱ, Γ')"
using step10 ‹s10' = _› invars10 ord_res_10_preserves_invars by metis
next
show "ord_res_11_invars N (U⇩e⇩r⇩1⇩1, ℱ, Γ', None)"
using step11 invars11 ‹𝒞 = None› ord_res_11_preserves_invars by metis
next
show "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0 |-| {|{#}|}"
unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0›
using ‹{#} |∉| N |∪| U⇩e⇩r⇩1⇩0› by simp
next
have "{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)"
using ‹{#} |∉| N |∪| U⇩e⇩r⇩1⇩0› by (simp add: iefac_def)
thus "if {#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0) then Γ' = [] ∧ None = Some {#} else None = None"
by argo
qed
qed
next
case step_hyps: (decide_pos A C Γ' ℱ')
have "𝒞 = None"
using step_hyps 𝒞_eq_None_if_no_false_cls by argo
have "{#} |∉| N |∪| U⇩e⇩r⇩1⇩0"
using step_hyps mempty_not_in_if by argo
have "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0"
using step_hyps U⇩e⇩r⇩1⇩1_eq_U⇩e⇩r⇩1⇩0_if by argo
show ?thesis
proof (intro exI conjI)
have step11: "ord_res_11 N (U⇩e⇩r⇩1⇩1, ℱ, Γ, None) (U⇩e⇩r⇩1⇩1, ℱ', Γ', None)"
proof (rule ord_res_11.decide_pos)
show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩1). trail_false_cls Γ C)"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "linorder_trm.is_least_in_fset
{|A⇩2 |∈| atms_of_clss (N |∪| U⇩e⇩r⇩1⇩1). ∀A⇩1|∈|trail_atms Γ. A⇩1 ≺⇩t A⇩2|} A"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "linorder_cls.is_least_in_fset
{|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩1). clause_could_propagate Γ C (Pos A)|} C"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "Γ' = (Pos A, None) # Γ"
using step_hyps by argo
next
show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩1). trail_false_cls Γ' C)"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "ℱ' = (if ord_res.is_strictly_maximal_lit (Pos A) C then ℱ else finsert C ℱ)"
using step_hyps by argo
qed
thus "(constant_context ord_res_11)⇧+⇧+ S11 (N, U⇩e⇩r⇩1⇩1, ℱ', Γ', None)"
unfolding S11_def ‹𝒞 = None› by (auto intro: constant_context.intros)
show "ord_res_10_matches_ord_res_11 S10' (N, U⇩e⇩r⇩1⇩1, ℱ', Γ', None)"
unfolding ‹S10' = (N, s10')› ‹s10' = _›
proof (rule ord_res_10_matches_ord_res_11.intros)
show "ord_res_10_invars N (U⇩e⇩r⇩1⇩0, ℱ', Γ')"
using step10 ‹s10' = _› invars10 ord_res_10_preserves_invars by metis
next
show "ord_res_11_invars N (U⇩e⇩r⇩1⇩1, ℱ', Γ', None)"
using step11 invars11 ‹𝒞 = None› ord_res_11_preserves_invars by metis
next
show "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0 |-| {|{#}|}"
unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0›
using ‹{#} |∉| N |∪| U⇩e⇩r⇩1⇩0› by simp
next
have "{#} |∉| iefac ℱ' |`| (N |∪| U⇩e⇩r⇩1⇩0)"
using ‹{#} |∉| N |∪| U⇩e⇩r⇩1⇩0› by (simp add: iefac_def)
thus "if {#} |∈| iefac ℱ' |`| (N |∪| U⇩e⇩r⇩1⇩0) then Γ' = [] ∧ None = Some {#} else None = None"
by argo
qed
qed
next
case step_hyps: (propagate A C Γ' ℱ')
have "𝒞 = None"
using step_hyps 𝒞_eq_None_if_no_false_cls by argo
have "{#} |∉| N |∪| U⇩e⇩r⇩1⇩0"
using step_hyps mempty_not_in_if by argo
have "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0"
using step_hyps U⇩e⇩r⇩1⇩1_eq_U⇩e⇩r⇩1⇩0_if by argo
show ?thesis
proof (intro exI conjI)
have step11: "ord_res_11 N (U⇩e⇩r⇩1⇩1, ℱ, Γ, None) (U⇩e⇩r⇩1⇩1, ℱ', Γ', None)"
proof (rule ord_res_11.propagate)
show "¬ (∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩1). trail_false_cls Γ C)"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "linorder_trm.is_least_in_fset
{|A⇩2 |∈| atms_of_clss (N |∪| U⇩e⇩r⇩1⇩1). ∀A⇩1|∈|trail_atms Γ. A⇩1 ≺⇩t A⇩2|} A"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "linorder_cls.is_least_in_fset
{|C |∈| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩1). clause_could_propagate Γ C (Pos A)|} C"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "Γ' = (Pos A, Some (efac C)) # Γ"
using step_hyps by argo
next
show "∃C|∈|iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩1). trail_false_cls Γ' C"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
next
show "ℱ' = (if ord_res.is_strictly_maximal_lit (Pos A) C then ℱ else finsert C ℱ)"
using step_hyps by argo
qed
thus "(constant_context ord_res_11)⇧+⇧+ S11 (N, U⇩e⇩r⇩1⇩1, ℱ', Γ', None)"
unfolding S11_def ‹𝒞 = None› by (auto intro: constant_context.intros)
show "ord_res_10_matches_ord_res_11 S10' (N, U⇩e⇩r⇩1⇩1, ℱ', Γ', None)"
unfolding ‹S10' = (N, s10')› ‹s10' = _›
proof (rule ord_res_10_matches_ord_res_11.intros)
show "ord_res_10_invars N (U⇩e⇩r⇩1⇩0, ℱ', Γ')"
using step10 ‹s10' = _› invars10 ord_res_10_preserves_invars by metis
next
show "ord_res_11_invars N (U⇩e⇩r⇩1⇩1, ℱ', Γ', None)"
using step11 invars11 ‹𝒞 = None› ord_res_11_preserves_invars by metis
next
show "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0 |-| {|{#}|}"
unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0›
using ‹{#} |∉| N |∪| U⇩e⇩r⇩1⇩0› by simp
next
have "{#} |∉| iefac ℱ' |`| (N |∪| U⇩e⇩r⇩1⇩0)"
using ‹{#} |∉| N |∪| U⇩e⇩r⇩1⇩0› by (simp add: iefac_def)
thus "if {#} |∈| iefac ℱ' |`| (N |∪| U⇩e⇩r⇩1⇩0) then Γ' = [] ∧ None = Some {#} else None = None"
by argo
qed
qed
next
case step_hyps: (resolution D A C U⇩e⇩r⇩1⇩0' Γ')
note D_max_lit = ‹ord_res.is_maximal_lit (Neg A) D›
have "{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)"
using ‹linorder_cls.is_least_in_fset _ D› ‹linorder_lit.is_maximal_in_mset D _›
unfolding linorder_cls.is_least_in_ffilter_iff linorder_lit.is_maximal_in_mset_iff
by (metis (no_types, lifting) empty_iff linorder_cls.leD mempty_lesseq_cls set_mset_empty
trail_false_cls_mempty)
have "𝒞 = None"
using match_hyps ‹{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)› by argo
have "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0"
using match_hyps ‹{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)› by force
have step11_conf: "ord_res_11 N (U⇩e⇩r⇩1⇩1, ℱ, Γ, None) (U⇩e⇩r⇩1⇩1, ℱ, Γ, Some D)"
proof (rule ord_res_11.conflict)
show "linorder_cls.is_least_in_fset
(ffilter (trail_false_cls Γ) (iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩1))) D"
using step_hyps unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› by argo
qed
have Γ_spec: "∀Ln Γ'. Γ = Ln # Γ' ⟶
(snd Ln ≠ None) = fBex (iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)) (trail_false_cls Γ) ∧
(∀x∈set Γ'. snd x = None)"
using invars10 by (simp add: ord_res_10_invars.simps)
then obtain Γ''' where "Γ = (Pos A, Some C) # Γ'''"
using ‹map_of Γ (Pos A) = Some (Some C)›
by (metis list.set_cases map_of_SomeD not_Some_eq snd_conv)
have C_max_lit: "linorder_lit.is_greatest_in_mset C (Pos A)"
using invars10 by (simp add: ord_res_10_invars.simps ‹Γ = (Pos A, Some C) # Γ'''›)
have "ord_res.ground_resolution D C ((D - {#Neg A#}) + (C - {#Pos A#}))"
proof (rule ord_res.ground_resolutionI)
show "D = add_mset (Neg A) (D - {#Neg A#})"
using D_max_lit unfolding linorder_lit.is_maximal_in_mset_iff by simp
next
show "C = add_mset (Pos A) (C - {#Pos A#})"
using C_max_lit unfolding linorder_lit.is_greatest_in_mset_iff by simp
next
show "C ≺⇩c D"
using C_max_lit D_max_lit
by (simp add: clause_lt_clause_if_max_lit_comp
linorder_lit.is_greatest_in_mset_iff_is_maximal_and_count_eq_one)
next
show "{#} = {#} ∧ ord_res.is_maximal_lit (Neg A) D ∨ Neg A ∈# {#}"
using D_max_lit by argo
next
show "{#} = {#}"
by argo
next
show "ord_res.is_strictly_maximal_lit (Pos A) C"
using C_max_lit .
next
show "remove1_mset (Neg A) D + remove1_mset (Pos A) C = (D - {#Neg A#}) + (C - {#Pos A#})"
..
qed
hence "eres C D ≠ D"
unfolding eres_ident_iff not_not ground_resolution_def by metis
have D_false: "trail_false_cls Γ D"
using step_hyps unfolding linorder_cls.is_least_in_ffilter_iff by argo
have "clause_could_propagate Γ''' C (Pos A)"
using invars10 ‹Γ = (Pos A, Some C) # Γ'''› by (simp add: ord_res_10_invars.simps)
hence "trail_false_cls Γ''' {#L ∈# C. L ≠ Pos A#}"
unfolding clause_could_propagate_def by argo
hence C_almost_false: "trail_false_cls Γ {#L ∈# C. L ≠ Pos A#}"
unfolding ‹Γ = (Pos A, Some C) # Γ'''›
by (meson suffix_ConsD suffix_order.dual_order.refl trail_false_cls_if_trail_false_suffix)
have eres_false: "trail_false_cls Γ (eres C D)"
unfolding trail_false_cls_def
proof (intro ballI)
fix K :: "'f gliteral"
assume "K ∈# eres C D"
hence "K ∈# C ∧ K ≠ Pos A ∨ K ∈# D"
using strong_lit_in_one_of_resolvents_if_in_eres[OF D_max_lit] by simp
thus "trail_false_lit Γ K"
proof (elim disjE conjE)
assume "K ∈# C" and "K ≠ Pos A"
thus "trail_false_lit Γ K"
using C_almost_false unfolding trail_false_cls_def by simp
next
assume "K ∈# D"
thus "trail_false_lit Γ K"
using D_false unfolding trail_false_cls_def by simp
qed
qed
have "∀Ln ∈ set Γ. ∀C. snd Ln = Some C ⟶ ord_res.is_strictly_maximal_lit (fst Ln) C"
using invars10 by (simp add: ord_res_10_invars.simps)
hence C_max_lit: "ord_res.is_strictly_maximal_lit (Pos A) C"
unfolding ‹Γ = (Pos A, Some C) # Γ'''› by simp
have steps11_reso: "(ord_res_11 N)⇧*⇧* (U⇩e⇩r⇩1⇩1, ℱ, Γ, Some D) (U⇩e⇩r⇩1⇩1, ℱ, Γ, Some (eres C D))"
unfolding ‹Γ = (Pos A, Some C) # Γ'''›
using C_max_lit rtrancl_ord_res_11_all_resolution_steps by metis
define strict_P :: "'f gterm literal × 'f gterm literal multiset option ⇒ bool" where
"strict_P ≡ λLn. ∀K. ord_res.is_maximal_lit K (eres C D) ⟶ atm_of K ≺⇩t atm_of (fst Ln)"
have "∀K ∈# eres C D. - K ∈ fst ` set Γ"
using eres_false unfolding trail_false_cls_def trail_false_lit_def .
moreover have Γ_sorted: "sorted_wrt (λx y. atm_of (fst y) ≺⇩t atm_of (fst x)) Γ"
using invars10 by (simp add: ord_res_10_invars.simps)
ultimately have "dropWhile strict_P Γ = dropWhile (λLn. - fst Ln ∉# eres C D) Γ"
proof (induction Γ)
case Nil
show ?case by simp
next
case (Cons Ln Γ)
show ?case
proof (cases "eres C D = {#}")
case True
thus ?thesis
unfolding strict_P_def linorder_lit.is_maximal_in_mset_iff by simp
next
case False
then obtain L where eres_max_lit: "ord_res.is_maximal_lit L (eres C D)"
using linorder_lit.ex_maximal_in_mset by metis
hence strict_P_Ln_iff: "strict_P Ln ⟷ atm_of L ≺⇩t atm_of (fst Ln)"
unfolding strict_P_def
by (metis linorder_lit.Uniq_is_maximal_in_mset the1_equality')
show ?thesis
proof (cases "atm_of L ≺⇩t atm_of (fst Ln)")
case True
moreover have "- fst Ln ∉# eres C D"
using True
by (smt (verit, best) atm_of_uminus eres_max_lit linorder_lit.dual_order.strict_trans
linorder_lit.is_maximal_in_mset_iff linorder_lit.neq_iff linorder_trm.order.irrefl
literal.exhaust_sel ord_res.less_lit_simps(4))
moreover have "dropWhile strict_P Γ = dropWhile (λLn. - fst Ln ∉# eres C D) Γ"
proof (rule Cons.IH)
show "∀K∈#eres C D. - K ∈ fst ` set Γ"
using Cons.prems True ‹- fst Ln ∉# eres C D›
using image_iff by fastforce
next
show "sorted_wrt (λx y. atm_of (fst y) ≺⇩t atm_of (fst x)) Γ"
using Cons.prems by simp
qed
ultimately show ?thesis
unfolding dropWhile.simps
unfolding strict_P_Ln_iff
by simp
next
case False
hence "atm_of (fst Ln) ≼⇩t atm_of L"
by order
hence "atm_of (fst Ln) = atm_of L"
using Cons.prems
using atm_of_eq_atm_of eres_max_lit linorder_lit.is_maximal_in_mset_iff by fastforce
hence "L = fst Ln ∨ L = - fst Ln"
by (metis atm_of_eq_atm_of)
moreover have "- fst Ln ∉ fst ` set Γ"
using ‹sorted_wrt (λx y. atm_of (fst y) ≺⇩t atm_of (fst x)) (Ln # Γ)›
by fastforce
moreover have "- L ∈ fst ` (set (Ln # Γ))"
using Cons.prems(1) eres_max_lit linorder_lit.is_maximal_in_mset_iff by blast
ultimately have "- fst Ln ∈# eres C D"
using eres_max_lit linorder_lit.is_maximal_in_mset_iff by auto
then show ?thesis
unfolding dropWhile.simps
unfolding strict_P_Ln_iff
using False
by argo
qed
qed
qed
hence steps11_skip: "(ord_res_11 N)⇧*⇧*
(U⇩e⇩r⇩1⇩1, ℱ, Γ, Some (eres C D))
(U⇩e⇩r⇩1⇩1, ℱ, dropWhile strict_P Γ, Some (eres C D))"
using rtrancl_ord_res_11_all_skip_steps by metis
have most_steps11: "(ord_res_11 N)⇧+⇧+ (U⇩e⇩r⇩1⇩1, ℱ, Γ, None)
(U⇩e⇩r⇩1⇩1, ℱ, dropWhile strict_P Γ, Some (eres C D))"
using step11_conf steps11_reso steps11_skip by simp
show ?thesis
proof (cases "eres C D = {#}")
case True
hence "dropWhile strict_P Γ = []"
unfolding strict_P_def ‹eres C D = {#}›
unfolding linorder_lit.is_maximal_in_mset_iff
by simp
have "Γ' = []"
unfolding ‹Γ' = dropWhile _ Γ› ‹eres C D = {#}›
unfolding linorder_lit.is_maximal_in_mset_iff
by simp
show ?thesis
proof (intro exI conjI)
show "(constant_context ord_res_11)⇧+⇧+ S11 (N, U⇩e⇩r⇩1⇩1, ℱ, [], Some {#})"
unfolding S11_def ‹𝒞 = None›
using most_steps11[unfolded ‹dropWhile strict_P Γ = []› ‹eres C D = {#}›]
using tranclp_constant_context by metis
show "ord_res_10_matches_ord_res_11 S10' (N, U⇩e⇩r⇩1⇩1, ℱ, [], Some {#})"
unfolding ‹S10' = (N, s10')› ‹s10' = _› ‹Γ' = []›
proof (rule ord_res_10_matches_ord_res_11.intros)
show "ord_res_10_invars N (U⇩e⇩r⇩1⇩0', ℱ, [])"
using step10 ‹s10' = _› ‹Γ' = []› invars10 ord_res_10_preserves_invars by metis
next
show "ord_res_11_invars N (U⇩e⇩r⇩1⇩1, ℱ, [], Some {#})"
using most_steps11 invars11
unfolding ‹𝒞 = None› ‹dropWhile strict_P Γ = []› ‹eres C D = {#}›
by (metis tranclp_ord_res_11_preserves_invars)
next
show "U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0' |-| {|{#}|}"
unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› ‹U⇩e⇩r⇩1⇩0' = finsert (eres C D) U⇩e⇩r⇩1⇩0› ‹eres C D = {#}›
using ‹{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)›
by force
next
show "if {#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0') then
[] = [] ∧ Some {#} = Some {#} else Some {#} = None"
unfolding ‹U⇩e⇩r⇩1⇩0' = finsert (eres C D) U⇩e⇩r⇩1⇩0› ‹eres C D = {#}›
by simp
qed
qed
next
case False
then obtain L where eres_max_lit: "ord_res.is_maximal_lit L (eres C D)"
using linorder_lit.ex_maximal_in_mset by metis
hence "L ∈# eres C D"
unfolding linorder_lit.is_maximal_in_mset_iff by argo
hence "L ∈# C ∧ L ≠ Pos A ∨ L ∈# D ∧ L ≠ Neg A"
using stronger_lit_in_one_of_resolvents_if_in_eres ‹eres C D ≠ D› D_max_lit
by (metis uminus_Neg)
hence "L ≺⇩l Neg A"
using C_max_lit D_max_lit
unfolding linorder_lit.is_greatest_in_mset_iff linorder_lit.is_maximal_in_mset_iff
by (metis C_max_lit D_max_lit ‹L ∈# eres C D› eres_lt_if ord_res.asymp_less_lit
ord_res.less_lit_simps(3,4) ord_res.transp_less_lit in_remove1_mset_neq
linorder_lit.less_than_maximal_if_multp⇩H⇩O linorder_lit.order.not_eq_order_implies_strict
literal.disc(2) multp_eq_multp⇩H⇩O uminus_Neg)
have "dropWhile strict_P Γ = dropWhile (λLn. atm_of L ≺⇩t atm_of (fst Ln)) Γ"
unfolding strict_P_def
using eres_max_lit
by (metis (no_types) Uniq_D linorder_lit.Uniq_is_maximal_in_mset)
also have "… = (- L, None) # dropWhile (λLn. atm_of L ≼⇩t atm_of (fst Ln)) Γ"
proof -
have "- L ≠ Pos A"
using ‹L ≺⇩l Neg A› by (cases L) simp_all
hence "- L ∈ fst ` set Γ'''"
using eres_false ‹L ∈# eres C D›
unfolding trail_false_cls_def trail_false_lit_def
unfolding ‹Γ = (Pos A, Some C) # Γ'''›
by auto
hence "(- L, None) ∈ set Γ'''"
using Γ_spec unfolding ‹Γ = (Pos A, Some C) # Γ'''› by auto
then obtain Γ⇩0 Γ⇩1 where "Γ''' = Γ⇩1 @ (- L, None) # Γ⇩0"
by (meson split_list)
hence "Γ = (Pos A, Some C) # Γ⇩1 @ (- L, None) # Γ⇩0"
unfolding ‹Γ = (Pos A, Some C) # Γ'''› by (simp only:)
have AAA: "∀Ln ∈ set ((Pos A, Some C) # Γ⇩1). atm_of L ≺⇩t atm_of (fst Ln)"
using Γ_sorted unfolding ‹Γ = (Pos A, Some C) # Γ⇩1 @ (- L, None) # Γ⇩0›
by (simp add: sorted_wrt_append)
hence BBB: "∀Ln ∈ set ((Pos A, Some C) # Γ⇩1 @ [(- L, None)]). atm_of L ≼⇩t atm_of (fst Ln)"
by simp
have "∀Ln ∈ set Γ⇩0. atm_of (fst Ln) ≺⇩t atm_of L"
using Γ_sorted unfolding ‹Γ = (Pos A, Some C) # Γ⇩1 @ (- L, None) # Γ⇩0›
by (simp add: sorted_wrt_append)
hence CCC: "∀Ln ∈ set Γ⇩0. ¬ atm_of L ≼⇩t atm_of (fst Ln)"
using linorder_trm.leD by blast
have "dropWhile (λLn. atm_of L ≺⇩t atm_of (fst Ln)) Γ = (- L, None) # Γ⇩0"
unfolding ‹Γ = (Pos A, Some C) # Γ⇩1 @ (- L, None) # Γ⇩0›
using dropWhile_append2 AAA by simp
also have "… = (- L, None) # dropWhile (λLn. atm_of L ≼⇩t atm_of (fst Ln)) Γ⇩0"
using CCC by (simp add: dropWhile_ident_if_pred_always_false)
also have "… = (- L, None) # dropWhile (λLn. atm_of L ≼⇩t atm_of (fst Ln)) Γ"
unfolding ‹Γ = (Pos A, Some C) # Γ⇩1 @ (- L, None) # Γ⇩0›
using dropWhile_append2 BBB by simp
finally show ?thesis .
qed
also have "… = (- L, None) # Γ'"
unfolding ‹Γ' = dropWhile _ Γ›
using eres_max_lit
by (metis (no_types) Uniq_D linorder_lit.Uniq_is_maximal_in_mset)
finally have "dropWhile strict_P Γ = (- L, None) # Γ'" .
have step10_back: "ord_res_11 N
(U⇩e⇩r⇩1⇩1, ℱ, dropWhile strict_P Γ, Some (eres C D))
(finsert (eres C D) U⇩e⇩r⇩1⇩1, ℱ, Γ', None)"
unfolding ‹dropWhile strict_P Γ = (- L, None) # Γ'›
proof (rule ord_res_11.backtrack)
show "(- L, None) # Γ' = (- L, None) # Γ'" ..
next
show "- (- L) ∈# eres C D"
unfolding uminus_of_uminus_id
using eres_max_lit
unfolding linorder_lit.is_maximal_in_mset_iff by argo
qed
hence all_steps11: "(ord_res_11 N)⇧+⇧+ (U⇩e⇩r⇩1⇩1, ℱ, Γ, None)
(finsert (eres C D) U⇩e⇩r⇩1⇩1, ℱ, Γ', None)"
using most_steps11 by simp
show ?thesis
proof (intro exI conjI)
show "(constant_context ord_res_11)⇧+⇧+ S11 (N, finsert (eres C D) U⇩e⇩r⇩1⇩1, ℱ, Γ', None)"
unfolding S11_def ‹𝒞 = None›
using all_steps11 tranclp_constant_context by metis
have "{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0')"
by (smt (verit, del_insts) False ‹{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0)› fimage.rep_eq
fimageE fimageI finsertE funion_iff iefac_def mempty_in_image_efac_iff step_hyps(5))
show "ord_res_10_matches_ord_res_11 S10' (N, finsert (eres C D) U⇩e⇩r⇩1⇩1, ℱ, Γ', None)"
unfolding ‹S10' = (N, s10')› ‹s10' = _›
proof (rule ord_res_10_matches_ord_res_11.intros)
show "ord_res_10_invars N (U⇩e⇩r⇩1⇩0', ℱ, Γ')"
using step10 ‹s10' = _› invars10 ord_res_10_preserves_invars by metis
next
show "ord_res_11_invars N (finsert (eres C D) U⇩e⇩r⇩1⇩1, ℱ, Γ', None)"
using all_steps11 invars11
unfolding ‹𝒞 = None›
by (metis tranclp_ord_res_11_preserves_invars)
next
show "finsert (eres C D) U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0' |-| {|{#}|}"
unfolding ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› ‹U⇩e⇩r⇩1⇩0' = finsert (eres C D) U⇩e⇩r⇩1⇩0›
using ‹{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0')›
using False ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0› ‹U⇩e⇩r⇩1⇩1 = U⇩e⇩r⇩1⇩0 |-| {|{#}|}› by auto
next
show "if {#} |∈| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0') then
Γ' = [] ∧ None = Some {#} else None = None"
using ‹{#} |∉| iefac ℱ |`| (N |∪| U⇩e⇩r⇩1⇩0')› by simp
qed
qed
qed
qed
qed
theorem bisimulation_ord_res_10_ord_res_11:
defines "match ≡ λ_. ord_res_10_matches_ord_res_11"
shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_10_state ⇒ 'f ord_res_11_state ⇒ bool) ℛ.
bisimulation
(constant_context ord_res_10) (constant_context ord_res_11)
ord_res_8_final ord_res_11_final
ℛ MATCH"
proof (rule ex_bisimulation_from_forward_simulation)
show "right_unique (constant_context ord_res_10)"
using right_unique_constant_context right_unique_ord_res_10 by metis
next
show "right_unique (constant_context ord_res_11)"
using right_unique_constant_context right_unique_ord_res_11 by metis
next
show "∀S. ord_res_8_final S ⟶ (∄S'. constant_context ord_res_10 S S')"
by (metis finished_def ord_res_10_semantics.final_finished)
next
show "∀S. ord_res_11_final S ⟶ (∄S'. constant_context ord_res_11 S S')"
by (metis finished_def ord_res_11_semantics.final_finished)
next
show "∀i S10 S11. match i S10 S11 ⟶ ord_res_8_final S10 ⟷ ord_res_11_final S11"
unfolding match_def
using ord_res_10_final_iff_ord_res_11_final by metis
next
show "∀i S10 S11. match i S10 S11 ⟶
safe_state (constant_context ord_res_10) ord_res_8_final S10 ∧
safe_state (constant_context ord_res_11) ord_res_11_final S11"
proof (intro allI impI conjI)
fix i S10 S11
assume match: "match i S10 S11"
show "safe_state (constant_context ord_res_10) ord_res_8_final S10"
using match[unfolded match_def]
using ord_res_10_safe_state_if_invars
unfolding ord_res_10_matches_ord_res_11.simps by auto
show "safe_state (constant_context ord_res_11) ord_res_11_final S11"
using match[unfolded match_def]
using ord_res_11_safe_state_if_invars
using ord_res_10_matches_ord_res_11.simps by auto
qed
next
show "wfp (λ_ _. False)"
by simp
next
show "∀i S10 S11 S10'. match i S10 S11 ⟶ constant_context ord_res_10 S10 S10' ⟶
(∃i' S11'. (constant_context ord_res_11)⇧+⇧+ S11 S11' ∧ match i' S10' S11') ∨
(∃i'. match i' S10' S11 ∧ False)"
unfolding match_def
using forward_simulation_between_10_and_11 by metis
qed
end
lemma forward_simulation_composition:
assumes
"forward_simulation step1 step2 final1 final2 order1 match1"
"forward_simulation step2 step3 final2 final3 order2 match2"
defines "ℛ ≡ λi i'. lex_prodp order2⇧+⇧+ order1 (prod.swap i) (prod.swap i')"
shows "forward_simulation step1 step3 final1 final3 ℛ (rel_comp match1 match2)"
proof intro_locales
show "semantics step1 final1"
using assms
by (auto intro: forward_simulation.axioms)
next
show "semantics step3 final3"
using assms
by (auto intro: forward_simulation.axioms)
next
have "well_founded order1" "well_founded order2"
using assms
by (auto intro: forward_simulation.axioms)
hence "wfp ℛ"
unfolding ℛ_def
by (metis (no_types, lifting) lex_prodp_wfP well_founded.wf wfP_trancl
wfp_if_convertible_to_wfp)
thus "well_founded ℛ"
by unfold_locales
next
show "forward_simulation_axioms step1 step3 final1 final3 ℛ (rel_comp match1 match2)"
proof unfold_locales
fix i s1 s3
assume
match: "rel_comp match1 match2 i s1 s3" and
final: "final1 s1"
obtain i1 i2 s2 where "match1 i1 s1 s2" and "match2 i2 s2 s3" and "i = (i1, i2)"
using match unfolding rel_comp_def by auto
thus "final3 s3"
using final assms(1,2)[THEN forward_simulation.match_final]
by simp
next
fix i s1 s3 s1'
assume
match: "rel_comp match1 match2 i s1 s3" and
step: "step1 s1 s1'"
obtain i1 i2 s2 where "match1 i1 s1 s2" and "match2 i2 s2 s3" and i_def: "i = (i1, i2)"
using match unfolding rel_comp_def by auto
from forward_simulation.simulation[OF assms(1) ‹match1 i1 s1 s2› step]
show "(∃i' s3'. step3⇧+⇧+ s3 s3' ∧ rel_comp match1 match2 i' s1' s3') ∨
(∃i'. rel_comp match1 match2 i' s1' s3 ∧ ℛ i' i)"
(is "(∃i' s1'. ?STEPS i' s1') ∨ (∃i'. ?STALL i')")
proof (elim disjE exE conjE)
fix i1' s2'
assume "step2⇧+⇧+ s2 s2'" and "match1 i1' s1' s2'"
from forward_simulation.lift_simulation_plus[OF assms(2) ‹step2⇧+⇧+ s2 s2'› ‹match2 i2 s2 s3›]
show ?thesis
proof (elim disjE exE conjE)
fix i2' s3'
assume "step3⇧+⇧+ s3 s3'" and "match2 i2' s2' s3'"
hence "?STEPS (i1', i2') s3'"
by (auto intro: ‹match1 i1' s1' s2'› simp: rel_comp_def)
thus ?thesis by auto
next
fix i2''
assume "match2 i2'' s2' s3" and "order2⇧+⇧+ i2'' i2"
hence "?STALL (i1', i2'')"
unfolding rel_comp_def i_def comp_def prod.swap_def prod.sel
proof (intro conjI)
show "(match1 i1' OO match2 i2'') s1' s3"
using ‹match1 i1' s1' s2'› ‹match2 i2'' s2' s3›
by (auto simp add: relcompp_apply)
next
show "ℛ (i1', i2'') (i1, i2)"
unfolding ℛ_def lex_prodp_def prod.swap_def prod.sel
using ‹order2⇧+⇧+ i2'' i2› by argo
qed
thus ?thesis
by metis
qed
next
fix i1'
assume "match1 i1' s1' s2" and "order1 i1' i1"
hence "?STALL (i1', i2)"
unfolding rel_comp_def i_def prod.sel
using ‹match2 i2 s2 s3› by (auto simp: ℛ_def lex_prodp_def)
thus ?thesis
by metis
qed
qed
qed
text ‹For AFP-devel, delete
@{thm Simulation_SCLFOL_ORDRES.forward_simulation_composition}
as it is available in \<^theory>‹VeriComp.Simulation›.›
type_synonym bisim_index_1_2 = "nat × nat"
type_synonym bisim_index_1_3 = "bisim_index_1_2 × (nat × nat)"
type_synonym bisim_index_1_4 = "bisim_index_1_3 × (nat × nat)"
type_synonym bisim_index_1_5 = "bisim_index_1_4 × (nat × nat)"
type_synonym bisim_index_1_6 = "bisim_index_1_5 × (nat × nat)"
type_synonym bisim_index_1_7 = "bisim_index_1_6 × (nat × nat)"
type_synonym bisim_index_1_8 = "bisim_index_1_7 × (nat × nat)"
type_synonym bisim_index_1_9 = "bisim_index_1_8 × (nat × nat)"
type_synonym bisim_index_1_10 = "bisim_index_1_9 × (nat × nat)"
type_synonym bisim_index_1_11 = "bisim_index_1_10 × (nat × nat)"
context simulation_SCLFOL_ground_ordered_resolution begin
theorem forward_simulation_ord_res_1_ord_res_11:
obtains
MATCH :: "bisim_index_1_11 ⇒ 'f ord_res_1_state ⇒ 'f ord_res_11_state ⇒ bool" and
ℛ :: "bisim_index_1_11 ⇒ bisim_index_1_11 ⇒ bool"
where
"forward_simulation
ord_res_1 (constant_context ord_res_11)
ord_res_1_final ord_res_11_final
ℛ MATCH"
proof -
have bi_to_fwd: "∃MATCH (ℛ :: 'a ⇒ 'a ⇒ bool).
forward_simulation step1 step2 final1 final2 ℛ MATCH"
if "∃MATCH (ℛ :: 'a ⇒ 'a ⇒ bool). bisimulation step1 step2 final1 final2 ℛ MATCH"
for step1 step2 final1 final2
using that by (auto intro: bisimulation.axioms)
show ?thesis
using that
using bisimulation_ord_res_1_ord_res_2[THEN bi_to_fwd]
bisimulation_ord_res_2_ord_res_3[THEN bi_to_fwd]
bisimulation_ord_res_3_ord_res_4[THEN bi_to_fwd]
bisimulation_ord_res_4_ord_res_5[THEN bi_to_fwd]
bisimulation_ord_res_5_ord_res_6[THEN bi_to_fwd]
bisimulation_ord_res_6_ord_res_7[THEN bi_to_fwd]
bisimulation_ord_res_7_ord_res_8[THEN bi_to_fwd]
bisimulation_ord_res_8_ord_res_9[THEN bi_to_fwd]
bisimulation_ord_res_9_ord_res_10[THEN bi_to_fwd]
bisimulation_ord_res_10_ord_res_11[THEN bi_to_fwd]
using forward_simulation_composition by meson
qed
theorem backward_simulation_ord_res_1_ord_res_11:
obtains
MATCH :: "bisim_index_1_11 ⇒ 'f ord_res_1_state ⇒ 'f ord_res_11_state ⇒ bool" and
ℛ :: "bisim_index_1_11 ⇒ bisim_index_1_11 ⇒ bool"
where
"backward_simulation
ord_res_1 (constant_context ord_res_11)
ord_res_1_final ord_res_11_final
ℛ MATCH"
proof -
have bi_to_bwd: "∃MATCH (ℛ :: 'a ⇒ 'a ⇒ bool).
backward_simulation step1 step2 final1 final2 ℛ MATCH"
if "∃MATCH (ℛ :: 'a ⇒ 'a ⇒ bool). bisimulation step1 step2 final1 final2 ℛ MATCH"
for step1 step2 final1 final2
using that by (auto intro: bisimulation.axioms)
show ?thesis
using that
using bisimulation_ord_res_1_ord_res_2[THEN bi_to_bwd]
bisimulation_ord_res_2_ord_res_3[THEN bi_to_bwd]
bisimulation_ord_res_3_ord_res_4[THEN bi_to_bwd]
bisimulation_ord_res_4_ord_res_5[THEN bi_to_bwd]
bisimulation_ord_res_5_ord_res_6[THEN bi_to_bwd]
bisimulation_ord_res_6_ord_res_7[THEN bi_to_bwd]
bisimulation_ord_res_7_ord_res_8[THEN bi_to_bwd]
bisimulation_ord_res_8_ord_res_9[THEN bi_to_bwd]
bisimulation_ord_res_9_ord_res_10[THEN bi_to_bwd]
bisimulation_ord_res_10_ord_res_11[THEN bi_to_bwd]
using backward_simulation_composition by meson
qed
section ‹ORD-RES-11 is a regular SCL strategy›
definition gtrailelem_of_trailelem where
"gtrailelem_of_trailelem ≡ λ(L, opt).
(lit_of_glit L, map_option (λC. (cls_of_gcls {#K ∈# C. K ≠ L#}, lit_of_glit L, Var)) opt)"
fun state_of_gstate :: "_ ⇒ ('f, 'v) SCL_FOL.state" where
"state_of_gstate (U⇩G, _, Γ⇩G, 𝒞⇩G) =
(let
Γ = map gtrailelem_of_trailelem Γ⇩G;
U = cls_of_gcls |`| U⇩G;
𝒞 = map_option (λC⇩G. (cls_of_gcls C⇩G, Var)) 𝒞⇩G
in (Γ, U, 𝒞))"
lemma fst_case_prod_simp: "fst (case p of (x, y) ⇒ (f x, g x y)) = f (fst p)"
by (cases p) simp
lemma trail_false_cls_nonground_iff_trail_false_cls_ground:
fixes Γ⇩G and D⇩G :: "'f gclause"
fixes Γ :: "('f, 'v) SCL_FOL.trail" and D :: "('f, 'v) term clause"
defines "Γ ≡ map gtrailelem_of_trailelem Γ⇩G" and "D ≡ cls_of_gcls D⇩G"
shows "trail_false_cls Γ D ⟷ trail_false_cls Γ⇩G D⇩G"
proof -
have "trail_false_cls Γ D ⟷ (∀L ∈# D. trail_false_lit Γ L)"
unfolding trail_false_cls_def ..
also have "… ⟷ (∀L⇩G ∈# D⇩G. trail_false_lit Γ (lit_of_glit L⇩G))"
unfolding D_def cls_of_gcls_def by simp
also have "… ⟷ (∀L⇩G ∈# D⇩G. trail_false_lit Γ⇩G L⇩G)"
proof -
have "trail_false_lit Γ (lit_of_glit L⇩G) ⟷ trail_false_lit Γ⇩G L⇩G"
for L⇩G :: "'f gterm literal"
proof -
have "trail_false_lit Γ (lit_of_glit L⇩G) ⟷ - lit_of_glit L⇩G ∈ fst ` set Γ"
unfolding trail_false_lit_def ..
also have "… ⟷
- (lit_of_glit L⇩G :: ('f, 'v) term literal) ∈ set (map (λx. lit_of_glit (fst x)) Γ⇩G)"
unfolding Γ_def image_set list.map_comp
unfolding gtrailelem_of_trailelem_def
unfolding list.map_comp
unfolding comp_def fst_case_prod_simp ..
also have "… ⟷ (lit_of_glit (- L⇩G) :: ('f, 'v) term literal) ∈ lit_of_glit ` fst ` set Γ⇩G"
by (cases L⇩G) (auto simp: lit_of_glit_def)
also have "… ⟷ - L⇩G ∈ fst ` set Γ⇩G"
using inj_image_mem_iff inj_lit_of_glit by metis
also have "… ⟷ trail_false_lit Γ⇩G L⇩G"
unfolding trail_false_lit_def ..
finally show "trail_false_lit Γ (lit_of_glit L⇩G) = trail_false_lit Γ⇩G L⇩G" .
qed
thus ?thesis by metis
qed
also have "… ⟷ trail_false_cls Γ⇩G D⇩G"
unfolding trail_false_cls_def ..
finally show ?thesis .
qed
theorem ord_res_11_is_strategy_for_regular_scl:
fixes
N⇩G :: "'f gclause fset" and
N :: "('f, 'v) term clause fset" and
β⇩G :: "'f gterm" and
β :: "('f, 'v) term" and
S⇩G S⇩G' :: "'f gclause fset × 'f gclause fset × ('f gliteral × 'f gclause option) list × 'f gclause option" and
S S' :: "('f, 'v) SCL_FOL.state"
defines
"N ≡ cls_of_gcls |`| N⇩G" and
"β ≡ term_of_gterm β⇩G" and
"S ≡ state_of_gstate S⇩G" and
"S' ≡ state_of_gstate S⇩G'"
assumes
ball_le_β⇩G: "∀A⇩G |∈| atms_of_clss N⇩G. A⇩G ≼⇩t β⇩G" and
run: "(ord_res_11 N⇩G)⇧*⇧* ({||}, {||}, [], None) S⇩G" and
step: "ord_res_11 N⇩G S⇩G S⇩G'"
shows
"scl_fol.regular_scl N β S S'"
proof -
have "ord_res_11_invars N⇩G ({||}, {||}, [], None)"
using ord_res_11_invars_initial_state .
hence "ord_res_11_invars N⇩G S⇩G"
using run rtranclp_ord_res_11_preserves_invars by metis
obtain U⇩G ℱ Γ⇩G 𝒞⇩G where S⇩G_def: "S⇩G = (U⇩G, ℱ, Γ⇩G, 𝒞⇩G)"
by (metis surj_pair)
obtain Γ U 𝒞 where S_def: "S = (Γ, U, 𝒞)"
by (metis surj_pair)
have Γ_def: "Γ = map gtrailelem_of_trailelem Γ⇩G"
using S_def S⇩G_def ‹S ≡ state_of_gstate S⇩G› by simp
have U_def: "U = cls_of_gcls |`| U⇩G"
using S_def S⇩G_def ‹S ≡ state_of_gstate S⇩G› by simp
have 𝒞_def: "𝒞 = map_option (λC⇩G. (cls_of_gcls C⇩G, Var)) 𝒞⇩G"
using S_def S⇩G_def ‹S ≡ state_of_gstate S⇩G› by simp
obtain ℱ' U⇩G' :: "'f gclause fset" and Γ⇩G' :: "_ list" and 𝒞⇩G' :: "_ option" where
S⇩G'_def: "S⇩G' = (U⇩G', ℱ', Γ⇩G', 𝒞⇩G')"
by (metis surj_pair)
obtain Γ' :: "_ list" and U' :: "_ fset" and 𝒞' :: "_ option" where
S'_def: "S' = (Γ', U', 𝒞')"
by (metis surj_pair)
have Γ'_def: "Γ' = map gtrailelem_of_trailelem Γ⇩G'"
using S'_def S⇩G'_def ‹S' ≡ state_of_gstate S⇩G'› by simp
have U'_def: "U' = cls_of_gcls |`| U⇩G'"
using S'_def S⇩G'_def ‹S' ≡ state_of_gstate S⇩G'› by simp
have 𝒞'_def: "𝒞' = map_option (λC⇩G. (cls_of_gcls C⇩G, Var)) 𝒞⇩G'"
using S'_def S⇩G'_def ‹S' ≡ state_of_gstate S⇩G'› by simp
have "atms_of_clss U⇩G |⊆| atms_of_clss N⇩G"
using ‹ord_res_11_invars N⇩G S⇩G›[unfolded S⇩G_def]
unfolding ord_res_11_invars.simps by simp
have "atms_of_clss (N⇩G |∪| U⇩G) = atms_of_clss N⇩G |∪| atms_of_clss U⇩G"
by (simp add: atms_of_clss_def fimage_funion)
also have "… = atms_of_clss N⇩G"
using ‹atms_of_clss U⇩G |⊆| atms_of_clss N⇩G› by auto
finally have "atms_of_clss (N⇩G |∪| U⇩G) = atms_of_clss N⇩G" .
have clauses_in_ℱ_have_pos_max_lit: "∀C|∈|ℱ. ∃L. is_pos L ∧ ord_res.is_maximal_lit L C"
using ‹ord_res_11_invars N⇩G S⇩G›[unfolded S⇩G_def ord_res_11_invars.simps]
by simp
have nex_conflict_if_nbex_trail_false:
"¬ fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls Γ⇩G) ⟹ ¬ Ex (scl_fol.conflict N β S)"
proof (elim contrapos_nn exE)
fix x :: "('f, 'v) state"
assume "scl_fol.conflict N β S x"
hence "fBex (N⇩G |∪| U⇩G) (trail_false_cls Γ⇩G)"
unfolding S_def
proof (cases N β "(Γ, U, 𝒞)" x rule: scl_fol.conflict.cases)
case (conflictI D γ)
obtain D⇩G where "D⇩G |∈| N⇩G |∪| U⇩G" and D_def: "D = cls_of_gcls D⇩G"
using ‹D |∈| N |∪| U›
unfolding N_def U_def by blast
moreover have "trail_false_cls Γ⇩G D⇩G"
proof -
have "is_ground_cls D"
using ‹D = cls_of_gcls D⇩G› by simp
hence "D ⋅ γ = D"
by simp
hence "trail_false_cls Γ D"
using conflictI
unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
by argo
thus ?thesis
unfolding Γ_def D_def
unfolding trail_false_cls_nonground_iff_trail_false_cls_ground .
qed
ultimately show ?thesis by metis
qed
thus "fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls Γ⇩G)"
unfolding bex_trail_false_cls_simp .
qed
have nex_conflict_if_alread_in_conflict: "𝒞⇩G = Some C⇩G ⟹ ¬ Ex (scl_fol.conflict N β S)" for C⇩G
unfolding S_def 𝒞_def by (simp add: scl_fol.conflict.simps)
have nex_conflict_if_no_clause_could_propagate_comp:
"¬ Ex (scl_fol.conflict N β ((lit_of_glit L⇩G, None) # Γ, U, 𝒞))"
if
nex_false_clause_wrt_Γ⇩G: "¬ fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls Γ⇩G)" and
ball_lt_atm_L⇩G: "∀x |∈| trail_atms Γ⇩G. x ≺⇩t atm_of L⇩G" and
nex_clause_that_propagate: "¬ (∃C |∈| iefac ℱ |`| (N⇩G |∪| U⇩G).
clause_could_propagate Γ⇩G C (- L⇩G))"
for L⇩G
proof (intro notI, elim exE)
fix S'' :: "('f, 'v) SCL_FOL.state"
assume "scl_fol.conflict N β ((lit_of_glit L⇩G, None) # Γ, U, 𝒞) S''"
thus "False"
proof (cases N β "((lit_of_glit L⇩G, None) # Γ, U, 𝒞)" S'' rule: scl_fol.conflict.cases)
case (conflictI D γ)
obtain D⇩G where "D⇩G |∈| N⇩G |∪| U⇩G" and D_def: "D = cls_of_gcls D⇩G"
using ‹D |∈| N |∪| U› N_def U_def by blast
have "(lit_of_glit L⇩G :: ('f, 'v) term literal, None) # Γ =
(map gtrailelem_of_trailelem ((L⇩G, None) # Γ⇩G) :: ('f, 'v) SCL_FOL.trail)"
by (simp add: Γ_def gtrailelem_of_trailelem_def)
moreover have "D ⋅ γ = cls_of_gcls D⇩G"
unfolding D_def by simp
ultimately have "trail_false_cls
(map gtrailelem_of_trailelem ((L⇩G, None) # Γ⇩G) :: ('f, 'v) SCL_FOL.trail) (cls_of_gcls D⇩G)"
using ‹SCL_FOL.trail_false_cls ((lit_of_glit L⇩G, None) # Γ) (D ⋅ γ)›
unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
by metis
hence "trail_false_cls ((L⇩G, None) # Γ⇩G) D⇩G"
using trail_false_cls_nonground_iff_trail_false_cls_ground by blast
hence "trail_false_cls Γ⇩G {#K⇩G ∈# D⇩G. K⇩G ≠ - L⇩G#}"
unfolding trail_false_cls_def trail_false_lit_def
by auto
moreover have "ord_res.is_maximal_lit (- L⇩G) D⇩G"
unfolding linorder_lit.is_maximal_in_mset_iff
proof (intro conjI ballI impI)
show "- L⇩G ∈# D⇩G"
using ‹D⇩G |∈| N⇩G |∪| U⇩G› ‹trail_false_cls ((L⇩G, None) # Γ⇩G) D⇩G› subtrail_falseI
nex_false_clause_wrt_Γ⇩G
unfolding bex_trail_false_cls_simp
by blast
next
fix K⇩G assume "K⇩G ∈# D⇩G" and "K⇩G ≠ - L⇩G"
hence "trail_false_lit Γ⇩G K⇩G"
using ‹trail_false_cls Γ⇩G {#K⇩G ∈# D⇩G. K⇩G ≠ - L⇩G#}›
unfolding trail_false_cls_def by simp
hence "trail_defined_lit Γ⇩G K⇩G"
by (simp add: trail_defined_lit_iff_true_or_false)
hence "atm_of K⇩G |∈| trail_atms Γ⇩G"
unfolding trail_defined_lit_iff_trail_defined_atm .
hence "atm_of K⇩G ≺⇩t atm_of L⇩G"
using ball_lt_atm_L⇩G by metis
hence "K⇩G ≺⇩l - L⇩G"
by (cases L⇩G; cases K⇩G) simp_all
thus "¬ - L⇩G ≺⇩l K⇩G"
by order
qed
moreover have "¬ trail_defined_lit Γ⇩G (- L⇩G)"
by (metis atm_of_uminus linorder_trm.less_irrefl that(2)
trail_defined_lit_iff_trail_defined_atm)
ultimately have "clause_could_propagate Γ⇩G D⇩G (- L⇩G)"
unfolding clause_could_propagate_def by argo
hence "∃C|∈| N⇩G |∪| U⇩G. clause_could_propagate Γ⇩G C (- L⇩G)"
using ‹D⇩G |∈| N⇩G |∪| U⇩G› by metis
hence False
using nex_clause_that_propagate
unfolding bex_clause_could_propagate_simp
by contradiction
thus ?thesis .
qed
qed
show ?thesis
using step unfolding S⇩G_def S⇩G'_def
proof (cases N⇩G "(U⇩G, ℱ, Γ⇩G, 𝒞⇩G)" "(U⇩G', ℱ', Γ⇩G', 𝒞⇩G')" rule: ord_res_11.cases)
case step_hyps: (decide_neg A⇩G)
define A :: "('f, 'v) term" where
"A = term_of_gterm A⇩G"
let ?f = "gtrailelem_of_trailelem"
have "Γ' = map ?f Γ⇩G'"
unfolding Γ'_def ..
also have "… = map ?f ((Neg A⇩G, None) # Γ⇩G)"
unfolding ‹Γ⇩G' = (Neg A⇩G, None) # Γ⇩G› ..
also have "… = ?f (Neg A⇩G, None) # map ?f Γ⇩G"
unfolding list.map ..
also have "… = ?f (Neg A⇩G, None) # Γ"
unfolding Γ_def ..
also have "… = (lit_of_glit (Neg A⇩G), None) # Γ"
unfolding gtrailelem_of_trailelem_def prod.case option.map ..
also have "… = (Neg (term_of_gterm A⇩G), None) # Γ"
unfolding lit_of_glit_def literal.map ..
also have "… = (Neg A, None) # Γ"
unfolding A_def ..
finally have "Γ' = decide_lit (Neg A) # Γ"
unfolding decide_lit_def .
have "U' = U"
unfolding U'_def ‹U⇩G' = U⇩G› U_def ..
have "¬ Ex (scl_fol.conflict N β S)"
using ‹¬ fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls Γ⇩G)› nex_conflict_if_nbex_trail_false
by metis
moreover have "scl_fol.reasonable_scl N β S S'"
unfolding scl_fol.reasonable_scl_def
proof (intro conjI impI notI ; (elim exE) ?)
have "scl_fol.decide N β S S'"
unfolding S_def S'_def ‹U' = U› 𝒞_def 𝒞'_def ‹𝒞⇩G = None› ‹𝒞⇩G' = None› option.map
proof (rule scl_fol.decideI')
show "is_ground_lit (Neg A ⋅l Var)"
by (simp add: A_def)
next
have "∀x |∈| trail_atms Γ⇩G. x ≺⇩t A⇩G"
using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
hence "A⇩G |∉| trail_atms Γ⇩G"
by blast
hence "A⇩G ∉ atm_of ` fst ` set Γ⇩G"
unfolding fset_trail_atms .
hence "term_of_gterm A⇩G ∉ term_of_gterm ` atm_of ` fst ` set Γ⇩G"
using inj_image_mem_iff inj_term_of_gterm by metis
hence "term_of_gterm A⇩G ∉ set (map (λx. term_of_gterm (atm_of (fst x))) Γ⇩G)"
unfolding image_set list.map_comp comp_def .
hence "A ∉ set (map (λx. atm_of (lit_of_glit (fst x))) Γ⇩G)"
unfolding A_def atm_of_lit_of_glit_conv .
hence "A ∉ atm_of ` fst ` set Γ"
unfolding image_set list.map_comp comp_def Γ_def gtrailelem_of_trailelem_def
fst_case_prod_simp .
hence "A |∉| trail_atms Γ"
unfolding fset_trail_atms .
hence "¬ trail_defined_lit Γ (Neg A ⋅l Var)"
by (simp add: trail_defined_lit_iff_trail_defined_atm)
thus "¬ SCL_FOL.trail_defined_lit Γ (Neg A ⋅l Var)"
by (simp add: SCL_FOL.trail_defined_lit_def trail_defined_lit_def)
next
have "A⇩G |∈| atms_of_clss (N⇩G |∪| U⇩G)"
using step_hyps linorder_trm.is_least_in_ffilter_iff by blast
hence "A⇩G |∈| atms_of_clss N⇩G"
unfolding ‹atms_of_clss (N⇩G |∪| U⇩G) = atms_of_clss N⇩G› .
hence "A⇩G ≼⇩t β⇩G"
using ball_le_β⇩G by metis
moreover have "gterm_of_term A = A⇩G"
by (simp add: A_def)
moreover have "gterm_of_term β = β⇩G"
by (simp add: β_def)
ultimately have "gterm_of_term A ≼⇩t gterm_of_term β"
by argo
thus "less_B⇧=⇧= (atm_of (Neg A) ⋅a Var) β"
using inj_term_of_gterm[THEN injD]
by (auto simp: less_B_def A_def β_def)
next
show "Γ' = trail_decide Γ (Neg A ⋅l Var)"
using ‹Γ' = decide_lit (Neg A) # Γ›
unfolding subst_lit_id_subst .
qed
thus "scl_fol.scl N β S S'"
unfolding scl_fol.scl_def by argo
next
fix S'' :: "('f, 'v) SCL_FOL.state"
assume "scl_fol.conflict N β S' S''"
moreover have "∄S''. scl_fol.conflict N β S' S''"
proof -
have "¬ Ex (scl_fol.conflict N β ((lit_of_glit (Neg A⇩G), None) # Γ, U, 𝒞))"
proof (rule nex_conflict_if_no_clause_could_propagate_comp)
show "¬ fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls Γ⇩G)"
using step_hyps by argo
next
show "∀x |∈| trail_atms Γ⇩G. x ≺⇩t atm_of (Neg A⇩G)"
unfolding literal.sel
using step_hyps linorder_trm.is_least_in_fset_iff by simp
next
show "¬ (∃C |∈| iefac ℱ |`| (N⇩G |∪| U⇩G). clause_could_propagate Γ⇩G C (- Neg A⇩G))"
using step_hyps by simp
qed
moreover have "lit_of_glit (Neg A⇩G) = Neg A"
unfolding A_def lit_of_glit_def literal.map ..
ultimately show ?thesis
unfolding S'_def ‹Γ' = decide_lit (Neg A) # Γ› decide_lit_def
using 𝒞'_def 𝒞_def ‹U' = U› step_hyps(1,4) by argo
qed
ultimately show False
by metis
qed
ultimately show ?thesis
unfolding scl_fol.regular_scl_def by argo
next
case step_hyps: (decide_pos A⇩G)
define A :: "('f, 'v) term" where
"A = term_of_gterm A⇩G"
let ?f = "gtrailelem_of_trailelem"
have "Γ' = map ?f Γ⇩G'"
unfolding Γ'_def ..
also have "… = map ?f ((Pos A⇩G, None) # Γ⇩G)"
unfolding ‹Γ⇩G' = (Pos A⇩G, None) # Γ⇩G› ..
also have "… = ?f (Pos A⇩G, None) # map ?f Γ⇩G"
unfolding list.map ..
also have "… = ?f (Pos A⇩G, None) # Γ"
unfolding Γ_def ..
also have "… = (lit_of_glit (Pos A⇩G), None) # Γ"
unfolding prod.case option.map gtrailelem_of_trailelem_def ..
also have "… = (Pos (term_of_gterm A⇩G), None) # Γ"
unfolding lit_of_glit_def literal.map ..
also have "… = (Pos A, None) # Γ"
unfolding A_def ..
finally have "Γ' = decide_lit (Pos A) # Γ"
unfolding decide_lit_def .
have "U' = U"
unfolding U'_def ‹U⇩G' = U⇩G› U_def ..
have "¬ Ex (scl_fol.conflict N β S)"
using step_hyps nex_conflict_if_nbex_trail_false by metis
moreover have "scl_fol.reasonable_scl N β S S'"
unfolding scl_fol.reasonable_scl_def
proof (intro conjI impI notI ; (elim exE) ?)
have "scl_fol.decide N β S S'"
unfolding S_def S'_def ‹U' = U› 𝒞_def 𝒞'_def ‹𝒞⇩G = None› ‹𝒞⇩G' = None› option.map
proof (rule scl_fol.decideI')
show "is_ground_lit (Pos A ⋅l Var)"
by (simp add: A_def)
next
have "∀x |∈| trail_atms Γ⇩G. x ≺⇩t A⇩G"
using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
hence "A⇩G |∉| trail_atms Γ⇩G"
by blast
hence "A⇩G ∉ atm_of ` fst ` set Γ⇩G"
unfolding fset_trail_atms .
hence "term_of_gterm A⇩G ∉ term_of_gterm ` atm_of ` fst ` set Γ⇩G"
using inj_image_mem_iff inj_term_of_gterm by metis
hence "term_of_gterm A⇩G ∉ set (map (λx. term_of_gterm (atm_of (fst x))) Γ⇩G)"
unfolding image_set list.map_comp comp_def .
hence "A ∉ set (map (λx. atm_of (lit_of_glit (fst x))) Γ⇩G)"
unfolding A_def atm_of_lit_of_glit_conv .
hence "A ∉ atm_of ` fst ` set Γ"
unfolding image_set list.map_comp comp_def Γ_def gtrailelem_of_trailelem_def
fst_case_prod_simp .
hence "A |∉| trail_atms Γ"
unfolding fset_trail_atms .
hence "¬ trail_defined_lit Γ (Pos A ⋅l Var)"
by (simp add: trail_defined_lit_iff_trail_defined_atm)
thus "¬ SCL_FOL.trail_defined_lit Γ (Pos A ⋅l Var)"
by (simp add: SCL_FOL.trail_defined_lit_def trail_defined_lit_def)
next
have "A⇩G |∈| atms_of_clss (N⇩G |∪| U⇩G)"
using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
hence "A⇩G |∈| atms_of_clss N⇩G"
unfolding ‹atms_of_clss (N⇩G |∪| U⇩G) = atms_of_clss N⇩G› .
hence "A⇩G ≼⇩t β⇩G"
using ball_le_β⇩G by metis
moreover have "gterm_of_term A = A⇩G"
by (simp add: A_def)
moreover have "gterm_of_term β = β⇩G"
by (simp add: β_def)
ultimately have "gterm_of_term A ≼⇩t gterm_of_term β"
by argo
thus "less_B⇧=⇧= (atm_of (Pos A) ⋅a Var) β"
using inj_term_of_gterm[THEN injD]
by (auto simp: less_B_def A_def β_def)
next
show "Γ' = trail_decide Γ (Pos A ⋅l Var)"
using ‹Γ' = decide_lit (Pos A) # Γ›
unfolding subst_lit_id_subst .
qed
thus "scl_fol.scl N β S S'"
unfolding scl_fol.scl_def by argo
next
fix S'' :: "('f, 'v) SCL_FOL.state"
assume "scl_fol.conflict N β S' S''"
moreover have "∄S''. scl_fol.conflict N β S' S''"
proof -
have "¬ Ex (scl_fol.conflict N β ((lit_of_glit (Pos A⇩G), None) # Γ, U, 𝒞))"
proof (rule nex_conflict_if_no_clause_could_propagate_comp)
show "¬ fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls Γ⇩G)"
using step_hyps by argo
next
show "∀x|∈| trail_atms Γ⇩G. x ≺⇩t atm_of (Pos A⇩G)"
unfolding literal.sel
using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
next
have "clause_could_propagate Γ⇩G C (Neg A⇩G) ⟹ trail_false_cls Γ⇩G' C" for C
unfolding ‹Γ⇩G' = (Pos A⇩G, None) # Γ⇩G›
using trail_false_if_could_have_propagatated by fastforce
thus "¬ (∃C |∈| iefac ℱ |`| (N⇩G |∪| U⇩G). clause_could_propagate Γ⇩G C (- Pos A⇩G))"
unfolding uminus_Pos
using step_hyps by metis
qed
moreover have "lit_of_glit (Pos A⇩G) = Pos A"
unfolding A_def lit_of_glit_def literal.map ..
ultimately show ?thesis
unfolding S'_def ‹Γ' = decide_lit (Pos A) # Γ› decide_lit_def
using 𝒞'_def 𝒞_def ‹U' = U› step_hyps(1) step_hyps(3) by argo
qed
ultimately show False by metis
qed
ultimately show ?thesis
unfolding scl_fol.regular_scl_def by argo
next
case step_hyps: (propagate A⇩G C⇩G)
have "C⇩G |∈| iefac ℱ |`| (N⇩G |∪| U⇩G)" and C⇩G_prop: "clause_could_propagate Γ⇩G C⇩G (Pos A⇩G)"
using step_hyps linorder_cls.is_least_in_fset_iff by simp_all
define A :: "('f, 'v) term" where
"A = term_of_gterm A⇩G"
define C :: "('f, 'v) term clause" where
"C = cls_of_gcls C⇩G"
have "ord_res.is_maximal_lit (Pos A⇩G) C⇩G" and "trail_false_cls Γ⇩G {#K ∈# C⇩G. K ≠ Pos A⇩G#}"
using ‹clause_could_propagate Γ⇩G C⇩G (Pos A⇩G)›
unfolding clause_could_propagate_def by metis+
then obtain C⇩G' where "C⇩G = add_mset (Pos A⇩G) C⇩G'"
by (metis linorder_lit.is_maximal_in_mset_iff mset_add)
define C' :: "('f, 'v) term clause" where
"C' = cls_of_gcls C⇩G'"
let ?f = "gtrailelem_of_trailelem"
have "Γ' = map ?f Γ⇩G'"
unfolding Γ'_def ..
also have "… = map ?f ((Pos A⇩G, Some (efac C⇩G)) # Γ⇩G)"
unfolding ‹Γ⇩G' = (Pos A⇩G, Some _) # Γ⇩G› ..
also have "… = ?f (Pos A⇩G, Some (efac C⇩G)) # map ?f Γ⇩G"
unfolding list.map ..
also have "… = ?f (Pos A⇩G, Some (efac C⇩G)) # Γ"
unfolding Γ_def ..
also have "… = (lit_of_glit (Pos A⇩G),
Some (cls_of_gcls {#K ∈# efac C⇩G. K ≠ Pos A⇩G#}, lit_of_glit (Pos A⇩G), Var)) # Γ"
unfolding gtrailelem_of_trailelem_def prod.case option.map ..
also have "… = (lit_of_glit (Pos A⇩G),
Some (cls_of_gcls {#K ∈# add_mset (Pos A⇩G) {#K ∈# C⇩G. K ≠ Pos A⇩G#}. K ≠ Pos A⇩G#},
lit_of_glit (Pos A⇩G), Var)) # Γ"
proof -
have "is_pos (Pos A⇩G)"
by simp
moreover have "linorder_lit.is_maximal_in_mset C⇩G (Pos A⇩G)"
using C⇩G_prop unfolding clause_could_propagate_def by argo
ultimately show ?thesis
using efac_spec_if_pos_lit_is_maximal by metis
qed
also have "… = (lit_of_glit (Pos A⇩G),
Some (cls_of_gcls {#K ∈# C⇩G. K ≠ Pos A⇩G#}, lit_of_glit (Pos A⇩G), Var)) # Γ"
by (simp add: filter_filter_mset)
also have "… = (Pos A, Some (cls_of_gcls {#K ∈# C⇩G. K ≠ Pos A⇩G#}, Pos A, Var)) # Γ"
by (simp add: A_def lit_of_glit_def)
also have "… = (Pos A, Some (cls_of_gcls {#L ∈# C⇩G. lit_of_glit L ≠ Pos A#}, Pos A, Var)) # Γ"
by (metis A_def glit_of_lit_lit_of_glit lit_of_glit_def literal.simps(9))
also have "… = (Pos A, Some ({#L ∈# cls_of_gcls C⇩G. L ≠ Pos A#}, Pos A, Var)) # Γ"
unfolding cls_of_gcls_def
unfolding image_mset_filter_mset_swap[of "lit_of_glit" "λL. L ≠ Pos A" C⇩G]
unfolding cls_of_gcls_def[symmetric] ..
also have "… = (Pos A ⋅l Var, Some ({#L ∈# cls_of_gcls C⇩G. L ≠ Pos A#}, Pos A, Var)) # Γ"
by simp
also have "… = (Pos A ⋅l Var, Some ({#L ∈# C. L ≠ Pos A#}, Pos A, Var)) # Γ"
unfolding C_def ..
finally have "Γ' = propagate_lit (Pos A) {#L ∈# C. L ≠ Pos A#} Var # Γ"
unfolding propagate_lit_def .
have "U' = U"
unfolding U'_def ‹U⇩G' = U⇩G› U_def ..
obtain C⇩G⇩0 where "C⇩G⇩0 |∈| N⇩G |∪| U⇩G" and "C⇩G = iefac ℱ C⇩G⇩0"
using ‹C⇩G |∈| iefac ℱ |`| (N⇩G |∪| U⇩G)› by blast
define C⇩0 :: "('f, 'v) term clause" where
"C⇩0 = cls_of_gcls C⇩G⇩0"
have "ord_res.is_maximal_lit (Pos A⇩G) C⇩G⇩0"
using ‹ord_res.is_maximal_lit (Pos A⇩G) C⇩G› ‹C⇩G = iefac ℱ C⇩G⇩0›
by (metis iefac_def linorder_lit.is_maximal_in_mset_iff set_mset_efac)
have "¬ Ex (scl_fol.conflict N β S)"
using step_hyps nex_conflict_if_nbex_trail_false by metis
moreover have "scl_fol.reasonable_scl N β S S'"
unfolding scl_fol.reasonable_scl_def
proof (intro conjI impI notI ; (elim exE) ?)
have "scl_fol.propagate N β S S'"
unfolding S_def S'_def ‹U' = U› 𝒞_def 𝒞'_def ‹𝒞⇩G = None› ‹𝒞⇩G' = None› option.map
proof (rule scl_fol.propagateI')
show "C⇩0 |∈| N |∪| U"
unfolding C⇩0_def N_def U_def
using ‹C⇩G⇩0 |∈| N⇩G |∪| U⇩G›
by blast
next
show "is_ground_cls (C⇩0 ⋅ Var)"
by (simp add: C⇩0_def)
have "Pos A ∈# C⇩0"
unfolding A_def C⇩0_def
by (metis (no_types, lifting) ‹C⇩G = iefac ℱ C⇩G⇩0› ‹ord_res.is_maximal_lit (Pos A⇩G) C⇩G›
cls_of_gcls_def iefac_def image_eqI linorder_lit.is_maximal_in_mset_iff
lit_of_glit_def literal.map(1) multiset.set_map set_mset_efac)
then show "C⇩0 = add_mset (Pos A) (C⇩0 - {#Pos A#})"
by simp
have "A⇩G |∈| atms_of_clss (N⇩G |∪| U⇩G)"
using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
hence "A⇩G |∈| atms_of_clss N⇩G"
unfolding ‹atms_of_clss (N⇩G |∪| U⇩G) = atms_of_clss N⇩G› .
hence "A⇩G ≼⇩t β⇩G"
using ball_le_β⇩G by metis
moreover have "gterm_of_term A = A⇩G"
by (simp add: A_def)
moreover have "gterm_of_term β = β⇩G"
by (simp add: β_def)
ultimately have "gterm_of_term A ≼⇩t gterm_of_term β"
by argo
hence "less_B⇧=⇧= A β"
by (auto simp: less_B_def A_def β_def)
show "∀K ∈# C⇩0 ⋅ Var. less_B⇧=⇧= (atm_of K) β"
unfolding subst_cls_id_subst
proof (intro ballI)
fix K :: "('f, 'v) Term.term literal"
assume "K ∈# C⇩0"
then obtain K⇩G where "K⇩G ∈# C⇩G⇩0" and K_def: "K = lit_of_glit K⇩G"
unfolding C⇩0_def cls_of_gcls_def by blast
have "K⇩G ≼⇩l Pos A⇩G"
using ‹ord_res.is_maximal_lit (Pos A⇩G) C⇩G⇩0› ‹K⇩G ∈# C⇩G⇩0›
unfolding linorder_lit.is_maximal_in_mset_iff by fastforce
hence "atm_of K⇩G ≼⇩t A⇩G"
by (metis literal.collapse(1) literal.collapse(2) literal.sel(1)
ord_res.less_lit_simps(1) ord_res.less_lit_simps(4) reflclp_iff)
hence "less_B⇧=⇧= (atm_of K) A"
by (auto simp: less_B_def K_def A_def atm_of_lit_of_glit_conv)
then show "less_B⇧=⇧= (atm_of K) β"
using ‹less_B⇧=⇧= A β› by order
qed
show "{#K ∈# C⇩0. K ≠ Pos A#} = {#K ∈# remove1_mset (Pos A) C⇩0. K ⋅l Var ≠ Pos A ⋅l Var#}"
by simp
show "{#K ∈# remove1_mset (Pos A) C⇩0. K = Pos A#} =
{#K ∈# remove1_mset (Pos A) C⇩0. K ⋅l Var = Pos A ⋅l Var#}"
by simp
have "trail_false_cls Γ⇩G ({#K⇩G ∈# C⇩G⇩0. K⇩G ≠ Pos A⇩G#})"
by (smt (verit, ccfv_threshold) ‹C⇩G = iefac ℱ C⇩G⇩0›
‹trail_false_cls Γ⇩G {#K ∈# C⇩G. K ≠ Pos A⇩G#}› iefac_def mem_Collect_eq set_mset_efac
set_mset_filter trail_false_cls_def)
moreover have "(cls_of_gcls {#K⇩G ∈# C⇩G⇩0. K⇩G ≠ Pos A⇩G#} :: ('f, 'v) term clause) =
{#K ∈# C⇩0. K ≠ Pos A#}"
by (smt (verit) A_def C⇩0_def cls_of_gcls_def filter_mset_cong0 glit_of_lit_lit_of_glit
image_mset_filter_mset_swap lit_of_glit_def literal.map(1))
ultimately have "trail_false_cls Γ ({#K ∈# C⇩0. K ≠ Pos A#} ⋅ Var)"
unfolding subst_cls_id_subst
using trail_false_cls_nonground_iff_trail_false_cls_ground[THEN iffD2]
by (metis Γ_def)
thus "SCL_FOL.trail_false_cls Γ ({#K ∈# C⇩0. K ≠ Pos A#} ⋅ Var)"
unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
by argo
have "∀x |∈| trail_atms Γ⇩G. x ≺⇩t A⇩G"
using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
hence "A⇩G |∉| trail_atms Γ⇩G"
by blast
hence "A⇩G ∉ atm_of ` fst ` set Γ⇩G"
unfolding fset_trail_atms .
hence "term_of_gterm A⇩G ∉ term_of_gterm ` atm_of ` fst ` set Γ⇩G"
using inj_image_mem_iff inj_term_of_gterm by metis
hence "term_of_gterm A⇩G ∉ set (map (λx. term_of_gterm (atm_of (fst x))) Γ⇩G)"
unfolding image_set list.map_comp comp_def .
hence "A ∉ set (map (λx. atm_of (lit_of_glit (fst x))) Γ⇩G)"
unfolding A_def atm_of_lit_of_glit_conv .
hence "A ∉ atm_of ` fst ` set Γ"
unfolding image_set list.map_comp comp_def Γ_def gtrailelem_of_trailelem_def
fst_case_prod_simp .
hence "A |∉| trail_atms Γ"
unfolding fset_trail_atms .
hence "¬ trail_defined_lit Γ (Pos A ⋅l Var)"
by (simp add: trail_defined_lit_iff_trail_defined_atm)
thus "¬ SCL_FOL.trail_defined_lit Γ (Pos A ⋅l Var)"
by (simp add: SCL_FOL.trail_defined_lit_def trail_defined_lit_def)
have "set_mset (add_mset (Pos A) {#K ∈# remove1_mset (Pos A) C⇩0. K = Pos A#}) =
{Pos A}"
by (smt (verit) Collect_cong insert_compr mem_Collect_eq set_mset_add_mset_insert
set_mset_filter singletonD)
hence "is_unifier Var (atm_of ` set_mset
(add_mset (Pos A) {#K ∈# remove1_mset (Pos A) C⇩0. K = Pos A#}))"
by (metis (no_types, lifting) finite_imageI finite_set_mset image_empty image_insert
is_unifier_alt singletonD)
hence "is_unifiers Var {atm_of ` set_mset
(add_mset (Pos A) {#K ∈# remove1_mset (Pos A) C⇩0. K = Pos A#})}"
unfolding SCL_FOL.is_unifiers_def by simp
thus "SCL_FOL.is_imgu Var {atm_of ` set_mset
(add_mset (Pos A) {#K ∈# remove1_mset (Pos A) C⇩0. K = Pos A#})}"
unfolding SCL_FOL.is_imgu_def by simp
have "{#K ∈# C⇩G⇩0. K ≠ Pos A⇩G#} = {#K ∈# C⇩G. K ≠ Pos A⇩G#}"
using ‹ord_res.is_maximal_lit (Pos A⇩G) C⇩G›
using ‹ord_res.is_maximal_lit (Pos A⇩G) C⇩G⇩0›
unfolding ‹C⇩G = iefac ℱ C⇩G⇩0›
by (metis add_mset_remove_trivial efac_spec_if_pos_lit_is_maximal
ex1_efac_eq_factoring_chain factorizable_if_neq_efac iefac_def literal.disc(1))
hence "{#K ∈# C⇩0. K ≠ Pos A#} = {#K ∈# C. K ≠ Pos A#}"
unfolding C⇩0_def C_def
by (smt (verit, ccfv_SIG) A_def cls_of_gcls_def filter_mset_cong0 glit_of_lit_lit_of_glit
image_mset_filter_mset_swap lit_of_glit_def literal.map(1))
thus "Γ' = trail_propagate Γ (Pos A ⋅l Var) ({#K ∈# C⇩0. K ≠ Pos A#} ⋅ Var) Var"
unfolding subst_cls_id_subst subst_lit_id_subst
using ‹Γ' = propagate_lit (Pos A) {#L ∈# C. L ≠ Pos A#} Var # Γ›
by argo
qed
thus "scl_fol.scl N β S S'"
unfolding scl_fol.scl_def by argo
next
fix S'' :: "('f, 'v) SCL_FOL.state"
assume "scl_fol.decide N β S S'"
thus False
unfolding S_def S'_def
proof (cases N β "(Γ, U, 𝒞)" "(Γ', U', 𝒞')" rule: scl_fol.decide.cases)
case (decideI L γ)
show False
using ‹Γ' = decide_lit (L ⋅l γ) # Γ›
using ‹Γ' = propagate_lit (Pos A) {#L ∈# C. L ≠ Pos A#} Var # Γ›
unfolding decide_lit_def propagate_lit_def
by blast
qed
qed
ultimately show ?thesis
unfolding scl_fol.regular_scl_def by argo
next
case step_hyps: (conflict C⇩G)
have "Γ' = Γ"
unfolding Γ'_def ‹Γ⇩G' = Γ⇩G› Γ_def ..
have "U' = U"
unfolding U'_def ‹U⇩G' = U⇩G› U_def ..
have
C⇩G_in: "C⇩G |∈| iefac ℱ |`| (N⇩G |∪| U⇩G)" and
C⇩G_false: "trail_false_cls Γ⇩G C⇩G" and
C⇩G_lt: "∀E⇩G |∈| iefac ℱ |`| (N⇩G |∪| U⇩G). E⇩G ≠ C⇩G ⟶ trail_false_cls Γ⇩G E⇩G ⟶ C⇩G ≺⇩c E⇩G"
using ‹linorder_cls.is_least_in_fset _ C⇩G›
unfolding atomize_conj linorder_cls.is_least_in_ffilter_iff by argo
have "∄L. is_pos L ∧ ord_res.is_maximal_lit L C⇩G"
proof (rule notI , elim exE conjE)
fix L :: "'f gliteral"
assume "is_pos L" and C⇩G_max_lit: "ord_res.is_maximal_lit L C⇩G"
have "{#} |∉| iefac ℱ |`| (N⇩G |∪| U⇩G)"
using C⇩G_lt
by (metis (full_types) C⇩G_max_lit bot_fset.rep_eq fBex_fempty linorder_cls.leD
linorder_lit.is_maximal_in_mset_iff mempty_lesseq_cls set_mset_empty
trail_false_cls_mempty)
have "trail_false_lit Γ⇩G L"
using C⇩G_max_lit C⇩G_false
unfolding linorder_lit.is_maximal_in_mset_iff trail_false_cls_def
by metis
then obtain Ln Γ⇩G⇩0 where "Γ⇩G = Ln # Γ⇩G⇩0"
unfolding trail_false_lit_def
by (metis (no_types) List.insert_def image_iff insert_Nil neq_Nil_conv)
moreover have
AAA: "∀x xs. Γ⇩G = x # xs ⟶
((snd x ≠ None) ⟷ fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls (x # xs))) ∧
(snd x ≠ None ⟶ is_pos (fst x)) ∧
(∀x ∈ set xs. snd x = None)" and
BBB: "(∀Γ⇩1 Ln Γ⇩0. Γ⇩G = Γ⇩1 @ Ln # Γ⇩0 ⟶ snd Ln = None ⟶
¬ fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls (Ln # Γ⇩0)))" and
Γ⇩G_sorted: "sorted_wrt (λx y. atm_of (fst y) ≺⇩t atm_of (fst x)) Γ⇩G"
using ‹ord_res_11_invars N⇩G S⇩G›[unfolded S⇩G_def ord_res_11_invars.simps
ord_res_10_invars.simps]
by simp_all
moreover have "fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls Γ⇩G)"
using C⇩G_in C⇩G_false by metis
ultimately have "snd Ln ≠ None" and "is_pos (fst Ln)" and "∀x ∈ set Γ⇩G⇩0. snd x = None"
unfolding atomize_conj by metis
have "¬ fBex (iefac ℱ |`| (N⇩G |∪| U⇩G)) (trail_false_cls Γ⇩G⇩0)"
proof (cases Γ⇩G⇩0)
case Nil
then show ?thesis
using ‹{#} |∉| iefac ℱ |`| (N⇩G |∪| U⇩G)›
unfolding trail_false_cls_def trail_false_lit_def
by simp
next
case (Cons x xs)
then show ?thesis
using ‹Γ⇩G = Ln # Γ⇩G⇩0›
using ‹∀x ∈ set Γ⇩G⇩0. snd x = None›
using BBB[rule_format, of "[Ln]", unfolded append_Cons append_Nil]
by simp
qed
hence "¬ trail_false_cls Γ⇩G⇩0 C⇩G"
using C⇩G_in by metis
hence "fst Ln = - L"
using C⇩G_false C⇩G_max_lit Γ⇩G_sorted[unfolded ‹Γ⇩G = Ln # Γ⇩G⇩0› sorted_wrt.simps]
by (smt (verit, ccfv_SIG) Neg_atm_of_iff ‹Γ⇩G = Ln # Γ⇩G⇩0› atm_of_uminus
ord_res.less_lit_simps(4) imageE image_insert insertE
linorder_lit.dual_order.strict_trans linorder_lit.is_maximal_in_mset_iff
linorder_lit.neq_iff linorder_trm.order.irrefl list.simps(15) literal.collapse(1)
ord_res.ground_ordered_resolution_calculus_axioms trail_false_cls_def
trail_false_lit_def)
hence "¬ is_pos L"
using ‹is_pos (fst Ln)›
by (simp add: is_pos_neg_not_is_pos)
thus False
using ‹is_pos L› by contradiction
qed
hence "C⇩G |∈| N⇩G |∪| U⇩G"
proof -
obtain C where "C |∈| N⇩G |∪| U⇩G" and "C⇩G = iefac ℱ C"
using C⇩G_in by blast
hence "C⇩G = C"
using ‹∄L. is_pos L ∧ ord_res.is_maximal_lit L C⇩G›
by (metis clauses_in_ℱ_have_pos_max_lit ex1_efac_eq_factoring_chain iefac_def
ord_res.ground_factorings_preserves_maximal_literal)
thus ?thesis
using ‹C |∈| N⇩G |∪| U⇩G› by simp
qed
have "scl_fol.conflict N β S S'"
unfolding S_def S'_def ‹Γ' = Γ› ‹U' = U› 𝒞_def 𝒞'_def ‹𝒞⇩G = None› ‹𝒞⇩G' = Some C⇩G› option.map
proof (rule scl_fol.conflictI)
show "cls_of_gcls C⇩G |∈| N |∪| U"
unfolding N_def U_def
using ‹C⇩G |∈| N⇩G |∪| U⇩G› by auto
next
show "is_ground_cls (cls_of_gcls C⇩G ⋅ (Var::'v ⇒ ('f, _) Term.term))"
by simp
next
have "trail_false_cls Γ⇩G C⇩G"
using ‹trail_false_cls Γ⇩G C⇩G› .
hence "trail_false_cls Γ (cls_of_gcls C⇩G ⋅ Var)"
unfolding Γ_def subst_cls_id_subst
using trail_false_cls_nonground_iff_trail_false_cls_ground by metis
thus "SCL_FOL.trail_false_cls Γ (cls_of_gcls C⇩G ⋅ Var)"
unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
by argo
qed
thus ?thesis
unfolding scl_fol.regular_scl_def by argo
next
case step_hyps: (skip L⇩G C⇩G n⇩G)
have "Γ = gtrailelem_of_trailelem (L⇩G, n⇩G) # Γ'"
unfolding Γ_def Γ'_def ‹Γ⇩G = (L⇩G, n⇩G) # Γ⇩G'› by simp
have "U' = U"
unfolding U'_def ‹U⇩G' = U⇩G› U_def ..
have "¬ Ex (scl_fol.conflict N β S)"
using ‹𝒞⇩G = Some C⇩G› nex_conflict_if_alread_in_conflict by metis
moreover have "scl_fol.reasonable_scl N β S S'"
unfolding scl_fol.reasonable_scl_def
proof (intro conjI impI notI ; (elim exE) ?)
have "scl_fol.skip N β S S'"
unfolding S_def S'_def ‹U' = U› 𝒞_def 𝒞'_def ‹𝒞⇩G = Some C⇩G› ‹𝒞⇩G' = Some C⇩G› option.map
unfolding ‹Γ = _ # Γ'› gtrailelem_of_trailelem_def prod.case
proof (rule scl_fol.skipI)
have "- lit_of_glit L⇩G = lit_of_glit (- L⇩G)"
by (cases L⇩G) (simp_all add: lit_of_glit_def)
show "- lit_of_glit L⇩G ∉# cls_of_gcls C⇩G ⋅ Var"
unfolding subst_cls_id_subst
unfolding ‹- lit_of_glit L⇩G = lit_of_glit (- L⇩G)›
unfolding cls_of_gcls_def
using ‹- L⇩G ∉# C⇩G› inj_image_mset_mem_iff[OF inj_lit_of_glit]
by metis
qed
thus "scl_fol.scl N β S S'"
unfolding scl_fol.scl_def by argo
next
fix S'' :: "('f, 'v) SCL_FOL.state"
assume "scl_fol.conflict N β S' S''"
moreover have "∄S''. scl_fol.conflict N β S' S''"
unfolding S'_def 𝒞'_def ‹𝒞⇩G' = Some C⇩G› by (simp add: scl_fol.conflict.simps)
ultimately show False
by metis
qed
ultimately show ?thesis
unfolding scl_fol.regular_scl_def by argo
next
case step_hyps: (resolution L⇩G C⇩G Γ⇩G'' D⇩G)
have "Γ' = Γ"
unfolding Γ'_def ‹Γ⇩G' = Γ⇩G› Γ_def ..
have "U' = U"
unfolding U'_def ‹U⇩G' = U⇩G› U_def ..
have "𝒞 = Some (cls_of_gcls D⇩G, Var)"
unfolding 𝒞_def ‹𝒞⇩G = Some D⇩G› option.map ..
hence 𝒞_eq: "𝒞 = Some
(add_mset (- lit_of_glit L⇩G) (remove1_mset (- lit_of_glit L⇩G) (cls_of_gcls D⇩G)), Var)"
by (smt (verit, best) add_mset_remove_trivial atm_of_eq_atm_of atm_of_lit_of_glit_conv
cls_of_gcls_def glit_of_lit_lit_of_glit image_mset_add_mset insert_DiffM step_hyps(7)
uminus_not_id')
have "𝒞' = Some (
remove1_mset (- (lit_of_glit L⇩G)) (cls_of_gcls D⇩G) +
remove1_mset (lit_of_glit L⇩G) (cls_of_gcls C⇩G), Var)"
unfolding 𝒞'_def ‹𝒞⇩G' = Some _› option.map
apply (simp add: cls_of_gcls_def)
by (smt (verit, ccfv_threshold) add_diff_cancel_right' add_mset_add_single atm_of_eq_atm_of
atm_of_lit_of_glit_conv diff_single_trivial glit_of_lit_lit_of_glit
image_mset_remove1_mset_if insert_DiffM is_pos_neg_not_is_pos msed_map_invR)
hence 𝒞'_eq: "𝒞' = Some (
(remove1_mset (- (lit_of_glit L⇩G)) (cls_of_gcls D⇩G) ⋅ Var +
remove1_mset (lit_of_glit L⇩G) (cls_of_gcls C⇩G) ⋅ Var) ⋅ Var, Var)"
by simp
have "linorder_lit.is_greatest_in_mset C⇩G L⇩G"
using ‹ord_res_11_invars N⇩G S⇩G›[unfolded S⇩G_def ‹Γ⇩G = (L⇩G, Some C⇩G) # Γ⇩G''›]
unfolding ord_res_11_invars.simps ord_res_10_invars.simps
by simp
hence "add_mset L⇩G {#y ∈# C⇩G. y ≠ L⇩G#} = C⇩G"
using linorder_lit.explode_greatest_in_mset by metis
hence "C⇩G - {#L⇩G#} = {#K ∈# C⇩G. K ≠ L⇩G#}"
by (metis add_mset_remove_trivial)
hence "cls_of_gcls (C⇩G - {#L⇩G#}) = cls_of_gcls {#K ∈# C⇩G. K ≠ L⇩G#}"
by argo
moreover have "cls_of_gcls (C⇩G - {#L⇩G#}) = cls_of_gcls C⇩G - cls_of_gcls {#L⇩G#}"
unfolding cls_of_gcls_def
proof (rule image_mset_Diff)
show "{#L⇩G#} ⊆# C⇩G"
by (metis ‹ord_res.is_strictly_maximal_lit L⇩G C⇩G› linorder_lit.is_greatest_in_mset_iff
single_subset_iff)
qed
ultimately have "cls_of_gcls C⇩G - {#lit_of_glit L⇩G#} = cls_of_gcls {#K ∈# C⇩G. K ≠ L⇩G#}"
by (metis ‹add_mset L⇩G {#y ∈# C⇩G. y ≠ L⇩G#} = C⇩G› cls_of_gcls_def image_mset_remove1_mset_if
union_single_eq_member)
have "¬ Ex (scl_fol.conflict N β S)"
using ‹𝒞⇩G = Some _› nex_conflict_if_alread_in_conflict by metis
moreover have "scl_fol.reasonable_scl N β S S'"
unfolding scl_fol.reasonable_scl_def
proof (intro conjI impI notI ; (elim exE) ?)
have "scl_fol.resolve N β S S'"
unfolding S_def S'_def ‹Γ' = Γ› ‹U' = U›
unfolding 𝒞_eq 𝒞'_eq
proof (rule scl_fol.resolveI)
show "Γ = trail_propagate (map gtrailelem_of_trailelem Γ⇩G'')
(lit_of_glit L⇩G) (remove1_mset (lit_of_glit L⇩G) (cls_of_gcls C⇩G)) Var"
unfolding Γ_def ‹Γ⇩G = (L⇩G, Some C⇩G) # Γ⇩G''› gtrailelem_of_trailelem_def list.map prod.case
unfolding propagate_lit_def subst_lit_id_subst option.map
unfolding ‹remove1_mset (lit_of_glit L⇩G) (cls_of_gcls C⇩G) = cls_of_gcls {#K ∈# C⇩G. K ≠ L⇩G#}›
by argo
next
show "lit_of_glit L⇩G ⋅l Var = - (- lit_of_glit L⇩G ⋅l Var)"
by simp
next
show "SCL_FOL.is_renaming Var"
by simp
next
show "SCL_FOL.is_renaming Var"
by simp
next
show "
vars_cls (add_mset (- lit_of_glit L⇩G)
(remove1_mset (- lit_of_glit L⇩G) (cls_of_gcls D⇩G)) ⋅ Var) ∩
vars_cls (add_mset (lit_of_glit L⇩G)
(remove1_mset (lit_of_glit L⇩G) (cls_of_gcls C⇩G)) ⋅ Var) =
{}"
by (metis (no_types, lifting) boolean_algebra.conj_zero_right cls_of_gcls_def
diff_single_trivial image_mset_add_mset insert_DiffM subst_cls_id_subst
vars_cls_cls_of_gcls)
next
show "SCL_FOL.is_imgu Var
{{atm_of (- lit_of_glit L⇩G) ⋅a Var, atm_of (lit_of_glit L⇩G) ⋅a Var}}"
by (simp add: SCL_FOL.is_imgu_def SCL_FOL.is_unifiers_def SCL_FOL.is_unifier_def)
next
show "is_grounding_merge Var
(vars_cls
(add_mset (- lit_of_glit L⇩G) (remove1_mset (- lit_of_glit L⇩G) (cls_of_gcls D⇩G)) ⋅ Var))
(rename_subst_domain Var Var)
(vars_cls
(add_mset (lit_of_glit L⇩G) (remove1_mset (lit_of_glit L⇩G) (cls_of_gcls C⇩G)) ⋅ Var))
(rename_subst_domain Var Var)"
by (simp add: is_grounding_merge_def)
qed
thus "scl_fol.scl N β S S'"
unfolding scl_fol.scl_def by argo
next
fix S'' :: "('f, 'v) SCL_FOL.state"
assume "scl_fol.conflict N β S' S''"
moreover have "∄S''. scl_fol.conflict N β S' S''"
unfolding S'_def 𝒞'_def ‹𝒞⇩G' = Some _› by (simp add: scl_fol.conflict.simps)
ultimately show False
by metis
qed
ultimately show ?thesis
unfolding scl_fol.regular_scl_def by argo
next
case step_hyps: (backtrack L⇩G C⇩G)
define K :: "('f, 'v) term literal" where
"K = - lit_of_glit L⇩G"
define D :: "('f, 'v) term clause" where
"D = cls_of_gcls C⇩G - {#K#}"
have "add_mset K D = cls_of_gcls C⇩G"
by (smt (verit, best) D_def K_def add_mset_remove_trivial atm_of_eq_atm_of
atm_of_lit_of_glit_conv cls_of_gcls_def glit_of_lit_lit_of_glit image_mset_add_mset
insert_DiffM step_hyps(6) uminus_not_id')
have "U' = finsert (add_mset K D) U"
unfolding U_def U'_def ‹U⇩G' = finsert C⇩G U⇩G›
by (smt (verit, ccfv_SIG) D_def K_def add_mset_remove_trivial atm_of_eq_atm_of
atm_of_lit_of_glit_conv cls_of_gcls_def fimage_finsert glit_of_lit_lit_of_glit
image_mset_add_mset insert_DiffM step_hyps(6) uminus_not_id')
have "𝒞 = Some (add_mset K D, Var)"
by (smt (verit) D_def K_def 𝒞_def add_mset_remove_trivial atm_of_eq_atm_of
atm_of_lit_of_glit_conv cls_of_gcls_def glit_of_lit_lit_of_glit image_mset_add_mset
insert_DiffM option.map(2) step_hyps(1,6) uminus_not_id')
have "𝒞' = None"
unfolding 𝒞'_def ‹𝒞⇩G' = None› option.map ..
have "¬ Ex (scl_fol.conflict N β S)"
using ‹𝒞⇩G = Some _› nex_conflict_if_alread_in_conflict by metis
moreover have "scl_fol.reasonable_scl N β S S'"
unfolding scl_fol.reasonable_scl_def
proof (intro conjI impI notI ; (elim exE) ?)
have "scl_fol.backtrack N β S S'"
unfolding S_def S'_def
unfolding ‹U' = finsert (add_mset K D) U› ‹𝒞 = Some (add_mset K D, Var)› ‹𝒞' = None›
proof (rule scl_fol.backtrackI)
show "Γ = trail_decide ([] @ Γ') (lit_of_glit L⇩G)"
unfolding append_Nil
unfolding decide_lit_def
unfolding Γ_def ‹Γ⇩G = _ # _› list.map Γ'_def[symmetric]
unfolding gtrailelem_of_trailelem_def prod.case option.map
..
next
show "lit_of_glit L⇩G = - (K ⋅l Var)"
unfolding K_def by simp
next
have "sorted_wrt (λx y. atm_of (fst y) ≺⇩t atm_of (fst x)) Γ⇩G"
using ‹ord_res_11_invars N⇩G S⇩G›[unfolded S⇩G_def]
unfolding ord_res_11_invars.simps ord_res_10_invars.simps
by fast
hence "trail_consistent Γ⇩G"
using trail_consistent_if_sorted_wrt_atoms by metis
hence "¬ trail_defined_lit Γ⇩G' (- L⇩G)"
by (metis trail_consistent.cases atm_of_eq_atm_of list.distinct(1) list.inject
prod.sel(1) step_hyps(5) trail_defined_lit_iff_trail_defined_atm)
hence "¬ trail_false_lit Γ⇩G' (- L⇩G)"
using trail_defined_lit_iff_true_or_false by metis
hence "¬ trail_false_cls Γ⇩G' C⇩G"
using ‹- L⇩G ∈# C⇩G›
unfolding trail_false_cls_def by metis
hence "¬ trail_false_cls Γ' (add_mset K D)"
unfolding Γ'_def ‹add_mset K D = cls_of_gcls C⇩G›
unfolding trail_false_cls_nonground_iff_trail_false_cls_ground .
moreover have "is_ground_cls (add_mset K D)"
using 𝒞_def ‹𝒞 = Some (add_mset K D, Var)› step_hyps(1) by auto
ultimately have "∄γ. is_ground_cls (add_mset K D ⋅ γ) ∧ trail_false_cls Γ' (add_mset K D ⋅ γ)"
by simp
thus "∄γ. is_ground_cls (add_mset K D ⋅ γ) ∧ SCL_FOL.trail_false_cls Γ' (add_mset K D ⋅ γ)"
unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
by argo
qed
thus "scl_fol.scl N β S S'"
unfolding scl_fol.scl_def by argo
next
fix S'' :: "('f, 'v) SCL_FOL.state"
assume "scl_fol.decide N β S S'"
thus False
unfolding S_def ‹𝒞 = Some _›
by (auto elim!: scl_fol.decide.cases)
qed
ultimately show ?thesis
unfolding scl_fol.regular_scl_def by argo
qed
qed
end
lemma wfp_on_antimono_stronger:
fixes
A :: "'a set" and B :: "'b set" and
f :: "'a ⇒ 'b" and
R :: "'b ⇒ 'b ⇒ bool" and Q :: "'a ⇒ 'a ⇒ bool"
assumes
wf: "wfp_on B R" and
sub: "f ` A ⊆ B" and
mono: "⋀x y. x ∈ A ⟹ y ∈ A ⟹ Q x y ⟹ R (f x) (f y)"
shows "wfp_on A Q"
unfolding wfp_on_iff_ex_minimal
proof (intro allI impI)
fix A' :: "'a set"
assume "A' ⊆ A" and "A' ≠ {}"
have "f ` A' ⊆ B"
using ‹A' ⊆ A› sub by blast
moreover have "f ` A' ≠ {}"
using ‹A' ≠ {}› by blast
ultimately have "∃z∈f ` A'. ∀y. R y z ⟶ y ∉ f ` A'"
using wf wfp_on_iff_ex_minimal by blast
hence "∃z∈A'. ∀y. R (f y) (f z) ⟶ y ∉ A'"
by blast
thus "∃z∈A'. ∀y. Q y z ⟶ y ∉ A'"
using ‹A' ⊆ A› mono by blast
qed
text ‹For AFP-devel, delete
@{thm Simulation_SCLFOL_ORDRES.wfp_on_antimono_stronger}
as it is available in \<^theory>‹HOL.Wellfounded›.›
corollary (in scl_fol_calculus) termination_projectable_strategy:
fixes
N :: "('f, 'v) Term.term clause fset" and
β :: "('f, 'v) Term.term" and
strategy and strategy_init and proj
assumes strategy_restricts_regular_scl:
"⋀S S'. strategy⇧*⇧* strategy_init S ⟹ strategy S S' ⟹ regular_scl N β (proj S) (proj S')" and
initial_state: "proj strategy_init = initial_state"
shows "wfp_on {S. strategy⇧*⇧* strategy_init S} strategy¯¯"
proof (rule wfp_on_antimono_stronger)
show "wfp_on {proj S | S. strategy⇧*⇧* strategy_init S} (regular_scl N β)¯¯"
proof (rule wfp_on_subset)
show "wfp_on {S. (regular_scl N β)⇧*⇧* initial_state S} (regular_scl N β)¯¯"
using termination_regular_scl by metis
next
show "{proj S | S. strategy⇧*⇧* strategy_init S} ⊆ {S. (regular_scl N β)⇧*⇧* initial_state S}"
proof (intro Collect_mono impI, elim exE conjE)
fix s S assume "s = proj S" and "strategy⇧*⇧* strategy_init S"
show "(regular_scl N β)⇧*⇧* initial_state s"
unfolding ‹s = proj S›
using ‹strategy⇧*⇧* strategy_init S›
proof (induction S rule: rtranclp_induct)
case base
thus ?case
unfolding initial_state by simp
next
case (step y z)
thus ?case
using strategy_restricts_regular_scl
by (meson rtranclp.simps)
qed
qed
qed
next
show "proj ` {S. strategy⇧*⇧* strategy_init S} ⊆ {proj S |S. strategy⇧*⇧* strategy_init S}"
by blast
next
show "⋀S' S. S ∈ {S. strategy⇧*⇧* strategy_init S} ⟹ strategy¯¯ S' S ⟹
(regular_scl N β)¯¯ (proj S') (proj S)"
using strategy_restricts_regular_scl by simp
qed
text ‹For AFP-devel, delete
@{thm Simulation_SCLFOL_ORDRES.scl_fol_calculus.termination_projectable_strategy}
as it is available in \<^theory>‹Simple_Clause_Learning.Termination›.›
corollary (in simulation_SCLFOL_ground_ordered_resolution) ord_res_11_termination:
fixes N :: "'f gclause fset"
shows "wfp_on {S. (ord_res_11 N)⇧*⇧* ({||}, {||}, [], None) S} (ord_res_11 N)¯¯"
proof (rule scl_fol.termination_projectable_strategy)
fix S S'
assume run: "(ord_res_11 N)⇧*⇧* ({||}, {||}, [], None) S" and step: "ord_res_11 N S S'"
define β :: "'f gterm" where
"β = (THE A. linorder_trm.is_greatest_in_fset (atms_of_clss N) A)"
show "scl_fol.regular_scl (cls_of_gcls |`| N) (term_of_gterm β) (state_of_gstate S) (state_of_gstate S')"
proof (rule ord_res_11_is_strategy_for_regular_scl)
show "∀A⇩G |∈| atms_of_clss N. A⇩G ≼⇩t β"
proof (cases "atms_of_clss N = {||}")
case True
thus ?thesis
by simp
next
case False
then show ?thesis
unfolding β_def
by (metis (full_types) linorder_trm.Uniq_is_greatest_in_fset
linorder_trm.ex_greatest_in_fset linorder_trm.is_greatest_in_fset_iff sup2CI
the1_equality')
qed
next
show "(ord_res_11 N)⇧*⇧* ({||}, {||}, [], None) S"
using run .
next
show "ord_res_11 N S S'"
using step .
qed
next
show "state_of_gstate ({||}, {||}, [], None) = SCL_FOL.initial_state"
by simp
qed
corollary (in scl_fol_calculus) static_non_subsumption_projectable_strategy:
fixes strategy and strategy_init and proj
assumes
run: "strategy⇧*⇧* strategy_init S" and
step: "backtrack N β (proj S) S'" and
strategy_restricts_regular_scl:
"⋀S S'. strategy⇧*⇧* strategy_init S ⟹ strategy S S' ⟹ regular_scl N β (proj S) (proj S')" and
initial_state: "proj strategy_init = initial_state"
defines
"U ≡ state_learned (proj S)"
shows "∃C γ. state_conflict (proj S) = Some (C, γ) ∧ ¬ (∃D |∈| N |∪| U. subsumes D C)"
proof -
have "(regular_scl N β)⇧*⇧* initial_state (proj S)"
using run
proof (induction S rule: rtranclp_induct)
case base
thus ?case
unfolding initial_state by simp
next
case (step y z)
thus ?case
using strategy_restricts_regular_scl
by (meson rtranclp.simps)
qed
moreover have "backtrack N β (proj S) S'"
using step by simp
ultimately show ?thesis
unfolding U_def
using static_non_subsumption_regular_scl
by simp
qed
text ‹For AFP-devel, delete
@{thm Simulation_SCLFOL_ORDRES.scl_fol_calculus.static_non_subsumption_projectable_strategy}
as it is available in \<^theory>‹Simple_Clause_Learning.Non_Redundancy›.›
corollary (in simulation_SCLFOL_ground_ordered_resolution) ord_res_11_non_subsumption:
fixes N⇩G :: "'f gclause fset" and s :: "_ × _ × _ × _"
defines
"β ≡ (THE A. linorder_trm.is_greatest_in_fset (atms_of_clss N⇩G) A)"
assumes
run: "(ord_res_11 N⇩G)⇧*⇧* ({||}, {||}, [], None) s" and
step: "scl_fol.backtrack (cls_of_gcls |`| N⇩G) (term_of_gterm β) (state_of_gstate s) s'"
shows "∃U⇩e⇩r ℱ Γ D. s = (U⇩e⇩r, ℱ, Γ, Some D) ∧ ¬ (∃C |∈| N⇩G |∪| U⇩e⇩r. C ⊆# D)"
proof -
have "∃C γ. state_conflict (state_of_gstate s) = Some (C, γ) ∧
¬ (∃D|∈| cls_of_gcls |`| N⇩G |∪| state_learned (state_of_gstate s). subsumes D C)"
proof (rule scl_fol.static_non_subsumption_projectable_strategy[
of "ord_res_11 N⇩G" _ _ _ _ state_of_gstate , OF run step])
fix S S'
assume run: "(ord_res_11 N⇩G)⇧*⇧* ({||}, {||}, [], None) S" and step: "ord_res_11 N⇩G S S'"
show "scl_fol.regular_scl (cls_of_gcls |`| N⇩G) (term_of_gterm β)
(state_of_gstate S) (state_of_gstate S')"
proof (intro ord_res_11_is_strategy_for_regular_scl ballI)
fix A⇩G :: "'f gterm"
assume "A⇩G |∈| atms_of_clss N⇩G"
show "A⇩G ≼⇩t β"
proof (cases "atms_of_clss N⇩G = {||}")
case True
thus ?thesis
using ‹A⇩G |∈| atms_of_clss N⇩G›
by simp
next
case False
then show ?thesis
using ‹A⇩G |∈| atms_of_clss N⇩G›
unfolding β_def
by (metis (full_types) linorder_trm.Uniq_is_greatest_in_fset
linorder_trm.ex_greatest_in_fset linorder_trm.is_greatest_in_fset_iff sup2CI
the1_equality')
qed
next
show "(ord_res_11 N⇩G)⇧*⇧* ({||}, {||}, [], None) S"
using run .
next
show "ord_res_11 N⇩G S S'"
using step .
qed
next
show "state_of_gstate ({||}, {||}, [], None) = initial_state"
by simp
qed
moreover obtain U⇩G ℱ Γ D where "s = (U⇩G, ℱ, Γ, Some D)"
proof atomize_elim
obtain U⇩G ℱ Γ 𝒞 where "s = (U⇩G, ℱ, Γ, 𝒞)"
by (metis prod.exhaust)
moreover obtain D where "𝒞 = Some D"
using step
by (auto simp: ‹s = (U⇩G, ℱ, Γ, 𝒞)› elim: scl_fol.backtrack.cases)
ultimately show "∃U⇩e⇩r ℱ Γ D. s = (U⇩e⇩r, ℱ, Γ, Some D)"
by metis
qed
ultimately have "¬ (∃C |∈| N⇩G |∪|U⇩G.
subsumes (cls_of_gcls C :: ('f, 'v) term clause) (cls_of_gcls D))"
by auto
hence "¬ (∃C |∈| N⇩G |∪|U⇩G. (cls_of_gcls C :: ('f, 'v) term clause) ⊆# (cls_of_gcls D))"
by (simp add: subsumes_def)
hence "¬ (∃C |∈| N⇩G |∪|U⇩G. C ⊆# D)"
by (metis cls_of_gcls_def image_mset_subseteq_mono)
thus ?thesis
unfolding ‹s = (U⇩G, ℱ, Γ, Some D)› by metis
qed
end