theory Watched_Literals_Transition_System
imports WB_More_Refinement CDCL.CDCL_W_Abstract_State
CDCL.CDCL_W_Restart
begin
chapter ‹Two-Watched Literals›
section ‹Rule-based system›
subsection ‹Types and Transitions System›
subsubsection ‹Types and accessing functions›
datatype 'v twl_clause =
TWL_Clause (watched: 'v) (unwatched: 'v)
fun clause :: ‹'a twl_clause ⇒ 'a :: {plus}› where
‹clause (TWL_Clause W UW) = W + UW›
abbreviation clauses :: ‹'a :: {plus} twl_clause multiset ⇒ 'a multiset› where
‹clauses C ≡ clause `# C›
type_synonym 'v twl_cls = ‹'v clause twl_clause›
type_synonym 'v twl_clss = ‹'v twl_cls multiset›
type_synonym 'v clauses_to_update = ‹('v literal × 'v twl_cls) multiset›
type_synonym 'v lit_queue = ‹'v literal multiset›
type_synonym 'v twl_st =
‹('v, 'v clause) ann_lits × 'v twl_clss × 'v twl_clss ×
'v clause option × 'v clauses × 'v clauses × 'v clauses_to_update × 'v lit_queue›
fun get_trail :: ‹'v twl_st ⇒ ('v, 'v clause) ann_lit list› where
‹get_trail (M, _, _, _, _, _, _, _) = M›
fun clauses_to_update :: ‹'v twl_st ⇒ ('v literal × 'v twl_cls) multiset› where
‹clauses_to_update (_, _, _, _, _, _, WS, _) = WS›
fun set_clauses_to_update :: ‹('v literal × 'v twl_cls) multiset ⇒ 'v twl_st ⇒ 'v twl_st› where
‹set_clauses_to_update WS (M, N, U, D, NE, UE, _, Q) = (M, N, U, D, NE, UE, WS, Q)›
fun literals_to_update :: ‹'v twl_st ⇒ 'v lit_queue› where
‹literals_to_update (_, _, _, _, _, _, _, Q) = Q›
fun set_literals_to_update :: ‹'v lit_queue ⇒ 'v twl_st ⇒ 'v twl_st› where
‹set_literals_to_update Q (M, N, U, D, NE, UE, WS, _) = (M, N, U, D, NE, UE, WS, Q)›
fun set_conflict :: ‹'v clause ⇒ 'v twl_st ⇒ 'v twl_st› where
‹set_conflict D (M, N, U, _, NE, UE, WS, Q) = (M, N, U, Some D, NE, UE, WS, Q)›
fun get_conflict :: ‹'v twl_st ⇒ 'v clause option› where
‹get_conflict (M, N, U, D, NE, UE, WS, Q) = D›
fun get_clauses :: ‹'v twl_st ⇒ 'v twl_clss› where
‹get_clauses (M, N, U, D, NE, UE, WS, Q) = N + U›
fun unit_clss :: ‹'v twl_st ⇒ 'v clause multiset› where
‹unit_clss (M, N, U, D, NE, UE, WS, Q) = NE + UE›
fun unit_init_clauses :: ‹'v twl_st ⇒ 'v clauses› where
‹unit_init_clauses (M, N, U, D, NE, UE, WS, Q) = NE›
fun get_all_init_clss :: ‹'v twl_st ⇒ 'v clause multiset› where
‹get_all_init_clss (M, N, U, D, NE, UE, WS, Q) = clause `# N + NE›
fun get_learned_clss :: ‹'v twl_st ⇒ 'v twl_clss› where
‹get_learned_clss (M, N, U, D, NE, UE, WS, Q) = U›
fun get_init_learned_clss :: ‹'v twl_st ⇒ 'v clauses› where
‹get_init_learned_clss (_, N, U, _, _, UE, _) = UE›
fun get_all_learned_clss :: ‹'v twl_st ⇒ 'v clauses› where
‹get_all_learned_clss (_, N, U, _, _, UE, _) = clause `# U + UE›
fun get_all_clss :: ‹'v twl_st ⇒ 'v clause multiset› where
‹get_all_clss (M, N, U, D, NE, UE, WS, Q) = clause `# N + NE + clause `# U + UE›
fun update_clause where
‹update_clause (TWL_Clause W UW) L L' =
TWL_Clause (add_mset L' (remove1_mset L W)) (add_mset L (remove1_mset L' UW))›
text ‹
When updating clause, we do it non-deterministically: in case of duplicate clause in the two
sets, one of the two can be updated (and it does not matter), contrary to an if-condition.
In later refinement, we know where the clause comes from and update it.
›
inductive update_clauses ::
‹'a multiset twl_clause multiset × 'a multiset twl_clause multiset ⇒
'a multiset twl_clause ⇒ 'a ⇒ 'a ⇒
'a multiset twl_clause multiset × 'a multiset twl_clause multiset ⇒ bool› where
‹D ∈# N ⟹ update_clauses (N, U) D L L' (add_mset (update_clause D L L') (remove1_mset D N), U)›
| ‹D ∈# U ⟹ update_clauses (N, U) D L L' (N, add_mset (update_clause D L L') (remove1_mset D U))›
inductive_cases update_clausesE: ‹update_clauses (N, U) D L L' (N', U')›
subsubsection ‹The Transition System›
text ‹We ensure that there are always ∗‹2› watched literals and that there are different. All
clauses containing a single literal are put in \<^term>‹NE› or \<^term>‹UE›.›
inductive cdcl_twl_cp :: ‹'v twl_st ⇒ 'v twl_st ⇒ bool› where
pop:
‹cdcl_twl_cp (M, N, U, None, NE, UE, {#}, add_mset L Q)
(M, N, U, None, NE, UE, {#(L, C)|C ∈# N + U. L ∈# watched C#}, Q)› |
propagate:
‹cdcl_twl_cp (M, N, U, None, NE, UE, add_mset (L, D) WS, Q)
(Propagated L' (clause D) # M, N, U, None, NE, UE, WS, add_mset (-L') Q)›
if
‹watched D = {#L, L'#}› and ‹undefined_lit M L'› and ‹∀L ∈# unwatched D. -L ∈ lits_of_l M› |
conflict:
‹cdcl_twl_cp (M, N, U, None, NE, UE, add_mset (L, D) WS, Q)
(M, N, U, Some (clause D), NE, UE, {#}, {#})›
if ‹watched D = {#L, L'#}› and ‹-L' ∈ lits_of_l M› and ‹∀L ∈# unwatched D. -L ∈ lits_of_l M› |
delete_from_working:
‹cdcl_twl_cp (M, N, U, None, NE, UE, add_mset (L, D) WS, Q) (M, N, U, None, NE, UE, WS, Q)›
if ‹L' ∈# clause D› and ‹L' ∈ lits_of_l M› |
update_clause:
‹cdcl_twl_cp (M, N, U, None, NE, UE, add_mset (L, D) WS, Q)
(M, N', U', None, NE, UE, WS, Q)›
if ‹watched D = {#L, L'#}› and ‹-L ∈ lits_of_l M› and ‹L' ∉ lits_of_l M› and
‹K ∈# unwatched D› and ‹undefined_lit M K ∨ K ∈ lits_of_l M› and
‹update_clauses (N, U) D L K (N', U')›
inductive_cases cdcl_twl_cpE: ‹cdcl_twl_cp S T›
text ‹We do not care about the \<^term>‹literals_to_update› literals.›
inductive cdcl_twl_o :: ‹'v twl_st ⇒ 'v twl_st ⇒ bool› where
decide:
‹cdcl_twl_o (M, N, U, None, NE, UE, {#}, {#}) (Decided L # M, N, U, None, NE, UE, {#}, {#-L#})›
if ‹undefined_lit M L› and ‹atm_of L ∈ atms_of_mm (clause `# N + NE)›
| skip:
‹cdcl_twl_o (Propagated L C' # M, N, U, Some D, NE, UE, {#}, {#})
(M, N, U, Some D, NE, UE, {#}, {#})›
if ‹-L ∉# D› and ‹D ≠ {#}›
| resolve:
‹cdcl_twl_o (Propagated L C # M, N, U, Some D, NE, UE, {#}, {#})
(M, N, U, Some (cdcl⇩W_restart_mset.resolve_cls L D C), NE, UE, {#}, {#})›
if ‹-L ∈# D› and
‹get_maximum_level (Propagated L C # M) (remove1_mset (-L) D) = count_decided M›
| backtrack_unit_clause:
‹cdcl_twl_o (M, N, U, Some D, NE, UE, {#}, {#})
(Propagated L {#L#} # M1, N, U, None, NE, add_mset {#L#} UE, {#}, {#-L#})›
if
‹L ∈# D› and
‹(Decided K # M1, M2) ∈ set (get_all_ann_decomposition M)› and
‹get_level M L = count_decided M› and
‹get_level M L = get_maximum_level M D'› and
‹get_maximum_level M (D' - {#L#}) ≡ i› and
‹get_level M K = i + 1›
‹D' = {#L#}› and
‹D' ⊆# D› and
‹clause `# (N + U) + NE + UE ⊨pm D'›
| backtrack_nonunit_clause:
‹cdcl_twl_o (M, N, U, Some D, NE, UE, {#}, {#})
(Propagated L D' # M1, N, add_mset (TWL_Clause {#L, L'#} (D' - {#L, L'#})) U, None, NE, UE,
{#}, {#-L#})›
if
‹L ∈# D› and
‹(Decided K # M1, M2) ∈ set (get_all_ann_decomposition M)› and
‹get_level M L = count_decided M› and
‹get_level M L = get_maximum_level M D'› and
‹get_maximum_level M (D' - {#L#}) ≡ i› and
‹get_level M K = i + 1›
‹D' ≠ {#L#}› and
‹D' ⊆# D› and
‹clause `# (N + U) + NE + UE ⊨pm D'› and
‹L ∈# D'›
‹L' ∈# D'› and
‹get_level M L' = i›
inductive_cases cdcl_twl_oE: ‹cdcl_twl_o S T›
inductive cdcl_twl_stgy :: ‹'v twl_st ⇒ 'v twl_st ⇒ bool› for S :: ‹'v twl_st› where
cp: ‹cdcl_twl_cp S S' ⟹ cdcl_twl_stgy S S'› |
other': ‹cdcl_twl_o S S' ⟹ cdcl_twl_stgy S S'›
inductive_cases cdcl_twl_stgyE: ‹cdcl_twl_stgy S T›
subsection ‹Definition of the Two-watched Literals Invariants›
subsubsection ‹Definitions›
text ‹The structural invariants states that there are at most two watched elements, that the watched
literals are distinct, and that there are 2 watched literals if there are at least than two
different literals in the full clauses.›
primrec struct_wf_twl_cls :: ‹'v multiset twl_clause ⇒ bool› where
‹struct_wf_twl_cls (TWL_Clause W UW) ⟷
size W = 2 ∧ distinct_mset (W + UW)›
fun state⇩W_of :: ‹'v twl_st ⇒ 'v cdcl⇩W_restart_mset› where
‹state⇩W_of (M, N, U, C, NE, UE, Q) =
(M, clause `# N + NE, clause `# U + UE, C)›
named_theorems twl_st ‹Conversions simp rules›
lemma [twl_st]: ‹trail (state⇩W_of S') = get_trail S'›
by (cases S') (auto simp: trail.simps)
lemma [twl_st]:
‹get_trail S' ≠ [] ⟹ cdcl⇩W_restart_mset.hd_trail (state⇩W_of S') = hd (get_trail S')›
by (cases S') (auto simp: trail.simps)
lemma [twl_st]: ‹conflicting (state⇩W_of S') = get_conflict S'›
by (cases S') (auto simp: conflicting.simps)
text ‹
The invariant on the clauses is the following:
▪ the structure is correct (the watched part is of length exactly two).
▪ if we do not have to update the clause, then the invariant holds.
›
definition twl_is_an_exception :: ‹'a multiset twl_clause ⇒ 'a multiset ⇒
('b × 'a multiset twl_clause) multiset ⇒ bool›
where
‹twl_is_an_exception C Q WS ⟷
(∃L. L ∈# Q ∧ L ∈# watched C) ∨ (∃L. (L, C) ∈# WS)›
definition is_blit :: ‹('a, 'b) ann_lits ⇒ 'a clause ⇒ 'a literal ⇒ bool›where
[simp]: ‹is_blit M D L ⟷ (L ∈# D ∧ L ∈ lits_of_l M)›
definition has_blit :: ‹('a, 'b) ann_lits ⇒ 'a clause ⇒ 'a literal ⇒ bool›where
‹has_blit M D L' ⟷ (∃L. is_blit M D L ∧ get_level M L ≤ get_level M L')›
text ‹This invariant state that watched literals are set at the end and are not swapped with an
unwatched literal later.›
fun twl_lazy_update :: ‹('a, 'b) ann_lits ⇒ 'a twl_cls ⇒ bool› where
‹twl_lazy_update M (TWL_Clause W UW) ⟷
(∀L. L ∈# W ⟶ -L ∈ lits_of_l M ⟶ ¬has_blit M (W+UW) L ⟶
(∀K ∈# UW. get_level M L ≥ get_level M K ∧ -K ∈ lits_of_l M))›
text ‹
If one watched literals has been assigned to false (\<^term>‹-L ∈ lits_of_l M›) and the clause
has not yet been updated (\<^term>‹L' ∉ lits_of_l M›: it should be removed either by
updating ‹L›, propagating ‹L'›, or marking the conflict), then the literals \<^term>‹L› is of
maximal level.
›
fun watched_literals_false_of_max_level :: ‹('a, 'b) ann_lits ⇒ 'a twl_cls ⇒ bool› where
‹watched_literals_false_of_max_level M (TWL_Clause W UW) ⟷
(∀L. L ∈# W ⟶ -L ∈ lits_of_l M ⟶ ¬has_blit M (W+UW) L ⟶
get_level M L = count_decided M)›
text ‹
This invariants talks about the enqueued literals:
▪ the working stack contains a single literal;
▪ the working stack and the \<^term>‹literals_to_update› literals are false with respect to the
trail and there are no duplicates;
▪ and the latter condition holds even when \<^term>‹WS = {#}›.›
fun no_duplicate_queued :: ‹'v twl_st ⇒ bool› where
‹no_duplicate_queued (M, N, U, D, NE, UE, WS, Q) ⟷
(∀C C'. C ∈# WS ⟶ C' ∈# WS ⟶ fst C = fst C') ∧
(∀C. C ∈# WS ⟶ add_mset (fst C) Q ⊆# uminus `# lit_of `# mset M) ∧
Q ⊆# uminus `# lit_of `# mset M›
lemma no_duplicate_queued_alt_def:
‹no_duplicate_queued S =
((∀C C'. C ∈# clauses_to_update S ⟶ C' ∈# clauses_to_update S ⟶ fst C = fst C') ∧
(∀C. C ∈# clauses_to_update S ⟶
add_mset (fst C) (literals_to_update S) ⊆# uminus `# lit_of `# mset (get_trail S)) ∧
literals_to_update S ⊆# uminus `# lit_of `# mset (get_trail S))›
by (cases S) auto
fun distinct_queued :: ‹'v twl_st ⇒ bool› where
‹distinct_queued (M, N, U, D, NE, UE, WS, Q) ⟷
distinct_mset Q ∧
(∀L C. count WS (L, C) ≤ count (N + U) C)›
text ‹
These are the conditions to indicate that the 2-WL invariant does not hold and is not
\<^term>‹literals_to_update›.
›
fun clauses_to_update_prop where
‹clauses_to_update_prop Q M (L, C) ⟷
(L ∈# watched C ∧ -L ∈ lits_of_l M ∧ L ∉# Q ∧ ¬has_blit M (clause C) L)›
declare clauses_to_update_prop.simps[simp del]
text ‹
This invariants talks about the enqueued literals:
▪ all clauses that should be updated are in \<^term>‹WS› and are repeated often enough in it.
▪ if \<^term>‹WS = {#}›, then there are no clauses to updated that is not enqueued;
▪ all clauses to updated are either in \<^term>‹WS› or \<^term>‹Q›.
The first two conditions are written that way to please Isabelle.›
fun clauses_to_update_inv :: ‹'v twl_st ⇒ bool› where
‹clauses_to_update_inv (M, N, U, None, NE, UE, WS, Q) ⟷
(∀L C. ((L, C) ∈# WS ⟶ {#(L, C)| C ∈# N + U. clauses_to_update_prop Q M (L, C)#} ⊆# WS)) ∧
(∀L. WS = {#} ⟶ {#(L, C)| C ∈# N + U. clauses_to_update_prop Q M (L, C)#} = {#}) ∧
(∀L C. C ∈# N + U ⟶ L ∈# watched C ⟶ -L ∈ lits_of_l M ⟶ ¬has_blit M (clause C) L ⟶
(L, C) ∉# WS ⟶ L ∈# Q)›
| ‹clauses_to_update_inv (M, N, U, D, NE, UE, WS, Q) ⟷ True›
text ‹
This is the invariant of the 2WL structure: if one watched literal is false, then all unwatched
are false.
›
fun twl_exception_inv :: ‹'v twl_st ⇒ 'v twl_cls ⇒ bool› where
‹twl_exception_inv (M, N, U, None, NE, UE, WS, Q) C ⟷
(∀L. L ∈# watched C ⟶ -L ∈ lits_of_l M ⟶ ¬has_blit M (clause C) L ⟶
L ∉# Q ⟶ (L, C) ∉# WS ⟶
(∀K ∈# unwatched C. -K ∈ lits_of_l M))›
| ‹twl_exception_inv (M, N, U, D, NE, UE, WS, Q) C ⟷ True›
declare twl_exception_inv.simps[simp del]
fun twl_st_exception_inv :: ‹'v twl_st ⇒ bool› where
‹twl_st_exception_inv (M, N, U, D, NE, UE, WS, Q) ⟷
(∀C ∈# N + U. twl_exception_inv (M, N, U, D, NE, UE, WS, Q) C)›
text ‹Candidats for propagation (i.e., the clause where only one literals is non
assigned) are enqueued.›
fun propa_cands_enqueued :: ‹'v twl_st ⇒ bool› where
‹propa_cands_enqueued (M, N, U, None, NE, UE, WS, Q) ⟷
(∀L C. C ∈# N+U ⟶ L ∈# clause C ⟶ M ⊨as CNot (remove1_mset L (clause C)) ⟶
undefined_lit M L ⟶
(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃L. (L, C) ∈# WS))›
| ‹propa_cands_enqueued (M, N, U, D, NE, UE, WS, Q) ⟷ True›
fun confl_cands_enqueued :: ‹'v twl_st ⇒ bool› where
‹confl_cands_enqueued (M, N, U, None, NE, UE, WS, Q) ⟷
(∀C ∈# N + U. M ⊨as CNot (clause C) ⟶
(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃L. (L, C) ∈# WS))›
| ‹confl_cands_enqueued (M, N, U, Some _, NE, UE, WS, Q) ⟷
True›
text ‹This invariant talk about the decomposition of the trail and the invariants that holds in
these states.›
fun past_invs :: ‹'v twl_st ⇒ bool› where
‹past_invs (M, N, U, D, NE, UE, WS, Q) ⟷
(∀M1 M2 K. M = M2 @ Decided K # M1 ⟶ (
(∀C ∈# N + U. twl_lazy_update M1 C ∧
watched_literals_false_of_max_level M1 C ∧
twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C) ∧
confl_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#}) ∧
propa_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#}) ∧
clauses_to_update_inv (M1, N, U, None, NE, UE, {#}, {#})))›
declare past_invs.simps[simp del]
fun twl_st_inv :: ‹'v twl_st ⇒ bool› where
‹twl_st_inv (M, N, U, D, NE, UE, WS, Q) ⟷
(∀C ∈# N + U. struct_wf_twl_cls C) ∧
(∀C ∈# N + U. D = None ⟶ ¬twl_is_an_exception C Q WS ⟶ (twl_lazy_update M C)) ∧
(∀C ∈# N + U. D = None ⟶ watched_literals_false_of_max_level M C)›
lemma twl_st_inv_alt_def:
‹twl_st_inv S ⟷
(∀C ∈# get_clauses S. struct_wf_twl_cls C) ∧
(∀C ∈# get_clauses S. get_conflict S = None ⟶
¬twl_is_an_exception C (literals_to_update S) (clauses_to_update S) ⟶
(twl_lazy_update (get_trail S) C)) ∧
(∀C ∈# get_clauses S. get_conflict S = None ⟶
watched_literals_false_of_max_level (get_trail S) C)›
by (cases S) (auto simp: twl_st_inv.simps)
text ‹All the unit clauses are all propagated initially except when we have found a conflict of
level \<^term>‹0::nat›.›
fun entailed_clss_inv :: ‹'v twl_st ⇒ bool› where
‹entailed_clss_inv (M, N, U, D, NE, UE, WS, Q) ⟷
(∀C ∈# NE + UE.
(∃L. L ∈# C ∧ (D = None ∨ count_decided M > 0 ⟶ get_level M L = 0 ∧ L ∈ lits_of_l M)))›
text ‹
\<^term>‹literals_to_update› literals are of maximum level and their negation is in the trail.
›
fun valid_enqueued :: ‹'v twl_st ⇒ bool› where
‹valid_enqueued (M, N, U, C, NE, UE, WS, Q) ⟷
(∀(L, C) ∈# WS. L ∈# watched C ∧ C ∈# N + U ∧ -L ∈ lits_of_l M ∧
get_level M L = count_decided M) ∧
(∀L ∈# Q. -L ∈ lits_of_l M ∧ get_level M L = count_decided M)›
text ‹Putting invariants together:›
definition twl_struct_invs :: ‹'v twl_st ⇒ bool› where
‹twl_struct_invs S ⟷
(twl_st_inv S ∧
valid_enqueued S ∧
cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S) ∧
cdcl⇩W_restart_mset.no_smaller_propa (state⇩W_of S) ∧
twl_st_exception_inv S ∧
no_duplicate_queued S ∧
distinct_queued S ∧
confl_cands_enqueued S ∧
propa_cands_enqueued S ∧
(get_conflict S ≠ None ⟶ clauses_to_update S = {#} ∧ literals_to_update S = {#}) ∧
entailed_clss_inv S ∧
clauses_to_update_inv S ∧
past_invs S)
›
definition twl_stgy_invs :: ‹'v twl_st ⇒ bool› where
‹twl_stgy_invs S ⟷
cdcl⇩W_restart_mset.cdcl⇩W_stgy_invariant (state⇩W_of S) ∧
cdcl⇩W_restart_mset.conflict_non_zero_unless_level_0 (state⇩W_of S)›
subsubsection ‹Initial properties›
lemma twl_is_an_exception_add_mset_to_queue: ‹twl_is_an_exception C (add_mset L Q) WS ⟷
(twl_is_an_exception C Q WS ∨ (L ∈# watched C))›
unfolding twl_is_an_exception_def by auto
lemma twl_is_an_exception_add_mset_to_clauses_to_update:
‹twl_is_an_exception C Q (add_mset (L, D) WS) ⟷ (twl_is_an_exception C Q WS ∨ C = D)›
unfolding twl_is_an_exception_def by auto
lemma twl_is_an_exception_empty[simp]: ‹¬twl_is_an_exception C {#} {#}›
unfolding twl_is_an_exception_def by auto
lemma twl_inv_empty_trail:
shows
‹watched_literals_false_of_max_level [] C› and
‹twl_lazy_update [] C›
by (solves ‹cases C; auto›)+
lemma clauses_to_update_inv_cases[case_names WS_nempty WS_empty Q]:
assumes
‹⋀L C. (L, C) ∈# WS ⟹ {#(L, C)| C ∈# N + U. clauses_to_update_prop Q M (L, C)#} ⊆# WS› and
‹⋀L. WS = {#} ⟹ {#(L, C)| C ∈# N + U. clauses_to_update_prop Q M (L, C)#} = {#}› and
‹⋀L C. C ∈# N + U ⟹ L ∈# watched C ⟹ -L ∈ lits_of_l M ⟹ ¬has_blit M (clause C) L ⟹
(L, C) ∉# WS ⟹ L ∈# Q›
shows
‹clauses_to_update_inv (M, N, U, None, NE, UE, WS, Q)›
using assms unfolding clauses_to_update_inv.simps by blast
lemma
assumes ‹⋀C. C ∈# N + U ⟹ struct_wf_twl_cls C›
shows
twl_st_inv_empty_trail: ‹twl_st_inv ([], N, U, C, NE, UE, WS, Q)›
by (auto simp: assms twl_inv_empty_trail)
lemma
shows
no_duplicate_queued_no_queued: ‹no_duplicate_queued (M, N, U, D, NE, UE, {#}, {#})› and
no_distinct_queued_no_queued: ‹distinct_queued ([], N, U, D, NE, UE, {#}, {#})›
by auto
lemma twl_st_inv_add_mset_clauses_to_update:
assumes ‹D ∈# N + U›
shows ‹twl_st_inv (M, N, U, None, NE, UE, WS, Q)
⟷ twl_st_inv (M, N, U, None, NE, UE, add_mset (L, D) WS, Q) ∧
(¬ twl_is_an_exception D Q WS ⟶twl_lazy_update M D)›
using assms by (auto simp: twl_is_an_exception_add_mset_to_clauses_to_update)
lemma twl_st_simps:
‹twl_st_inv (M, N, U, D, NE, UE, WS, Q) ⟷
(∀C ∈# N + U. struct_wf_twl_cls C ∧
(D = None ⟶ (¬twl_is_an_exception C Q WS ⟶ twl_lazy_update M C) ∧
watched_literals_false_of_max_level M C))›
unfolding twl_st_inv.simps by fast
lemma propa_cands_enqueued_unit_clause:
‹propa_cands_enqueued (M, N, U, C, add_mset L NE, UE, WS, Q) ⟷
propa_cands_enqueued (M, N, U, C, {#}, {#}, WS, Q)›
‹propa_cands_enqueued (M, N, U, C, NE, add_mset L UE, WS, Q) ⟷
propa_cands_enqueued (M, N, U, C, {#}, {#}, WS, Q)›
by (cases C; auto)+
lemma past_invs_enqueud: ‹past_invs (M, N, U, D, NE, UE, WS, Q) ⟷
past_invs (M, N, U, D, NE, UE, {#}, {#})›
unfolding past_invs.simps by simp
lemma confl_cands_enqueued_unit_clause:
‹confl_cands_enqueued (M, N, U, C, add_mset L NE, UE, WS, Q) ⟷
confl_cands_enqueued (M, N, U, C, {#}, {#}, WS, Q)›
‹confl_cands_enqueued (M, N, U, C, NE, add_mset L UE, WS, Q) ⟷
confl_cands_enqueued (M, N, U, C, {#}, {#}, WS, Q)›
by (cases C; auto)+
lemma twl_inv_decomp:
assumes
lazy: ‹twl_lazy_update M C› and
decomp: ‹(Decided K # M1, M2) ∈ set (get_all_ann_decomposition M)› and
n_d: ‹no_dup M›
shows
‹twl_lazy_update M1 C›
proof -
obtain W UW where C: ‹C = TWL_Clause W UW› by (cases C)
obtain M3 where M: ‹M = M3 @ M2 @ Decided K # M1›
using decomp by blast
define M' where M': ‹M' = M3 @ M2 @ [Decided K]›
have MM': ‹M = M' @ M1›
by (auto simp: M M')
have lev_M_M1: ‹get_level M L = get_level M1 L› if ‹L ∈ lits_of_l M1› for L
proof -
have LM: ‹L ∈ lits_of_l M›
using that unfolding M by auto
have ‹undefined_lit M' L›
by (rule no_dup_append_in_atm_notin)
(use that n_d in ‹auto simp: M M' defined_lit_map›)
then show lev_L_M1: ‹get_level M L = get_level M1 L›
using that n_d by (auto simp: M image_Un M')
qed
show ‹twl_lazy_update M1 C›
unfolding C twl_lazy_update.simps
proof (intro allI impI)
fix L
assume
W: ‹L ∈# W› and
uL: ‹- L ∈ lits_of_l M1› and
L': ‹¬has_blit M1 (W+UW) L›
then have lev_L_M1: ‹get_level M L = get_level M1 L›
using uL n_d lev_M_M1[of ‹-L›] by auto
have L'M: ‹¬has_blit M (W+UW) L›
proof (rule ccontr)
assume ‹¬ ?thesis›
then obtain L' where
b: ‹is_blit M (W+UW) L'› and
lev_L'_L: ‹get_level M L' ≤ get_level M L›unfolding has_blit_def by auto
then have L'M': ‹L' ∈ lits_of_l M'›
using L' MM' W lev_L_M1 lev_M_M1 unfolding has_blit_def by auto
moreover {
have ‹atm_of L' ∈ atm_of ` lits_of_l M'›
using L'M' by (simp add: atm_of_in_atm_of_set_iff_in_set_or_uminus_in_set)
moreover have ‹Decided K ∈set (dropWhile (λS. atm_of (lit_of S) ≠ atm_of K') M')›
if ‹K' ∈ lits_of_l M'› for K'
unfolding M' append_assoc[symmetric] by (rule last_in_set_dropWhile)
(use that in ‹auto simp: lits_of_def M' MM'›)
ultimately have ‹get_level M L' > count_decided M1›
unfolding MM' by (force simp: filter_empty_conv get_level_def count_decided_def
lits_of_def) }
ultimately show False
using lev_M_M1[of ‹-L›] uL count_decided_ge_get_level[of M1 ‹-L›] lev_L'_L by auto
qed
show ‹∀K∈#UW. get_level M1 K ≤ get_level M1 L ∧ -K ∈ lits_of_l M1›
proof clarify
fix K''
assume ‹K'' ∈# UW›
then have
lev_K'_L: ‹get_level M K'' ≤ get_level M L› and
uK'_M: ‹-K'' ∈ lits_of_l M›
using lazy W uL L'M unfolding C MM' by auto
then have uK'_M1: ‹- K'' ∈ lits_of_l M1›
using uK'_M unfolding M apply (auto simp: get_level_append_if
split: if_splits)
using M' MM' n_d uL count_decided_ge_get_level[of M1 L]
by (auto dest: defined_lit_no_dupD in_lits_of_l_defined_litD
simp: get_level_cons_if atm_of_eq_atm_of
split: if_splits)
have ‹get_level M K'' = get_level M1 K''›
proof (rule ccontr, cases ‹defined_lit M' K''›)
case False
moreover assume ‹get_level M K'' ≠ get_level M1 K''›
ultimately show False unfolding MM' by auto
next
case True
assume K'': ‹get_level M K'' ≠ get_level M1 K''›
have ‹get_level M' K'' = 0›
proof -
have a1: ‹get_level M' K'' + count_decided M1 ≤ get_level M1 L›
using lev_K'_L unfolding lev_L_M1 unfolding MM' get_level_skip_end[OF True] .
then have ‹count_decided M1 ≤ get_level M1 L›
by linarith
then have ‹get_level M1 L = count_decided M1›
using count_decided_ge_get_level le_antisym by blast
then show ?thesis
using a1 by linarith
qed
moreover have ‹Decided K ∈ set (dropWhile (λS. atm_of (lit_of S) ≠ atm_of K'') M')›
unfolding M' append_assoc[symmetric] by (rule last_in_set_dropWhile)
(use True in ‹auto simp: lits_of_def M' MM' defined_lit_map›)
ultimately show False
by (auto simp: M' filter_empty_conv get_level_def)
qed
then show ‹get_level M1 K'' ≤ get_level M1 L ∧ -K'' ∈ lits_of_l M1›
using lev_M_M1[OF uL] lev_K'_L uK'_M uK'_M1 by auto
qed
qed
qed
declare twl_st_inv.simps[simp del]
lemma has_blit_Cons[simp]:
assumes blit: ‹has_blit M C L› and n_d: ‹no_dup (K # M)›
shows ‹has_blit (K # M) C L›
proof -
obtain L' where
‹is_blit M C L'› and
‹get_level M L' ≤ get_level M L›
using blit unfolding has_blit_def by auto
then have
‹is_blit (K # M) C L'› and
‹get_level (K # M) L' ≤ get_level (K # M) L›
using n_d by (auto simp add: has_blit_def get_level_cons_if atm_of_eq_atm_of
dest: in_lits_of_l_defined_litD)
then show ?thesis
unfolding has_blit_def by blast
qed
lemma is_blit_Cons:
‹is_blit (K # M) C L ⟷ (L = lit_of K ∧ lit_of K ∈# C) ∨ is_blit M C L›
by (auto simp: has_blit_def)
lemma no_has_blit_propagate:
‹¬has_blit (Propagated L D # M) (W + UW) La ⟹
undefined_lit M L ⟹ no_dup M ⟹ ¬has_blit M (W + UW) La›
apply (auto simp: has_blit_def get_level_cons_if
dest: in_lits_of_l_defined_litD
split: cong: if_cong)
apply (smt atm_lit_of_set_lits_of_l count_decided_ge_get_level defined_lit_map image_eqI)
by (smt atm_lit_of_set_lits_of_l count_decided_ge_get_level defined_lit_map image_eqI)
lemma no_has_blit_propagate':
‹¬has_blit (Propagated L D # M) (clause C) La ⟹
undefined_lit M L ⟹ no_dup M ⟹ ¬has_blit M (clause C) La›
using no_has_blit_propagate[of L D M ‹watched C› ‹unwatched C›]
by (cases C) auto
lemma no_has_blit_decide:
‹¬has_blit (Decided L # M) (W + UW) La ⟹
undefined_lit M L ⟹ no_dup M ⟹ ¬has_blit M (W + UW) La›
apply (auto simp: has_blit_def get_level_cons_if
dest: in_lits_of_l_defined_litD
split: cong: if_cong)
apply (smt count_decided_ge_get_level defined_lit_map in_lits_of_l_defined_litD le_SucI)
apply (smt count_decided_ge_get_level defined_lit_map in_lits_of_l_defined_litD le_SucI)
done
lemma no_has_blit_decide':
‹¬has_blit (Decided L # M) (clause C) La ⟹
undefined_lit M L ⟹ no_dup M ⟹ ¬has_blit M (clause C) La›
using no_has_blit_decide[of L M ‹watched C› ‹unwatched C›]
by (cases C) auto
lemma twl_lazy_update_Propagated:
assumes
W: ‹L ∈# W› and n_d: ‹no_dup (Propagated L D # M)› and
lazy: ‹twl_lazy_update M (TWL_Clause W UW)›
shows
‹twl_lazy_update (Propagated L D # M) (TWL_Clause W UW)›
unfolding twl_lazy_update.simps
proof (intro conjI impI allI)
fix La
assume
La: ‹La ∈# W› and
uL_M: ‹- La ∈ lits_of_l (Propagated L D # M)› and
b: ‹¬ has_blit (Propagated L D # M) (W + UW) La›
have b': ‹¬has_blit M (W + UW) La›
apply (rule no_has_blit_propagate[OF b])
using assms by auto
have ‹- La ∈ lits_of_l M ⟶ (∀K∈#UW. get_level M K ≤ get_level M La ∧ - K ∈ lits_of_l M)›
using lazy assms b' uL_M La unfolding twl_lazy_update.simps
by blast
then consider
‹∀K∈#UW. get_level M K ≤ get_level M La ∧ -K ∈ lits_of_l M› and ‹La ≠ -L› |
‹La = -L›
using b' uL_M La
by (simp only: list.set(2) lits_of_insert insert_iff uminus_lit_swap)
fastforce
then show ‹∀K∈#UW. get_level (Propagated L D # M) K ≤ get_level (Propagated L D # M) La ∧
-K ∈ lits_of_l (Propagated L D # M)›
proof cases
case 1
have [simp]: ‹has_blit (Propagated L D # M) (W + UW) L› if ‹L ∈# W+UW›
using that unfolding has_blit_def apply -
by (rule exI[of _ L]) (auto simp: get_level_cons_if atm_of_eq_atm_of)
show ?thesis
using n_d b 1 b' uL_M
by (auto simp: get_level_cons_if atm_of_eq_atm_of
count_decided_ge_get_level Decided_Propagated_in_iff_in_lits_of_l
dest!: multi_member_split)
next
case 2
have [simp]: ‹has_blit (Propagated L D # M) (W + UW) (-L)›
using 2 La W unfolding has_blit_def apply -
by (rule exI[of _ L])
(auto simp: get_level_cons_if atm_of_eq_atm_of)
show ?thesis
using 2 b count_decided_ge_get_level[of ‹Propagated L D # M›]
by (auto simp: uminus_lit_swap split: if_splits)
qed
qed
lemma pair_in_image_Pair:
‹(La, C) ∈ Pair L ` D ⟷ La = L ∧ C ∈ D›
by auto
lemma image_Pair_subset_mset:
‹Pair L `# A ⊆# Pair L `# B ⟷ A ⊆# B›
proof -
have [simp]: ‹remove1_mset (L, x) (Pair L `# B) = Pair L `# (remove1_mset x B)› for x :: 'b and B
proof -
have ‹(L, x) ∈# Pair L `# B ⟶ x ∈# B›
by force
then show ?thesis
by (metis (no_types) diff_single_trivial image_mset_remove1_mset_if)
qed
show ?thesis
by (induction A arbitrary: B) (auto simp: insert_subset_eq_iff)
qed
lemma count_image_mset_Pair2:
‹count {#(L, x). L ∈# M x#} (L, C) = (if x = C then count (M x) L else 0)›
proof -
have ‹count (M C) L = count {#L. L∈#M C#} L›
by simp
also have ‹… = count ((λL. Pair L C) `# {#L. L∈#M C#}) ((λL. Pair L C) L)›
by (subst (2) count_image_mset_inj) (simp_all add: inj_on_def)
finally have C: ‹count {#(L, C). L ∈# {#L. L ∈# M C#}#} (L, C) = count (M C) L› ..
show ?thesis
apply (cases ‹x ≠ C›)
apply (auto simp: not_in_iff[symmetric] count_image_mset; fail)[]
using C by simp
qed
lemma lit_of_inj_on_no_dup: ‹no_dup M ⟹ inj_on (λx. - lit_of x) (set M)›
by (induction M) (auto simp: no_dup_def)
lemma
assumes
cdcl: ‹cdcl_twl_cp S T› and
twl: ‹twl_st_inv S› and
twl_excep: ‹twl_st_exception_inv S› and
valid: ‹valid_enqueued S› and
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)› and
no_dup: ‹no_duplicate_queued S› and
dist_q: ‹distinct_queued S› and
ws: ‹clauses_to_update_inv S›
shows twl_cp_twl_st_exception_inv: ‹twl_st_exception_inv T› and
twl_cp_clauses_to_update: ‹clauses_to_update_inv T›
using cdcl twl twl_excep valid inv no_dup ws
proof (induction rule: cdcl_twl_cp.induct)
case (pop M N U NE UE L Q)
case 1 note _ = this(2)
then show ?case unfolding twl_st_inv.simps twl_is_an_exception_def
by (fastforce simp add: pair_in_image_Pair image_constant_conv uminus_lit_swap
twl_exception_inv.simps)
case 2 note twl = this(1) and ws = this(6)
have struct: ‹struct_wf_twl_cls C› if ‹C ∈# N + U› for C
using twl that by (simp add: twl_st_inv.simps)
have H: ‹count (watched C) L ≤ 1› if ‹C ∈# N + U› for C L
using struct[OF that] by (cases C) (auto simp add: twl_st_inv.simps size_2_iff)
have sum_le_count: ‹(∑x∈#N+U. count {#(L, x). L ∈# watched x#} (a, b)) ≤ count (N+U) b›
for a b
apply (subst (2) count_sum_mset_if_1_0)
apply (rule sum_mset_mono)
using H apply (auto simp: count_image_mset_Pair2)
done
define NU where NU[symmetric]: ‹NU = N + U›
show ?case
using ws by (fastforce simp add: pair_in_image_Pair multiset_filter_mono2 image_Pair_subset_mset
clauses_to_update_prop.simps NU filter_mset_empty_conv)
next
case (propagate D L L' M N U NE UE WS Q) note watched = this(1) and undef = this(2) and
unw = this(3)
case 1
note twl = this(1) and twl_excep = this(2) and valid = this(3) and inv = this(4) and
no_dup = this(5) and ws = this(6)
have [simp]: ‹- L' ∉ lits_of_l M›
using Decided_Propagated_in_iff_in_lits_of_l propagate.hyps(2) by blast
have D_N_U: ‹D ∈# N + U› and lev_L: ‹get_level M L = count_decided M›
using valid by auto
then have wf_D: ‹struct_wf_twl_cls D›
using twl by (simp add: twl_st_inv.simps)
have ‹∀s∈#clause `# U. ¬ tautology s›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.distinct_cdcl⇩W_state_def by (simp_all add: cdcl⇩W_restart_mset_state)
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
have [simp]: ‹L ≠ L'›
using wf_D watched by (cases D) auto
have [simp]: ‹- L ∈ lits_of_l M›
using valid by auto
then have [simp]: ‹L ∉ lits_of_l M›
using n_d no_dup_consistentD by blast
obtain NU where NU: ‹N + U = add_mset D NU›
by (metis D_N_U insert_DiffM)
have [simp]: ‹has_blit (Propagated L' (add_mset L (add_mset L' x2)) # M)
(add_mset L (add_mset L' x2)) L› for x2
unfolding has_blit_def
by (rule exI[of _ L'])
(use lev_L in ‹auto simp: get_level_cons_if›)
have HH: ‹¬clauses_to_update_prop (add_mset (-L') Q) (Propagated L' (clause D) # M) (L, D)›
using watched unfolding clauses_to_update_prop.simps by (cases D) (auto simp: watched)
have ‹add_mset L Q ⊆# {#- lit_of x. x ∈# mset M#}›
using no_dup by (auto)
moreover have ‹distinct_mset {#- lit_of x. x ∈# mset M#}›
by (subst distinct_image_mset_inj)
(use n_d in ‹auto simp: lit_of_inj_on_no_dup distinct_map no_dup_def›)
ultimately have [simp]: ‹L ∉# Q›
by (metis distinct_mset_add_mset distinct_mset_union subset_mset.le_iff_add)
have ‹¬has_blit M (clause D) L›
using watched undef unw n_d by (cases D)
(auto simp: has_blit_def Decided_Propagated_in_iff_in_lits_of_l dest: no_dup_consistentD)
then have w_q_p_D: ‹clauses_to_update_prop Q M (L, D)›
by (auto simp: clauses_to_update_prop.simps watched)
have ‹Pair L `# {#C ∈# add_mset D NU. clauses_to_update_prop Q M (L, C)#} ⊆# add_mset (L, D) WS›
using ws no_dup unfolding clauses_to_update_inv.simps NU
by (auto simp: all_conj_distrib)
then have IH: ‹Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#} ⊆# WS›
using w_q_p_D by auto
have IH_Q: ‹∀La C. C ∈# add_mset D NU ⟶ La ∈# watched C ⟶ - La ∈ lits_of_l M ⟶
¬ has_blit M (clause C) La ⟶ (La, C) ∉# add_mset (L, D) WS ⟶ La ∈# Q›
using ws no_dup unfolding clauses_to_update_inv.simps NU
by (auto simp: all_conj_distrib)
show ?case
unfolding Ball_def twl_st_exception_inv.simps twl_exception_inv.simps
proof (intro allI conjI impI)
fix C J K
assume C: ‹C ∈# N + U› and
watched_C: ‹J ∈# watched C› and
J: ‹- J ∈ lits_of_l (Propagated L' (clause D) # M)› and
J': ‹¬ has_blit (Propagated L' (clause D) # M) (clause C) J› and
J_notin: ‹J ∉# add_mset (- L') Q› and
C_WS: ‹(J, C) ∉# WS› and
‹K ∈# unwatched C›
moreover have ‹¬ has_blit M (clause C) J›
using no_has_blit_propagate'[OF J'] n_d undef by fast
ultimately have ‹- K ∈ lits_of_l (Propagated L' (clause D) # M)› if ‹C ≠ D›
using twl_excep that by (auto simp add: uminus_lit_swap twl_exception_inv.simps)
moreover have CD: False if ‹C = D›
using J J' watched_C watched that J_notin
by (cases D) (auto simp: add_mset_eq_add_mset)
ultimately show ‹- K ∈ lits_of_l (Propagated L' (clause D) # M)›
by blast
qed
case 2
show ?case
proof (induction rule: clauses_to_update_inv_cases)
case (WS_nempty L'' C)
then have [simp]: ‹L'' = L›
using ws no_dup unfolding clauses_to_update_inv.simps NU by (auto simp: all_conj_distrib)
have *: ‹Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#} ⊇#
Pair L `# {#C ∈# NU.
clauses_to_update_prop (add_mset (- L') Q) (Propagated L' (clause D) # M) (L'', C)#}›
using undef n_d
unfolding image_Pair_subset_mset multiset_filter_mono2 clauses_to_update_prop.simps
by (auto dest!: no_has_blit_propagate')
show ?case
using subset_mset.dual_order.trans[OF IH *] HH
unfolding NU ‹L'' = L›
by simp
next
case (WS_empty K)
then show ?case
using IH IH_Q watched undef n_d unfolding NU
by (cases D) (auto simp: filter_mset_empty_conv
clauses_to_update_prop.simps watched add_mset_eq_add_mset
dest!: no_has_blit_propagate')
next
case (Q LC' C)
then show ?case
using watched "1.prems"(6) HH Q.hyps HH IH_Q undef n_d
apply (cases D)
apply (cases C)
apply (auto simp: add_mset_eq_add_mset NU)
by (metis HH Q.IH(2) Q.IH(3) Q.hyps clauses_to_update_prop.simps insert_iff
no_has_blit_propagate' set_mset_add_mset_insert)
qed
next
case (conflict D L L' M N U NE UE WS Q)
case 1
note twl = this(5)
show ?case by (auto simp: twl_st_inv.simps twl_exception_inv.simps)
case 2
show ?case
by (auto simp: twl_st_inv.simps twl_exception_inv.simps)
next
case (delete_from_working L' D M N U NE UE L WS Q) note watched = this(1) and L' = this(2)
case 1 note twl = this(1) and twl_excep = this(2) and valid = this(3) and inv = this(4) and
no_dup = this(5) and ws = this(6)
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
have D_N_U: ‹D ∈# N + U›
using valid by auto
then have wf_D: ‹struct_wf_twl_cls D›
using twl by (simp add: twl_st_inv.simps)
obtain NU where NU: ‹N + U = add_mset D NU›
by (metis D_N_U insert_DiffM)
have D_N_U: ‹D ∈# N + U› and lev_L: ‹get_level M L = count_decided M›
using valid by auto
have [simp]: ‹has_blit M (clause D) L›
unfolding has_blit_def
by (rule exI[of _ L'])
(use watched L' lev_L in ‹auto simp: count_decided_ge_get_level›)
have [simp]: ‹¬clauses_to_update_prop Q M (L, D)›
using L' by (auto simp: clauses_to_update_prop.simps watched)
have IH_WS: ‹Pair L `# {#C ∈# N + U. clauses_to_update_prop Q M (L, C)#} ⊆# add_mset (L, D) WS›
using ws by (auto simp del: filter_union_mset simp: NU)
then have IH_WS_NU: ‹Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#} ⊆#
add_mset (L, D) WS›
using ws by (auto simp del: filter_union_mset simp: NU)
have IH_WS': ‹Pair L `# {#C ∈# N + U. clauses_to_update_prop Q M (L, C)#} ⊆# WS›
by (rule subset_add_mset_notin_subset_mset[OF IH_WS]) auto
have IH_Q: ‹∀La C. C ∈# add_mset D NU ⟶ La ∈# watched C ⟶ - La ∈ lits_of_l M ⟶
¬has_blit M (clause C) La ⟶ (La, C) ∉# add_mset (L, D) WS ⟶ La ∈# Q›
using ws no_dup unfolding clauses_to_update_inv.simps NU
by (auto simp: all_conj_distrib)
show ?case
unfolding Ball_def twl_st_exception_inv.simps twl_exception_inv.simps
proof (intro allI conjI impI)
fix C J K
assume C: ‹C ∈# N + U› and
watched_C: ‹J ∈# watched C› and
J: ‹- J ∈ lits_of_l M› and
J': ‹¬has_blit M (clause C) J› and
J_notin: ‹J ∉# Q› and
C_WS: ‹(J, C) ∉# WS› and
‹K ∈# unwatched C›
then have ‹- K ∈ lits_of_l M› if ‹C ≠ D›
using twl_excep that by (simp add: uminus_lit_swap twl_exception_inv.simps)
moreover {
from n_d have False if ‹ - L' ∈ lits_of_l M› ‹L' ∈ lits_of_l M›
using that consistent_interp_def distinct_consistent_interp by blast
then have CD: False if ‹C = D›
using J J' watched_C watched L' C_WS IH_Q J_notin ‹¬ clauses_to_update_prop Q M (L, D)› that
apply (auto simp: add_mset_eq_add_mset)
by (metis C_WS J_notin ‹¬ clauses_to_update_prop Q M (L, D)›
clauses_to_update_prop.simps that)
}
ultimately show ‹- K ∈ lits_of_l M›
by blast
qed
case 2
show ?case
proof (induction rule: clauses_to_update_inv_cases)
case (WS_nempty K C) note KC = this
have LK: ‹L = K›
using no_dup KC by auto
from subset_add_mset_notin_subset_mset[OF IH_WS]
have 1: ‹Pair K `# {#C ∈# N + U. clauses_to_update_prop Q M (L, C)#} ⊆# WS›
using L' LK ‹has_blit M (clause D) L›
by (auto simp del: filter_union_mset simp: pair_in_image_Pair watched add_mset_eq_add_mset
all_conj_distrib clauses_to_update_prop.simps)
show ?case
by (metis (no_types, lifting) 1 LK)
next
case (WS_empty K) note [simp] = this(1)
have [simp]: ‹¬clauses_to_update_prop Q M (K, D)›
using IH_Q WS_empty.IH watched ‹has_blit M (clause D) L›
using IH_WS' IH_Q watched by (auto simp: add_mset_eq_add_mset NU filter_mset_empty_conv
all_conj_distrib clauses_to_update_prop.simps)
show ?case
using IH_WS' IH_Q watched by (auto simp: add_mset_eq_add_mset NU filter_mset_empty_conv
all_conj_distrib clauses_to_update_prop.simps)
next
case (Q K C)
then show ?case
using ‹¬ clauses_to_update_prop Q M (L, D)› ws
unfolding clauses_to_update_inv.simps(1) clauses_to_update_prop.simps member_add_mset
is_blit_def
by blast
qed
next
case (update_clause D L L' M K N U N' U' NE UE WS Q) note watched = this(1) and uL = this(2) and
L' = this(3) and K = this(4) and undef = this(5) and N'U' = this(6)
case 1 note twl = this(1) and twl_excep = this(2) and valid = this(3) and inv = this(4) and
no_dup = this(5) and ws = this(6)
obtain WD UWD where D: ‹D = TWL_Clause WD UWD› by (cases D)
have L: ‹L ∈# watched D› and D_N_U: ‹D ∈# N + U› and lev_L: ‹get_level M L = count_decided M›
using valid by auto
then have struct_D: ‹struct_wf_twl_cls D›
using twl by (auto simp: twl_st_inv.simps)
have L'_UWD: ‹L ∉# remove1_mset L' UWD› if ‹L ∈# WD› for L
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD L ≥ 1›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) L ≥ 2›
using D that by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have L'_L'_UWD: ‹K ∉# remove1_mset K UWD›
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD K ≥ 2›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) K ≥ 2›
using D L' by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps)
let ?D = ‹update_clause D L K›
have *: ‹C ∈# N + U› if ‹C ≠ ?D› and C: ‹C ∈# N' + U'› for C
using C N'U' that by (auto elim!: update_clausesE dest: in_diffD)
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def
by (auto simp: trail.simps)
then have uK_M: ‹- K ∉ lits_of_l M›
using undef Decided_Propagated_in_iff_in_lits_of_l consistent_interp_def
distinct_consistent_interp by blast
have add_remove_WD: ‹add_mset K (remove1_mset L WD) ≠ WD›
using uK_M uL by (auto simp: add_mset_remove_trivial_iff trivial_add_mset_remove_iff)
obtain NU where NU: ‹N + U = add_mset D NU›
by (metis D_N_U insert_DiffM)
have L_M: ‹L ∉ lits_of_l M›
using n_d uL by (fastforce dest!: distinct_consistent_interp
simp: consistent_interp_def lits_of_def uminus_lit_swap)
have w_max_D: ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps)
have lev_L': ‹get_level M L' = count_decided M›
if ‹- L' ∈ lits_of_l M› ‹¬has_blit M (clause D) L'›
using L_M w_max_D D watched L' uL that by auto
have D_ne_D: ‹D ≠ update_clause D L K›
using D add_remove_WD by auto
have N'U': ‹N' + U' = add_mset ?D (remove1_mset D (N + U))›
using N'U' D_N_U by (auto elim!: update_clausesE)
define NU where ‹NU = remove1_mset D (N + U)›
then have NU: ‹N + U = add_mset D NU›
using D_N_U by auto
have watched_D: ‹watched ?D = {#K, L'#}›
using D add_remove_WD watched by auto
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
have D_N_U: ‹D ∈# N + U› and lev_L: ‹get_level M L = count_decided M›
using valid by auto
have ‹has_blit (Propagated L' C # M)
(add_mset L (add_mset L' x2)) L› for C x2
unfolding has_blit_def
by (rule exI[of _ L'])
(use lev_L in ‹auto simp: count_decided_ge_get_level get_level_cons_if›)
then have HH: ‹¬clauses_to_update_prop (add_mset (-L') Q) (Propagated L' (clause D) # M) (L, D)›
using watched unfolding clauses_to_update_prop.simps by (cases D) (auto simp: watched)
have ‹add_mset L Q ⊆# {#- lit_of x. x ∈# mset M#}›
using no_dup by (auto)
moreover have ‹distinct_mset {#- lit_of x. x ∈# mset M#}›
by (subst distinct_image_mset_inj)
(use n_d in ‹auto simp: lit_of_inj_on_no_dup distinct_map no_dup_def›)
ultimately have LQ: ‹L ∉# Q›
by (metis distinct_mset_add_mset distinct_mset_union subset_mset.le_iff_add)
have w_q_p_D: ‹¬has_blit M (clause D) L ⟹ clauses_to_update_prop Q M (L, D)›
using watched uL L' by (cases D) (auto simp: LQ clauses_to_update_prop.simps)
have ‹Pair L `# {#C ∈# add_mset D NU. clauses_to_update_prop Q M (L, C)#} ⊆# add_mset (L, D) WS›
using ws no_dup unfolding clauses_to_update_inv.simps NU
by (auto simp: all_conj_distrib)
then have IH: ‹¬has_blit M (clause D) L ⟹ Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#} ⊆# WS›
using w_q_p_D by auto
have IH_Q: ‹⋀La C. C ∈# add_mset D NU ⟹ La ∈# watched C ⟹ - La ∈ lits_of_l M ⟹
¬has_blit M (clause C) La ⟹ (La, C) ∉# add_mset (L, D) WS ⟹ La ∈# Q›
using ws no_dup unfolding clauses_to_update_inv.simps NU
by (auto simp: all_conj_distrib)
have blit_clss_to_upd: ‹has_blit M (clause D) L ⟹ ¬ clauses_to_update_prop Q M (L, D)›
by (auto simp: clauses_to_update_prop.simps)
have
‹Pair L `# {#C ∈# N + U. clauses_to_update_prop Q M (L, C)#} ⊆# add_mset (L, D) WS›
using ws by (auto simp del: filter_union_mset)
moreover have ‹has_blit M (clause D) L ⟹
(L, D) ∉# Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#}›
by (auto simp: clauses_to_update_prop.simps)
ultimately have Q_M_L_WS:
‹Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#} ⊆# WS›
by (auto simp del: filter_union_mset simp: NU w_q_p_D blit_clss_to_upd
intro: subset_add_mset_notin_subset_mset split: if_splits)
have L_ne_L': ‹L ≠ L'›
using struct_D D watched by auto
have clss_upd_D[simp]: ‹clause ?D = clause D›
using D K watched by auto
show ?case
unfolding Ball_def twl_st_exception_inv.simps twl_exception_inv.simps
proof (intro allI conjI impI)
fix C J K''
assume C: ‹C ∈# N' + U'› and
watched_C: ‹J∈# watched C› and
J: ‹- J ∈ lits_of_l M› and
J': ‹¬has_blit M (clause C) J› and
J_notin: ‹J ∉# Q› and
C_WS: ‹(J, C) ∉# WS› and
K'': ‹K'' ∈# unwatched C›
then have ‹- K'' ∈ lits_of_l M› if ‹C ≠ D› ‹C ≠ ?D›
using twl_excep that *[OF _ C] N'U' by (simp add: uminus_lit_swap twl_exception_inv.simps)
moreover have ‹- K'' ∈ lits_of_l M› if CD: ‹C = D›
proof (rule ccontr)
assume uK''_M: ‹- K'' ∉ lits_of_l M›
have ‹Pair L `# {#C ∈# N + U. clauses_to_update_prop Q M (L, C)#} ⊆# add_mset (L, D) WS›
using ws by (auto simp: all_conj_distrib
simp del: filter_union_mset)
show False
proof cases
assume [simp]: ‹J = L›
have w_q_p_L: ‹clauses_to_update_prop Q M (L, C)›
unfolding clauses_to_update_prop.simps watched_C J J' K'' uK''_M
apply (auto simp add: add_mset_eq_add_mset conj_disj_distribR ex_disj_distrib)
using watched watched_C CD J J' J_notin K'' uK''_M uL L' L_M
by (auto simp: clauses_to_update_prop.simps add_mset_eq_add_mset)
then have ‹Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#} ⊆# WS›
using ws by (auto simp: all_conj_distrib NU CD simp del: filter_union_mset)
moreover have ‹(L, C) ∈# Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#}›
using C w_q_p_L D_ne_D by (auto simp: pair_in_image_Pair N'U' NU CD)
ultimately have ‹(L, C) ∈# WS›
by blast
then show ‹False›
using C_WS by simp
next
assume ‹J ≠ L›
then have ‹clauses_to_update_prop Q M (L, C)›
unfolding clauses_to_update_prop.simps watched_C J J' K'' uK''_M
apply (auto simp add: add_mset_eq_add_mset conj_disj_distribR ex_disj_distrib)
using watched watched_C CD J J' J_notin K'' uK''_M uL L' L_M
apply (auto simp: clauses_to_update_prop.simps add_mset_eq_add_mset)
using C_WS D_N_U clauses_to_update_prop.simps ws by auto
then show ‹False›
using C_WS D_N_U J J' J_notin ‹J ≠ L› that watched_C ws by auto
qed
qed
moreover {
assume CD: ‹C = ?D›
have JL[simp]: ‹J = L'›
using CD J J' watched_C watched L' D uK_M undef
by (auto simp: add_mset_eq_add_mset)
have ‹K'' ≠ K›
using K'' uK_M uL D L'_L'_UWD unfolding CD
by (cases D) auto
have K''_unwatched_L: ‹K'' ∈# remove1_mset K (unwatched D) ∨ K'' = L›
using K'' unfolding CD by (cases D) auto
have ‹clause C = clause D›
using D K watched unfolding CD by auto
then have blit: ‹¬ has_blit M (clause D) L'›
using J' unfolding CD by simp
have False if ‹- L' ∈ lits_of_l M› ‹L' ∈ lits_of_l M›
using n_d that consistent_interp_def distinct_consistent_interp by blast
have H: ‹⋀x La xa. x ∈# N + U ⟹
La ∈# watched x ⟹ - La ∈ lits_of_l M ⟹
¬has_blit M (clause x) La ⟹ La ∉# Q ⟹ (La, x) ∉# add_mset (L, D) WS ⟹
xa ∈# unwatched x ⟹ - xa ∈ lits_of_l M›
using twl_excep[unfolded twl_st_exception_inv.simps Ball_def twl_exception_inv.simps]
unfolding has_blit_def is_blit_def
by blast
have LL': ‹L ≠ L'›
using struct_D watched by (cases D) auto
have L'D_WS: ‹(L', D) ∉# WS›
using no_dup LL' by (auto dest: multi_member_split)
have ‹xa ∈# unwatched D ⟹ - xa ∈ lits_of_l M›
if ‹- L' ∈ lits_of_l M› and ‹L' ∉# Q› and ‹¬ has_blit M (clause D) L'› for xa
by (rule H[of D L'])
(use D_N_U watched LL' that L'D_WS K'' that in ‹auto simp: add_mset_eq_add_mset L_M›)
consider
(unwatched_unqueued) ‹K'' ∈# remove1_mset K (unwatched D)› |
(KL) ‹K'' = L›
using K''_unwatched_L by blast
then have ‹- K'' ∈ lits_of_l M›
proof cases
case KL
then show ?thesis
using uL by simp
next
case unwatched_unqueued
moreover have ‹L' ∉# Q›
using JL J_notin by blast
ultimately show ?thesis
using blit H[of D L'] D_N_U watched LL' L'D_WS K'' J J'
by (auto simp: add_mset_eq_add_mset L_M dest: in_diffD)
qed
}
ultimately show ‹- K'' ∈ lits_of_l M›
by blast
qed
case 2
show ?case
proof (induction rule: clauses_to_update_inv_cases)
case (WS_nempty K'' C) note KC = this(1)
have LK: ‹L = K''›
using no_dup KC by auto
have [simp]: ‹¬clauses_to_update_prop Q M (K'', update_clause D K'' K)›
using watched uK_M struct_D
by (cases D) (auto simp: clauses_to_update_prop.simps add_mset_eq_add_mset LK)
have 1: ‹Pair L `# {#C ∈# N' + U'. clauses_to_update_prop Q M (L, C)#} ⊆#
Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#}›
unfolding image_Pair_subset_mset LK
using LK N'U' by (auto simp del: filter_union_mset simp: pair_in_image_Pair watched NU
add_mset_eq_add_mset all_conj_distrib)
then show ‹Pair K'' `# {#C ∈# N' + U'. clauses_to_update_prop Q M (K'', C)#} ⊆# WS›
using Q_M_L_WS unfolding LK by auto
next
case (WS_empty K'')
then show ?case
using IH IH_Q uL uK_M L_M watched L_ne_L' unfolding N'U' NU
by (force simp: filter_mset_empty_conv clauses_to_update_prop.simps
add_mset_eq_add_mset watched_D all_conj_distrib)
next
case (Q K' C) note C = this(1) and uK'_M = this(2) and uK''_M = this(3) and KC_WS = this(4)
and watched_C = this(5)
have ?case if CD: ‹C ≠ D› ‹C ≠ ?D›
using IH_Q[of C K'] CD watched uK_M L' L_ne_L' L_M uK'_M uK''_M
Q unfolding N'U' NU
by auto
moreover have ?case if CD: ‹C = D›
proof -
consider
(KL) ‹K' = L› |
(K'L') ‹K' = L'›
using watched watched_C CD by (auto simp: add_mset_eq_add_mset)
then show ?thesis
proof cases
case KL note [simp] = this
have ‹(L, C) ∈# Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#}›
using CD C w_q_p_D uK''_M unfolding NU N'U' by (auto simp: pair_in_image_Pair D_ne_D)
then have ‹(L, C) ∈# WS›
using Q_M_L_WS by blast
then have False using KC_WS unfolding CD by simp
then show ?thesis by fast
next
case K'L' note [simp] = this
show ?thesis
by (rule IH_Q[of C]) (use CD watched_C uK'_M uK''_M KC_WS L_ne_L' in auto)
qed
qed
moreover {
have ‹(L', D) ∉# WS›
using no_dup L_ne_L' by (auto simp: all_conj_distrib)
then have ?case if CD: ‹C = ?D›
using IH_Q[of D L] IH_Q[of D L'] CD watched watched_D watched_C watched uK_M L'
L_ne_L' L_M uK'_M uK''_M D_ne_D C unfolding NU N'U'
by (auto simp: add_mset_eq_add_mset all_conj_distrib imp_conjR)
}
ultimately show ?case
by blast
qed
qed
lemma twl_cp_twl_inv:
assumes
cdcl: ‹cdcl_twl_cp S T› and
twl: ‹twl_st_inv S› and
valid: ‹valid_enqueued S› and
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)› and
twl_excep: ‹twl_st_exception_inv S› and
no_dup: ‹no_duplicate_queued S› and
wq: ‹clauses_to_update_inv S›
shows ‹twl_st_inv T›
using cdcl twl valid inv twl_excep no_dup wq
proof (induction rule: cdcl_twl_cp.induct)
case (pop M N U NE UE L Q) note inv = this(1)
then show ?case unfolding twl_st_inv.simps twl_is_an_exception_def
by (fastforce simp add: pair_in_image_Pair)
next
case (propagate D L L' M N U NE UE WS Q) note watched = this(1) and undef = this(2) and
unw = this(3) and twl = this(4) and valid = this(5) and inv = this(6) and exception = this(7)
have uL'_M[simp]: ‹- L' ∉ lits_of_l M›
using Decided_Propagated_in_iff_in_lits_of_l propagate.hyps(2) by blast
have D_N_U: ‹D ∈# N + U› and lev_L: ‹get_level M L = count_decided M›
using valid by auto
then have wf_D: ‹struct_wf_twl_cls D›
using twl by (auto simp add: twl_st_inv.simps)
have [simp]: ‹- L ∈ lits_of_l M›
using valid by auto
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
show ?case unfolding twl_st_simps Ball_def
proof (intro allI conjI impI)
fix C
assume C: ‹C ∈# N + U›
show ‹struct_wf_twl_cls C›
using twl C by (auto simp: twl_st_inv.simps)[]
have watched_max: ‹watched_literals_false_of_max_level M C›
using twl C by (auto simp: twl_st_inv.simps)
then show ‹watched_literals_false_of_max_level (Propagated L' (clause D) # M) C›
using undef n_d
by (cases C) (auto simp: get_level_cons_if dest!: no_has_blit_propagate')
assume excep: ‹¬twl_is_an_exception C (add_mset (- L') Q) WS›
have excep_C: ‹¬ twl_is_an_exception C Q (add_mset (L, D) WS)› if ‹C ≠ D›
using excep that by (auto simp add: twl_is_an_exception_def)
then
have ‹twl_lazy_update M C› if ‹C ≠ D›
using twl C D_N_U that by (cases ‹C = D›) (auto simp add: twl_st_inv.simps)
then show ‹twl_lazy_update (Propagated L' (clause D) # M) C›
using twl C excep uL'_M twl undef n_d uL'_M unw watched_max
apply (cases C)
apply (auto simp: get_level_cons_if count_decided_ge_get_level
twl_is_an_exception_add_mset_to_queue atm_of_eq_atm_of
dest!: no_has_blit_propagate' no_has_blit_propagate)
apply (metis twl_clause.sel(2) uL'_M unw)
apply (metis twl_clause.sel(2) uL'_M unw)
apply (metis twl_clause.sel(2) uL'_M unw)
apply (metis twl_clause.sel(2) uL'_M unw)
done
qed
next
case (conflict D L L' M N U NE UE WS Q) note twl = this(4)
then show ?case
by (auto simp: twl_st_inv.simps)
next
case (delete_from_working L' D M N U NE UE L WS Q) note watched = this(1) and L' = this(2) and
twl = this(3) and valid = this(4) and inv = this(5) and tauto = this(6)
show ?case unfolding twl_st_simps Ball_def
proof (intro allI conjI impI)
fix C
assume C: ‹C ∈# N + U›
show ‹struct_wf_twl_cls C›
using twl C by (auto simp: twl_st_inv.simps)[]
show ‹watched_literals_false_of_max_level M C›
using twl C by (auto simp: twl_st_inv.simps)
assume excep: ‹¬twl_is_an_exception C Q WS›
have ‹get_level M L = count_decided M› and L: ‹-L ∈ lits_of_l M› and D: ‹D ∈# N + U›
using valid by auto
have ‹watched_literals_false_of_max_level M D›
using twl D by (auto simp: twl_st_inv.simps)
have ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: trail.simps)
then have [simp]: ‹- L' ∉ lits_of_l M›
using L' consistent_interp_def distinct_consistent_interp by blast
have ‹¬ twl_is_an_exception C Q (add_mset (L, D) WS)› if ‹C ≠ D›
using excep that by (auto simp add: twl_is_an_exception_def)
have twl_D: ‹twl_lazy_update M D›
using twl C excep twl watched L' ‹watched_literals_false_of_max_level M D›
by (cases D)
(auto simp: get_level_cons_if count_decided_ge_get_level has_blit_def
twl_is_an_exception_add_mset_to_queue atm_of_eq_atm_of count_decided_ge_get_level
dest!: no_has_blit_propagate' no_has_blit_propagate)
have twl_C: ‹twl_lazy_update M C› if ‹C ≠ D›
using twl C excep that by (auto simp add: twl_st_inv.simps
twl_is_an_exception_add_mset_to_clauses_to_update)
show ‹twl_lazy_update M C›
using twl_C twl_D by blast
qed
next
case (update_clause D L L' M K N U N' U' NE UE WS Q) note watched = this(1) and uL = this(2) and
L' = this(3) and K = this(4) and undef = this(5) and N'U' = this(6) and twl = this(7) and
valid = this(8) and inv = this(9) and twl_excep = this(10) and
no_dup = this(11) and wq = this(12)
obtain WD UWD where D: ‹D = TWL_Clause WD UWD› by (cases D)
have L: ‹L ∈# watched D› and D_N_U: ‹D ∈# N + U› and lev_L: ‹get_level M L = count_decided M›
using valid by auto
then have struct_D: ‹struct_wf_twl_cls D›
using twl by (auto simp: twl_st_inv.simps)
have L'_UWD: ‹L ∉# remove1_mset L' UWD› if ‹L ∈# WD› for L
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD L ≥ 1›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) L ≥ 2›
using D that by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have L'_L'_UWD: ‹K ∉# remove1_mset K UWD›
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD K ≥ 2›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) K ≥ 2›
using D L' by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps)
let ?D = ‹update_clause D L K›
have *: ‹C ∈# N + U› if ‹C ≠ ?D› and C: ‹C ∈# N' + U'› for C
using C N'U' that by (auto elim!: update_clausesE dest: in_diffD)
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
then have uK_M: ‹- K ∉ lits_of_l M›
using undef Decided_Propagated_in_iff_in_lits_of_l consistent_interp_def
distinct_consistent_interp by blast
have add_remove_WD: ‹add_mset K (remove1_mset L WD) ≠ WD›
using uK_M uL by (auto simp: add_mset_remove_trivial_iff trivial_add_mset_remove_iff)
have cls_D_D: ‹clause ?D = clause D›
by (cases D) (use watched K in auto)
have L_M: ‹L ∉ lits_of_l M›
using n_d uL by (fastforce dest!: distinct_consistent_interp
simp: consistent_interp_def lits_of_def uminus_lit_swap)
have w_max_D: ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps)
show ?case unfolding twl_st_simps Ball_def
proof (intro allI conjI impI)
fix C
assume C: ‹C ∈# N' + U'›
moreover have ‹L ≠ L'›
using struct_D watched by (auto simp: D dest: multi_member_split)
ultimately have struct_D': ‹struct_wf_twl_cls ?D›
using L K struct_D watched by (auto simp: D L'_UWD L'_L'_UWD dest: in_diffD)
have struct_C: ‹struct_wf_twl_cls C› if ‹C ≠ ?D›
using twl C that N'U' by (fastforce simp: twl_st_inv.simps elim!: update_clausesE
split: if_splits dest: in_diffD)
show ‹struct_wf_twl_cls C›
using struct_D' struct_C by blast
have H: ‹⋀C. C ∈# N+U ⟹ ¬ twl_is_an_exception C Q WS ⟹ C ≠ D ⟹
twl_lazy_update M C›
using twl
by (auto simp add: twl_st_inv.simps twl_is_an_exception_add_mset_to_clauses_to_update)
have ‹watched_literals_false_of_max_level M C› if ‹C ≠ ?D›
using twl C that N'U' by (fastforce simp: twl_st_inv.simps elim!: update_clausesE
dest: in_diffD)
moreover have ‹watched_literals_false_of_max_level M ?D›
using w_max_D D watched L' uK_M distinct_consistent_interp[OF n_d] uL K
apply (cases D)
apply (simp_all add: add_mset_eq_add_mset consistent_interp_def)
by (metis add_mset_eq_add_mset)
ultimately show ‹watched_literals_false_of_max_level M C›
by blast
assume excep: ‹¬twl_is_an_exception C Q WS›
have ‹get_level M L = count_decided M› and L: ‹-L ∈ lits_of_l M› and D_N_U: ‹D ∈# N + U›
using valid by auto
have excep_WS: ‹¬ twl_is_an_exception C Q WS›
using excep C by (force simp: twl_is_an_exception_def)
have excep_inv_D: ‹twl_exception_inv (M, N, U, None, NE, UE, add_mset (L, D) WS, Q) D›
using twl_excep D_N_U unfolding twl_st_exception_inv.simps
by blast
then have ‹¬ has_blit M (clause D) L ⟹
L ∉# Q ⟹ (L, D) ∉# add_mset (L, D) WS ⟹ (∀K∈#unwatched D. - K ∈ lits_of_l M)›
using watched L
unfolding twl_exception_inv.simps
apply auto
done
have NU_WS: ‹Pair L `# {#C ∈# N+U. clauses_to_update_prop Q M (L, C)#} ⊆# add_mset (L, D) WS›
using wq by auto
have ‹distinct_mset {#- lit_of x. x ∈# mset M#}›
by (subst distinct_image_mset_inj)
(use n_d in ‹auto simp: lit_of_inj_on_no_dup distinct_map no_dup_def›)
moreover have ‹add_mset L Q ⊆# {#- lit_of x. x ∈# mset M#}›
using no_dup by auto
ultimately have LQ[simp]: ‹L ∉# Q›
by (metis distinct_mset_add_mset distinct_mset_union subset_mset.le_iff_add)
have ‹twl_lazy_update M C› if CD: ‹C = D›
unfolding twl_lazy_update.simps CD D
proof (intro conjI impI allI)
fix K'
assume ‹K' ∈# WD› ‹- K' ∈ lits_of_l M›‹¬ has_blit M (WD + UWD) K'›
have C_D': ‹C ≠ update_clause D L K›
using D add_remove_WD that by auto
have H: ‹¬ has_blit M (add_mset L (add_mset L' UWD)) L' ⟹
has_blit M (add_mset L (add_mset L' UWD)) L ⟹ False›
using ‹- K' ∈ lits_of_l M› ‹K' ∈# WD› ‹¬ has_blit M (WD + UWD) K'›
lev_L w_max_D
using L_M by (auto simp: has_blit_def D)
obtain NU where NU: ‹N+U = add_mset D NU›
using multi_member_split[OF D_N_U] by auto
have ‹C ∈# remove1_mset D (N + U)›
using C C_D' N'U' unfolding NU
apply (auto simp: update_clauses.simps NU[symmetric])
using C by auto
then obtain NU' where ‹N+U = add_mset C (add_mset D NU')›
using NU multi_member_split by force
moreover have ‹clauses_to_update_prop Q M (L, D)›
using watched uL ‹¬ has_blit M (WD + UWD) K'› ‹K' ∈# WD› LQ
by (auto simp: clauses_to_update_prop.simps D dest: H)
ultimately have ‹(L, D) ∈# WS›
using NU_WS by (auto simp: CD split: if_splits)
then have False
using excep unfolding CD
by (auto simp: twl_is_an_exception_def)
then show ‹∀K∈#UWD. get_level M K ≤ get_level M K' ∧ - K ∈ lits_of_l M›
by fast
qed
moreover have ‹twl_lazy_update M C› if ‹C ≠ ?D› ‹C ≠ D›
using H[of C] that excep_WS * C
by (auto simp add: twl_st_inv.simps)[]
moreover {
have D': ‹?D = TWL_Clause {#K, L'#} (add_mset L (remove1_mset K UWD))› and
mset_D': ‹{#K, L'#} + add_mset L (remove1_mset K UWD) = clause D›
using D watched cls_D_D by auto
have lev_L': ‹get_level M L' = count_decided M› if ‹- L' ∈ lits_of_l M › and
‹¬ has_blit M (clause D) L'›
using L_M w_max_D D watched L' uL that
by simp
have ‹∀C. C ∈# WS ⟶ fst C = L›
using no_dup
using watched uL L' undef D
by (auto simp del: set_mset_union simp: )
then have ‹(L', TWL_Clause {#L, L'#} UWD) ∉# WS›
using wq multi_member_split[OF D_N_U] struct_D
using watched uL L' undef D
by auto
then have ‹- L' ∈ lits_of_l M ⟹ ¬ has_blit M (add_mset L (add_mset L' UWD)) L' ⟹
L' ∈# Q ›
using wq multi_member_split[OF D_N_U] struct_D
using watched uL L' undef D
by (auto simp del: set_mset_union simp: )
then have
H: ‹- L' ∈ lits_of_l M ⟹ ¬ has_blit M (add_mset L (add_mset L' UWD)) L' ⟹
False› if ‹C = ?D›
using excep multi_member_split[OF D_N_U] struct_D
using watched uL L' undef D that
by (auto simp del: set_mset_union simp: twl_is_an_exception_def)
have in_remove1_mset: ‹K' ∈# remove1_mset K UWD ⟷ K' ≠ K ∧ K' ∈# UWD› for K'
using struct_D L'_L'_UWD by (auto simp: D in_remove1_mset_neq dest: in_diffD)
have ‹twl_lazy_update M ?D› if ‹C = ?D›
using watched uL L' undef D w_max_D H
unfolding twl_lazy_update.simps D' mset_D' that
by (auto simp: uK_M D add_mset_eq_add_mset lev_L count_decided_ge_get_level
in_remove1_mset twl_is_an_exception_def)
}
ultimately show ‹twl_lazy_update M C›
by blast
qed
qed
lemma twl_cp_no_duplicate_queued:
assumes
cdcl: ‹cdcl_twl_cp S T› and
no_dup: ‹no_duplicate_queued S›
shows ‹no_duplicate_queued T›
using cdcl no_dup
proof (induction rule: cdcl_twl_cp.induct)
case (pop M N U NE UE L Q)
then show ?case
by (auto simp: image_Un image_image subset_mset.less_imp_le
dest: mset_subset_eq_insertD)
qed auto
lemma distinct_mset_Pair: ‹distinct_mset (Pair L `# C) ⟷ distinct_mset C›
by (induction C) auto
lemma distinct_image_mset_clause:
‹distinct_mset (clause `# C) ⟹ distinct_mset C›
by (induction C) auto
lemma twl_cp_distinct_queued:
assumes
cdcl: ‹cdcl_twl_cp S T› and
twl: ‹twl_st_inv S› and
valid: ‹valid_enqueued S› and
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)› and
no_dup: ‹no_duplicate_queued S› and
dist: ‹distinct_queued S›
shows ‹distinct_queued T›
using cdcl twl valid inv no_dup dist
proof (induction rule: cdcl_twl_cp.induct)
case (pop M N U NE UE L Q) note c_dist = this(4) and dist = this(5)
show ?case
using dist by (auto simp: distinct_mset_Pair count_image_mset_Pair simp del: image_mset_union)
next
case (propagate D L L' M N U NE UE WS Q) note watched = this(1) and undef = this(2) and
twl = this(4) and valid = this(5) and inv = this(6) and no_dup = this(7)
and dist = this(8)
have ‹L' ∉ lits_of_l M›
using Decided_Propagated_in_iff_in_lits_of_l propagate.hyps(2) by auto
then have ‹-L' ∉# Q›
using no_dup by (fastforce simp: lits_of_def dest!: mset_subset_eqD)
then show ?case
using dist by (auto simp: all_conj_distrib split: if_splits dest!: Suc_leD)
next
case (conflict D L L' M N U NE UE WS Q) note dist = this(8)
then show ?case
by auto
next
case (delete_from_working D L L' M N U NE UE WS Q) note dist = this(7)
show ?case using dist by (auto simp: all_conj_distrib split: if_splits dest!: Suc_leD)
next
case (update_clause D L L' M K N U N' U' NE UE WS Q) note watched = this(1) and uL = this(2) and
L' = this(3) and K = this(4) and undef = this(5) and N'U' = this(6) and twl = this(7) and
valid = this(8) and inv = this(9) and no_dup = this(10) and dist = this(11)
show ?case
unfolding distinct_queued.simps
proof (intro conjI allI)
show ‹distinct_mset Q›
using dist N'U' by (auto simp: all_conj_distrib split: if_splits intro: le_SucI)
fix K'' C
have LD: ‹Suc (count WS (L, D)) ≤ count N D + count U D›
using dist N'U' by (auto split: if_splits)
have LC: ‹count WS (La, Ca) ≤ count N Ca + count U Ca›
if ‹(La , Ca) ≠ (L, D)› for Ca La
using dist N'U' by (force simp: all_conj_distrib split: if_splits intro: le_SucI)
show ‹count WS (K'', C) ≤ count (N' + U') C›
proof (cases ‹K'' ≠ L›)
case True
then have ‹count WS (K'', C) = 0›
using no_dup by auto
then show ?thesis by arith
next
case False
then show ?thesis
apply (cases ‹C = D›)
using LD N'U' apply (auto simp: all_conj_distrib elim!: update_clausesE intro: le_SucI;
fail)
using LC[of L C] N'U' by (auto simp: all_conj_distrib elim!: update_clausesE intro: le_SucI)
qed
qed
qed
lemma twl_cp_valid:
assumes
cdcl: ‹cdcl_twl_cp S T› and
twl: ‹twl_st_inv S› and
valid: ‹valid_enqueued S› and
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)› and
no_dup: ‹no_duplicate_queued S› and
dist: ‹distinct_queued S›
shows ‹valid_enqueued T›
using cdcl twl valid inv no_dup dist
proof (induction rule: cdcl_twl_cp.induct)
case (pop M N U NE UE L Q) note valid = this(2)
then show ?case
by (auto simp del: filter_union_mset)
next
case (propagate D L L' M N U NE UE WS Q) note watched = this(1) and twl = this(4) and
valid = this(5) and inv = this(6) and no_taut = this(7)
show ?case
using valid by (auto dest: mset_subset_eq_insertD simp: get_level_cons_if)
next
case (conflict D L L' M N U NE UE WS Q) note valid = this(5)
then show ?case
by auto
next
case (delete_from_working D L L' M N U NE UE WS Q) note watched = this(1) and L' = this(2) and
twl = this(3) and valid = this(4) and inv = this(5)
show ?case unfolding twl_st_simps Ball_def
using valid by (auto dest: mset_subset_eq_insertD)
next
case (update_clause D L L' M K N U N' U' NE UE WS Q) note watched = this(1) and uL = this(2) and
L' = this(3) and K = this(4) and undef = this(5) and N'U' = this(6) and twl = this(7) and
valid = this(8) and inv = this(9) and no_dup = this(10) and dist = this(11)
show ?case
unfolding valid_enqueued.simps Ball_def
proof (intro allI impI conjI)
fix L :: ‹'a literal›
assume L: ‹L ∈# Q›
then show ‹-L ∈ lits_of_l M›
using valid by auto
show ‹get_level M L = count_decided M›
using L valid by auto
next
fix KC :: ‹'a literal × 'a twl_cls›
assume LC_WS: ‹KC ∈# WS›
obtain K'' C where LC: ‹KC = (K'', C)› by (cases KC)
have ‹K'' ∈# watched C›
using LC_WS valid LC by auto
have C_ne_D: ‹case KC of (L, C) ⇒ L ∈# watched C ∧ C ∈# N' + U' ∧ - L ∈ lits_of_l M ∧
get_level M L = count_decided M› if ‹C ≠ D›
by (cases ‹C = D›)
(use valid LC LC_WS N'U' that in ‹auto simp: in_remove1_mset_neq elim!: update_clausesE›)
have K''_L: ‹K'' = L›
using no_dup LC_WS LC by auto
have ‹Suc (count WS (L, D)) ≤ count N D + count U D›
using dist by (auto simp: all_conj_distrib split: if_splits)
then have D_DN_U: ‹D ∈# remove1_mset D (N+U)› if [simp]: ‹C = D›
using LC_WS unfolding count_greater_zero_iff[symmetric]
by (auto simp del: count_greater_zero_iff simp: LC K''_L)
have D_D_N: ‹D ∈# remove1_mset D N› if ‹D ∈# N› and ‹D ∉# U› and [simp]: ‹C = D›
proof -
have ‹D ∈# remove1_mset D (U + N)›
using D_DN_U by (simp add: union_commute)
then have ‹D ∈# U + remove1_mset D N›
using that(1) by (metis (no_types) add_mset_remove_trivial insert_DiffM
union_mset_add_mset_right)
then show ‹D ∈# remove1_mset D N›
using that(2) by (meson union_iff)
qed
have D_D_U: ‹D ∈# remove1_mset D U› if ‹D ∈# U› and ‹D ∉# N› and [simp]: ‹C = D›
proof -
have ‹D ∈# remove1_mset D (U + N)›
using D_DN_U by (simp add: union_commute)
then have ‹D ∈# N + remove1_mset D U›
using D_DN_U that(1) by fastforce
then show ‹D ∈# remove1_mset D U›
using that(2) by (meson union_iff)
qed
have CD: ‹case KC of (L, C) ⇒ L ∈# watched C ∧ C ∈# N' + U' ∧ - L ∈ lits_of_l M ∧
get_level M L = count_decided M› if ‹C = D›
by (use valid LC_WS N'U' in ‹auto simp: LC D_D_N that in_remove1_mset_neq
dest!: D_D_U elim!: update_clausesE›)
show ‹case KC of (L, C) ⇒ L ∈# watched C ∧ C ∈# N' + U' ∧ - L ∈ lits_of_l M ∧
get_level M L = count_decided M›
using CD C_ne_D by blast
qed
qed
lemma twl_cp_propa_cands_enqueued:
assumes
cdcl: ‹cdcl_twl_cp S T› and
twl: ‹twl_st_inv S› and
valid: ‹valid_enqueued S› and
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)› and
twl_excep: ‹twl_st_exception_inv S› and
no_dup: ‹no_duplicate_queued S› and
cands: ‹propa_cands_enqueued S› and
ws: ‹clauses_to_update_inv S›
shows ‹propa_cands_enqueued T›
using cdcl twl valid inv twl_excep no_dup cands ws
proof (induction rule: cdcl_twl_cp.induct)
case (pop M N U NE UE L Q) note inv = this(1) and valid = this(2) and cands = this(6)
show ?case unfolding propa_cands_enqueued.simps
proof (intro allI conjI impI)
fix C K
assume C: ‹C ∈# N + U› and
‹K ∈# clause C› and
‹M ⊨as CNot (remove1_mset K (clause C))› and
‹undefined_lit M K›
then have ‹(∃L'. L' ∈# watched C ∧ L' ∈# add_mset L Q)›
using cands by auto
then show
‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨
(∃La. (La, C) ∈# Pair L `# {#C ∈# N + U. L ∈# watched C#})›
using C by auto
qed
next
case (propagate D L L' M N U NE UE WS Q) note watched = this(1) and undef = this(2) and
false = this(3) and
twl = this(4) and valid = this(5) and inv = this(6) and excep = this(7)
and no_dup = this(8) and cands = this(9) and to_upd = this(10)
have uL'_M: ‹- L' ∉ lits_of_l M›
using Decided_Propagated_in_iff_in_lits_of_l propagate.hyps(2) by blast
have D_N_U: ‹D ∈# N + U›
using valid by auto
then have wf_D: ‹struct_wf_twl_cls D›
using twl by (simp add: twl_st_inv.simps)
show ?case unfolding propa_cands_enqueued.simps
proof (intro allI conjI impI)
fix C K
assume C: ‹C ∈# N + U› and
K: ‹K ∈# clause C› and
L'_M_C: ‹Propagated L' (clause D) # M ⊨as CNot (remove1_mset K (clause C))› and
undef_K: ‹undefined_lit (Propagated L' (clause D) # M) K›
then have wf_C: ‹struct_wf_twl_cls C›
using twl by (simp add: twl_st_inv.simps)
have undef_K_M: ‹undefined_lit M K›
using undef_K by (simp add: Decided_Propagated_in_iff_in_lits_of_l)
consider
(no_L') ‹M ⊨as CNot (remove1_mset K (clause C))› |
(L') ‹-L' ∈# remove1_mset K (clause C)›
using L'_M_C ‹- L' ∉ lits_of_l M›
by (metis insertE list.simps(15) lit_of.simps(2) lits_of_insert
true_annots_CNot_lit_of_notin_skip true_annots_true_cls_def_iff_negation_in_model)
then show ‹(∃L'a. L'a ∈# watched C ∧ L'a ∈# add_mset (- L') Q) ∨ (∃L. (L, C) ∈# WS)›
proof cases
case no_L'
then have ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃La. (La, C) ∈# add_mset (L, D) WS)›
using cands C K undef_K_M by auto
moreover {
have ‹K = L'› if ‹C = D›
by (metis ‹- L' ∉ lits_of_l M› add_mset_add_single clause.simps in_CNot_implies_uminus(2)
in_remove1_mset_neq multi_member_this no_L' that twl_clause.exhaust twl_clause.sel(1)
union_iff watched)
then have False if ‹C = D›
using undef_K by (simp add: Decided_Propagated_in_iff_in_lits_of_l that)
}
ultimately show ?thesis by auto
next
case L'
have ?thesis if ‹L' ∈# watched C›
proof -
have ‹K = L'›
using that L'_M_C ‹- L' ∉ lits_of_l M› L' undef
by (metis clause.simps in_CNot_implies_uminus(2) in_lits_of_l_defined_litD
in_remove1_mset_neq insert_iff list.simps(15) lits_of_insert
twl_clause.exhaust_sel uminus_not_id' uminus_of_uminus_id union_iff)
then have False
using Decided_Propagated_in_iff_in_lits_of_l undef_K by force
then show ?thesis
by fastforce
qed
moreover have ?thesis if L'_C: ‹L' ∉# watched C›
proof (rule ccontr, clarsimp)
assume
Q: ‹∀L'a. L'a ∈# watched C ⟶ L'a ≠ - L' ∧ L'a ∉# Q› and
WS: ‹∀L. (L, C) ∉# WS›
then have ‹¬ twl_is_an_exception C (add_mset (- L') Q) WS›
by (auto simp: twl_is_an_exception_def)
moreover have
‹twl_st_inv (Propagated L' (clause D) # M, N, U, None, NE, UE, WS, add_mset (- L') Q)›
using twl_cp_twl_inv[OF _ twl valid inv excep no_dup to_upd]
cdcl_twl_cp.propagate[OF propagate(1-3)] by fast
ultimately have ‹twl_lazy_update (Propagated L' (clause D) # M) C›
using C by (auto simp: twl_st_inv.simps)
have CD: ‹C ≠ D›
using that watched by auto
have struct: ‹struct_wf_twl_cls C›
using twl C by (simp add: twl_st_inv.simps)
obtain a b W UW where
C_W_UW: ‹C = TWL_Clause W UW› and
W: ‹W = {#a, b#}›
using struct by (cases C, auto simp: size_2_iff)
have ua_or_ub: ‹-a ∈ lits_of_l M ∨ -b ∈ lits_of_l M›
using L'_M_C C_W_UW W ‹∀L'a. L'a ∈# watched C ⟶ L'a ≠ - L' ∧ L'a ∉# Q›
apply (cases ‹K = a›) by fastforce+
have ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: trail.simps)
then have [dest]: False if ‹a ∈ lits_of_l M› and ‹-a ∈ lits_of_l M› for a
using consistent_interp_def distinct_consistent_interp that(1) that(2) by blast
have uab: ‹a ∉ lits_of_l M› if ‹-b ∈ lits_of_l M›
using L'_M_C C_W_UW W that undef_K_M uL'_M
by (cases ‹K = a›) (fastforce simp: Decided_Propagated_in_iff_in_lits_of_l
simp del: uL'_M)+
have uba: ‹b ∉ lits_of_l M› if ‹-a ∈ lits_of_l M›
using L'_M_C C_W_UW W that undef_K_M uL'_M
by (cases ‹K = b›) (fastforce simp: Decided_Propagated_in_iff_in_lits_of_l
add_mset_commute[of a b])+
have [simp]: ‹-a ≠ L'› ‹-b ≠ L'›
using Q W C_W_UW by fastforce+
have H': ‹∀La L'. watched C = {#La, L'#} ⟶ - La ∈ lits_of_l M ⟶
¬has_blit M (clause C) La ⟶ L' ∉ lits_of_l M ⟶
(∀K∈#unwatched C. - K ∈ lits_of_l M)›
using excep C CD Q W WS uab uba by (auto simp: twl_exception_inv.simps simp del: set_mset_union
dest: multi_member_split)
moreover have ‹watched C = {#La, L''#} ⟶- La ∈ lits_of_l M ⟶ ¬has_blit M (clause C) La› for La L''
using in_CNot_implies_uminus[OF _ L'_M_C] wf_C L' uL'_M undef_K_M undef uab uba
unfolding C_W_UW has_blit_def apply -
apply (cases ‹La = K›)
apply (auto simp: has_blit_def Decided_Propagated_in_iff_in_lits_of_l W
add_mset_eq_add_mset in_remove1_mset_neq)
apply (metis ‹⋀a. ⟦a ∈ lits_of_l M; - a ∈ lits_of_l M⟧ ⟹ False› add_mset_remove_trivial
defined_lit_uminus in_lits_of_l_defined_litD in_remove1_mset_neq undef)
apply (metis ‹⋀a. ⟦a ∈ lits_of_l M; - a ∈ lits_of_l M⟧ ⟹ False› add_mset_remove_trivial
defined_lit_uminus in_lits_of_l_defined_litD in_remove1_mset_neq undef)
done
ultimately have ‹∀K∈#unwatched C. - K ∈ lits_of_l M›
using uab uba W C_W_UW ua_or_ub wf_C unfolding C_W_UW
by (auto simp: add_mset_eq_add_mset )
then show False
by (metis Decided_Propagated_in_iff_in_lits_of_l L' uminus_lit_swap
Q clause.simps in_diffD propagate.hyps(2) twl_clause.collapse union_iff)
qed
ultimately show ?thesis by fast
qed
qed
next
case (conflict D L L' M N U NE UE WS Q) note cands = this(10)
then show ?case
by auto
next
case (delete_from_working L' D M N U NE UE L WS Q) note watched = this(1) and L' = this(2) and
twl = this(3) and valid = this(4) and inv = this(5) and cands = this(8) and ws = this(9)
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: trail.simps)
show ?case unfolding propa_cands_enqueued.simps
proof (intro allI conjI impI)
fix C K
assume C: ‹C ∈# N + U› and
K: ‹K ∈# clause C› and
L'_M_C: ‹M ⊨as CNot (remove1_mset K (clause C))› and
undef_K: ‹undefined_lit M K›
then have ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃La. La = L ∧ C = D ∨ (La, C) ∈# WS)›
using cands by auto
moreover have False if [simp]: ‹C = D›
using L' L'_M_C undef_K watched
using Decided_Propagated_in_iff_in_lits_of_l consistent_interp_def distinct_consistent_interp
local.K n_d K
by (cases D)
(auto 5 5 simp: true_annots_true_cls_def_iff_negation_in_model add_mset_eq_add_mset
dest: in_lits_of_l_defined_litD no_dup_consistentD dest!: multi_member_split)
ultimately show ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃L. (L, C) ∈# WS)›
by auto
qed
next
case (update_clause D L L' M K N U N' U' NE UE WS Q) note watched = this(1) and uL = this(2) and
L' = this(3) and K = this(4) and undef = this(5) and N'U' = this(6) and twl = this(7) and
valid = this(8) and inv = this(9) and twl_excep = this(10) and no_dup = this(11) and
cands = this(12) and ws = this(13)
obtain WD UWD where D: ‹D = TWL_Clause WD UWD› by (cases D)
have L: ‹L ∈# watched D› and D_N_U: ‹D ∈# N + U› and lev_L: ‹get_level M L = count_decided M›
using valid by auto
then have struct_D: ‹struct_wf_twl_cls D›
using twl by (auto simp: twl_st_inv.simps)
have L'_UWD: ‹L ∉# remove1_mset L' UWD› if ‹L ∈# WD› for L
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD L ≥ 1›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) L ≥ 2›
using D that by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have L'_L'_UWD: ‹K ∉# remove1_mset K UWD›
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD K ≥ 2›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) K ≥ 2›
using D L' by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps)
let ?D = ‹update_clause D L K›
have *: ‹C ∈# N + U› if ‹C ≠ ?D› and C: ‹C ∈# N' + U'› for C
using C N'U' that by (auto elim!: update_clausesE dest: in_diffD)
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
then have uK_M: ‹- K ∉ lits_of_l M›
using undef Decided_Propagated_in_iff_in_lits_of_l consistent_interp_def
distinct_consistent_interp by blast
have add_remove_WD: ‹add_mset K (remove1_mset L WD) ≠ WD›
using uK_M uL by (auto simp: add_mset_remove_trivial_iff trivial_add_mset_remove_iff)
have D_N_U: ‹D ∈# N + U›
using N'U' D uK_M uL D_N_U by (auto simp: add_mset_remove_trivial_iff split: if_splits)
have D_ne_D: ‹D ≠ update_clause D L K›
using D add_remove_WD by auto
have L_M: ‹L ∉ lits_of_l M›
using n_d uL by (fastforce dest!: distinct_consistent_interp
simp: consistent_interp_def lits_of_def uminus_lit_swap)
have w_max_D: ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps)
have clause_D: ‹clause ?D = clause D›
using D K watched by auto
show ?case unfolding propa_cands_enqueued.simps
proof (intro allI conjI impI)
fix C K2
assume C: ‹C ∈# N' + U'› and
K: ‹K2 ∈# clause C› and
L'_M_C: ‹M ⊨as CNot (remove1_mset K2 (clause C))› and
undef_K: ‹undefined_lit M K2›
then have ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃La. (La, C) ∈# WS)› if ‹C ≠ ?D› ‹C ≠ D›
using cands *[OF that(1) C] that(2) by auto
moreover have ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃L. (L, C) ∈# WS)› if [simp]: ‹C = ?D›
proof (rule ccontr)
have ‹K ∉ lits_of_l M›
by (metis D Decided_Propagated_in_iff_in_lits_of_l L'_M_C add_diff_cancel_left'
clause.simps clause_D in_diffD in_remove1_mset_neq that
true_annots_true_cls_def_iff_negation_in_model twl_clause.sel(2) uK_M undef_K
update_clause.hyps(4))
moreover have ‹∀L∈#remove1_mset K2 (clause ?D). defined_lit M L›
using L'_M_C unfolding true_annots_true_cls_def_iff_negation_in_model
by (auto simp: clause_D Decided_Propagated_in_iff_in_lits_of_l)
ultimately have [simp]: ‹K2 = K›
using undef undef_K K unfolding that clause_D
by (metis D clause.simps in_remove1_mset_neq twl_clause.sel(2) union_iff
update_clause.hyps(4))
have uL'_M: ‹- L' ∈ lits_of_l M›
using D watched L'_M_C by auto
have [simp]: ‹L ≠ L'› ‹L' ≠ L›
using struct_D D watched by auto
assume ‹¬ ((∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃L. (L, C) ∈# WS))›
then have [simp]: ‹L' ∉# Q› and L'_C_WS: ‹(L', C) ∉# WS›
using watched D by auto
have ‹C ∈# add_mset (L, TWL_Clause WD UWD) WS ⟶
C' ∈# add_mset (L, TWL_Clause WD UWD) WS ⟶
fst C = fst C'› for C C'
using no_dup unfolding D no_duplicate_queued.simps
by blast
from this[of ‹(L, TWL_Clause WD UWD)› ‹(L', TWL_Clause {#L, L'#} UWD)›]
have notin: ‹False› if ‹(L', TWL_Clause {#L, L'#} UWD) ∈# WS›
using struct_D watched that unfolding D
by auto
have ‹?D ≠ D›
using C D watched L K uK_M uL by auto
then have excep: ‹twl_exception_inv (M, N, U, None, NE, UE, add_mset (L, D) WS, Q) D›
using twl_excep *[of D] D_N_U by (auto simp: twl_st_inv.simps)
moreover have ‹D = TWL_Clause {#L, L'#} UWD ⟹
WD = {#L, L'#} ⟹
∀L∈#remove1_mset K UWD.
- L ∈ lits_of_l M ⟹
¬has_blit M (add_mset L (add_mset L' UWD)) L'›
using uL uL'_M n_d ‹K ∉ lits_of_l M› unfolding has_blit_def
apply (auto dest:no_dup_consistentD simp: in_remove1_mset_neq Ball_def)
by (metis in_remove1_mset_neq no_dup_consistentD)
ultimately have ‹∀K ∈# unwatched D. -K ∈ lits_of_l M›
using D watched L'_M_C L'_C_WS
by (auto simp: add_mset_eq_add_mset uL'_M L_M uL twl_exception_inv.simps
true_annots_true_cls_def_iff_negation_in_model dest: in_diffD notin)
then show False
using uK_M update_clause.hyps(4) by blast
qed
moreover have ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃L. (L, C) ∈# WS)› if [simp]: ‹C = D›
unfolding that
proof -
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
obtain NU where NU: ‹N + U = add_mset D NU›
by (metis D_N_U insert_DiffM)
have N'U': ‹N' + U' = add_mset ?D (remove1_mset D (N + U))›
using N'U' D_N_U by (auto elim!: update_clausesE)
have ‹add_mset L Q ⊆# {#- lit_of x. x ∈# mset M#}›
using no_dup by (auto)
moreover have ‹distinct_mset {#- lit_of x. x ∈# mset M#}›
by (subst distinct_image_mset_inj)
(use n_d in ‹auto simp: lit_of_inj_on_no_dup distinct_map no_dup_def›)
ultimately have [simp]: ‹L ∉# Q›
by (metis distinct_mset_add_mset distinct_mset_union subset_mset.le_iff_add)
have ‹has_blit M (clause D) L ⟹ False›
by (smt K L'_M_C has_blit_def in_lits_of_l_defined_litD insert_DiffM insert_iff
is_blit_def n_d no_dup_consistentD set_mset_add_mset_insert that
true_annots_true_cls_def_iff_negation_in_model undef_K)
then have w_q_p_D: ‹clauses_to_update_prop Q M (L, D)›
by (auto simp: clauses_to_update_prop.simps watched)
(use uL undef L' in ‹auto simp: Decided_Propagated_in_iff_in_lits_of_l›)
have ‹Pair L `# {#C ∈# add_mset D NU. clauses_to_update_prop Q M (L, C)#} ⊆#
add_mset (L, D) WS›
using ws no_dup unfolding clauses_to_update_inv.simps NU
by (auto simp: all_conj_distrib)
then have IH: ‹Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#} ⊆# WS›
using w_q_p_D by auto
moreover have ‹(L, D) ∈# Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#}›
using C D_ne_D w_q_p_D unfolding NU N'U' by (auto simp: pair_in_image_Pair)
ultimately show ‹(∃L'. L' ∈# watched D ∧ L' ∈# Q) ∨ (∃L. (L, D) ∈# WS)›
by blast
qed
ultimately show ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃L. (L, C) ∈# WS)›
by auto
qed
qed
lemma twl_cp_confl_cands_enqueued:
assumes
cdcl: ‹cdcl_twl_cp S T› and
twl: ‹twl_st_inv S› and
valid: ‹valid_enqueued S› and
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)› and
excep: ‹twl_st_exception_inv S› and
no_dup: ‹no_duplicate_queued S› and
cands: ‹confl_cands_enqueued S› and
ws: ‹clauses_to_update_inv S›
shows
‹confl_cands_enqueued T›
using cdcl
proof (induction rule: cdcl_twl_cp.cases)
case (pop M N U NE UE L Q) note S = this(1) and T = this(2)
show ?case unfolding confl_cands_enqueued.simps Ball_def S T
proof (intro allI conjI impI)
fix C K
assume C: ‹C ∈# N + U› and
‹M ⊨as CNot (clause C)›
then have ‹(∃L'. L' ∈# watched C ∧ L' ∈# add_mset L Q)›
using cands S by auto
then show
‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨
(∃La. (La, C) ∈# Pair L `# {#C ∈# N + U. L ∈# watched C#})›
using C by auto
qed
next
case (propagate D L L' M N U NE UE WS Q) note S = this(1) and T = this(2) and watched = this(3)
and undef = this(4)
have uL'_M: ‹- L' ∉ lits_of_l M›
using Decided_Propagated_in_iff_in_lits_of_l undef by blast
have D_N_U: ‹D ∈# N + U›
using valid S by auto
then have wf_D: ‹struct_wf_twl_cls D›
using twl by (simp add: twl_st_inv.simps S)
show ?case unfolding confl_cands_enqueued.simps Ball_def S T
proof (intro allI conjI impI)
fix C K
assume C: ‹C ∈# N + U› and
L'_M_C: ‹Propagated L' (clause D) # M ⊨as CNot (clause C)›
consider
(no_L') ‹M ⊨as CNot (clause C)›
| (L') ‹-L' ∈# clause C›
using L'_M_C ‹- L' ∉ lits_of_l M›
by (metis insertE list.simps(15) lit_of.simps(2) lits_of_insert
true_annots_CNot_lit_of_notin_skip true_annots_true_cls_def_iff_negation_in_model)
then show ‹(∃L'a. L'a ∈# watched C ∧ L'a ∈# add_mset (- L') Q) ∨ (∃L. (L, C) ∈# WS)›
proof cases
case no_L'
then have ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃La. (La, C) ∈# add_mset (L, D) WS)›
using cands C by (auto simp: S)
moreover {
have ‹C ≠ D›
by (metis ‹- L' ∉ lits_of_l M› add_mset_add_single clause.simps in_CNot_implies_uminus(2)
multi_member_this no_L' twl_clause.exhaust twl_clause.sel(1)
union_iff watched)
}
ultimately show ?thesis by auto
next
case L'
have L'_C: ‹L' ∉# watched C›
using L'_M_C ‹- L' ∉ lits_of_l M›
by (metis (no_types, hide_lams) Decided_Propagated_in_iff_in_lits_of_l L' clause.simps
in_CNot_implies_uminus(2) insertE list.simps(15) lits_of_insert twl_clause.exhaust_sel
uminus_not_id' uminus_of_uminus_id undef union_iff)
moreover have ?thesis
proof (rule ccontr, clarsimp)
assume
Q: ‹∀L'a. L'a ∈# watched C ⟶ L'a ≠ - L' ∧ L'a ∉# Q› and
WS: ‹∀L. (L, C) ∉# WS›
then have ‹¬ twl_is_an_exception C (add_mset (- L') Q) WS›
by (auto simp: twl_is_an_exception_def)
moreover have
‹twl_st_inv (Propagated L' (clause D) # M, N, U, None, NE, UE, WS, add_mset (- L') Q)›
using twl_cp_twl_inv[OF _ twl valid inv excep no_dup ws] cdcl unfolding S T by fast
ultimately have ‹twl_lazy_update (Propagated L' (clause D) # M) C›
using C by (auto simp: twl_st_inv.simps)
have struct: ‹struct_wf_twl_cls C›
using twl C by (simp add: twl_st_inv.simps S)
have CD: ‹C ≠ D›
using L'_C watched by auto
have struct: ‹struct_wf_twl_cls C›
using twl C by (simp add: twl_st_inv.simps S)
obtain a b W UW where
C_W_UW: ‹C = TWL_Clause W UW› and
W: ‹W = {#a, b#}›
using struct by (cases C) (auto simp: size_2_iff)
have ua_ub: ‹-a ∈ lits_of_l M ∨ -b ∈ lits_of_l M›
using L'_M_C C_W_UW W ‹∀L'a. L'a ∈# watched C ⟶ L'a ≠ - L' ∧ L'a ∉# Q›
by (cases ‹K = a›) fastforce+
have ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: trail.simps S)
then have [dest]: False if ‹a ∈ lits_of_l M› and ‹-a ∈ lits_of_l M› for a
using consistent_interp_def distinct_consistent_interp that(1) that(2) by blast
have uab: ‹a ∉ lits_of_l M› if ‹-b ∈ lits_of_l M›
using L'_M_C C_W_UW W that uL'_M by (cases ‹K = a›) auto
have uba: ‹b ∉ lits_of_l M› if ‹-a ∈ lits_of_l M›
using L'_M_C C_W_UW W that uL'_M by (cases ‹K = b›) auto
have [simp]: ‹-a ≠ L'› ‹-b ≠ L'›
using ‹∀L'a. L'a ∈# watched C ⟶ L'a ≠ - L' ∧ L'a ∉# Q› W C_W_UW
by fastforce+
have H': ‹∀La L'. watched C = {#La, L'#} ⟶ - La ∈ lits_of_l M ⟶ L' ∉ lits_of_l M ⟶
¬ has_blit M (clause C) La ⟶(∀K∈#unwatched C. - K ∈ lits_of_l M)›
using excep C CD Q W WS uab uba
by (auto simp: twl_exception_inv.simps S dest: multi_member_split)
moreover have ‹¬ has_blit M (clause C) a› ‹¬ has_blit M (clause C) b›
using multi_member_split[OF C]
using watched L' undef L'_M_C
unfolding has_blit_def
by (metis (no_types, lifting) Clausal_Logic.uminus_lit_swap
‹⋀a. ⟦a ∈ lits_of_l M; - a ∈ lits_of_l M⟧ ⟹ False› in_CNot_implies_uminus(2)
in_lits_of_l_defined_litD insert_iff is_blit_def list.set(2) lits_of_insert uL'_M)+
ultimately have ‹∀K∈#unwatched C. - K ∈ lits_of_l M›
using uab uba W C_W_UW ua_ub struct
by (auto simp: add_mset_eq_add_mset)
then show False
by (metis Decided_Propagated_in_iff_in_lits_of_l L' uminus_lit_swap
Q clause.simps undef twl_clause.collapse union_iff)
qed
ultimately show ?thesis by fast
qed
qed
next
case (conflict D L L' M N U NE UE WS Q)
then show ?case
by auto
next
case (delete_from_working L' D M N U NE UE L WS Q) note S = this(1) and T = this(2) and
watched = this(3) and L' = this(4)
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: trail.simps S)
show ?case unfolding confl_cands_enqueued.simps Ball_def S T
proof (intro allI conjI impI)
fix C
assume C: ‹C ∈# N + U› and
L'_M_C: ‹M ⊨as CNot (clause C)›
then have ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃La. La = L ∧ C = D ∨ (La, C) ∈# WS)›
using cands S by auto
moreover have False if [simp]: ‹C = D›
using L'_M_C watched L' n_d by (cases D) (auto dest!: distinct_consistent_interp
simp: consistent_interp_def dest!: multi_member_split)
ultimately show ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃L. (L, C) ∈# WS)›
by auto
qed
next
case (update_clause D L L' M K N U N' U' NE UE WS Q) note S = this(1) and T = this(2) and
watched = this(3) and uL = this(4) and L' = this(5) and K = this(6) and undef = this(7) and
N'U' = this(8)
obtain WD UWD where D: ‹D = TWL_Clause WD UWD› by (cases D)
have L: ‹L ∈# watched D› and D_N_U: ‹D ∈# N + U› and lev_L: ‹get_level M L = count_decided M›
using valid S by auto
then have struct_D: ‹struct_wf_twl_cls D›
using twl by (auto simp: twl_st_inv.simps S)
have L'_UWD: ‹L ∉# remove1_mset L' UWD› if ‹L ∈# WD› for L
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD L ≥ 1›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) L ≥ 2›
using D that by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have L'_L'_UWD: ‹K ∉# remove1_mset K UWD›
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD K ≥ 2›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) K ≥ 2›
using D L' by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps S)
let ?D = ‹update_clause D L K›
have *: ‹C ∈# N + U› if ‹C ≠ ?D› and C: ‹C ∈# N' + U'› for C
using C N'U' that by (auto elim!: update_clausesE dest: in_diffD)
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps S)
then have uK_M: ‹- K ∉ lits_of_l M›
using undef Decided_Propagated_in_iff_in_lits_of_l consistent_interp_def
distinct_consistent_interp by blast
have add_remove_WD: ‹add_mset K (remove1_mset L WD) ≠ WD›
using uK_M uL by (auto simp: add_mset_remove_trivial_iff trivial_add_mset_remove_iff)
have D_N_U: ‹D ∈# N + U›
using N'U' D uK_M uL D_N_U by (auto simp: add_mset_remove_trivial_iff split: if_splits)
have D_ne_D: ‹D ≠ update_clause D L K›
using D add_remove_WD by auto
have L_M: ‹L ∉ lits_of_l M›
using n_d uL by (fastforce dest!: distinct_consistent_interp
simp: consistent_interp_def lits_of_def uminus_lit_swap)
have w_max_D: ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps S)
have clause_D: ‹clause ?D = clause D›
using D K watched by auto
show ?case unfolding confl_cands_enqueued.simps Ball_def S T
proof (intro allI conjI impI)
fix C
assume C: ‹C ∈# N' + U'› and
L'_M_C: ‹M ⊨as CNot (clause C)›
then have ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃La. (La, C) ∈# WS)› if ‹C ≠ ?D› ‹C ≠ D›
using cands *[OF that(1) C] that(2) S by auto
moreover have ‹C ≠ ?D›
by (metis D L'_M_C add_diff_cancel_left' clause.simps clause_D in_diffD
true_annots_true_cls_def_iff_negation_in_model twl_clause.sel(2) uK_M K)
moreover have ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃La. (La, C) ∈# WS)› if [simp]: ‹C = D›
unfolding that
proof -
obtain NU where NU: ‹N + U = add_mset D NU›
by (metis D_N_U insert_DiffM)
have N'U': ‹N' + U' = add_mset ?D (remove1_mset D (N + U))›
using N'U' D_N_U by (auto elim!: update_clausesE)
have ‹add_mset L Q ⊆# {#- lit_of x. x ∈# mset M#}›
using no_dup by (auto simp: S)
moreover have ‹distinct_mset {#- lit_of x. x ∈# mset M#}›
by (subst distinct_image_mset_inj)
(use n_d in ‹auto simp: lit_of_inj_on_no_dup distinct_map no_dup_def›)
ultimately have [simp]: ‹L ∉# Q›
by (metis distinct_mset_add_mset distinct_mset_union subset_mset.le_iff_add)
have ‹has_blit M (clause D) L ⟹ False›
by (smt K L'_M_C has_blit_def in_lits_of_l_defined_litD insert_DiffM insert_iff
is_blit_def n_d no_dup_consistentD set_mset_add_mset_insert that
true_annots_true_cls_def_iff_negation_in_model)
then have w_q_p_D: ‹clauses_to_update_prop Q M (L, D)›
by (auto simp: clauses_to_update_prop.simps watched)
(use uL undef L' in ‹auto simp: Decided_Propagated_in_iff_in_lits_of_l›)
have ‹Pair L `# {#C ∈# add_mset D NU. clauses_to_update_prop Q M (L, C)#} ⊆#
add_mset (L, D) WS›
using ws no_dup unfolding clauses_to_update_inv.simps NU S
by (auto simp: all_conj_distrib)
then have IH: ‹Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#} ⊆# WS›
using w_q_p_D by auto
moreover have ‹(L, D) ∈# Pair L `# {#C ∈# NU. clauses_to_update_prop Q M (L, C)#}›
using C D_ne_D w_q_p_D unfolding NU N'U' by (auto simp: pair_in_image_Pair)
ultimately show ‹(∃L'. L' ∈# watched D ∧ L' ∈# Q) ∨ (∃L. (L, D) ∈# WS)›
by blast
qed
ultimately show ‹(∃L'. L' ∈# watched C ∧ L' ∈# Q) ∨ (∃L. (L, C) ∈# WS)›
by auto
qed
qed
lemma twl_cp_past_invs:
assumes
cdcl: ‹cdcl_twl_cp S T› and
twl: ‹twl_st_inv S› and
valid: ‹valid_enqueued S› and
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)› and
twl_excep: ‹twl_st_exception_inv S› and
no_dup: ‹no_duplicate_queued S› and
past_invs: ‹past_invs S›
shows ‹past_invs T›
using cdcl twl valid inv twl_excep no_dup past_invs
proof (induction rule: cdcl_twl_cp.induct)
case (pop M N U NE UE L Q) note past_invs = this(6)
then show ?case
by (subst past_invs_enqueud, subst (asm) past_invs_enqueud)
next
case (propagate D L L' M N U NE UE WS Q) note watched = this(1) and twl = this(4) and
valid = this(5) and inv = this(6) and past_invs = this(9)
have [simp]: ‹- L' ∉ lits_of_l M›
using Decided_Propagated_in_iff_in_lits_of_l propagate.hyps(2) by blast
have D_N_U: ‹D ∈# N + U›
using valid by auto
then have wf_D: ‹struct_wf_twl_cls D›
using twl by (simp add: twl_st_inv.simps)
show ?case unfolding past_invs.simps Ball_def
proof (intro allI conjI impI)
fix C
assume C: ‹C ∈# N + U›
fix M1 M2 :: ‹('a, 'a clause) ann_lits› and K
assume ‹Propagated L' (clause D) # M = M2 @ Decided K # M1›
then have M: ‹M = tl M2 @ Decided K # M1›
by (meson cdcl⇩W_restart_mset.propagated_cons_eq_append_decide_cons)
then show
‹twl_lazy_update M1 C› and
‹watched_literals_false_of_max_level M1 C› and
‹twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using C past_invs by (auto simp add: past_invs.simps)
next
fix M1 M2 :: ‹('a, 'a clause) ann_lits› and K
assume ‹Propagated L' (clause D) # M = M2 @ Decided K # M1›
then have M: ‹M = tl M2 @ Decided K # M1›
by (meson cdcl⇩W_restart_mset.propagated_cons_eq_append_decide_cons)
then show ‹confl_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
‹propa_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
‹clauses_to_update_inv (M1, N, U, None, NE, UE, {#}, {#})›
using past_invs by (auto simp add: past_invs.simps)
qed
next
case (conflict D L L' M N U NE UE WS Q) note twl = this(9)
then show ?case
by (auto simp: past_invs.simps)
next
case (delete_from_working L' D M N U NE UE L WS Q) note watched = this(1) and L' = this(2) and
twl = this(3) and valid = this(4) and inv = this(5) and past_invs = this(8)
show ?case unfolding past_invs.simps Ball_def
proof (intro allI conjI impI)
fix C
assume C: ‹C ∈# N + U›
fix M1 M2 :: ‹('a, 'a clause) ann_lits› and K
assume ‹M = M2 @ Decided K # M1›
then show ‹twl_lazy_update M1 C› and
‹watched_literals_false_of_max_level M1 C› and
‹twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using C past_invs by (auto simp add: past_invs.simps)
next
fix M1 M2 :: ‹('a, 'a clause) ann_lits› and K
assume ‹M = M2 @ Decided K # M1›
then show ‹confl_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
‹propa_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
‹clauses_to_update_inv (M1, N, U, None, NE, UE, {#}, {#})›
using past_invs by (auto simp add: past_invs.simps)
qed
next
case (update_clause D L L' M K N U N' U' NE UE WS Q) note watched = this(1) and uL = this(2) and
L' = this(3) and K = this(4) and undef = this(5) and N'U' = this(6) and twl = this(7) and
valid = this(8) and inv = this(9) and twl_excep = this(10) and no_dup = this(11) and
past_invs = this(12)
obtain WD UWD where D: ‹D = TWL_Clause WD UWD› by (cases D)
have L: ‹L ∈# watched D› and D_N_U: ‹D ∈# N + U› and lev_L: ‹get_level M L = count_decided M›
using valid by auto
then have struct_D: ‹struct_wf_twl_cls D›
using twl by (auto simp: twl_st_inv.simps)
have L'_UWD: ‹L ∉# remove1_mset L' UWD› if ‹L ∈# WD› for L
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD L ≥ 1›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) L ≥ 2›
using D that by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have L'_L'_UWD: ‹K ∉# remove1_mset K UWD›
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹count UWD K ≥ 2›
by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
then have ‹count (clause D) K ≥ 2›
using D L' by (auto simp del: count_greater_zero_iff simp: count_greater_zero_iff[symmetric]
split: if_splits)
moreover have ‹distinct_mset (clause D)›
using struct_D D by (auto simp: distinct_mset_union)
ultimately show False
unfolding distinct_mset_count_less_1 by (metis Suc_1 not_less_eq_eq)
qed
have ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps)
let ?D = ‹update_clause D L K›
have *: ‹C ∈# N + U› if ‹C ≠ ?D› and C: ‹C ∈# N' + U'› for C
using C N'U' that by (auto elim!: update_clausesE dest: in_diffD)
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
then have uK_M: ‹- K ∉ lits_of_l M›
using undef Decided_Propagated_in_iff_in_lits_of_l consistent_interp_def
distinct_consistent_interp by blast
have add_remove_WD: ‹add_mset K (remove1_mset L WD) ≠ WD›
using uK_M uL by (auto simp: add_mset_remove_trivial_iff trivial_add_mset_remove_iff)
have cls_D_D: ‹clause ?D = clause D›
by (cases D) (use watched K in auto)
have L_M: ‹L ∉ lits_of_l M›
using n_d uL by (fastforce dest!: distinct_consistent_interp
simp: consistent_interp_def lits_of_def uminus_lit_swap)
have w_max_D: ‹watched_literals_false_of_max_level M D›
using D_N_U twl by (auto simp: twl_st_inv.simps)
show ?case unfolding past_invs.simps Ball_def
proof (intro allI conjI impI)
fix C
assume C: ‹C ∈# N' + U'›
fix M1 M2 :: ‹('a, 'a clause) ann_lits› and K'
assume M: ‹M = M2 @ Decided K' # M1›
have lev_L_M1: ‹get_level M1 L = 0›
using lev_L n_d unfolding M
apply (auto simp: get_level_append_if get_level_cons_if
atm_of_notin_get_level_eq_0 split: if_splits dest: defined_lit_no_dupD)
using atm_of_notin_get_level_eq_0 defined_lit_no_dupD(1) apply blast
apply (simp add: defined_lit_map)
by (metis Suc_count_decided_gt_get_level add_Suc_right not_add_less2)
have ‹twl_lazy_update M1 D›
using past_invs D_N_U unfolding past_invs.simps M twl_lazy_update.simps C
by fast
then have
lazy_L': ‹- L' ∈ lits_of_l M1 ⟹ ¬ has_blit M1 (add_mset L (add_mset L' UWD)) L' ⟹
(∀K∈#UWD. get_level M1 K ≤ get_level M1 L' ∧ - K ∈ lits_of_l M1)›
using watched unfolding D twl_lazy_update.simps
by (simp_all add: all_conj_distrib)
have excep_inv: ‹twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C› if ‹C ≠ ?D›
using * C past_invs that M by (auto simp add: past_invs.simps)
then have ‹twl_exception_inv (M1, N', U', None, NE, UE, {#}, {#}) C› if ‹C ≠ ?D›
using N'U' that by (auto simp add: twl_st_inv.simps twl_exception_inv.simps)
moreover have ‹twl_lazy_update M1 C› ‹watched_literals_false_of_max_level M1 C›
if ‹C ≠ ?D›
using * C twl past_invs M N'U' that
by (auto simp add: past_invs.simps twl_exception_inv.simps)
moreover {
have ‹twl_lazy_update M1 ?D›
using D watched uK_M K lazy_L'
by (auto simp add: M add_mset_eq_add_mset twl_exception_inv.simps lev_L_M1
all_conj_distrib add_mset_commute dest!: multi_member_split[of K])
}
moreover have ‹watched_literals_false_of_max_level M1 ?D›
using D watched uK_M K lazy_L'
by (auto simp add: M add_mset_eq_add_mset twl_exception_inv.simps lev_L_M1
all_conj_distrib add_mset_commute dest!: multi_member_split[of K])
moreover have ‹twl_exception_inv (M1, N', U', None, NE, UE, {#}, {#}) ?D›
using D watched uK_M K lazy_L'
by (auto simp add: M add_mset_eq_add_mset twl_exception_inv.simps lev_L_M1
all_conj_distrib add_mset_commute dest!: multi_member_split[of K])
ultimately show ‹twl_lazy_update M1 C› ‹watched_literals_false_of_max_level M1 C›
‹twl_exception_inv (M1, N', U', None, NE, UE, {#}, {#}) C›
by blast+
next
have [dest!]: ‹C ∈# N' ⟹ C ∈# N ∨ C = ?D› ‹C ∈# U' ⟹ C ∈# U ∨ C = ?D› for C
using N'U' by (auto elim!: update_clausesE dest: in_diffD)
fix M1 M2 :: ‹('a, 'a clause) ann_lits› and K'
assume M: ‹M = M2 @ Decided K' # M1›
then have ‹confl_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
‹propa_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
w_q: ‹clauses_to_update_inv (M1, N, U, None, NE, UE, {#}, {#})›
using past_invs by (auto simp add: past_invs.simps)
moreover have ‹¬M1 ⊨as CNot (clause ?D)›
using K uK_M unfolding true_annots_true_cls_def_iff_negation_in_model cls_D_D M
by (cases D) auto
moreover {
have lev_L_M: ‹get_level M L = count_decided M› and uL_M: ‹-L ∈ lits_of_l M›
using valid by auto
have ‹-L ∉ lits_of_l M1›
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹undefined_lit (M2 @ [Decided K']) L›
using uL_M n_d unfolding M
by (auto simp: lits_of_def uminus_lit_swap no_dup_def defined_lit_map
dest: mk_disjoint_insert)
then show False
using lev_L_M count_decided_ge_get_level[of M1 L]
by (auto simp: lits_of_def uminus_lit_swap M)
qed
then have ‹¬M1 ⊨as CNot (remove1_mset K'' (clause ?D))› for K''
using K uK_M watched D unfolding M by (cases ‹K'' = L›) auto }
ultimately show ‹confl_cands_enqueued (M1, N', U', None, NE, UE, {#}, {#})› and
‹propa_cands_enqueued (M1, N', U', None, NE, UE, {#}, {#})›
by (auto simp add: twl_st_inv.simps split: if_splits)
obtain NU where NU: ‹N + U = add_mset D NU›
by (metis D_N_U insert_DiffM)
then have NU_remove: ‹NU = remove1_mset D (N + U)›
by auto
have ‹N' + U' = add_mset ?D (remove1_mset D (N + U))›
using N'U' D_N_U by (auto elim!: update_clausesE)
then have N'U': ‹N'+U' = add_mset ?D NU›
unfolding NU_remove .
have watched_D: ‹watched ?D = {#K, L'#}›
using D add_remove_WD watched by auto
have ‹twl_lazy_update M1 D›
using past_invs D_N_U unfolding past_invs.simps M twl_lazy_update.simps
by fast
then have
lazy_L': ‹- L' ∈ lits_of_l M1 ⟹ ¬ has_blit M1 (add_mset L (add_mset L' UWD)) L' ⟹
(∀K∈#UWD. get_level M1 K ≤ get_level M1 L' ∧ - K ∈ lits_of_l M1)›
using watched unfolding D twl_lazy_update.simps
by (simp_all add: all_conj_distrib)
have uL'_M1: ‹has_blit M1 (clause (update_clause D L K)) L'› if ‹- L' ∈ lits_of_l M1›
proof -
show ?thesis
using K uK_M lazy_L' that D watched unfolding cls_D_D
by (force simp: M dest!: multi_member_split[of K UWD])
qed
show ‹clauses_to_update_inv (M1, N', U', None, NE, UE, {#}, {#})›
proof (induction rule: clauses_to_update_inv_cases)
case (WS_nempty L C)
then show ?case by simp
next
case (WS_empty K'')
have uK_M1: ‹- K ∉ lits_of_l M1›
using uK_M unfolding M by auto
have ‹¬clauses_to_update_prop {#} M1 (K'', ?D)›
using uK_M1 uL'_M1 by (auto simp: clauses_to_update_prop.simps watched_D
add_mset_eq_add_mset)
then show ?case
using w_q unfolding clauses_to_update_inv.simps N'U' NU
by (auto split: if_splits simp: all_conj_distrib watched_D add_mset_eq_add_mset)
next
case (Q J C)
moreover have ‹- K ∉ lits_of_l M1›
using uK_M unfolding M by auto
moreover have ‹clauses_to_update_prop {#} M1 (L', D)› if ‹- L' ∈ lits_of_l M1›
using watched that uL'_M1 Q.hyps calculation(1,2,3,6) cls_D_D
insert_DiffM w_q watched_D by auto
ultimately show ?case
using w_q watched_D unfolding clauses_to_update_inv.simps N'U' NU
by (fastforce split: if_splits simp: all_conj_distrib add_mset_eq_add_mset)
qed
qed
qed
subsection ‹Invariants and the Transition System›
subsubsection ‹Conflict and propagate›
fun literals_to_update_measure :: ‹'v twl_st ⇒ nat list› where
‹literals_to_update_measure S = [size (literals_to_update S), size (clauses_to_update S)]›
lemma twl_cp_propagate_or_conflict:
assumes
cdcl: ‹cdcl_twl_cp S T› and
twl: ‹twl_st_inv S› and
valid: ‹valid_enqueued S› and
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)›
shows
‹cdcl⇩W_restart_mset.propagate (state⇩W_of S) (state⇩W_of T) ∨
cdcl⇩W_restart_mset.conflict (state⇩W_of S) (state⇩W_of T) ∨
(state⇩W_of S = state⇩W_of T ∧ (literals_to_update_measure T, literals_to_update_measure S) ∈
lexn less_than 2)›
using cdcl twl valid inv
proof (induction rule: cdcl_twl_cp.induct)
case (pop M N U L Q)
then show ?case by (simp add: lexn2_conv)
next
case (propagate D L L' M N U NE UE WS Q) note watched = this(1) and undef = this(2) and
no_upd = this(3) and twl = this(4) and valid = this(5) and inv = this(6)
let ?S = ‹state⇩W_of (M, N, U, None, NE, UE, add_mset (L, D) WS, Q)›
let ?T = ‹state⇩W_of (Propagated L' (clause D) # M, N, U, None, NE, UE, WS, add_mset (- L') Q)›
have ‹∀s∈#clause `# U. ¬ tautology s›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.distinct_cdcl⇩W_state_def by (simp_all add: cdcl⇩W_restart_mset_state)
have D_N_U: ‹D ∈# N + U›
using valid by auto
have ‹cdcl⇩W_restart_mset.propagate ?S ?T›
apply (rule cdcl⇩W_restart_mset.propagate.intros[of _ ‹clause D› L'])
apply (simp add: cdcl⇩W_restart_mset_state; fail)
apply (metis ‹D ∈# N + U› clauses_def state⇩W_of.simps image_eqI
in_image_mset union_iff)
using watched apply (cases D, simp add: clauses_def; fail)
using no_upd watched valid apply (cases D;
simp add: trail.simps true_annots_true_cls_def_iff_negation_in_model; fail)
using undef apply (simp add: trail.simps)
by (simp add: cdcl⇩W_restart_mset_state del: cdcl⇩W_restart_mset.state_simp)
then show ?case by blast
next
case (conflict D L L' M N U NE UE WS Q) note watched = this(1) and defined = this(2)
and no_upd = this(3) and twl = this(3) and valid = this(5) and inv = this(6)
let ?S = ‹state⇩W_of (M, N, U, None, NE, UE, add_mset (L, D) WS, Q)›
let ?T = ‹state⇩W_of (M, N, U, Some (clause D), NE, UE, {#}, {#})›
have D_N_U: ‹D ∈# N + U›
using valid by auto
have ‹distinct_mset (clause D)›
using inv valid ‹D ∈# N + U› unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.distinct_cdcl⇩W_state_def distinct_mset_set_def
by (auto simp: cdcl⇩W_restart_mset_state)
then have ‹L ≠ L'›
using watched by (cases D) simp
have ‹M ⊨as CNot (unwatched D)›
using no_upd by (auto simp: true_annots_true_cls_def_iff_negation_in_model)
have ‹cdcl⇩W_restart_mset.conflict ?S ?T›
apply (rule cdcl⇩W_restart_mset.conflict.intros[of _ ‹clause D›])
apply (simp add: cdcl⇩W_restart_mset_state)
apply (metis ‹D ∈# N + U› clauses_def state⇩W_of.simps image_eqI
in_image_mset union_iff)
using watched defined valid ‹M ⊨as CNot (unwatched D)›
apply (cases D; auto simp add: clauses_def
trail.simps twl_st_inv.simps; fail)
by (simp add: cdcl⇩W_restart_mset_state del: cdcl⇩W_restart_mset.state_simp)
then show ?case by fast
next
case (delete_from_working D L L' M N U NE UE WS Q)
then show ?case by (simp add: lexn2_conv)
next
case (update_clause D L L' M K N U N' U' NE UE WS Q) note unwatched = this(4) and
valid = this(8)
have ‹D ∈# N + U›
using valid by auto
have [simp]: ‹clause (update_clause D L K) = clause D›
using valid unwatched by (cases D) (auto simp: diff_union_swap2[symmetric]
simp del: diff_union_swap2)
have ‹state⇩W_of (M, N, U, None, NE, UE, add_mset (L, D) WS, Q) =
state⇩W_of (M, N', U', None, NE, UE, WS, Q)›
‹(literals_to_update_measure (M, N', U', None, NE, UE, WS, Q),
literals_to_update_measure (M, N, U, None, NE, UE, add_mset (L, D) WS, Q))
∈ lexn less_than 2›
using update_clause ‹D ∈# N + U› by (cases ‹D ∈# N›)
(fastforce simp: image_mset_remove1_mset_if elim!: update_clausesE
simp add: lexn2_conv)+
then show ?case by fast
qed
lemma cdcl_twl_o_cdcl⇩W_o:
assumes
cdcl: ‹cdcl_twl_o S T› and
twl: ‹twl_st_inv S› and
valid: ‹valid_enqueued S› and
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)›
shows ‹cdcl⇩W_restart_mset.cdcl⇩W_o (state⇩W_of S) (state⇩W_of T)›
using cdcl twl valid inv
proof (induction rule: cdcl_twl_o.induct)
case (decide M L N NE U UE) note undef = this(1) and atm = this(2)
have ‹cdcl⇩W_restart_mset.decide (state⇩W_of (M, N, U, None, NE, UE, {#}, {#}))
(state⇩W_of (Decided L # M, N, U, None, NE, UE, {#}, {#-L#}))›
apply (rule cdcl⇩W_restart_mset.decide_rule)
apply (simp add: cdcl⇩W_restart_mset_state; fail)
using undef apply (simp add: trail.simps; fail)
using atm apply (simp add: cdcl⇩W_restart_mset_state; fail)
by (simp add: state_eq_def cdcl⇩W_restart_mset_state del: cdcl⇩W_restart_mset.state_simp)
then show ?case
by (blast dest: cdcl⇩W_restart_mset.cdcl⇩W_o.intros)
next
case (skip L D C' M N U NE UE) note LD = this(1) and D = this(2)
show ?case
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_o.bj)
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_bj.skip)
apply (rule cdcl⇩W_restart_mset.skip_rule)
apply (simp add: trail.simps; fail)
apply (simp add: cdcl⇩W_restart_mset_state; fail)
using LD apply (simp; fail)
using D apply (simp; fail)
by (simp add: state_eq_def cdcl⇩W_restart_mset_state del: cdcl⇩W_restart_mset.state_simp)
next
case (resolve L D C M N U NE UE) note LD = this(1) and lev = this(2) and inv = this(5)
have ‹∀La mark a b. a @ Propagated La mark # b = Propagated L C # M ⟶
b ⊨as CNot (remove1_mset La mark) ∧ La ∈# mark›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_conflicting_def
by (auto simp: trail.simps)
then have LC: ‹L ∈# C›
by blast
show ?case
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_o.bj)
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_bj.resolve)
apply (rule cdcl⇩W_restart_mset.resolve_rule)
apply (simp add: trail.simps; fail)
apply (simp add: trail.simps; fail)
using LC apply (simp add: trail.simps; fail)
apply (simp add: cdcl⇩W_restart_mset_state; fail)
using LD apply (simp; fail)
using lev apply (simp add: cdcl⇩W_restart_mset_state; fail)
by (simp add: state_eq_def cdcl⇩W_restart_mset_state del: cdcl⇩W_restart_mset.state_simp)
next
case (backtrack_unit_clause L D K M1 M2 M D' i N U NE UE) note L_D = this(1) and
decomp = this(2) and lev_L = this(3) and max_D'_L = this(4) and lev_D = this(5) and
lev_K = this(6) and D'_D = this(8) and NU_D' = this(9) and inv = this(12) and
D'[simp] = this(7)
let ?S = ‹state⇩W_of (M, N, U, Some {#L#}, NE, UE, {#}, {#})›
let ?T = ‹state⇩W_of (Propagated L {#L#} # M1, N, U, None, NE, add_mset {#L#} UE, {#}, {#L#})›
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def
by (simp add: cdcl⇩W_restart_mset_state)
have ‹undefined_lit M1 L›
apply (rule cdcl⇩W_restart_mset.backtrack_lit_skiped[of ?S _ K _ M2 i])
subgoal using lev_L inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def
by (simp add: cdcl⇩W_restart_mset_state; fail)
subgoal using decomp by (simp add: trail.simps; fail)
subgoal
using lev_L inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def
by (simp add: cdcl⇩W_restart_mset_state; fail)
subgoal using lev_K by (simp add: trail.simps; fail)
done
obtain M3 where M3: ‹M = M3 @ M2 @ Decided K # M1›
using decomp by (blast dest!: get_all_ann_decomposition_exists_prepend)
have D: ‹D = add_mset L (remove1_mset L D)›
using L_D by auto
have ‹undefined_lit (M3 @ M2) K›
using n_d unfolding M3 by auto
then have [simp]: ‹count_decided M1 = 0›
using lev_D lev_K by (auto simp: M3 image_Un)
show ?case
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_o.bj)
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_bj.backtrack)
apply (rule cdcl⇩W_restart_mset.backtrack_rule[of _ L ‹remove1_mset L D› K M1 M2
‹remove1_mset L D'› i])
subgoal using L_D by (simp add: cdcl⇩W_restart_mset_state)
subgoal using decomp by (simp add: cdcl⇩W_restart_mset_state)
subgoal using lev_L by (simp add: cdcl⇩W_restart_mset_state)
subgoal using max_D'_L L_D by (simp add: cdcl⇩W_restart_mset_state)
subgoal using lev_D L_D by (simp add: cdcl⇩W_restart_mset_state)
subgoal using lev_K by (simp add: cdcl⇩W_restart_mset_state)
subgoal using D'_D by (simp add: cdcl⇩W_restart_mset_state)
subgoal using NU_D' by (simp add: cdcl⇩W_restart_mset_state clauses_def ac_simps)
subgoal using decomp unfolding state_eq_def state_def prod.inject
by (simp add: cdcl⇩W_restart_mset_state)
done
next
case (backtrack_nonunit_clause L D K M1 M2 M D' i N U NE UE L') note LD = this(1) and
decomp = this(2) and lev_L = this(3) and max_lev = this(4) and i = this(5) and lev_K = this(6)
and D'_D = this(8) and NU_D' = this(9) and L_D' = this(10) and L' = this(11-12) and
inv = this(15)
let ?S = ‹state⇩W_of (M, N, U, Some D, NE, UE, {#}, {#})›
let ?T = ‹state⇩W_of (Propagated L D # M1, N, U, None, NE, add_mset {#L#} UE, {#}, {#L#})›
have n_d: ‹no_dup M›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def
by (simp add: cdcl⇩W_restart_mset_state)
have ‹undefined_lit M1 L›
apply (rule cdcl⇩W_restart_mset.backtrack_lit_skiped[of ?S _ K _ M2 i])
subgoal
using lev_L inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def
by (simp add: cdcl⇩W_restart_mset_state; fail)
subgoal using decomp by (simp add: trail.simps; fail)
subgoal using lev_L inv
unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def
by (simp add: cdcl⇩W_restart_mset_state; fail)
subgoal using lev_K by (simp add: trail.simps; fail)
done
obtain M3 where M3: ‹M = M3 @ M2 @ Decided K # M1›
using decomp by (blast dest!: get_all_ann_decomposition_exists_prepend)
have ‹undefined_lit (M3 @ M2) K›
using n_d unfolding M3 by (auto simp: lits_of_def)
then have count_M1: ‹count_decided M1 = i›
using lev_K unfolding M3 by (auto simp: image_Un)
have ‹L ≠ L'›
using L' lev_L lev_K count_decided_ge_get_level[of M K] L' by auto
then have D: ‹add_mset L (add_mset L' (D' - {#L, L'#})) = D'›
using L' L_D'
by (metis add_mset_diff_bothsides diff_single_eq_union insert_noteq_member mset_add)
have D': ‹remove1_mset L D' = add_mset L' (D' - {#L, L'#})›
by (subst D[symmetric]) auto
show ?case
apply (subst D[symmetric])
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_o.bj)
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_bj.backtrack)
apply (rule cdcl⇩W_restart_mset.backtrack_rule[of _ L ‹remove1_mset L D› K M1 M2
‹remove1_mset L D'› i])
subgoal using LD by (simp add: cdcl⇩W_restart_mset_state)
subgoal using decomp by (simp add: trail.simps)
subgoal using lev_L by (simp add: cdcl⇩W_restart_mset_state; fail)
subgoal using max_lev L_D' by (simp add: cdcl⇩W_restart_mset_state get_maximum_level_add_mset)
subgoal using i by (simp add: cdcl⇩W_restart_mset_state)
subgoal using lev_K i unfolding D' by (simp add: trail.simps)
subgoal using D'_D by (simp add: mset_le_subtract)
subgoal using NU_D' L_D' by (simp add: mset_le_subtract clauses_def ac_simps)
subgoal
using decomp unfolding state_eq_def state_def prod.inject
using i lev_K count_M1 L_D' by (simp add: cdcl⇩W_restart_mset_state D)
done
qed
lemma cdcl_twl_cp_cdcl⇩W_stgy:
‹cdcl_twl_cp S T ⟹ twl_struct_invs S ⟹
cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T) ∨
(state⇩W_of S = state⇩W_of T ∧ (literals_to_update_measure T, literals_to_update_measure S)
∈ lexn less_than 2)›
by (auto dest!: twl_cp_propagate_or_conflict
cdcl⇩W_restart_mset.cdcl⇩W_stgy.conflict'
cdcl⇩W_restart_mset.cdcl⇩W_stgy.propagate'
simp: twl_struct_invs_def)
lemma cdcl_twl_cp_conflict:
‹cdcl_twl_cp S T ⟹ get_conflict T ≠ None ⟶
clauses_to_update T = {#} ∧ literals_to_update T = {#}›
by (induction rule: cdcl_twl_cp.induct) auto
lemma cdcl_twl_cp_entailed_clss_inv:
‹cdcl_twl_cp S T ⟹ entailed_clss_inv S ⟹ entailed_clss_inv T›
proof (induction rule: cdcl_twl_cp.induct)
case (pop M N U NE UE L Q)
then show ?case by auto
next
case (propagate D L L' M N U NE UE WS Q) note undef = this(2) and _ = this
then have unit: ‹entailed_clss_inv (M, N, U, None, NE, UE, add_mset (L, D) WS, Q)›
by auto
show ?case
unfolding entailed_clss_inv.simps Ball_def
proof (intro allI impI conjI)
fix C
assume ‹C ∈# NE + UE›
then obtain L where
C: ‹L ∈# C› and lev_L: ‹get_level M L = 0› and L_M: ‹L ∈ lits_of_l M›
using unit by auto
have ‹atm_of L' ≠ atm_of L›
using undef L_M by (auto simp: defined_lit_map lits_of_def)
then show ‹∃L. L ∈# C ∧ (None = None ∨ 0 < count_decided (Propagated L' (clause D) # M) ⟶
get_level (Propagated L' (clause D) # M) L = 0 ∧
L ∈ lits_of_l (Propagated L' (clause D) # M))›
using lev_L L_M C by auto
qed
next
case (conflict D L L' M N U NE UE WS Q)
then show ?case by auto
next
case (delete_from_working D L L' M N U NE UE WS Q)
then show ?case by auto
next
case (update_clause D L L' M K N' U' N U NE UE WS Q)
then show ?case by auto
qed
lemma cdcl_twl_cp_init_clss:
‹cdcl_twl_cp S T ⟹ twl_struct_invs S ⟹ init_clss (state⇩W_of T) = init_clss (state⇩W_of S)›
by (metis cdcl⇩W_restart_mset.cdcl⇩W_stgy_no_more_init_clss cdcl_twl_cp_cdcl⇩W_stgy)
lemma cdcl_twl_cp_twl_struct_invs:
‹cdcl_twl_cp S T ⟹ twl_struct_invs S ⟹ twl_struct_invs T›
apply (subst twl_struct_invs_def)
apply (intro conjI)
subgoal by (rule twl_cp_twl_inv; auto simp add: twl_struct_invs_def twl_cp_twl_inv)
subgoal by (simp add: twl_cp_valid twl_struct_invs_def)
subgoal by (metis cdcl_twl_cp_cdcl⇩W_stgy cdcl⇩W_restart_mset.cdcl⇩W_stgy_cdcl⇩W_all_struct_inv
twl_struct_invs_def)
subgoal by (metis cdcl_twl_cp_cdcl⇩W_stgy twl_struct_invs_def
cdcl⇩W_restart_mset.cdcl⇩W_stgy_no_smaller_propa)
subgoal by (rule twl_cp_twl_st_exception_inv; auto simp add: twl_struct_invs_def; fail)
subgoal by (use twl_struct_invs_def twl_cp_no_duplicate_queued in blast)
subgoal by (rule twl_cp_distinct_queued; auto simp add: twl_struct_invs_def)
subgoal by (rule twl_cp_confl_cands_enqueued; auto simp add: twl_struct_invs_def; fail)
subgoal by (rule twl_cp_propa_cands_enqueued; auto simp add: twl_struct_invs_def; fail)
subgoal by (simp add: cdcl_twl_cp_conflict; fail)
subgoal by (simp add: cdcl_twl_cp_entailed_clss_inv twl_struct_invs_def; fail)
subgoal by (simp add: twl_struct_invs_def twl_cp_clauses_to_update; fail)
subgoal by (simp add: twl_cp_past_invs twl_struct_invs_def; fail)
done
lemma twl_struct_invs_no_false_clause:
assumes ‹twl_struct_invs S›
shows ‹cdcl⇩W_restart_mset.no_false_clause (state⇩W_of S)›
proof -
obtain M N U D NE UE WS Q where
S: ‹S = (M, N, U, D, NE, UE, WS, Q)›
by (cases S) auto
have wf: ‹⋀C. C ∈# N + U ⟹ struct_wf_twl_cls C› and entailed: ‹entailed_clss_inv S›
using assms unfolding twl_struct_invs_def twl_st_inv.simps S by fast+
have ‹{#} ∉# NE + UE›
using entailed unfolding S entailed_clss_inv.simps
by (auto simp del: set_mset_union)
moreover have ‹clause C = {#} ⟹ C ∈# N + U ⟹ False› for C
using wf[of C] by (cases C) (auto simp del: set_mset_union)
ultimately show ?thesis
by (fastforce simp: S clauses_def cdcl⇩W_restart_mset.no_false_clause_def)
qed
lemma cdcl_twl_cp_twl_stgy_invs:
‹cdcl_twl_cp S T ⟹ twl_struct_invs S ⟹ twl_stgy_invs S ⟹ twl_stgy_invs T›
using cdcl⇩W_restart_mset.cdcl⇩W_stgy_cdcl⇩W_stgy_invariant[of ‹state⇩W_of S› ‹state⇩W_of S›]
unfolding twl_stgy_invs_def
by (metis cdcl⇩W_restart_mset.cdcl⇩W_restart_conflict_non_zero_unless_level_0
cdcl⇩W_restart_mset.cdcl⇩W_stgy_cdcl⇩W_stgy_invariant
cdcl_twl_cp_cdcl⇩W_stgy cdcl⇩W_restart_mset.conflict
cdcl⇩W_restart_mset.propagate twl_cp_propagate_or_conflict
twl_struct_invs_def twl_struct_invs_no_false_clause)
subsubsection ‹The other rules›
lemma
assumes
cdcl: ‹cdcl_twl_o S T› and
twl: ‹twl_struct_invs S›
shows
cdcl_twl_o_twl_st_inv: ‹twl_st_inv T› and
cdcl_twl_o_past_invs: ‹past_invs T›
using cdcl twl
proof (induction rule: cdcl_twl_o.induct)
case (decide M K N NE U UE) note undef = this(1) and atm = this(2)
case 1 note invs = this(1)
let ?S = ‹(M, N, U, None, NE, UE, {#}, {#})›
have inv: ‹twl_st_inv ?S› and excep: ‹twl_st_exception_inv ?S› and past: ‹past_invs ?S› and
w_q: ‹clauses_to_update_inv ?S›
using invs unfolding twl_struct_invs_def by blast+
have n_d: ‹no_dup M›
using invs unfolding twl_struct_invs_def cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: cdcl⇩W_restart_mset_state)
have n_d': ‹no_dup (Decided K # M)›
using defined_lit_map n_d undef by auto
have propa_cands: ‹propa_cands_enqueued ?S› and
confl_cands: ‹confl_cands_enqueued ?S›
using invs unfolding twl_struct_invs_def by blast+
show ?case
unfolding twl_st_inv.simps Ball_def
proof (intro conjI allI impI)
fix C :: ‹'a twl_cls›
assume C: ‹C ∈# N + U›
show struct: ‹struct_wf_twl_cls C›
using inv C by (auto simp: twl_st_inv.simps)
have watched: ‹watched_literals_false_of_max_level M C› and
lazy: ‹twl_lazy_update M C›
using C inv by (auto simp: twl_st_inv.simps)
obtain W UW where C_W: ‹C = TWL_Clause W UW›
by (cases C)
have H: False if
W: ‹L ∈# W› and
uL: ‹- L ∈ lits_of_l (Decided K # M)› and
L': ‹¬has_blit (Decided K # M) (W + UW) L› and
False: ‹-L ≠ K› for L
proof -
have H: ‹- L ∈ lits_of_l M ⟹ ¬ has_blit M (W + UW) L ⟹ get_level M L = count_decided M ›
using watched W unfolding C_W
by auto
obtain L' where W': ‹W = {#L, L'#}›
using struct W size_2_iff[of W] unfolding C_W
by (auto simp: add_mset_eq_single add_mset_eq_add_mset dest!: multi_member_split)
have no_has_blit: ‹¬has_blit M (W + UW) L›
using no_has_blit_decide'[of K M C] L' n_d C_W W undef by auto
then have ‹∀K ∈# UW. -K ∈ lits_of_l M›
using uL L' False excep C W C_W L' W n_d undef
by (auto simp: twl_exception_inv.simps all_conj_distrib
dest!: multi_member_split[of _ N])
then have M_CNot_C: ‹M ⊨as CNot (remove1_mset L' (clause C))›
using uL False W' unfolding true_annots_true_cls_def_iff_negation_in_model
by (auto simp: C_W W)
moreover have L'_C: ‹L' ∈# clause C›
unfolding C_W W' by auto
ultimately have ‹defined_lit M L'›
using propa_cands C by auto
then have ‹-L' ∈ lits_of_l M›
using L' W' False uL C_W L'_C H no_has_blit
apply (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
by (metis C_W L'_C no_has_blit clause.simps
count_decided_ge_get_level has_blit_def is_blit_def)
then have ‹M ⊨as CNot (clause C)›
using M_CNot_C W' unfolding true_annots_true_cls_def_iff_negation_in_model
by (auto simp: C_W)
then show False
using confl_cands C by auto
qed
show ‹watched_literals_false_of_max_level (Decided K # M) C›
unfolding C_W watched_literals_false_of_max_level.simps
proof (intro allI impI)
fix L
assume
W: ‹L ∈# W› and
uL: ‹- L ∈ lits_of_l (Decided K # M)› and
L': ‹¬has_blit (Decided K # M) (W + UW) L›
then have ‹-L = K›
using H[OF W uL L'] by fast
then show ‹get_level (Decided K # M) L = count_decided (Decided K # M)›
by auto
qed
{
assume exception: ‹¬ twl_is_an_exception C {#-K#} {#}›
have ‹twl_lazy_update M C›
using C inv by (auto simp: twl_st_inv.simps)
have lev_le_Suc: ‹get_level M Ka ≤ Suc (count_decided M)› for Ka
using count_decided_ge_get_level le_Suc_eq by blast
show ‹twl_lazy_update (Decided K # M) C›
unfolding C_W twl_lazy_update.simps Ball_def
proof (intro allI impI)
fix L K' :: ‹'a literal›
assume
W: ‹L ∈# W› and
uL: ‹- L ∈ lits_of_l (Decided K # M)› and
L': ‹¬has_blit (Decided K # M) (W + UW) L› and
K': ‹K' ∈# UW›
then have ‹-L = K›
using H[OF W uL L'] by fast
then have False
using exception W
by (auto simp: C_W twl_is_an_exception_def)
then show ‹get_level (Decided K # M) K' ≤ get_level (Decided K # M) L ∧
-K' ∈ lits_of_l (Decided K # M)›
by fast
qed
}
qed
case 2
show ?case
unfolding past_invs.simps Ball_def
proof (intro allI impI conjI)
fix M1 M2 K' C
assume ‹Decided K # M = M2 @ Decided K' # M1› and C: ‹C ∈# N + U›
then have M: ‹M = tl M2 @ Decided K' # M1 ∨ M = M1›
by (cases M2) auto
have IH: ‹∀M1 M2 K. M = M2 @ Decided K # M1 ⟶
twl_lazy_update M1 C ∧ watched_literals_false_of_max_level M1 C ∧
twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using past C unfolding past_invs.simps by blast
have ‹twl_lazy_update M C›
using inv C unfolding twl_st_inv.simps by auto
then show ‹twl_lazy_update M1 C›
using IH M by blast
have ‹watched_literals_false_of_max_level M C›
using inv C unfolding twl_st_inv.simps by auto
then show ‹watched_literals_false_of_max_level M1 C›
using IH M by blast
have ‹twl_exception_inv (M, N, U, None, NE, UE, {#}, {#}) C›
using excep inv C unfolding twl_st_inv.simps by auto
then show ‹twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using IH M by blast
next
fix M1 M2 :: ‹('a, 'a clause) ann_lits› and K'
assume ‹Decided K # M = M2 @ Decided K' # M1›
then have M: ‹M = tl M2 @ Decided K' # M1 ∨ M = M1›
by (cases M2) auto
then show ‹confl_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
‹propa_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
‹clauses_to_update_inv (M1, N, U, None, NE, UE, {#}, {#})›
using confl_cands past propa_cands w_q unfolding past_invs.simps by blast+
qed
next
case (skip L D C' M N U NE UE)
case 1
then show ?case
by (auto simp: twl_st_inv.simps twl_struct_invs_def)
case 2
then show ?case
by (auto simp: past_invs.simps twl_struct_invs_def)
next
case (resolve L D C M N U NE UE)
case 1
then show ?case
by (auto simp: twl_st_inv.simps twl_struct_invs_def)
case 2
then show ?case
by (auto simp: past_invs.simps twl_struct_invs_def)
next
case (backtrack_unit_clause K' D K M1 M2 M D' i N U NE UE) note decomp = this(2) and
lev = this(3-5)
case 1 note invs = this(1)
let ?S = ‹(M, N, U, Some D, NE, UE, {#}, {#})›
let ?T = ‹(Propagated K' {#K'#} # M1, N, U, None, NE, add_mset {#K'#} UE, {#}, {#- K'#})›
let ?M1 = ‹Propagated K' {#K'#} # M1›
have bt_twl: ‹cdcl_twl_o ?S ?T›
using cdcl_twl_o.backtrack_unit_clause[OF backtrack_unit_clause.hyps] .
then have ‹cdcl⇩W_restart_mset.cdcl⇩W_o (state⇩W_of ?S) (state⇩W_of ?T)›
by (rule cdcl_twl_o_cdcl⇩W_o) (use invs in ‹simp_all add: twl_struct_invs_def›)
then have struct_inv_T: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?T)›
using cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_inv cdcl⇩W_restart_mset.other invs
unfolding twl_struct_invs_def by blast
have inv: ‹twl_st_inv ?S› and w_q: ‹clauses_to_update_inv ?S› and past: ‹past_invs ?S›
using invs unfolding twl_struct_invs_def by blast+
have n_d: ‹no_dup M›
using invs unfolding twl_struct_invs_def cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: cdcl⇩W_restart_mset_state)
have n_d': ‹no_dup ?M1›
using struct_inv_T unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: trail.simps)
have propa_cands: ‹propa_cands_enqueued ?S› and
confl_cands: ‹confl_cands_enqueued ?S›
using invs unfolding twl_struct_invs_def by blast+
have excep: ‹twl_st_exception_inv ?S›
using invs unfolding twl_struct_invs_def by fast
obtain M3 where M: ‹M = M3 @ M2 @ Decided K # M1›
using decomp by blast
define M2' where ‹M2' = M3 @ M2›
have M': ‹M = M2' @ Decided K # M1›
unfolding M M2'_def by simp
have propa_cands_M1:
‹propa_cands_enqueued (M1, N, U, None, NE, add_mset {#K'#} UE, {#}, {#- K'#})›
unfolding propa_cands_enqueued.simps
proof (intro allI impI)
fix L C
assume
C: ‹C ∈# N + U› and
L: ‹L ∈# clause C› and
M1_CNot: ‹M1 ⊨as CNot (remove1_mset L (clause C))› and
undef: ‹undefined_lit M1 L›
define D where ‹D = remove1_mset L (clause C)›
have ‹add_mset L D ∈# clause `# (N + U)› and ‹M1 ⊨as CNot D›
using C L M1_CNot unfolding D_def by auto
moreover have ‹cdcl⇩W_restart_mset.no_smaller_propa (state⇩W_of ?S)›
using invs unfolding twl_struct_invs_def by blast
ultimately have False
using undef M'
by (fastforce simp: cdcl⇩W_restart_mset.no_smaller_propa_def trail.simps clauses_def)
then show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- K'#}) ∨ (∃L. (L, C) ∈# {#})›
by fast
qed
have excep_M1: ‹twl_st_exception_inv (M1, N, U, None, NE, UE, {#}, {#})›
using past unfolding past_invs.simps M' by auto
show ?case
unfolding twl_st_inv.simps Ball_def
proof (intro conjI allI impI)
fix C :: ‹'a twl_cls›
assume C: ‹C ∈# N + U›
show struct: ‹struct_wf_twl_cls C›
using inv C by (auto simp: twl_st_inv.simps)
obtain CW CUW where C_W: ‹C = TWL_Clause CW CUW›
by (cases C)
{
assume exception: ‹¬ twl_is_an_exception C {#-K'#} {#}›
have
lazy: ‹twl_lazy_update M1 C› and
watched_max: ‹watched_literals_false_of_max_level M1 C›
using C past M by (auto simp: past_invs.simps)
have lev_le_Suc: ‹get_level M Ka ≤ Suc (count_decided M)› for Ka
using count_decided_ge_get_level le_Suc_eq by blast
have Lev_M1: ‹get_level (?M1) K ≤ count_decided M1› for K
by (auto simp: count_decided_ge_get_level get_level_cons_if)
show ‹twl_lazy_update ?M1 C›
proof -
show ?thesis
using Lev_M1
using twl C exception twl n_d' watched_max
unfolding C_W
apply (auto simp: count_decided_ge_get_level
twl_is_an_exception_add_mset_to_queue atm_of_eq_atm_of
dest!: no_has_blit_propagate' no_has_blit_propagate)
apply (metis count_decided_ge_get_level get_level_skip_beginning get_level_take_beginning)
using lazy unfolding C_W twl_lazy_update.simps apply blast
apply (metis count_decided_ge_get_level get_level_skip_beginning get_level_take_beginning)
using lazy unfolding C_W twl_lazy_update.simps apply blast
done
qed
}
have ‹watched_literals_false_of_max_level M1 C›
using past C unfolding M' past_invs.simps by blast
then show ‹watched_literals_false_of_max_level ?M1 C›
using has_blit_Cons n_d'
by (auto simp: C_W get_level_cons_if)
qed
case 2
show ?case
unfolding past_invs.simps Ball_def
proof (intro allI impI conjI)
fix M1'' M2'' K'' C
assume ‹?M1 = M2'' @ Decided K'' # M1''› and C: ‹C ∈# N + U›
then have M1: ‹M1 = tl M2'' @ Decided K'' # M1''›
by (cases M2'') auto
have ‹twl_lazy_update M1'' C›‹watched_literals_false_of_max_level M1'' C›
using past C unfolding past_invs.simps M M1 twl_exception_inv.simps by auto
moreover {
have ‹twl_exception_inv (M1'', N, U, None, NE, UE, {#}, {#}) C›
using past C unfolding past_invs.simps M M1 by auto
then have ‹twl_exception_inv (M1'', N, U, None, NE, add_mset {#K'#} UE, {#}, {#}) C›
using C unfolding twl_exception_inv.simps by auto }
ultimately show ‹twl_lazy_update M1'' C›‹watched_literals_false_of_max_level M1'' C›
‹twl_exception_inv (M1'', N, U, None, NE, add_mset {#K'#} UE, {#}, {#}) C›
by fast+
next
fix M1'' M2'' K''
assume ‹?M1 = M2'' @ Decided K'' # M1''›
then have M1: ‹M1 = tl M2'' @ Decided K'' # M1''›
by (cases M2'') auto
then show
‹confl_cands_enqueued (M1'', N, U, None, NE, add_mset {#K'#} UE, {#}, {#})› and
‹propa_cands_enqueued (M1'', N, U, None, NE, add_mset {#K'#} UE, {#}, {#})› and
‹clauses_to_update_inv (M1'', N, U, None, NE, add_mset {#K'#} UE, {#}, {#})›
using past by (auto simp add: past_invs.simps M)
qed
next
case (backtrack_nonunit_clause K' D K M1 M2 M D' i N U NE UE K'') note K'_D = this(1) and
decomp = this(2) and lev_K' = this(3) and i = this(5) and lev_K = this(6) and K'_D' = this(10)
and K'' = this(11) and lev_K'' = this(12)
case 1 note invs = this(1)
let ?S = ‹(M, N, U, Some D, NE, UE, {#}, {#})›
let ?M1 = ‹Propagated K' D' # M1›
let ?T = ‹(?M1, N, add_mset (TWL_Clause {#K', K''#} (D' - {#K', K''#})) U, None, NE, UE, {#},
{#- K'#})›
let ?D = ‹TWL_Clause {#K', K''#} (D' - {#K', K''#})›
have bt_twl: ‹cdcl_twl_o ?S ?T›
using cdcl_twl_o.backtrack_nonunit_clause[OF backtrack_nonunit_clause.hyps] .
then have ‹cdcl⇩W_restart_mset.cdcl⇩W_o (state⇩W_of ?S) (state⇩W_of ?T)›
by (rule cdcl_twl_o_cdcl⇩W_o) (use invs in ‹simp_all add: twl_struct_invs_def›)
then have struct_inv_T: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?T)›
using cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_inv cdcl⇩W_restart_mset.other invs
unfolding twl_struct_invs_def by blast
have inv: ‹twl_st_inv ?S› and
w_q: ‹clauses_to_update_inv ?S› and
past: ‹past_invs ?S›
using invs unfolding twl_struct_invs_def by blast+
have n_d: ‹no_dup M›
using invs unfolding twl_struct_invs_def cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: cdcl⇩W_restart_mset_state)
have n_d': ‹no_dup ?M1›
using struct_inv_T unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: trail.simps)
have propa_cands: ‹propa_cands_enqueued ?S› and
confl_cands: ‹confl_cands_enqueued ?S›
using invs unfolding twl_struct_invs_def by blast+
obtain M3 where M: ‹M = M3 @ M2 @ Decided K # M1›
using decomp by blast
define M2' where ‹M2' = M3 @ M2›
have M': ‹M = M2' @ Decided K # M1›
unfolding M M2'_def by simp
have struct_inv_S: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?S)›
using invs unfolding twl_struct_invs_def by blast
then have ‹distinct_mset D›
unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.distinct_cdcl⇩W_state_def
by (auto simp: conflicting.simps)
have ‹undefined_lit (M3 @ M2) K›
using n_d unfolding M by auto
then have count_M1: ‹count_decided M1 = i›
using lev_K unfolding M by (auto simp: image_Un)
then have K''_ne_K: ‹K' ≠ K''›
using lev_K lev_K' lev_K'' count_decided_ge_get_level[of M K''] unfolding M by auto
then have D:
‹add_mset K' (add_mset K'' (D' - {#K', K''#})) = D'›
‹add_mset K'' (add_mset K' (D' - {#K', K''#})) = D'›
using K'' K'_D' multi_member_split by fastforce+
have propa_cands_M1: ‹propa_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#- K''#})›
unfolding propa_cands_enqueued.simps
proof (intro allI impI)
fix L C
assume
C: ‹C ∈# N + U› and
L: ‹L ∈# clause C› and
M1_CNot: ‹M1 ⊨as CNot (remove1_mset L (clause C))› and
undef: ‹undefined_lit M1 L›
define D where ‹D = remove1_mset L (clause C)›
have ‹add_mset L D ∈# clause `# (N + U)› and ‹M1 ⊨as CNot D›
using C L M1_CNot unfolding D_def by auto
moreover have ‹cdcl⇩W_restart_mset.no_smaller_propa (state⇩W_of ?S)›
using invs unfolding twl_struct_invs_def by blast
ultimately have False
using undef M'
by (fastforce simp: cdcl⇩W_restart_mset.no_smaller_propa_def trail.simps clauses_def)
then show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- K''#}) ∨ (∃L. (L, C) ∈# {#})›
by fast
qed
have ‹cdcl⇩W_restart_mset.cdcl⇩W_conflicting (state⇩W_of ?T)›
using struct_inv_T unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def twl_struct_invs_def
by (auto simp: conflicting.simps)
then have M1_CNot_D: ‹M1 ⊨as CNot (remove1_mset K' D')›
unfolding cdcl⇩W_restart_mset.cdcl⇩W_conflicting_def
by (auto simp: conflicting.simps trail.simps)
then have uK''_M1: ‹-K'' ∈ lits_of_l M1›
using K'' K''_ne_K unfolding true_annots_true_cls_def_iff_negation_in_model
by (metis in_remove1_mset_neq)
then have ‹undefined_lit (M3 @ M2 @ Decided K # []) K''›
using n_d M by (auto simp: atm_of_eq_atm_of dest: in_lits_of_l_defined_litD defined_lit_no_dupD)
then have lev_M1_K'': ‹get_level M1 K'' = count_decided M1›
using lev_K'' count_M1 unfolding M by (auto simp: image_Un)
have excep_M1: ‹twl_st_exception_inv (M1, N, U, None, NE, UE, {#}, {#})›
using past unfolding past_invs.simps M' by auto
show ?case
unfolding twl_st_inv.simps Ball_def
proof (intro conjI allI impI)
fix C :: ‹'a twl_cls›
assume C: ‹C ∈# N + add_mset ?D U›
have ‹cdcl⇩W_restart_mset.distinct_cdcl⇩W_state (state⇩W_of ?T)›
using struct_inv_T unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def by blast
then have ‹distinct_mset D'›
unfolding cdcl⇩W_restart_mset.distinct_cdcl⇩W_state_def
by (auto simp: cdcl⇩W_restart_mset_state)
then show struct: ‹struct_wf_twl_cls C›
using inv C by (auto simp: twl_st_inv.simps D)
obtain CW CUW where C_W: ‹C = TWL_Clause CW CUW›
by (cases C)
have
lazy: ‹twl_lazy_update M1 C› and
watched_max: ‹watched_literals_false_of_max_level M1 C› if ‹C ≠ ?D›
using C past M' that by (auto simp: past_invs.simps)
from M1_CNot_D have in_D_M1: ‹L ∈# remove1_mset K' D' ⟹ - L ∈ lits_of_l M1› for L
by (auto simp: true_annots_true_cls_def_iff_negation_in_model)
then have in_K_D_M1: ‹L ∈# D' - {#K', K''#} ⟹ - L ∈ lits_of_l M1› for L
by (metis K'_D' add_mset_diff_bothsides add_mset_remove_trivial in_diffD mset_add)
have ‹- K' ∉ lits_of_l M1›
using n_d' by (simp add: Decided_Propagated_in_iff_in_lits_of_l)
have def_K'': ‹defined_lit M1 K''›
using n_d' uK''_M1
using Decided_Propagated_in_iff_in_lits_of_l uK''_M1 by blast
have
lazy_D: ‹twl_lazy_update ?M1 C› if ‹C = ?D›
using that n_d' uK''_M1 def_K'' ‹- K' ∉ lits_of_l M1› in_K_D_M1 lev_M1_K''
by (auto simp: add_mset_eq_add_mset count_decided_ge_get_level get_level_cons_if
atm_of_eq_atm_of)
have
watched_max_D: ‹watched_literals_false_of_max_level ?M1 C› if ‹C = ?D›
using that in_D_M1 by (auto simp add: add_mset_eq_add_mset lev_M1_K'' get_level_cons_if
dest: in_K_D_M1)
{
assume excep: ‹¬ twl_is_an_exception C {#-K'#} {#}›
have lev_le_Suc: ‹get_level M Ka ≤ Suc (count_decided M)› for Ka
using count_decided_ge_get_level le_Suc_eq by blast
have Lev_M1: ‹get_level (?M1) K ≤ count_decided M1› for K
by (auto simp: count_decided_ge_get_level get_level_cons_if)
have ‹twl_lazy_update ?M1 C› if ‹C ≠ ?D›
proof -
have 1: ‹get_level (Propagated K' D' # M1) K ≤ get_level (Propagated K' D' # M1) L›
if
‹∀L. L ∈# CW ⟶ - L ∈ lits_of_l M1 ⟶ ¬ has_blit M1 (CW + CUW) L ⟶
get_level M1 L = count_decided M1› and
‹L ∈# CW› and
‹- L ∈ lits_of_l M1› and
‹K ∈# CUW› and
‹¬ has_blit M1 (CW + CUW) L›
for L :: ‹'a literal› and K :: ‹'a literal›
using that Lev_M1
by (metis count_decided_ge_get_level get_level_skip_beginning get_level_take_beginning)
have 2: False
if
‹L ∈# CW› and
‹TWL_Clause CW CUW ∈# N› and
‹CW ≠ {#K', K''#}› and
‹- L ∈ lits_of_l M1› and
‹K ∈# CUW› and
‹- K ∉ lits_of_l M1› and
‹¬ has_blit M1 (CW + CUW) L›
for L :: ‹'a literal› and K :: ‹'a literal›
using lazy that unfolding C_W twl_lazy_update.simps by blast
show ?thesis
using Lev_M1 C_W that
using twl C excep twl n_d' watched_max 1
unfolding C_W
apply (auto simp: count_decided_ge_get_level
twl_is_an_exception_add_mset_to_queue atm_of_eq_atm_of that
dest!: no_has_blit_propagate' no_has_blit_propagate dest: 2)
using lazy unfolding C_W twl_lazy_update.simps apply blast
using lazy unfolding C_W twl_lazy_update.simps apply blast
using lazy unfolding C_W twl_lazy_update.simps apply blast
done
qed
then show ‹twl_lazy_update ?M1 C›
using lazy_D by blast
}
have ‹watched_literals_false_of_max_level M1 C› if ‹C ≠ ?D›
using past C that unfolding M past_invs.simps by auto
then have ‹watched_literals_false_of_max_level ?M1 C› if ‹C ≠ ?D›
using has_blit_Cons n_d' C_W that by (auto simp: get_level_cons_if)
then show ‹watched_literals_false_of_max_level ?M1 C›
using watched_max_D by blast
qed
case 2
show ?case
unfolding past_invs.simps Ball_def
proof (intro allI impI conjI)
fix M1'' M2'' K''' C
assume M1: ‹?M1 = M2'' @ Decided K''' # M1''› and C: ‹C ∈# N + add_mset ?D U›
then have M1: ‹M1 = tl M2'' @ Decided K''' # M1''›
by (cases M2'') auto
have ‹twl_lazy_update M1'' C›‹watched_literals_false_of_max_level M1'' C›
if ‹C ≠ ?D›
using past C that unfolding past_invs.simps M M1 twl_exception_inv.simps by auto
moreover {
have ‹twl_exception_inv (M1'', N, U, None, NE, UE, {#}, {#}) C› if ‹C ≠ ?D›
using past C unfolding past_invs.simps M M1 by (auto simp: that)
then have ‹twl_exception_inv (M1'', N, add_mset ?D U, None, NE, UE, {#}, {#}) C›
if ‹C ≠ ?D›
using C unfolding twl_exception_inv.simps by (auto simp: that) }
moreover {
have n_d_M1: ‹no_dup ?M1›
using struct_inv_T unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: cdcl⇩W_restart_mset_state)
then have ‹undefined_lit M1'' K'›
unfolding M1 by auto
moreover {
have ‹- K'' ∉ lits_of_l M1''›
proof (rule ccontr)
assume ‹¬ - K'' ∉ lits_of_l M1''›
then have ‹undefined_lit (tl M2'' @ Decided K''' # []) K''›
using n_d_M1 unfolding M1 by (auto simp: atm_lit_of_set_lits_of_l
atm_of_in_atm_of_set_iff_in_set_or_uminus_in_set
defined_lit_map atm_of_eq_atm_of image_Un
dest: no_dup_uminus_append_in_atm_notin)
then show False
using lev_M1_K'' count_decided_ge_get_level[of M1'' K''] unfolding M1
by (auto simp: image_Un Int_Un_distrib)
qed }
ultimately have ‹twl_lazy_update M1'' ?D› and
‹watched_literals_false_of_max_level M1'' ?D› and
‹twl_exception_inv (M1'', N, add_mset (TWL_Clause {#K', K''#} (D' - {#K', K''#})) U, None,
NE, UE, {#}, {#}) ?D›
by (auto simp: add_mset_eq_add_mset twl_exception_inv.simps get_level_cons_if
Decided_Propagated_in_iff_in_lits_of_l) }
ultimately show ‹twl_lazy_update M1'' C›
‹watched_literals_false_of_max_level M1'' C›
‹twl_exception_inv (M1'', N, add_mset (TWL_Clause {#K', K''#} (D' - {#K', K''#})) U, None,
NE, UE, {#}, {#}) C›
by blast+
next
fix M1'' M2'' K'''
assume M1: ‹?M1 = M2'' @ Decided K''' # M1''›
then have M1: ‹M1 = tl M2'' @ Decided K''' # M1''›
by (cases M2'') auto
then have confl_cands: ‹confl_cands_enqueued (M1'', N, U, None, NE, UE, {#}, {#})› and
propa_cands: ‹propa_cands_enqueued (M1'', N, U, None, NE, UE, {#}, {#})› and
w_q: ‹clauses_to_update_inv (M1'', N, U, None, NE, UE, {#}, {#})›
using past by (auto simp add: M M1 past_invs.simps simp del: propa_cands_enqueued.simps
confl_cands_enqueued.simps)
have uK''_M1'': ‹- K'' ∉ lits_of_l M1''›
proof (rule ccontr)
assume K''_M1'': ‹¬ ?thesis›
have ‹undefined_lit (tl M2'' @ Decided K''' # []) (-K'')›
apply (rule no_dup_append_in_atm_notin)
prefer 2 using K''_M1'' apply (simp; fail)
by (use n_d in ‹auto simp: M M1 no_dup_def; fail›)[]
then show False
using lev_M1_K'' count_decided_ge_get_level[of M1'' K''] unfolding M M1
by (auto simp: image_Un)
qed
have uK'_M1'': ‹- K' ∉ lits_of_l M1''›
proof (rule ccontr)
assume K'_M1'': ‹¬ ?thesis›
have ‹undefined_lit (M3 @ M2 @ Decided K # tl M2'' @ Decided K''' # []) (-K')›
apply (rule no_dup_append_in_atm_notin)
prefer 2 using K'_M1'' apply (simp; fail)
by (use n_d in ‹auto simp: M M1; fail›)[]
then show False
using lev_K' count_decided_ge_get_level[of M1'' K'] unfolding M M1
by (auto simp: image_Un)
qed
have [simp]: ‹¬clauses_to_update_prop {#} M1'' (L, ?D)› for L
using uK'_M1'' uK''_M1'' by (auto simp: clauses_to_update_prop.simps add_mset_eq_add_mset)
show ‹confl_cands_enqueued (M1'', N, add_mset ?D U, None, NE, UE, {#}, {#})› and
‹propa_cands_enqueued (M1'', N, add_mset ?D U, None, NE, UE, {#}, {#})› and
‹clauses_to_update_inv (M1'', N, add_mset ?D U, None, NE, UE, {#}, {#})›
using confl_cands propa_cands w_q uK'_M1'' uK''_M1''
by (fastforce simp add: twl_st_inv.simps add_mset_eq_add_mset)+
qed
qed
lemma
assumes
cdcl: ‹cdcl_twl_o S T›
shows
cdcl_twl_o_valid: ‹valid_enqueued T› and
cdcl_twl_o_conflict_None_queue:
‹get_conflict T ≠ None ⟹ clauses_to_update T = {#} ∧ literals_to_update T = {#}› and
cdcl_twl_o_no_duplicate_queued: ‹no_duplicate_queued T› and
cdcl_twl_o_distinct_queued: ‹distinct_queued T›
using cdcl by (induction rule: cdcl_twl_o.induct) auto
lemma cdcl_twl_o_twl_st_exception_inv:
assumes
cdcl: ‹cdcl_twl_o S T› and
twl: ‹twl_struct_invs S›
shows
‹twl_st_exception_inv T›
using cdcl twl
proof (induction rule: cdcl_twl_o.induct)
case (decide M L N U NE UE) note undef = this(1) and in_atms = this(2) and twl = this(3)
then have excep: ‹twl_st_exception_inv (M, N, NE, None, U, UE, {#}, {#})›
unfolding twl_struct_invs_def
by (auto simp: twl_exception_inv.simps)
let ?S = ‹(M, N, NE, None, U, UE, {#}, {#})›
have struct_inv_T: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?S)›
using cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_inv cdcl⇩W_restart_mset.other twl
unfolding twl_struct_invs_def by blast
have n_d: ‹no_dup M›
using twl unfolding twl_struct_invs_def cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: cdcl⇩W_restart_mset_state)
show ?case
using decide.hyps n_d excep
unfolding twl_struct_invs_def
by (auto simp: twl_exception_inv.simps dest!: no_has_blit_decide')
next
case (skip L D C' M N U NE UE)
then show ?case
unfolding twl_struct_invs_def by (auto simp: twl_exception_inv.simps)
next
case (resolve L D C M N U NE UE)
then show ?case
unfolding twl_struct_invs_def by (auto simp: twl_exception_inv.simps)
next
case (backtrack_unit_clause L D K M1 M2 M D' i N U NE UE) note decomp = this(2) and
invs = this(10)
let ?S = ‹(M, N, U, Some D, NE, UE, {#}, {#})›
let ?S' = ‹state⇩W_of S›
let ?T = ‹(M1, N, U, None, NE, UE, {#}, {#})›
let ?T' = ‹state⇩W_of T›
let ?U = ‹(Propagated L {#L#} # M1, N, U, None, NE, add_mset {#L#} UE, {#}, {#- L#})›
let ?U' = ‹state⇩W_of ?U›
have ‹twl_st_inv ?S› and past: ‹past_invs ?S› and valid: ‹valid_enqueued ?S›
using invs decomp unfolding twl_struct_invs_def by fast+
then have excep: ‹twl_exception_inv ?T C› if ‹C ∈# N + U› for C
using decomp that unfolding past_invs.simps by auto
have struct_inv_T: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?S)›
using invs unfolding twl_struct_invs_def by blast
have n_d: ‹no_dup M›
using invs unfolding twl_struct_invs_def cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: cdcl⇩W_restart_mset_state)
then have n_d: ‹no_dup M1›
using decomp by (auto dest: no_dup_appendD)
have struct_inv_U: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?U)›
using cdcl_twl_o_cdcl⇩W_o[OF cdcl_twl_o.backtrack_unit_clause[OF backtrack_unit_clause.hyps]
‹twl_st_inv ?S› valid struct_inv_T]
cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_inv cdcl⇩W_restart_mset.cdcl⇩W_restart.intros(3)
struct_inv_T by blast
then have undef: ‹undefined_lit M1 L›
unfolding twl_struct_invs_def cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: cdcl⇩W_restart_mset_state)
show ?case
using n_d excep undef
unfolding twl_struct_invs_def
by (auto simp: twl_exception_inv.simps dest!: no_has_blit_propagate')
next
case (backtrack_nonunit_clause L D K M1 M2 M D' i N U NE UE L') note decomp = this(2) and
lev_K = this(6) and lev_L' = this(12) and invs = this(13)
let ?S = ‹(M, N, U, Some D, NE, UE, {#}, {#})›
let ?D = ‹TWL_Clause {#L, L'#} (D' - {#L, L'#})›
let ?T = ‹(M1, N, U, None, NE, UE, {#}, {#})›
let ?U = ‹(Propagated L D' # M1, N, add_mset ?D U, None, NE, UE, {#}, {#- L#})›
have ‹twl_st_inv ?S› and past: ‹past_invs ?S› and valid: ‹valid_enqueued ?S›
using invs decomp unfolding twl_struct_invs_def by fast+
then have excep: ‹twl_exception_inv ?T C› if ‹C ∈# N + U› for C
using decomp that unfolding past_invs.simps by auto
have struct_inv_T: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?S)›
using invs unfolding twl_struct_invs_def by blast
have n_d_M: ‹no_dup M›
using invs unfolding twl_struct_invs_def cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: cdcl⇩W_restart_mset_state)
then have n_d: ‹no_dup M1›
using decomp by (auto dest: no_dup_appendD)
have struct_inv_U: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?U)›
using cdcl_twl_o_cdcl⇩W_o[OF cdcl_twl_o.backtrack_nonunit_clause[OF backtrack_nonunit_clause.hyps]
‹twl_st_inv ?S› valid struct_inv_T]
cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_inv cdcl⇩W_restart_mset.cdcl⇩W_restart.intros(3)
struct_inv_T by blast
then have undef: ‹undefined_lit M1 L›
unfolding twl_struct_invs_def cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (simp add: cdcl⇩W_restart_mset_state)
have n_d: ‹no_dup (Propagated L D' # M1)›
using struct_inv_U unfolding cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
by (simp add: trail.simps)
have ‹i = count_decided M1›
using decomp lev_K n_d_M by (auto dest!: get_all_ann_decomposition_exists_prepend
simp: get_level_append_if get_level_cons_if
split: if_splits)
then have lev_L'_M1: ‹get_level (Propagated L D' # M1) L' = count_decided M1›
using decomp lev_L' n_d_M by (auto dest!: get_all_ann_decomposition_exists_prepend
simp: get_level_append_if get_level_cons_if
split: if_splits)
have ‹- L ∉ lits_of_l M1›
using n_d by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
moreover have ‹has_blit (Propagated L D' # M1) (add_mset L (add_mset L' (D' - {#L, L'#}))) L'›
unfolding has_blit_def
apply (rule exI[of _ L])
using lev_L' lev_L'_M1
by auto
ultimately show ?case
using n_d excep undef
unfolding twl_struct_invs_def
by (auto simp: twl_exception_inv.simps dest!: no_has_blit_propagate')
qed
lemma
assumes
cdcl: ‹cdcl_twl_o S T› and
twl: ‹twl_struct_invs S›
shows
cdcl_twl_o_confl_cands_enqueued: ‹confl_cands_enqueued T› and
cdcl_twl_o_propa_cands_enqueued: ‹propa_cands_enqueued T› and
twl_o_clauses_to_update: ‹clauses_to_update_inv T›
using cdcl twl
proof (induction rule: cdcl_twl_o.induct)
case (decide M L N NE U UE)
let ?S = ‹(M, N, U, None, NE, UE, {#}, {#})›
let ?T = ‹(Decided L # M, N, U, None, NE, UE, {#}, {#-L#})›
case 1
then have confl_cand: ‹confl_cands_enqueued ?S› and
twl_st_inv: ‹twl_st_inv ?S› and
excep: ‹twl_st_exception_inv ?S› and
propa_cands: ‹propa_cands_enqueued ?S› and
confl_cands: ‹confl_cands_enqueued ?S› and
w_q: ‹clauses_to_update_inv ?S›
unfolding twl_struct_invs_def by fast+
have ‹cdcl⇩W_restart_mset.cdcl⇩W_o (state⇩W_of ?S) (state⇩W_of ?T)›
by (rule cdcl_twl_o_cdcl⇩W_o) (use cdcl_twl_o.decide[OF decide.hyps] 1 in
‹simp_all add: twl_struct_invs_def›)
then have ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?T)›
using 1 cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_inv cdcl⇩W_restart_mset.other twl_struct_invs_def
by blast
then have n_d: ‹no_dup (Decided L # M)›
unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def
by (auto simp: trail.simps)
show ?case
unfolding confl_cands_enqueued.simps Ball_def
proof (intro allI impI)
fix C
assume
C: ‹C ∈# N + U› and
LM_C: ‹Decided L # M ⊨as CNot (clause C)›
have struct_C: ‹struct_wf_twl_cls C›
using twl_st_inv C unfolding twl_st_inv.simps by blast
then have dist_C: ‹distinct_mset (clause C)›
by (cases C) auto
obtain W UW K K' where
C_W: ‹C = TWL_Clause W UW› and
W: ‹W = {#K, K'#}›
using struct_C by (cases C) (auto simp: size_2_iff)
have ‹¬M ⊨as CNot (clause C)›
using confl_cand C by auto
then have uL_C: ‹-L ∈# clause C› and neg_C: ‹∀K ∈# clause C. -K ∈ lits_of_l (Decided L # M)›
using LM_C unfolding true_annots_true_cls_def_iff_negation_in_model by auto
have ‹twl_exception_inv (M, N, U, None, NE, UE, {#}, {#}) C›
using excep C by auto
then have H: ‹L ∈# watched (TWL_Clause {#K, K'#} UW) ⟶
- L ∈ lits_of_l M ⟶ ¬ has_blit M (clause (TWL_Clause {#K, K'#} UW)) L ⟶
L ∉# {#} ⟶
(L, TWL_Clause {#K, K'#} UW) ∉# {#} ⟶
(∀K∈#unwatched (TWL_Clause {#K, K'#} UW).
- K ∈ lits_of_l M)› for L
unfolding twl_exception_inv.simps C_W W by blast
have excep: ‹L ∈# watched (TWL_Clause {#K, K'#} UW) ⟶
- L ∈ lits_of_l M ⟶ ¬ has_blit M (clause (TWL_Clause {#K, K'#} UW)) L ⟶
(∀K∈#unwatched (TWL_Clause {#K, K'#} UW). - K ∈ lits_of_l M)› for L
using H[of L] by simp
have ‹-L ∈# watched C›
proof (rule ccontr)
assume uL_W: ‹-L ∉# watched C›
then have uL_UW: ‹-L ∈# UW›
using uL_C unfolding C_W by auto
have ‹K ≠ -L ∨ K' ≠ -L›
using dist_C C_W W by auto
moreover have ‹K ∉ lits_of_l M› and ‹K' ∉ lits_of_l M› and L_M: ‹L ∉ lits_of_l M›
using neg_C uL_W n_d unfolding C_W W by (auto simp: lits_of_def uminus_lit_swap
no_dup_cannot_not_lit_and_uminus Decided_Propagated_in_iff_in_lits_of_l)
ultimately have disj: ‹(-K ∈ lits_of_l M ∧ K' ∉ lits_of_l M) ∨
(-K' ∈ lits_of_l M ∧ K ∉ lits_of_l M)›
using neg_C by (auto simp: C_W W)
have ‹¬has_blit M (clause C) K›
using ‹K ∉ lits_of_l M› ‹K' ∉ lits_of_l M›
using uL_C neg_C n_d unfolding has_blit_def by (auto dest!: multi_member_split
dest!: no_dup_consistentD
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
moreover have ‹¬has_blit M (clause C) K'›
using ‹K' ∉ lits_of_l M› ‹ K ∉ lits_of_l M›
using uL_C neg_C n_d unfolding has_blit_def by (auto dest!: multi_member_split
dest!: no_dup_consistentD
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
ultimately have ‹∀K ∈# unwatched C. -K ∈ lits_of_l M›
apply -
apply (rule disjE[OF disj])
subgoal
using excep[of K]
unfolding C_W twl_clause.sel member_add_mset W
by auto
subgoal
using excep[of K']
unfolding C_W twl_clause.sel member_add_mset W
by auto
done
then show False
using uL_W uL_C L_M unfolding C_W W by auto
qed
then show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- L#}) ∨ (∃L. (L, C) ∈# {#})›
by auto
qed
case 2
show ?case
unfolding propa_cands_enqueued.simps Ball_def
proof (intro allI impI)
fix FK C
assume
C: ‹C ∈# N + U› and
K: ‹FK ∈# clause C› and
LM_C: ‹Decided L # M ⊨as CNot (remove1_mset FK (clause C))› and
undef: ‹undefined_lit (Decided L # M) FK›
have undef_M_K: ‹undefined_lit M FK›
using undef by (auto simp: defined_lit_map)
then have ‹¬ M ⊨as CNot (remove1_mset FK (clause C))›
using propa_cands C K undef by auto
then have ‹-L ∈# clause C› and
neg_C: ‹∀K ∈# remove1_mset FK (clause C). -K ∈ lits_of_l (Decided L # M)›
using LM_C undef_M_K by (force simp: true_annots_true_cls_def_iff_negation_in_model
dest: in_diffD)+
have struct_C: ‹struct_wf_twl_cls C›
using twl_st_inv C unfolding twl_st_inv.simps by blast
then have dist_C: ‹distinct_mset (clause C)›
by (cases C) auto
have ‹-L ∈# watched C›
proof (rule ccontr)
assume uL_W: ‹-L ∉# watched C›
then obtain W UW K K' where
C_W: ‹C = TWL_Clause W UW› and
W: ‹W = {#K, K'#}› and
uK_M: ‹-K ∈ lits_of_l M›
using struct_C neg_C by (cases C) (auto simp: size_2_iff remove1_mset_add_mset_If
add_mset_commute split: if_splits)
have FK_F: ‹FK ≠ K›
using Decided_Propagated_in_iff_in_lits_of_l uK_M undef_M_K by blast
have L_M: ‹undefined_lit M L›
using neg_C uL_W n_d unfolding C_W W by auto
then have ‹K ≠ -L›
using uK_M by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
moreover have ‹K ∉ lits_of_l M›
using neg_C uL_W n_d uK_M by (auto simp: lits_of_def uminus_lit_swap
no_dup_cannot_not_lit_and_uminus)
ultimately have ‹K' ∉ lits_of_l M›
apply (cases ‹K' = FK›)
using Decided_Propagated_in_iff_in_lits_of_l undef_M_K apply blast
using neg_C C_W W FK_F n_d uL_W by (auto simp add: remove1_mset_add_mset_If uminus_lit_swap
lits_of_def no_dup_cannot_not_lit_and_uminus)
moreover have ‹twl_exception_inv (M, N, U, None, NE, UE, {#}, {#}) C›
using excep C by auto
moreover have ‹¬has_blit M (clause C) K›
using ‹K ∉ lits_of_l M› ‹K' ∉ lits_of_l M›
using K in_lits_of_l_defined_litD neg_C undef_M_K n_d unfolding has_blit_def
by (force dest!: multi_member_split
dest!: no_dup_consistentD
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
moreover have ‹¬has_blit M (clause C) K'›
using ‹K' ∉ lits_of_l M› ‹ K ∉ lits_of_l M› K in_lits_of_l_defined_litD neg_C undef_M_K
using n_d unfolding has_blit_def by (force dest!: multi_member_split
dest!: no_dup_consistentD
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
ultimately have ‹∀K ∈# unwatched C. -K ∈ lits_of_l M›
using uK_M
by (auto simp: twl_exception_inv.simps C_W W add_mset_eq_add_mset all_conj_distrib)
then show False
using C_W L_M(1) ‹- L ∈# clause C› uL_W
by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
qed
then show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- L#}) ∨ (∃L. (L, C) ∈# {#})›
by auto
qed
case 3
show ?case
proof (induction rule: clauses_to_update_inv_cases)
case (WS_nempty L C)
then show ?case by simp
next
case (WS_empty K)
then show ?case
using w_q n_d unfolding clauses_to_update_prop.simps
by (auto simp add: filter_mset_empty_conv
dest!: no_has_blit_decide')
next
case (Q K C)
then show ?case
using w_q n_d by (auto dest!: no_has_blit_decide')
qed
next
case (skip L D C' M N U NE UE)
case 1 then show ?case by auto
case 2 then show ?case by auto
case 3 then show ?case by auto
next
case (resolve L D C M N U NE UE)
case 1 then show ?case by auto
case 2 then show ?case by auto
case 3 then show ?case by auto
next
case (backtrack_unit_clause L D K M1 M2 M D' i N U NE UE) note decomp = this(2)
let ?S = ‹(M, N, U, Some D, NE, UE, {#}, {#})›
let ?U = ‹(Propagated L {#L#} # M1, N, U, None, NE, add_mset {#L#} UE, {#}, {#- L#})›
obtain M3 where
M: ‹M = M3 @ M2 @ Decided K # M1›
using decomp by blast
case 1
then have twl_st_inv: ‹twl_st_inv ?S› and
struct_inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?S)› and
excep: ‹twl_st_exception_inv ?S› and
past: ‹past_invs ?S›
using decomp unfolding twl_struct_invs_def by fast+
then have
confl_cands: ‹confl_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
propa_cands: ‹propa_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})›and
w_q: ‹clauses_to_update_inv (M1, N, U, None, NE, UE, {#}, {#})›
using decomp unfolding past_invs.simps by (auto simp del: clauses_to_update_inv.simps)
have n_d: ‹no_dup M›
using struct_inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
have ‹cdcl⇩W_restart_mset.cdcl⇩W_o (state⇩W_of ?S) (state⇩W_of ?U)›
using cdcl_twl_o.backtrack_unit_clause[OF backtrack_unit_clause.hyps]
by (meson "1.prems" twl_struct_invs_def cdcl_twl_o_cdcl⇩W_o)
then have struct_inv_T: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?U)›
using struct_inv cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_inv cdcl⇩W_restart_mset.other by blast
then have n_d_L_M1: ‹no_dup (Propagated L {#L#} # M1)›
using struct_inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
then have uL_M1: ‹undefined_lit M1 L›
by (simp_all add: atm_lit_of_set_lits_of_l atm_of_in_atm_of_set_iff_in_set_or_uminus_in_set)
have excep_M1: ‹∀C ∈# N + U. twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using past unfolding past_invs.simps M by auto
show ?case
unfolding confl_cands_enqueued.simps Ball_def
proof (intro allI impI)
fix C
assume
C: ‹C ∈# N + U› and
LM_C: ‹Propagated L {#L#} # M1 ⊨as CNot (clause C)›
have struct_C: ‹struct_wf_twl_cls C›
using twl_st_inv C unfolding twl_st_inv.simps by auto
then have dist_C: ‹distinct_mset (clause C)›
by (cases C) auto
obtain W UW K K' where
C_W: ‹C = TWL_Clause W UW› and
W: ‹W = {#K, K'#}›
using struct_C by (cases C) (auto simp: size_2_iff)
have ‹¬M1 ⊨as CNot (clause C)›
using confl_cands C by auto
then have uL_C: ‹-L ∈# clause C› and neg_C: ‹∀K ∈# clause C. -K ∈ lits_of_l (Decided L # M1)›
using LM_C unfolding true_annots_true_cls_def_iff_negation_in_model by auto
have K_L: ‹K ≠ L› and K'_L: ‹K' ≠ L›
apply (metis C_W LM_C W add_diff_cancel_right' clause.simps consistent_interp_def
distinct_consistent_interp in_CNot_implies_uminus(2) in_diffD n_d_L_M1 uL_C
union_single_eq_member)
using C_W LM_C W uL_M1 by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
have ‹-L ∈# watched C›
proof (rule ccontr)
assume uL_W: ‹-L ∉# watched C›
have ‹K ≠ -L ∨ K' ≠ -L›
using dist_C C_W W by auto
moreover have ‹K ∉ lits_of_l M1› and ‹K' ∉ lits_of_l M1› and L_M: ‹L ∉ lits_of_l M1›
proof -
have f2: ‹consistent_interp (lits_of_l M1)›
using distinct_consistent_interp n_d_L_M1 by auto
have undef_L: ‹undefined_lit M1 L›
using atm_lit_of_set_lits_of_l n_d_L_M1 by force
then show ‹K ∉ lits_of_l M1›
using f2 neg_C unfolding C_W W by (metis (no_types) C_W W add_diff_cancel_right'
atm_of_eq_atm_of clause.simps
consistent_interp_def in_diffD insertE list.simps(15) lits_of_insert uL_C
union_single_eq_member Decided_Propagated_in_iff_in_lits_of_l)
show ‹K' ∉ lits_of_l M1›
using consistent_interp_def distinct_consistent_interp n_d_L_M1
using neg_C uL_W n_d unfolding C_W W by auto
show ‹L ∉ lits_of_l M1›
using undef_L by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
qed
ultimately have ‹(-K ∈ lits_of_l M1 ∧ K' ∉ lits_of_l M1) ∨
(-K' ∈ lits_of_l M1 ∧ K ∉ lits_of_l M1)›
using neg_C by (auto simp: C_W W)
moreover have ‹twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using excep_M1 C by auto
have ‹¬has_blit M1 (clause C) K›
using ‹K ∉ lits_of_l M1› ‹K' ∉ lits_of_l M1› ‹L ∉ lits_of_l M1› uL_M1
n_d_L_M1 no_dup_cons
using uL_C neg_C n_d unfolding has_blit_def apply (auto dest!: multi_member_split
dest!: no_dup_consistentD[OF n_d_L_M1]
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
using n_d_L_M1 no_dup_cons no_dup_consistentD by blast
moreover have ‹¬has_blit M1 (clause C) K'›
using ‹K' ∉ lits_of_l M1› ‹ K ∉ lits_of_l M1› ‹L ∉ lits_of_l M1› uL_M1
n_d_L_M1 no_dup_cons no_dup_consistentD
using uL_C neg_C n_d unfolding has_blit_def apply (auto 10 10 dest!: multi_member_split
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
using n_d_L_M1 no_dup_cons no_dup_consistentD by auto
ultimately have ‹∀K ∈# unwatched C. -K ∈ lits_of_l M1›
using C twl_clause.sel(1) union_single_eq_member w_q
by (fastforce simp: twl_exception_inv.simps C_W W add_mset_eq_add_mset all_conj_distrib L_M)
then show False
using uL_W uL_C L_M K_L uL_M1 unfolding C_W W by auto
qed
then show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- L#}) ∨ (∃L. (L, C) ∈# {#})›
by auto
qed
case 2
then show ?case
unfolding propa_cands_enqueued.simps Ball_def
proof (intro allI impI)
fix FK C
assume
C: ‹C ∈# N + U› and
K: ‹FK ∈# clause C› and
LM_C: ‹Propagated L {#L#} # M1 ⊨as CNot (remove1_mset FK (clause C))› and
undef: ‹undefined_lit (Propagated L {#L#} # M1) FK›
have undef_M_K: ‹undefined_lit (Propagated L D # M1) FK›
using undef by (auto simp: defined_lit_map)
then have ‹¬ M1 ⊨as CNot (remove1_mset FK (clause C))›
using propa_cands C K undef by (auto simp: defined_lit_map)
then have uL_C: ‹-L ∈# clause C› and
neg_C: ‹∀K ∈# remove1_mset FK (clause C). -K ∈ lits_of_l (Propagated L D # M1)›
using LM_C undef_M_K by (force simp: true_annots_true_cls_def_iff_negation_in_model
dest: in_diffD)+
have struct_C: ‹struct_wf_twl_cls C›
using twl_st_inv C unfolding twl_st_inv.simps by blast
then have dist_C: ‹distinct_mset (clause C)›
by (cases C) auto
moreover have ‹-L ∈# watched C›
proof (rule ccontr)
assume uL_W: ‹-L ∉# watched C›
then obtain W UW K K' where
C_W: ‹C = TWL_Clause W UW› and
W: ‹W = {#K, K'#}› and
uK_M: ‹-K ∈ lits_of_l M1›
using struct_C neg_C by (cases C) (auto simp: size_2_iff remove1_mset_add_mset_If
add_mset_commute split: if_splits)
have ‹K ∉ lits_of_l M1› and L_M: ‹L ∉ lits_of_l M1›
proof -
have f2: ‹consistent_interp (lits_of_l M1)›
using distinct_consistent_interp n_d_L_M1 by auto
have undef_L: ‹undefined_lit M1 L›
using atm_lit_of_set_lits_of_l n_d_L_M1 by force
then show ‹K ∉ lits_of_l M1›
using f2 neg_C unfolding C_W W
using n_d_L_M1 no_dup_cons no_dup_consistentD uK_M by blast
show ‹L ∉ lits_of_l M1›
using undef_L by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
qed
have FK_F: ‹FK ≠ K›
using uK_M undef_M_K unfolding Decided_Propagated_in_iff_in_lits_of_l by auto
have ‹K ≠ -L›
using uK_M uL_M1 by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
moreover have ‹K ∉ lits_of_l M1›
using neg_C uL_W n_d uK_M n_d_L_M1 by (auto simp: lits_of_def uminus_lit_swap
no_dup_cannot_not_lit_and_uminus dest: no_dup_cannot_not_lit_and_uminus)
ultimately have ‹K' ∉ lits_of_l M1›
apply (cases ‹K' = FK›)
using undef_M_K apply (force simp: Decided_Propagated_in_iff_in_lits_of_l)
using neg_C C_W W FK_F n_d uL_W n_d_L_M1 by (auto simp add: remove1_mset_add_mset_If
uminus_lit_swap lits_of_def no_dup_cannot_not_lit_and_uminus
dest: no_dup_cannot_not_lit_and_uminus)
moreover have ‹twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using excep_M1 C by auto
moreover have ‹¬has_blit M1 (clause C) K›
using ‹K ∉ lits_of_l M1› ‹K' ∉ lits_of_l M1› ‹L ∉ lits_of_l M1› uL_M1
n_d_L_M1 no_dup_cons K undef
using uL_C neg_C n_d unfolding has_blit_def apply (auto dest!: multi_member_split
dest!: no_dup_consistentD[OF n_d_L_M1]
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
by (smt add_mset_commute add_mset_eq_add_mset defined_lit_uminus in_lits_of_l_defined_litD
insert_DiffM no_dup_consistentD set_subset_Cons true_annot_mono true_annot_singleton)+
moreover have ‹¬has_blit M1 (clause C) K'›
using ‹K' ∉ lits_of_l M1› ‹ K ∉ lits_of_l M1› ‹L ∉ lits_of_l M1› uL_M1
n_d_L_M1 no_dup_cons no_dup_consistentD K undef
using uL_C neg_C n_d unfolding has_blit_def apply (auto 10 10 dest!: multi_member_split
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
by (smt add_mset_commute add_mset_eq_add_mset defined_lit_uminus in_lits_of_l_defined_litD
insert_DiffM no_dup_consistentD set_subset_Cons true_annot_mono true_annot_singleton)+
ultimately have ‹∀K ∈# unwatched C. -K ∈ lits_of_l M1›
using uK_M
by (auto simp: twl_exception_inv.simps C_W W add_mset_eq_add_mset all_conj_distrib)
then show False
using C_W uL_M1 ‹- L ∈# clause C› uL_W
by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
qed
then show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- L#}) ∨ (∃L. (L, C) ∈# {#})›
by auto
qed
case 3
have
2: ‹⋀L. Pair L `# {#C ∈# N + U. clauses_to_update_prop {#} M1 (L, C)#} = {#}› and
3: ‹⋀L C. C ∈# N + U ⟹ L ∈# watched C ⟹ - L ∈ lits_of_l M1 ⟹
¬ has_blit M1 (clause C) L ⟹ (L, C) ∉# {#} ⟹ L ∈# {#}›
using w_q unfolding clauses_to_update_inv.simps by auto
show ?case
proof (induction rule: clauses_to_update_inv_cases)
case (WS_nempty L C)
then show ?case by simp
next
case (WS_empty K)
then show ?case
using 2[of K] n_d_L_M1
apply (simp only: filter_mset_empty_conv Ball_def image_mset_is_empty_iff)
by (auto simp add: clauses_to_update_prop.simps)
next
case (Q K C)
then show ?case
using 3[of C K] has_blit_Cons n_d_L_M1 by (fastforce simp add: clauses_to_update_prop.simps)
qed
next
case (backtrack_nonunit_clause L D K M1 M2 M D' i N U NE UE L') note LD = this(1) and
decomp = this(2) and lev_L = this(3) and lev_max_L = this(4) and i = this(5) and lev_K = this(6)
and LD' = this(11) and lev_L' = this(12)
let ?S = ‹(M, N, U, Some D, NE, UE, {#}, {#})›
let ?D = ‹TWL_Clause {#L, L'#} (D' - {#L, L'#})›
let ?U = ‹(Propagated L D' # M1, N, add_mset ?D U, None, NE,
UE, {#}, {#- L#})›
obtain M3 where
M: ‹M = M3 @ M2 @ Decided K # M1›
using decomp by blast
case 1
then have twl_st_inv: ‹twl_st_inv ?S› and
struct_inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?S)› and
excep: ‹twl_st_exception_inv ?S› and
past: ‹past_invs ?S›
using decomp unfolding twl_struct_invs_def by fast+
then have
confl_cands: ‹confl_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
propa_cands: ‹propa_cands_enqueued (M1, N, U, None, NE, UE, {#}, {#})› and
w_q: ‹clauses_to_update_inv (M1, N, U, None, NE, UE, {#}, {#})›
using decomp unfolding past_invs.simps by auto
have n_d: ‹no_dup M›
using struct_inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
have ‹undefined_lit (M3 @ M2 @ M1) K›
by (rule no_dup_append_in_atm_notin[of _ ‹[Decided K]›])
(use n_d M in ‹auto simp: no_dup_def›)
then have L_uL': ‹L ≠ - L'›
using lev_L lev_L' lev_K unfolding M by (auto simp: image_Un)
have ‹cdcl⇩W_restart_mset.cdcl⇩W_o (state⇩W_of ?S) (state⇩W_of ?U)›
using cdcl_twl_o.backtrack_nonunit_clause[OF backtrack_nonunit_clause.hyps]
by (meson "1.prems" twl_struct_invs_def cdcl_twl_o_cdcl⇩W_o)
then have struct_inv_T: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of ?U)›
using struct_inv cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_inv cdcl⇩W_restart_mset.other by blast
then have n_d_L_M1: ‹no_dup (Propagated L D' # M1)›
using struct_inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
then have uL_M1: ‹undefined_lit M1 L›
by simp
have M1_CNot_L_D: ‹M1 ⊨as CNot (remove1_mset L D')›
using struct_inv_T unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_conflicting_def by (auto simp: trail.simps)
have L_M1: ‹- L ∉ lits_of_l M1› ‹L ∉ lits_of_l M1›
using n_d n_d_L_M1 uL_M1 by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
have excep_M1: ‹∀C ∈# N + U. twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using past unfolding past_invs.simps M by auto
show ?case
unfolding confl_cands_enqueued.simps Ball_def
proof (intro allI impI)
fix C
assume
C: ‹C ∈# N + add_mset ?D U› and
LM_C: ‹Propagated L D' # M1 ⊨as CNot (clause C)›
have ‹twl_st_inv ?U›
using cdcl_twl_o.backtrack_nonunit_clause[OF backtrack_nonunit_clause.hyps] "1.prems"
cdcl_twl_o_twl_st_inv by blast
then have ‹struct_wf_twl_cls ?D›
unfolding twl_st_inv.simps by auto
show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- L#}) ∨ (∃L. (L, C) ∈# {#})›
proof (cases ‹C = ?D›)
case True
then have False
using LM_C L_uL' uL_M1 by (auto simp: true_annots_true_cls_def_iff_negation_in_model
Decided_Propagated_in_iff_in_lits_of_l)
then show ?thesis by fast
next
case False
have struct_C: ‹struct_wf_twl_cls C›
using twl_st_inv C False unfolding twl_st_inv.simps by auto
then have dist_C: ‹distinct_mset (clause C)›
by (cases C) auto
have C: ‹C ∈# N + U›
using C False by auto
obtain W UW K K' where
C_W: ‹C = TWL_Clause W UW› and
W: ‹W = {#K, K'#}›
using struct_C by (cases C) (auto simp: size_2_iff)
have ‹¬M1 ⊨as CNot (clause C)›
using confl_cands C by auto
then have uL_C: ‹-L ∈# clause C› and neg_C: ‹∀K ∈# clause C. -K ∈ lits_of_l (Decided L # M1)›
using LM_C unfolding true_annots_true_cls_def_iff_negation_in_model by auto
have K_L: ‹K ≠ L› and K'_L: ‹K' ≠ L›
apply (metis C_W LM_C W add_diff_cancel_right' clause.simps consistent_interp_def
distinct_consistent_interp in_CNot_implies_uminus(2) in_diffD n_d_L_M1 uL_C
union_single_eq_member)
using C_W LM_C W uL_M1 by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
have ‹-L ∈# watched C›
proof (rule ccontr)
assume uL_W: ‹-L ∉# watched C›
have ‹K ≠ -L ∨ K' ≠ -L›
using dist_C C_W W by auto
moreover have ‹K ∉ lits_of_l M1› and ‹K' ∉ lits_of_l M1› and L_M: ‹L ∉ lits_of_l M1›
proof -
have f2: ‹consistent_interp (lits_of_l M1)›
using distinct_consistent_interp n_d_L_M1 by auto
have undef_L: ‹undefined_lit M1 L›
using atm_lit_of_set_lits_of_l n_d_L_M1 by force
then show ‹K ∉ lits_of_l M1›
using f2 neg_C unfolding C_W W by (metis (no_types) C_W W add_diff_cancel_right'
atm_of_eq_atm_of clause.simps consistent_interp_def in_diffD insertE list.simps(15)
lits_of_insert uL_C union_single_eq_member Decided_Propagated_in_iff_in_lits_of_l)
show ‹K' ∉ lits_of_l M1›
using consistent_interp_def distinct_consistent_interp n_d_L_M1
using neg_C uL_W n_d unfolding C_W W by auto
show ‹L ∉ lits_of_l M1›
using undef_L by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
qed
ultimately have ‹(-K ∈ lits_of_l M1 ∧ K' ∉ lits_of_l M1) ∨
(-K' ∈ lits_of_l M1 ∧ K ∉ lits_of_l M1)›
using neg_C by (auto simp: C_W W)
moreover have ‹¬has_blit M1 (clause C) K›
using ‹K ∉ lits_of_l M1› ‹K' ∉ lits_of_l M1› ‹L ∉ lits_of_l M1› uL_M1
n_d_L_M1 no_dup_cons
using uL_C neg_C n_d unfolding has_blit_def apply (auto dest!: multi_member_split
dest!: no_dup_consistentD[OF n_d_L_M1]
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
using n_d_L_M1 no_dup_cons no_dup_consistentD by blast
moreover have ‹¬has_blit M1 (clause C) K'›
using ‹K' ∉ lits_of_l M1› ‹ K ∉ lits_of_l M1› ‹L ∉ lits_of_l M1› uL_M1
n_d_L_M1 no_dup_cons no_dup_consistentD
using uL_C neg_C n_d unfolding has_blit_def apply (auto 10 10 dest!: multi_member_split
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
using n_d_L_M1 no_dup_cons no_dup_consistentD by auto
moreover have ‹twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using excep_M1 C by auto
ultimately have ‹∀K ∈# unwatched C. -K ∈ lits_of_l M1›
using C twl_clause.sel(1) union_single_eq_member w_q
by (fastforce simp: twl_exception_inv.simps C_W W add_mset_eq_add_mset all_conj_distrib
L_M)
then show False
using uL_W uL_C L_M K_L uL_M1 unfolding C_W W by auto
qed
then show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- L#}) ∨ (∃L. (L, C) ∈# {#})›
by auto
qed
qed
case 2
then show ?case
unfolding propa_cands_enqueued.simps Ball_def
proof (intro allI impI)
fix FK C
assume
C: ‹C ∈# N + add_mset ?D U› and
K: ‹FK ∈# clause C› and
LM_C: ‹Propagated L D' # M1 ⊨as CNot (remove1_mset FK (clause C))› and
undef: ‹undefined_lit (Propagated L D' # M1) FK›
show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- L#}) ∨ (∃L. (L, C) ∈# {#})›
proof (cases ‹C = ?D›)
case False
then have C: ‹C ∈# N + U›
using C by auto
have undef_M_K: ‹undefined_lit (Propagated L D # M1) FK›
using undef by (auto simp: defined_lit_map)
then have ‹¬ M1 ⊨as CNot (remove1_mset FK (clause C))›
using propa_cands C K undef by (auto simp: defined_lit_map)
then have ‹-L ∈# clause C› and
neg_C: ‹∀K ∈# remove1_mset FK (clause C). -K ∈ lits_of_l (Propagated L D # M1)›
using LM_C undef_M_K by (force simp: true_annots_true_cls_def_iff_negation_in_model
dest: in_diffD)+
have struct_C: ‹struct_wf_twl_cls C›
using twl_st_inv C unfolding twl_st_inv.simps by blast
then have dist_C: ‹distinct_mset (clause C)›
by (cases C) auto
have ‹-L ∈# watched C›
proof (rule ccontr)
assume uL_W: ‹-L ∉# watched C›
then obtain W UW K K' where
C_W: ‹C = TWL_Clause W UW› and
W: ‹W = {#K, K'#}› and
uK_M: ‹-K ∈ lits_of_l M1›
using struct_C neg_C by (cases C) (auto simp: size_2_iff remove1_mset_add_mset_If
add_mset_commute split: if_splits)
have FK_F: ‹FK ≠ K›
using uK_M undef_M_K unfolding Decided_Propagated_in_iff_in_lits_of_l by auto
have ‹K ≠ -L›
using uK_M uL_M1 by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
moreover have ‹K ∉ lits_of_l M1›
using neg_C uL_W n_d uK_M n_d_L_M1 by (auto simp: lits_of_def uminus_lit_swap
no_dup_cannot_not_lit_and_uminus dest: no_dup_cannot_not_lit_and_uminus)
ultimately have ‹K' ∉ lits_of_l M1›
apply (cases ‹K' = FK›)
using undef_M_K apply (force simp: Decided_Propagated_in_iff_in_lits_of_l)
using neg_C C_W W FK_F n_d uL_W n_d_L_M1 by (auto simp add: remove1_mset_add_mset_If
uminus_lit_swap lits_of_def no_dup_cannot_not_lit_and_uminus
dest: no_dup_cannot_not_lit_and_uminus)
moreover have ‹twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using excep_M1 C by auto
moreover have ‹¬has_blit M1 (clause C) K›
using ‹K ∉ lits_of_l M1› ‹K' ∉ lits_of_l M1› uL_M1
n_d_L_M1 no_dup_cons
using n_d_L_M1 no_dup_cons no_dup_consistentD
using K in_lits_of_l_defined_litD undef
using neg_C n_d unfolding has_blit_def by (fastforce dest!: multi_member_split
dest!: no_dup_consistentD[OF n_d_L_M1]
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
moreover have ‹¬has_blit M1 (clause C) K'›
using ‹K' ∉ lits_of_l M1› ‹ K ∉ lits_of_l M1› uL_M1
n_d_L_M1 no_dup_cons no_dup_consistentD
using n_d_L_M1 no_dup_cons no_dup_consistentD
using K in_lits_of_l_defined_litD undef
using neg_C n_d unfolding has_blit_def by (fastforce dest!: multi_member_split
dest!: in_lits_of_l_defined_litD[of ‹-L›] simp: add_mset_eq_add_mset)
moreover have ‹twl_exception_inv (M1, N, U, None, NE, UE, {#}, {#}) C›
using excep_M1 C by auto
ultimately have ‹∀K ∈# unwatched C. -K ∈ lits_of_l M1›
using uK_M
by (auto simp: twl_exception_inv.simps C_W W add_mset_eq_add_mset all_conj_distrib)
then show False
using C_W uL_M1 ‹- L ∈# clause C› uL_W
by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
qed
then show ‹(∃L'. L' ∈# watched C ∧ L' ∈# {#- L#}) ∨ (∃L. (L, C) ∈# {#})›
by auto
next
case True
then have ‹∀K∈#remove1_mset L D'. - K ∈ lits_of_l (Propagated L D' # M1)›
using M1_CNot_L_D by (auto simp: true_annots_true_cls_def_iff_negation_in_model)
then have ‹∀K∈#remove1_mset L D'. defined_lit (Propagated L D' # M1) K›
using Decided_Propagated_in_iff_in_lits_of_l by blast
moreover have ‹defined_lit (Propagated L D' # M1) L›
by (auto simp: defined_lit_map)
ultimately have ‹∀K∈#D'. defined_lit (Propagated L D' # M1) K›
by (metis in_remove1_mset_neq)
then have ‹∀K∈#clause ?D. defined_lit (Propagated L D' # M1) K›
using LD' ‹defined_lit (Propagated L D' # M1) L› by (auto dest: in_diffD)
then have False
using K undef unfolding True by (auto simp: Decided_Propagated_in_iff_in_lits_of_l)
then show ?thesis by fast
qed
qed
case 3
then have
2: ‹⋀L. Pair L `# {#C ∈# N + U. clauses_to_update_prop {#} M1 (L, C)#} = {#}› and
3: ‹⋀L C. C ∈# N + U ⟹ L ∈# watched C ⟹ - L ∈ lits_of_l M1 ⟹
¬ has_blit M1 (clause C) L ⟹ (L, C) ∉# {#} ⟹ L ∈# {#}›
using w_q unfolding clauses_to_update_inv.simps by auto
have ‹i = count_decided M1›
using decomp lev_K n_d by (auto dest!: get_all_ann_decomposition_exists_prepend
simp: get_level_append_if get_level_cons_if
split: if_splits)
then have lev_L'_M1: ‹get_level (Propagated L D' # M1) L' = count_decided M1›
using decomp lev_L' n_d by (auto dest!: get_all_ann_decomposition_exists_prepend
simp: get_level_append_if get_level_cons_if
split: if_splits)
have blit_L': ‹has_blit (Propagated L D' # M1) (add_mset L (add_mset L' (D' - {#L, L'#}))) L'›
unfolding has_blit_def
by (rule_tac x=L in exI) (auto simp: lev_L'_M1)
show ?case
proof (induction rule: clauses_to_update_inv_cases)
case (WS_nempty L C)
then show ?case by simp
next
case (WS_empty K')
show ?case
using 2[of K] "3" n_d_L_M1 L_M1 blit_L'
apply (simp only: filter_mset_empty_conv Ball_def image_mset_is_empty_iff)
by (fastforce simp add: clauses_to_update_prop.simps )
next
case (Q K' C)
then show ?case
using 3[of C K'] uL_M1 blit_L' n_d_L_M1 has_blit_Cons
by (fastforce simp add: clauses_to_update_prop.simps
add_mset_eq_add_mset Decided_Propagated_in_iff_in_lits_of_l)
qed
qed
lemma no_dup_append_decided_Cons_lev:
assumes ‹no_dup (M2 @ Decided K # M1)›
shows ‹count_decided M1 = get_level (M2 @ Decided K # M1) K - 1›
proof -
have ‹undefined_lit (M2 @ M1) K›
by (rule no_dup_append_in_atm_notin[of _
‹[Decided K]›])
(use assms in auto)
then show ?thesis
by (auto)
qed
lemma cdcl_twl_o_entailed_clss_inv:
assumes
cdcl: ‹cdcl_twl_o S T› and
unit: ‹twl_struct_invs S›
shows ‹entailed_clss_inv T›
using cdcl unit
proof (induction rule: cdcl_twl_o.induct)
case (decide M L N NE U UE) note undef = this(1) and twl = this(3)
then have unit: ‹entailed_clss_inv (M, N, U, None, NE, UE, {#}, {#})›
unfolding twl_struct_invs_def by fast
show ?case
unfolding entailed_clss_inv.simps Ball_def
proof (intro allI impI)
fix C
assume ‹C ∈# NE + UE›
then obtain K where ‹K ∈# C› and K: ‹K ∈ lits_of_l M› and ‹get_level M K = 0›
using unit by auto
moreover have ‹atm_of L ≠ atm_of K›
using undef K by (auto simp: defined_lit_map lits_of_def)
ultimately show ‹∃La. La ∈# C ∧ (None = None ∨ 0 < count_decided (Decided L # M) ⟶
get_level (Decided L # M) La = 0 ∧ La ∈ lits_of_l (Decided L # M))›
by auto
qed
next
case (skip L D C' M N U NE UE) note twl = this(3)
let ?M = ‹Propagated L C' # M›
have unit: ‹entailed_clss_inv (?M, N, U, Some D, NE, UE, {#}, {#})›
using twl unfolding twl_struct_invs_def by fast
show ?case
unfolding entailed_clss_inv.simps Ball_def
proof (intro allI impI, cases ‹count_decided M = 0›)
case True note [simp] = this
fix C
assume ‹C ∈# NE + UE›
then obtain K where ‹K ∈# C›
using unit by auto
then show ‹∃L. L ∈# C ∧ (Some D = None ∨ 0 < count_decided M ⟶
get_level M L = 0 ∧ L ∈ lits_of_l M)›
by auto
next
case False
fix C
assume ‹C ∈# NE + UE›
then obtain K where ‹K ∈# C› and K: ‹K ∈ lits_of_l ?M› and lev_K: ‹get_level ?M K = 0›
using unit False by auto
moreover {
have ‹get_level ?M L > 0›
using False by auto
then have ‹atm_of L ≠ atm_of K›
using lev_K by fastforce }
ultimately show ‹∃L. L ∈# C ∧ (Some D = None ∨ 0 < count_decided M ⟶
get_level M L = 0 ∧ L ∈ lits_of_l M)›
using False by auto
qed
next
case (resolve L D C M N U NE UE) note twl = this(3)
let ?M = ‹Propagated L C # M›
let ?D = ‹Some (remove1_mset (- L) D ∪# remove1_mset L C)›
have unit: ‹entailed_clss_inv (?M, N, U, Some D, NE, UE, {#}, {#})›
using twl unfolding twl_struct_invs_def by fast
show ?case
unfolding entailed_clss_inv.simps Ball_def
proof (intro allI impI, cases ‹count_decided M = 0›)
case True note [simp] = this
fix E
assume ‹E ∈# NE + UE›
then obtain K where ‹K ∈# E›
using unit by auto
then show ‹∃La. La ∈# E ∧ (?D = None ∨ 0 < count_decided M ⟶
get_level M La = 0 ∧ La ∈ lits_of_l M)›
by auto
next
case False
fix E
assume ‹E ∈# NE + UE›
then obtain K where ‹K ∈# E› and K: ‹K ∈ lits_of_l ?M› and lev_K: ‹get_level ?M K = 0›
using unit False by auto
moreover {
have ‹get_level ?M L > 0›
using False by auto
then have ‹atm_of L ≠ atm_of K›
using lev_K by fastforce }
ultimately show ‹∃La. La ∈# E ∧ (?D = None ∨ 0 < count_decided M ⟶
get_level M La = 0 ∧ La ∈ lits_of_l M)›
using False by auto
qed
next
case (backtrack_unit_clause L D K M1 M2 M D' i N U NE UE) note decomp = this(2) and
lev_L = this(3) and i = this(5) and lev_K = this(6) and D'[simp] = this(7) and twl = this(10)
let ?S = ‹(M, N, U, Some D, NE, UE, {#}, {#})›
let ?T = ‹(Propagated L {#L#} # M1, N, U, None, NE, add_mset {#L#} UE, {#}, {#- L#})›
let ?M = ‹Propagated L {#L#} # M1›
have unit: ‹entailed_clss_inv ?S›
using twl unfolding twl_struct_invs_def by fast
obtain M3 where M: ‹M = M3 @ M2 @ Decided K # M1›
using decomp by auto
define M2' where ‹M2' = (M3 @ M2) @ Decided K # []›
have M2': ‹M = M2' @ M1›
unfolding M M2'_def by simp
have count_dec_M2': ‹count_decided M2' ≠ 0›
unfolding M2'_def by auto
have lev_M: ‹count_decided M > 0›
unfolding M by auto
have n_d: ‹no_dup M›
using twl unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def twl_struct_invs_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
have count_dec_M1: ‹count_decided M1 = 0›
using no_dup_append_decided_Cons_lev[of ‹M3 @ M2› K M1]
lev_K n_d i unfolding M by simp
show ?case
unfolding entailed_clss_inv.simps Ball_def
proof (intro allI impI)
fix C
assume C: ‹C ∈# NE + add_mset {#L#} UE›
show ‹∃La. La ∈# C∧ (None = None ∨ 0 < count_decided ?M ⟶ get_level ?M La = 0 ∧
La ∈ lits_of_l ?M)›
proof (cases ‹C ∈# NE + UE›)
case True
then obtain K'' where C_K: ‹K'' ∈# C› and K: ‹K'' ∈ lits_of_l M› and
lev_K'': ‹get_level M K'' = 0›
using unit lev_M by auto
have ‹K'' ∈ lits_of_l M1›
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹K'' ∈ lits_of_l M2'›
using K unfolding M2' by auto
then have ex_L: ‹∃L∈set ((M3 @ M2) @ [Decided K]). ¬ atm_of (lit_of L) ≠ atm_of K''›
by (metis M2'_def image_iff lits_of_def)
have ‹get_level (M2' @ M1) K'' = get_level M2' K'' + count_decided M1›
using ‹K'' ∈ lits_of_l M2'› Decided_Propagated_in_iff_in_lits_of_l get_level_skip_end
by blast
with last_in_set_dropWhile[OF ex_L, unfolded M2'_def[symmetric]]
have ‹¬get_level M K'' = 0›
unfolding M2' using ‹K'' ∈ lits_of_l M2'› by (force simp: filter_empty_conv get_level_def)
then show False
using lev_K'' by arith
qed
then have K: ‹K'' ∈ lits_of_l ?M›
unfolding M by auto
moreover {
have ‹atm_of L ≠ atm_of K''›
using lev_L lev_K'' lev_M by (auto simp: atm_of_eq_atm_of)
then have ‹get_level ?M K'' = 0›
using count_dec_M1 count_decided_ge_get_level[of ?M K''] by auto }
ultimately show ?thesis
using C_K by auto
next
case False
then have ‹C = {#L#}›
using C by auto
then show ?thesis
using count_dec_M1 by auto
qed
qed
next
case (backtrack_nonunit_clause L D K M1 M2 M D' i N U NE UE L') note decomp = this(2) and
lev_L_M = this(3) and lev_K = this(6) and twl = this(13)
let ?S = ‹(M, N, U, Some D, NE, UE, {#}, {#})›
let ?T = ‹(Propagated L D' # M1, N, add_mset (TWL_Clause {#L, L'#} (D' - {#L, L'#})) U, None,
NE, UE, {#}, {#-L#})›
let ?M = ‹Propagated L D' # M1›
have unit: ‹entailed_clss_inv ?S›
using twl unfolding twl_struct_invs_def by fast
obtain M3 where M: ‹M = M3 @ M2 @ Decided K # M1›
using decomp by auto
define M2' where ‹M2' = (M3 @ M2) @ Decided K # []›
have M2': ‹M = M2' @ M1›
unfolding M M2'_def by simp
have count_dec_M2': ‹count_decided M2' ≠ 0›
unfolding M2'_def by auto
have lev_M: ‹count_decided M > 0›
unfolding M by auto
have n_d: ‹no_dup M›
using twl unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def twl_struct_invs_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps)
have count_dec_M1: ‹count_decided M1 = i›
using no_dup_append_decided_Cons_lev[of ‹M3 @ M2› K M1]
lev_K n_d unfolding M by simp
show ?case
unfolding entailed_clss_inv.simps Ball_def
proof (intro allI impI)
fix C
assume C: ‹C ∈# NE + UE›
then obtain K'' where C_K: ‹K'' ∈# C› and K: ‹K'' ∈ lits_of_l M› and
lev_K'': ‹get_level M K'' = 0›
using unit lev_M by auto
have K''_M1: ‹K'' ∈ lits_of_l M1›
proof (rule ccontr)
assume ‹¬ ?thesis›
then have ‹K'' ∈ lits_of_l M2'›
using K unfolding M2' by auto
then have ‹∃L∈set ((M3 @ M2) @ [Decided K]). ¬ atm_of (lit_of L) ≠ atm_of K''›
by (metis M2'_def image_iff lits_of_def)
then have ex_L: ‹∃L∈set ((M3 @ M2) @ [Decided K]). ¬ atm_of (lit_of L) ≠ atm_of K''›
by (metis M2'_def image_iff lits_of_def)
have ‹get_level (M2' @ M1) K'' = get_level M2' K'' + count_decided M1›
using ‹K'' ∈ lits_of_l M2'› Decided_Propagated_in_iff_in_lits_of_l get_level_skip_end
by blast
with last_in_set_dropWhile[OF ex_L, unfolded M2'_def[symmetric]] have ‹¬get_level M K'' = 0›
unfolding M2' using ‹K'' ∈ lits_of_l M2'› by (force simp: filter_empty_conv get_level_def)
then show False
using lev_K'' by arith
qed
then have K: ‹K'' ∈ lits_of_l ?M›
unfolding M by auto
moreover {
have ‹undefined_lit (M3 @ M2 @ [Decided K]) K''›
by (rule no_dup_append_in_atm_notin[of _ ‹M1›])
(use n_d M K''_M1 in auto)
then have ‹get_level M1 K'' = 0›
using lev_K'' unfolding M by (auto simp: image_Un)
moreover have ‹atm_of L ≠ atm_of K''›
using lev_K'' lev_M lev_L_M by (metis atm_of_eq_atm_of get_level_uminus not_gr_zero)
ultimately have ‹get_level ?M K'' = 0›
by auto }
ultimately show ‹∃La. La ∈# C ∧ (None = None ∨ 0 < count_decided ?M ⟶
get_level ?M La = 0 ∧ La ∈ lits_of_l ?M)›
using C_K by auto
qed
qed
subsubsection ‹The Strategy›
lemma no_literals_to_update_no_cp:
assumes
WS: ‹clauses_to_update S = {#}› and Q: ‹literals_to_update S = {#}› and
twl: ‹twl_struct_invs S›
shows
‹no_step cdcl⇩W_restart_mset.propagate (state⇩W_of S)› and
‹no_step cdcl⇩W_restart_mset.conflict (state⇩W_of S)›
proof -
obtain M N U NE UE D where
S: ‹S = (M, N, U, D, NE, UE, {#}, {#})›
using WS Q by (cases S) auto
{
assume confl: ‹get_conflict S = None›
then have S: ‹S = (M, N, U, None, NE, UE, {#}, {#})›
using WS Q S by auto
have twl_st_inv: ‹twl_st_inv S› and
struct_inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)› and
excep: ‹twl_st_exception_inv S› and
confl_cands: ‹confl_cands_enqueued S› and
propa_cands: ‹propa_cands_enqueued S› and
unit: ‹entailed_clss_inv S›
using twl unfolding twl_struct_invs_def by fast+
have n_d: ‹no_dup M›
using struct_inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by (auto simp: trail.simps S)
then have L_uL: ‹L ∈ lits_of_l M ⟹ -L ∉ lits_of_l M› for L
using consistent_interp_def distinct_consistent_interp by blast
have ‹∀C ∈# N + U. ¬M⊨as CNot (clause C)›
using confl_cands unfolding S by auto
moreover have ‹¬M⊨as CNot C› if C: ‹C ∈# NE + UE› for C
proof -
obtain L where L: ‹L ∈# C› and ‹L ∈ lits_of_l M›
using unit C unfolding S by auto
then have ‹M ⊨a C›
by (auto simp: true_annot_def dest!: multi_member_split)
then show ?thesis
using L ‹L ∈ lits_of_l M› by (auto simp: true_annots_true_cls_def_iff_negation_in_model
dest: L_uL multi_member_split)
qed
ultimately have ns_confl: ‹no_step cdcl⇩W_restart_mset.conflict (state⇩W_of S)›
by (auto elim!: cdcl⇩W_restart_mset.conflictE simp: S trail.simps clauses_def)
have ns_propa: ‹no_step cdcl⇩W_restart_mset.propagate (state⇩W_of S)›
proof (rule ccontr)
assume ‹¬ ?thesis›
then obtain C L where
C: ‹C ∈# clause `# (N + U) + NE + UE› and
L: ‹L ∈# C› and
M: ‹M ⊨as CNot (remove1_mset L C)› and
undef: ‹undefined_lit M L›
by (auto elim!: cdcl⇩W_restart_mset.propagateE simp: S trail.simps clauses_def) blast+
show False
proof (cases ‹C ∈# clause `# (N + U)›)
case True
then show ?thesis
using propa_cands L M undef by (auto simp: S)
next
case False
then have ‹C ∈# NE + UE›
using C by auto
then obtain L'' where L'': ‹L'' ∈# C› and L''_def: ‹L'' ∈ lits_of_l M›
using unit unfolding S by auto
then show ?thesis
using undef L'' L''_def L M L_uL
by (auto simp: S true_annots_true_cls_def_iff_negation_in_model
add_mset_eq_add_mset
Decided_Propagated_in_iff_in_lits_of_l dest!: multi_member_split)
qed
qed
note ns_confl ns_propa
}
moreover {
assume ‹get_conflict S ≠ None›
then have ‹no_step cdcl⇩W_restart_mset.propagate (state⇩W_of S)›
‹no_step cdcl⇩W_restart_mset.conflict (state⇩W_of S)›
by (auto elim!: cdcl⇩W_restart_mset.propagateE cdcl⇩W_restart_mset.conflictE
simp: S conflicting.simps)
}
ultimately show ‹no_step cdcl⇩W_restart_mset.propagate (state⇩W_of S)›
‹no_step cdcl⇩W_restart_mset.conflict (state⇩W_of S)›
by blast+
qed
text ‹
When popping a literal from \<^term>‹literals_to_update› to the \<^term>‹clauses_to_update›,
we do not do any transition in the abstract transition system. Therefore, we use
\<^term>‹rtranclp› or a case distinction.
›
lemma cdcl_twl_stgy_cdcl⇩W_stgy2:
assumes ‹cdcl_twl_stgy S T› and twl: ‹twl_struct_invs S›
shows ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T) ∨
(state⇩W_of S = state⇩W_of T ∧ (literals_to_update_measure T, literals_to_update_measure S)
∈ lexn less_than 2)›
using assms(1)
proof (induction rule: cdcl_twl_stgy.induct)
case (cp S')
then show ?case
using twl by (auto dest!: cdcl_twl_cp_cdcl⇩W_stgy)
next
case (other' S') note o = this(1)
have wq: ‹clauses_to_update S = {#}› and p: ‹literals_to_update S = {#}›
using o by (cases rule: cdcl_twl_o.cases; auto)+
show ?case
apply (rule disjI1)
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_stgy.other')
using no_literals_to_update_no_cp[OF wq p twl] apply (simp; fail)
using no_literals_to_update_no_cp[OF wq p twl] apply (simp; fail)
using cdcl_twl_o_cdcl⇩W_o[of S S', OF o] twl apply (simp add: twl_struct_invs_def; fail)
done
qed
lemma cdcl_twl_stgy_cdcl⇩W_stgy:
assumes ‹cdcl_twl_stgy S T› and twl: ‹twl_struct_invs S›
shows ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy⇧*⇧* (state⇩W_of S) (state⇩W_of T)›
using cdcl_twl_stgy_cdcl⇩W_stgy2[OF assms] by auto
lemma cdcl_twl_o_twl_struct_invs:
assumes
cdcl: ‹cdcl_twl_o S T› and
twl: ‹twl_struct_invs S›
shows ‹twl_struct_invs T›
proof -
have cdcl⇩W: ‹cdcl⇩W_restart_mset.cdcl⇩W_restart (state⇩W_of S) (state⇩W_of T)›
using twl unfolding twl_struct_invs_def
by (meson cdcl cdcl⇩W_restart_mset.other cdcl_twl_o_cdcl⇩W_o)
have wq: ‹clauses_to_update S = {#}› and p: ‹literals_to_update S = {#}›
using cdcl by (cases rule: cdcl_twl_o.cases; auto)+
have cdcl⇩W_stgy: ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T)›
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_stgy.other')
using no_literals_to_update_no_cp[OF wq p twl] apply (simp; fail)
using no_literals_to_update_no_cp[OF wq p twl] apply (simp; fail)
using cdcl_twl_o_cdcl⇩W_o[of S T, OF cdcl] twl apply (simp add: twl_struct_invs_def; fail)
done
have init: ‹init_clss (state⇩W_of T) = init_clss (state⇩W_of S)›
using cdcl⇩W by (auto simp: cdcl⇩W_restart_mset.cdcl⇩W_restart_init_clss)
show ?thesis
unfolding twl_struct_invs_def
apply (intro conjI)
subgoal by (use cdcl cdcl_twl_o_twl_st_inv twl in ‹blast; fail›)
subgoal by (use cdcl cdcl_twl_o_valid in ‹blast; fail›)
subgoal by (use cdcl⇩W cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_inv twl twl_struct_invs_def in
‹blast; fail›)
subgoal by (rule cdcl⇩W_restart_mset.cdcl⇩W_stgy_no_smaller_propa[OF cdcl⇩W_stgy])
((use twl in ‹simp add: init twl_struct_invs_def; fail›)+)[2]
subgoal by (use cdcl cdcl_twl_o_twl_st_exception_inv twl in ‹blast; fail›)
subgoal by (use cdcl cdcl_twl_o_no_duplicate_queued in ‹blast; fail›)
subgoal by (use cdcl cdcl_twl_o_distinct_queued in ‹blast; fail›)
subgoal by (use cdcl cdcl_twl_o_confl_cands_enqueued twl twl_struct_invs_def in ‹blast; fail›)
subgoal by (use cdcl cdcl_twl_o_propa_cands_enqueued twl twl_struct_invs_def in ‹blast; fail›)
subgoal by (use cdcl twl cdcl_twl_o_conflict_None_queue in ‹blast; fail›)
subgoal by (use cdcl cdcl_twl_o_entailed_clss_inv twl twl_struct_invs_def in blast)
subgoal by (use cdcl twl_o_clauses_to_update twl in blast)
subgoal by (use cdcl cdcl_twl_o_past_invs twl twl_struct_invs_def in blast)
done
qed
lemma cdcl_twl_stgy_twl_struct_invs:
assumes
cdcl: ‹cdcl_twl_stgy S T› and
twl: ‹twl_struct_invs S›
shows ‹twl_struct_invs T›
using cdcl by (induction rule: cdcl_twl_stgy.induct)
(simp_all add: cdcl_twl_cp_twl_struct_invs cdcl_twl_o_twl_struct_invs twl)
lemma rtranclp_cdcl_twl_stgy_twl_struct_invs:
assumes
cdcl: ‹cdcl_twl_stgy⇧*⇧* S T› and
twl: ‹twl_struct_invs S›
shows ‹twl_struct_invs T›
using cdcl by (induction rule: rtranclp_induct) (simp_all add: cdcl_twl_stgy_twl_struct_invs twl)
lemma rtranclp_cdcl_twl_stgy_cdcl⇩W_stgy:
assumes ‹cdcl_twl_stgy⇧*⇧* S T› and twl: ‹twl_struct_invs S›
shows ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy⇧*⇧* (state⇩W_of S) (state⇩W_of T)›
using assms by (induction rule: rtranclp_induct)
(auto dest!: cdcl_twl_stgy_cdcl⇩W_stgy intro: rtranclp_cdcl_twl_stgy_twl_struct_invs)
lemma no_step_cdcl_twl_cp_no_step_cdcl⇩W_cp:
assumes ns_cp: ‹no_step cdcl_twl_cp S› and twl: ‹twl_struct_invs S›
shows ‹literals_to_update S = {#} ∧ clauses_to_update S = {#}›
proof (cases ‹get_conflict S›)
case (Some a)
then show ?thesis
using twl unfolding twl_struct_invs_def by simp
next
case None note confl = this(1)
then obtain M N U UE NE WS Q where S: ‹S = (M, N, U, None, NE, UE, WS, Q)›
by (cases S) auto
have valid: ‹valid_enqueued S› and twl: ‹twl_st_inv S›
using twl unfolding twl_struct_invs_def by fast+
have wq: ‹clauses_to_update S = {#}›
proof (rule ccontr)
assume ‹clauses_to_update S ≠ {#}›
then obtain L C WS' where LC: ‹(L, C) ∈# clauses_to_update S› and
WS': ‹WS = add_mset (L, C) WS'›
by (cases WS) (auto simp: S)
have C_N_U: ‹C ∈# N + U› and L_C: ‹L ∈# watched C› and uL_M: ‹- L ∈ lits_of_l M›
using valid LC unfolding S by auto
have ‹struct_wf_twl_cls C›
using C_N_U twl unfolding S by (auto simp: twl_st_inv.simps)
then obtain L' where watched: ‹watched C = {#L, L'#}›
using L_C by (cases C) (auto simp: size_2_iff)
then have ‹L ∈# clause C›
by (cases C) auto
then have L'_M: ‹L' ∉ lits_of_l M›
using cdcl_twl_cp.delete_from_working[of L' C M N U NE UE L WS' Q] watched
ns_cp unfolding S WS' by (cases C) auto
then have ‹undefined_lit M L' ∨ -L' ∈ lits_of_l M›
using Decided_Propagated_in_iff_in_lits_of_l by blast
then have ‹¬ (∀L ∈# unwatched C. -L ∈ lits_of_l M)›
using cdcl_twl_cp.conflict[of C L L' M N U NE UE WS' Q]
cdcl_twl_cp.propagate[of C L L' M N U NE UE WS' Q] watched
ns_cp unfolding S WS' by fast
then obtain K where K: ‹K ∈# unwatched C› and uK_M: ‹-K ∉ lits_of_l M›
by auto
then have undef_K_K_M: ‹undefined_lit M K ∨ K ∈ lits_of_l M›
using Decided_Propagated_in_iff_in_lits_of_l by blast
define NU where ‹NU = (if C ∈# N then (add_mset (update_clause C L K) (remove1_mset C N), U)
else (N, add_mset (update_clause C L K) (remove1_mset C U)))›
have upd: ‹update_clauses (N, U) C L K NU›
using C_N_U unfolding NU_def by (auto simp: update_clauses.intros)
have NU: ‹NU = (fst NU, snd NU)›
by simp
show False
using cdcl_twl_cp.update_clause[of C L L' M K N U ‹fst NU› ‹snd NU› NE UE WS' Q]
watched uL_M L'_M K undef_K_K_M upd ns_cp unfolding S WS' by simp
qed
then have p: ‹literals_to_update S = {#}›
using cdcl_twl_cp.pop[of M N U NE UE] S ns_cp by (cases ‹Q›) fastforce+
show ?thesis using wq p by blast
qed
lemma no_step_cdcl_twl_o_no_step_cdcl⇩W_o:
assumes
ns_o: ‹no_step cdcl_twl_o S› and
twl: ‹twl_struct_invs S› and
p: ‹literals_to_update S = {#}› and
w_q: ‹clauses_to_update S = {#}›
shows ‹no_step cdcl⇩W_restart_mset.cdcl⇩W_o (state⇩W_of S)›
proof (rule ccontr)
assume ‹¬ ?thesis›
then obtain T where T: ‹cdcl⇩W_restart_mset.cdcl⇩W_o (state⇩W_of S) T›
by blast
obtain M N U D NE UE where S: ‹S = (M, N, U, D, NE, UE, {#}, {#})›
using p w_q by (cases S) auto
have unit: ‹entailed_clss_inv S›
using twl unfolding twl_struct_invs_def by fast+
show False
using T
proof (cases rule: cdcl⇩W_restart_mset.cdcl⇩W_o_induct)
case (decide L T) note confl = this(1) and undef = this(2) and atm = this(3) and T = this(4)
show ?thesis
using cdcl_twl_o.decide[of M L N NE U UE] confl undef atm ns_o unfolding S
by (auto simp: cdcl⇩W_restart_mset_state)
next
case (skip L C' M' E T) note M = this and confl = this(2) and uL_E = this(3) and E = this(4) and
T = this(5)
show ?thesis
using cdcl_twl_o.skip[of L E C' M' N U NE UE] M uL_E E ns_o unfolding S
by (auto simp: cdcl⇩W_restart_mset_state)
next
case (resolve L E M' D T) note M = this(1) and L_E = this(2) and hd = this(3) and
confl = this(4) and uL_D = this(5) and max_lvl = this(6)
show ?thesis
using cdcl_twl_o.resolve[of L D E M' N U NE UE] M L_E ns_o max_lvl uL_D confl unfolding S
by (auto simp: cdcl⇩W_restart_mset_state)
next
case (backtrack L C K i M1 M2 T D') note confl = this(1) and decomp = this(2) and
lev_L_bt = this(3) and lev_L = this(4) and i = this(5) and lev_K = this(6) and D'_C = this(7)
show ?thesis
proof (cases ‹D' = {#}›)
case True
show ?thesis
using cdcl_twl_o.backtrack_unit_clause[of L ‹add_mset L C› K M1 M2 M
‹add_mset L D'› i N U NE UE]
decomp True lev_L_bt lev_L i lev_K ns_o confl backtrack unfolding S
by (auto simp: cdcl⇩W_restart_mset_state clauses_def inf_sup_aci(6) sup.left_commute)
next
case False
then obtain L' where
L'_C: ‹L' ∈# D'› and lev_L': ‹get_level M L' = i›
using i get_maximum_level_exists_lit_of_max_level[of D' M] confl S
by (auto simp: cdcl⇩W_restart_mset_state S dest: in_diffD)
show ?thesis
using cdcl_twl_o.backtrack_nonunit_clause[of L ‹add_mset L C› K M1 M2 M ‹add_mset L D'›
i N U NE UE L']
using decomp lev_L_bt lev_L i lev_K False L'_C lev_L' ns_o confl backtrack
by (auto simp: cdcl⇩W_restart_mset_state S inf_sup_aci(6) sup.left_commute clauses_def
dest: in_diffD)
qed
qed
qed
lemma no_step_cdcl_twl_stgy_no_step_cdcl⇩W_stgy:
assumes ns: ‹no_step cdcl_twl_stgy S› and twl: ‹twl_struct_invs S›
shows ‹no_step cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S)›
proof -
have ns_cp: ‹no_step cdcl_twl_cp S› and ns_o: ‹no_step cdcl_twl_o S›
using ns by (auto simp: cdcl_twl_stgy.simps)
then have w_q: ‹clauses_to_update S = {#}› and p: ‹literals_to_update S = {#}›
using ns_cp no_step_cdcl_twl_cp_no_step_cdcl⇩W_cp twl by blast+
then have
‹no_step cdcl⇩W_restart_mset.propagate (state⇩W_of S)› and
‹no_step cdcl⇩W_restart_mset.conflict (state⇩W_of S)›
using no_literals_to_update_no_cp twl by blast+
moreover have ‹no_step cdcl⇩W_restart_mset.cdcl⇩W_o (state⇩W_of S)›
using w_q p ns_o no_step_cdcl_twl_o_no_step_cdcl⇩W_o twl by blast
ultimately show ?thesis
by (auto simp: cdcl⇩W_restart_mset.cdcl⇩W_stgy.simps)
qed
lemma full_cdcl_twl_stgy_cdcl⇩W_stgy:
assumes ‹full cdcl_twl_stgy S T› and twl: ‹twl_struct_invs S›
shows ‹full cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T)›
by (metis (no_types, hide_lams) assms(1) full_def no_step_cdcl_twl_stgy_no_step_cdcl⇩W_stgy
rtranclp_cdcl_twl_stgy_cdcl⇩W_stgy rtranclp_cdcl_twl_stgy_twl_struct_invs twl)
definition init_state_twl where
‹init_state_twl N ≡ ([], N, {#}, None, {#}, {#}, {#}, {#})›
lemma
assumes
struct: ‹∀C ∈# N. struct_wf_twl_cls C› and
tauto: ‹∀C ∈# N. ¬tautology (clause C)›
shows
twl_stgy_invs_init_state_twl: ‹twl_stgy_invs (init_state_twl N)› and
twl_struct_invs_init_state_twl: ‹twl_struct_invs (init_state_twl N)›
proof -
have [simp]: ‹twl_lazy_update [] C› ‹watched_literals_false_of_max_level [] C›
‹twl_exception_inv ([], N, {#}, None, {#}, {#}, {#}, {#}) C› for C
by (cases C; solves ‹auto simp: twl_exception_inv.simps›)+
have size_C: ‹size (clause C) ≥ 2› if ‹C ∈# N› for C
proof -
have ‹struct_wf_twl_cls C›
using that struct by auto
then show ?thesis by (cases C) auto
qed
have
[simp]: ‹clause C ≠ {#}› (is ?G1) and
[simp]: ‹remove1_mset L (clause C) ≠ {#}› (is ?G2) if ‹C ∈# N› for C L
by (rule size_ne_size_imp_ne[of _ ‹{#}›]; use size_C[OF that] in
‹auto simp: remove1_mset_empty_iff union_is_single›)+
have ‹distinct_mset (clause C)› if ‹C ∈# N› for C
using struct that by (cases C) (auto)
then have dist: ‹distinct_mset_mset (clause `# N)›
by (auto simp: distinct_mset_set_def)
then have [simp]: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv ([], clause `# N, {#}, None)›
using struct unfolding init_state.simps[symmetric]
by (auto simp: cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def)
have [simp]: ‹cdcl⇩W_restart_mset.no_smaller_propa ([], clause `# N, {#}, None)›
by(auto simp: cdcl⇩W_restart_mset.no_smaller_propa_def cdcl⇩W_restart_mset_state)
show stgy_invs: ‹twl_stgy_invs (init_state_twl N)›
by (auto simp: twl_stgy_invs_def cdcl⇩W_restart_mset.cdcl⇩W_stgy_invariant_def
cdcl⇩W_restart_mset.conflict_non_zero_unless_level_0_def
cdcl⇩W_restart_mset_state cdcl⇩W_restart_mset.no_smaller_confl_def init_state_twl_def)
show ‹twl_struct_invs (init_state_twl N)›
using struct tauto
by (auto simp: twl_struct_invs_def twl_st_inv.simps clauses_to_update_prop.simps
past_invs.simps cdcl⇩W_restart_mset_state init_state_twl_def
cdcl⇩W_restart_mset.no_strange_atm_def)
qed
lemma full_cdcl_twl_stgy_cdcl⇩W_stgy_conclusive_from_init_state:
fixes N :: ‹'v twl_clss›
assumes
full_cdcl_twl_stgy: ‹full cdcl_twl_stgy (init_state_twl N) T› and
struct: ‹∀C ∈# N. struct_wf_twl_cls C› and
no_tauto: ‹∀C ∈# N. ¬tautology (clause C)›
shows ‹conflicting (state⇩W_of T) = Some {#} ∧ unsatisfiable (set_mset (clause `# N)) ∨
(conflicting (state⇩W_of T) = None ∧ trail (state⇩W_of T) ⊨asm clause `# N ∧
satisfiable (set_mset (clause `# N)))›
proof -
have ‹distinct_mset (clause C)› if ‹C ∈# N› for C
using struct that by (cases C) auto
then have dist: ‹distinct_mset_mset (clause `# N)›
using struct by (auto simp: distinct_mset_set_def)
have ‹twl_struct_invs (init_state_twl N)›
using struct no_tauto by (rule twl_struct_invs_init_state_twl)
with full_cdcl_twl_stgy
have ‹full cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of (init_state_twl N)) (state⇩W_of T)›
by (rule full_cdcl_twl_stgy_cdcl⇩W_stgy)
then have ‹full cdcl⇩W_restart_mset.cdcl⇩W_stgy (init_state (clause `# N)) (state⇩W_of T)›
by (simp add: init_state.simps init_state_twl_def)
then show ?thesis
by (rule cdcl⇩W_restart_mset.full_cdcl⇩W_stgy_final_state_conclusive_from_init_state)
(use dist in auto)
qed
lemma cdcl_twl_o_twl_stgy_invs:
‹cdcl_twl_o S T ⟹ twl_struct_invs S ⟹ twl_stgy_invs S ⟹ twl_stgy_invs T›
using cdcl⇩W_restart_mset.rtranclp_cdcl⇩W_stgy_cdcl⇩W_stgy_invariant cdcl_twl_stgy_cdcl⇩W_stgy
other' cdcl⇩W_restart_mset.cdcl⇩W_restart_conflict_non_zero_unless_level_0
unfolding twl_struct_invs_def twl_stgy_invs_def
apply (intro conjI)
apply blast
by (smt cdcl⇩W_restart_mset.cdcl⇩W_restart_conflict_non_zero_unless_level_0 cdcl⇩W_restart_mset.other
cdcl_twl_o_cdcl⇩W_o twl_struct_invs_def twl_struct_invs_no_false_clause)
paragraph ‹Well-foundedness›
lemma wf_cdcl⇩W_stgy_state⇩W_of:
‹wf {(T, S). cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S) ∧
cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T)}›
using wf_if_measure_f[OF cdcl⇩W_restart_mset.wf_cdcl⇩W_stgy, of state⇩W_of] by simp
lemma wf_cdcl_twl_cp:
‹wf {(T, S). twl_struct_invs S ∧ cdcl_twl_cp S T}› (is ‹wf ?TWL›)
proof -
let ?CDCL = ‹{(T, S). cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S) ∧
cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T)}›
let ?P = ‹{(T, S). state⇩W_of S = state⇩W_of T ∧
(literals_to_update_measure T, literals_to_update_measure S) ∈ lexn less_than 2}›
have wf_p_m:
‹wf {(T, S). (literals_to_update_measure T, literals_to_update_measure S) ∈ lexn less_than 2}›
using wf_if_measure_f[of ‹lexn less_than 2› literals_to_update_measure] by (auto simp: wf_lexn)
have ‹wf ?CDCL›
by (rule wf_subset[OF wf_cdcl⇩W_stgy_state⇩W_of])
(auto simp: twl_struct_invs_def)
moreover have ‹wf ?P›
by (rule wf_subset[OF wf_p_m]) auto
moreover have ‹?CDCL O ?P ⊆ ?CDCL› by auto
ultimately have ‹wf (?CDCL ∪ ?P)›
by (rule wf_union_compatible)
moreover have ‹?TWL ⊆ ?CDCL ∪ ?P›
proof
fix x
assume x_TWL: ‹x ∈ ?TWL›
then obtain S T where x: ‹x = (T, S)› by auto
have twl: ‹twl_struct_invs S› and cdcl: ‹cdcl_twl_cp S T›
using x_TWL x by auto
have ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)›
using twl by (auto simp: twl_struct_invs_def)
moreover have ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T) ∨
(state⇩W_of S = state⇩W_of T ∧
(literals_to_update_measure T, literals_to_update_measure S) ∈ lexn less_than 2)›
using cdcl cdcl_twl_cp_cdcl⇩W_stgy twl by blast
ultimately show ‹x ∈ ?CDCL ∪ ?P›
unfolding x by blast
qed
ultimately show ?thesis
using wf_subset[of ‹?CDCL ∪ ?P›] by blast
qed
lemma tranclp_wf_cdcl_twl_cp:
‹wf {(T, S). twl_struct_invs S ∧ cdcl_twl_cp⇧+⇧+ S T}›
proof -
have H: ‹{(T, S). twl_struct_invs S ∧ cdcl_twl_cp⇧+⇧+ S T} ⊆
{(T, S). twl_struct_invs S ∧ cdcl_twl_cp S T}⇧+›
proof -
{ fix T S :: ‹'v twl_st›
assume ‹cdcl_twl_cp⇧+⇧+ S T› ‹twl_struct_invs S›
then have ‹(T, S) ∈ {(T, S). twl_struct_invs S ∧ cdcl_twl_cp S T}⇧+› (is ‹_ ∈ ?S⇧+›)
proof (induction rule: tranclp_induct)
case (base y)
then show ?case by auto
next
case (step T U) note st = this(1) and cp = this(2) and IH = this(3)[OF this(4)] and
twl = this(4)
have ‹twl_struct_invs T›
by (metis (no_types, lifting) IH Nitpick.tranclp_unfold cdcl_twl_cp_twl_struct_invs
converse_tranclpE)
then have ‹(U, T) ∈ ?S⇧+›
using cp by auto
then show ?case using IH by auto
qed
}
then show ?thesis by blast
qed
show ?thesis using wf_trancl[OF wf_cdcl_twl_cp] wf_subset[OF _ H] by blast
qed
lemma wf_cdcl_twl_stgy:
‹wf {(T, S). twl_struct_invs S ∧ cdcl_twl_stgy S T}› (is ‹wf ?TWL›)
proof -
let ?CDCL = ‹{(T, S). cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S) ∧
cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T)}›
let ?P = ‹{(T, S). state⇩W_of S = state⇩W_of T ∧
(literals_to_update_measure T, literals_to_update_measure S) ∈ lexn less_than 2}›
have wf_p_m:
‹wf {(T, S). (literals_to_update_measure T, literals_to_update_measure S) ∈ lexn less_than 2}›
using wf_if_measure_f[of ‹lexn less_than 2› literals_to_update_measure] by (auto simp: wf_lexn)
have ‹wf ?CDCL›
by (rule wf_subset[OF wf_cdcl⇩W_stgy_state⇩W_of])
(auto simp: twl_struct_invs_def)
moreover have ‹wf ?P›
by (rule wf_subset[OF wf_p_m]) auto
moreover have ‹?CDCL O ?P ⊆ ?CDCL› by auto
ultimately have ‹wf (?CDCL ∪ ?P)›
by (rule wf_union_compatible)
moreover have ‹?TWL ⊆ ?CDCL ∪ ?P›
proof
fix x
assume x_TWL: ‹x ∈ ?TWL›
then obtain S T where x: ‹x = (T, S)› by auto
have twl: ‹twl_struct_invs S› and cdcl: ‹cdcl_twl_stgy S T›
using x_TWL x by auto
have ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (state⇩W_of S)›
using twl by (auto simp: twl_struct_invs_def)
moreover have ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T) ∨
(state⇩W_of S = state⇩W_of T ∧
(literals_to_update_measure T, literals_to_update_measure S) ∈ lexn less_than 2)›
using cdcl cdcl_twl_stgy_cdcl⇩W_stgy2 twl by blast
ultimately show ‹x ∈ ?CDCL ∪ ?P›
unfolding x by blast
qed
ultimately show ?thesis
using wf_subset[of ‹?CDCL ∪ ?P›] by blast
qed
lemma tranclp_wf_cdcl_twl_stgy:
‹wf {(T, S). twl_struct_invs S ∧ cdcl_twl_stgy⇧+⇧+ S T}›
proof -
have H: ‹{(T, S). twl_struct_invs S ∧ cdcl_twl_stgy⇧+⇧+ S T} ⊆
{(T, S). twl_struct_invs S ∧ cdcl_twl_stgy S T}⇧+›
proof -
{ fix T S :: ‹'v twl_st›
assume ‹cdcl_twl_stgy⇧+⇧+ S T› ‹twl_struct_invs S›
then have ‹(T, S) ∈ {(T, S). twl_struct_invs S ∧ cdcl_twl_stgy S T}⇧+› (is ‹_ ∈ ?S⇧+›)
proof (induction rule: tranclp_induct)
case (base y)
then show ?case by auto
next
case (step T U) note st = this(1) and stgy = this(2) and IH = this(3)[OF this(4)] and
twl = this(4)
have ‹twl_struct_invs T›
by (metis (no_types, lifting) IH Nitpick.tranclp_unfold cdcl_twl_stgy_twl_struct_invs
converse_tranclpE)
then have ‹(U, T) ∈ ?S⇧+›
using stgy by auto
then show ?case using IH by auto
qed
}
then show ?thesis by blast
qed
show ?thesis using wf_trancl[OF wf_cdcl_twl_stgy] wf_subset[OF _ H] by blast
qed
lemma rtranclp_cdcl_twl_o_stgyD: ‹cdcl_twl_o⇧*⇧* S T ⟹ cdcl_twl_stgy⇧*⇧* S T›
using rtranclp_mono[of cdcl_twl_o cdcl_twl_stgy] cdcl_twl_stgy.intros(2)
by blast
lemma rtranclp_cdcl_twl_cp_stgyD: ‹cdcl_twl_cp⇧*⇧* S T ⟹ cdcl_twl_stgy⇧*⇧* S T›
using rtranclp_mono[of cdcl_twl_cp cdcl_twl_stgy] cdcl_twl_stgy.intros(1)
by blast
lemma tranclp_cdcl_twl_o_stgyD: ‹cdcl_twl_o⇧+⇧+ S T ⟹ cdcl_twl_stgy⇧+⇧+ S T›
using tranclp_mono[of cdcl_twl_o cdcl_twl_stgy] cdcl_twl_stgy.intros(2)
by blast
lemma tranclp_cdcl_twl_cp_stgyD: ‹cdcl_twl_cp⇧+⇧+ S T ⟹ cdcl_twl_stgy⇧+⇧+ S T›
using tranclp_mono[of cdcl_twl_cp cdcl_twl_stgy] cdcl_twl_stgy.intros(1)
by blast
lemma wf_cdcl_twl_o:
‹wf {(T, S::'v twl_st). twl_struct_invs S ∧ cdcl_twl_o S T}›
by (rule wf_subset[OF wf_cdcl_twl_stgy]) (auto intro: cdcl_twl_stgy.intros)
lemma tranclp_wf_cdcl_twl_o:
‹wf {(T, S::'v twl_st). twl_struct_invs S ∧ cdcl_twl_o⇧+⇧+ S T}›
by (rule wf_subset[OF tranclp_wf_cdcl_twl_stgy]) (auto dest: tranclp_cdcl_twl_o_stgyD)
lemma (in -)propa_cands_enqueued_mono:
‹U' ⊆# U ⟹ N' ⊆# N ⟹
propa_cands_enqueued (M, N, U, D, NE, UE, WS, Q) ⟹
propa_cands_enqueued (M, N', U', D, NE', UE', WS, Q)›
by (cases D) (auto 5 5)
lemma (in -)confl_cands_enqueued_mono:
‹U' ⊆# U ⟹ N' ⊆# N ⟹
confl_cands_enqueued (M, N, U, D, NE, UE, WS, Q) ⟹
confl_cands_enqueued (M, N', U', D, NE', UE', WS, Q)›
by (cases D) auto
lemma (in -)twl_st_exception_inv_mono:
‹U' ⊆# U ⟹ N' ⊆# N ⟹
twl_st_exception_inv (M, N, U, D, NE, UE, WS, Q) ⟹
twl_st_exception_inv (M, N', U', D, NE', UE', WS, Q)›
by (cases D) (fastforce simp: twl_exception_inv.simps)+
lemma (in -)twl_st_inv_mono:
‹U' ⊆# U ⟹ N' ⊆# N ⟹
twl_st_inv (M, N, U, D, NE, UE, WS, Q) ⟹
twl_st_inv (M, N', U', D, NE', UE', WS, Q)›
by (cases D) (fastforce simp: twl_st_inv.simps)+
lemma (in -) rtranclp_cdcl_twl_stgy_twl_stgy_invs:
assumes
‹cdcl_twl_stgy⇧*⇧* S T› and
‹twl_struct_invs S› and
‹twl_stgy_invs S›
shows ‹twl_stgy_invs T›
using assms cdcl⇩W_restart_mset.rtranclp_cdcl⇩W_stgy_cdcl⇩W_stgy_invariant
rtranclp_cdcl_twl_stgy_cdcl⇩W_stgy
by (metis cdcl⇩W_restart_mset.rtranclp_cdcl⇩W_restart_conflict_non_zero_unless_level_0
cdcl⇩W_restart_mset.rtranclp_cdcl⇩W_stgy_rtranclp_cdcl⇩W_restart twl_stgy_invs_def
twl_struct_invs_def twl_struct_invs_no_false_clause)
lemma after_fast_restart_replay:
assumes
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (M', N, U, None)› and
stgy_invs: ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy_invariant (M', N, U, None)› and
smaller_propa: ‹cdcl⇩W_restart_mset.no_smaller_propa (M', N, U, None)› and
kept: ‹∀L E. Propagated L E ∈ set (drop (length M' - n) M') ⟶ E ∈# N + U'› and
U'_U: ‹U' ⊆# U›
shows
‹cdcl⇩W_restart_mset.cdcl⇩W_stgy⇧*⇧* ([], N, U', None) (drop (length M' - n) M', N, U', None)›
proof -
let ?S = ‹λn. (drop (length M' - n) M', N, U', None)›
note cdcl⇩W_restart_mset_state[simp]
have
M_lev: ‹cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv (M', N, U, None)› and
alien: ‹cdcl⇩W_restart_mset.no_strange_atm (M', N, U, None)› and
confl: ‹cdcl⇩W_restart_mset.cdcl⇩W_conflicting (M', N, U, None)› and
learned: ‹cdcl⇩W_restart_mset.cdcl⇩W_learned_clause (M', N, U, None)›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def by fast+
have smaller_confl: ‹cdcl⇩W_restart_mset.no_smaller_confl (M', N, U, None)›
using stgy_invs unfolding cdcl⇩W_restart_mset.cdcl⇩W_stgy_invariant_def by blast
have n_d: ‹no_dup M'›
using M_lev unfolding cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by simp
let ?L = ‹λm. M' ! (length M' - Suc m)›
have undef_nth_Suc:
‹undefined_lit (drop (length M' - m) M') (lit_of (?L m))›
if ‹m < length M'›
for m
proof -
define k where
‹k = length M' - Suc m›
then have Sk: ‹length M' - m = Suc k›
using that by linarith
have k_le_M': ‹k < length M'›
using that unfolding k_def by linarith
have n_d': ‹no_dup (take k M' @ ?L m # drop (Suc k) M')›
using n_d
apply (subst (asm) append_take_drop_id[symmetric, of _ ‹Suc k›])
apply (subst (asm) take_Suc_conv_app_nth)
apply (rule k_le_M')
apply (subst k_def[symmetric])
by simp
show ?thesis
using n_d'
apply (subst (asm) no_dup_append_cons)
apply (subst (asm) k_def[symmetric])+
apply (subst k_def[symmetric])+
apply (subst Sk)+
by blast
qed
have atm_in:
‹atm_of (lit_of (M' ! m)) ∈ atms_of_mm N›
if ‹m < length M'›
for m
using alien that
by (auto simp: cdcl⇩W_restart_mset.no_strange_atm_def lits_of_def)
show ?thesis
using kept
proof (induction n)
case 0
then show ?case by simp
next
case (Suc m) note IH = this(1) and kept = this(2)
consider
(le) ‹m < length M'› |
(ge) ‹m ≥ length M'›
by linarith
then show ?case
proof (cases)
case ge
then show ?thesis
using Suc by auto
next
case le
define k where
‹k = length M' - Suc m›
then have Sk: ‹length M' - m = Suc k›
using le by linarith
have k_le_M': ‹k < length M'›
using le unfolding k_def by linarith
have kept': ‹∀L E. Propagated L E ∈ set (drop (length M' - m) M') ⟶ E ∈# N + U'›
using kept k_le_M' unfolding k_def[symmetric] Sk
by (subst (asm) Cons_nth_drop_Suc[symmetric]) auto
have M': ‹M' = take (length M' - Suc m) M' @ ?L m # trail (?S m)›
apply (subst append_take_drop_id[symmetric, of _ ‹Suc k›])
apply (subst take_Suc_conv_app_nth)
apply (rule k_le_M')
apply (subst k_def[symmetric])
unfolding k_def[symmetric] Sk
by auto
have ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy (?S m) (?S (Suc m))›
proof (cases ‹?L (m)›)
case (Decided K) note K = this
have dec: ‹cdcl⇩W_restart_mset.decide (?S m) (?S (Suc m))›
apply (rule cdcl⇩W_restart_mset.decide_rule[of _ ‹lit_of (?L m)›])
subgoal by simp
subgoal using undef_nth_Suc[of m] le by simp
subgoal using le by (auto simp: atm_in)
subgoal using le k_le_M' K unfolding k_def[symmetric] Sk
by (auto simp: state_eq_def state_def Cons_nth_drop_Suc[symmetric])
done
have Dec: ‹M' ! k = Decided K›
using K unfolding k_def[symmetric] Sk .
have H: ‹D + {#L#} ∈# N + U ⟶ undefined_lit (trail (?S m)) L ⟶
¬ (trail (?S m)) ⊨as CNot D› for D L
using smaller_propa unfolding cdcl⇩W_restart_mset.no_smaller_propa_def
trail.simps clauses_def
cdcl⇩W_restart_mset_state
apply (subst (asm) M')
unfolding Dec Sk k_def[symmetric]
by (auto simp: clauses_def state_eq_def)
have ‹D ∈# N ⟶ undefined_lit (trail (?S m)) L ⟶ L ∈# D ⟶
¬ (trail (?S m)) ⊨as CNot (remove1_mset L D)› and
‹D ∈# U' ⟶ undefined_lit (trail (?S m)) L ⟶ L ∈# D ⟶
¬ (trail (?S m)) ⊨as CNot (remove1_mset L D)›for D L
using H[of ‹remove1_mset L D› L] U'_U by auto
then have nss: ‹no_step cdcl⇩W_restart_mset.propagate (?S m)›
by (auto simp: cdcl⇩W_restart_mset.propagate.simps clauses_def
state_eq_def k_def[symmetric] Sk)
have H: ‹D ∈# N + U' ⟶ ¬ (trail (?S m)) ⊨as CNot D› for D
using smaller_confl U'_U unfolding cdcl⇩W_restart_mset.no_smaller_confl_def
trail.simps clauses_def cdcl⇩W_restart_mset_state
apply (subst (asm) M')
unfolding Dec Sk k_def[symmetric]
by (auto simp: clauses_def state_eq_def)
then have nsc: ‹no_step cdcl⇩W_restart_mset.conflict (?S m)›
by (auto simp: cdcl⇩W_restart_mset.conflict.simps clauses_def state_eq_def
k_def[symmetric] Sk)
show ?thesis
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_stgy.other')
apply (rule nsc)
apply (rule nss)
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_o.decide)
apply (rule dec)
done
next
case K: (Propagated K C)
have Propa: ‹M' ! k = Propagated K C›
using K unfolding k_def[symmetric] Sk .
have
M_C: ‹trail (?S m) ⊨as CNot (remove1_mset K C)› and
K_C: ‹K ∈# C›
using confl unfolding cdcl⇩W_restart_mset.cdcl⇩W_conflicting_def trail.simps
by (subst (asm)(3) M'; auto simp: k_def[symmetric] Sk Propa)+
have [simp]: ‹k - min (length M') k = 0›
unfolding k_def by auto
have C_N_U: ‹C ∈# N + U'›
using learned kept unfolding cdcl⇩W_restart_mset.cdcl⇩W_learned_clause_alt_def Sk
k_def[symmetric] cdcl⇩W_restart_mset.reasons_in_clauses_def
apply (subst (asm)(4)M')
apply (subst (asm)(10)M')
unfolding K
by (auto simp: K k_def[symmetric] Sk Propa clauses_def)
have ‹cdcl⇩W_restart_mset.propagate (?S m) (?S (Suc m))›
apply (rule cdcl⇩W_restart_mset.propagate_rule[of _ C K])
subgoal by simp
subgoal using C_N_U by (simp add: clauses_def)
subgoal using K_C .
subgoal using M_C .
subgoal using undef_nth_Suc[of m] le K by (simp add: k_def[symmetric] Sk)
subgoal
using le k_le_M' K unfolding k_def[symmetric] Sk
by (auto simp: state_eq_def
state_def Cons_nth_drop_Suc[symmetric])
done
then show ?thesis
by (rule cdcl⇩W_restart_mset.cdcl⇩W_stgy.propagate')
qed
then show ?thesis
using IH[OF kept'] by simp
qed
qed
qed
lemma after_fast_restart_replay_no_stgy:
assumes
inv: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv (M', N, U, None)› and
kept: ‹∀L E. Propagated L E ∈ set (drop (length M' - n) M') ⟶ E ∈# N + U'› and
U'_U: ‹U' ⊆# U›
shows
‹cdcl⇩W_restart_mset.cdcl⇩W⇧*⇧* ([], N, U', None) (drop (length M' - n) M', N, U', None)›
proof -
let ?S = ‹λn. (drop (length M' - n) M', N, U', None)›
note cdcl⇩W_restart_mset_state[simp]
have
M_lev: ‹cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv (M', N, U, None)› and
alien: ‹cdcl⇩W_restart_mset.no_strange_atm (M', N, U, None)› and
confl: ‹cdcl⇩W_restart_mset.cdcl⇩W_conflicting (M', N, U, None)› and
learned: ‹cdcl⇩W_restart_mset.cdcl⇩W_learned_clause (M', N, U, None)›
using inv unfolding cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def by fast+
have n_d: ‹no_dup M'›
using M_lev unfolding cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def by simp
let ?L = ‹λm. M' ! (length M' - Suc m)›
have undef_nth_Suc:
‹undefined_lit (drop (length M' - m) M') (lit_of (?L m))›
if ‹m < length M'›
for m
proof -
define k where
‹k = length M' - Suc m›
then have Sk: ‹length M' - m = Suc k›
using that by linarith
have k_le_M': ‹k < length M'›
using that unfolding k_def by linarith
have n_d': ‹no_dup (take k M' @ ?L m # drop (Suc k) M')›
using n_d
apply (subst (asm) append_take_drop_id[symmetric, of _ ‹Suc k›])
apply (subst (asm) take_Suc_conv_app_nth)
apply (rule k_le_M')
apply (subst k_def[symmetric])
by simp
show ?thesis
using n_d'
apply (subst (asm) no_dup_append_cons)
apply (subst (asm) k_def[symmetric])+
apply (subst k_def[symmetric])+
apply (subst Sk)+
by blast
qed
have atm_in:
‹atm_of (lit_of (M' ! m)) ∈ atms_of_mm N›
if ‹m < length M'›
for m
using alien that
by (auto simp: cdcl⇩W_restart_mset.no_strange_atm_def lits_of_def)
show ?thesis
using kept
proof (induction n)
case 0
then show ?case by simp
next
case (Suc m) note IH = this(1) and kept = this(2)
consider
(le) ‹m < length M'› |
(ge) ‹m ≥ length M'›
by linarith
then show ?case
proof cases
case ge
then show ?thesis
using Suc by auto
next
case le
define k where
‹k = length M' - Suc m›
then have Sk: ‹length M' - m = Suc k›
using le by linarith
have k_le_M': ‹k < length M'›
using le unfolding k_def by linarith
have kept': ‹∀L E. Propagated L E ∈ set (drop (length M' - m) M') ⟶ E ∈# N + U'›
using kept k_le_M' unfolding k_def[symmetric] Sk
by (subst (asm) Cons_nth_drop_Suc[symmetric]) auto
have M': ‹M' = take (length M' - Suc m) M' @ ?L m # trail (?S m)›
apply (subst append_take_drop_id[symmetric, of _ ‹Suc k›])
apply (subst take_Suc_conv_app_nth)
apply (rule k_le_M')
apply (subst k_def[symmetric])
unfolding k_def[symmetric] Sk
by auto
have ‹cdcl⇩W_restart_mset.cdcl⇩W (?S m) (?S (Suc m))›
proof (cases ‹?L (m)›)
case (Decided K) note K = this
have dec: ‹cdcl⇩W_restart_mset.decide (?S m) (?S (Suc m))›
apply (rule cdcl⇩W_restart_mset.decide_rule[of _ ‹lit_of (?L m)›])
subgoal by simp
subgoal using undef_nth_Suc[of m] le by simp
subgoal using le by (auto simp: atm_in)
subgoal using le k_le_M' K unfolding k_def[symmetric] Sk
by (auto simp: state_eq_def state_def Cons_nth_drop_Suc[symmetric])
done
have Dec: ‹M' ! k = Decided K›
using K unfolding k_def[symmetric] Sk .
show ?thesis
apply (rule cdcl⇩W_restart_mset.cdcl⇩W.intros(3))
apply (rule cdcl⇩W_restart_mset.cdcl⇩W_o.decide)
apply (rule dec)
done
next
case K: (Propagated K C)
have Propa: ‹M' ! k = Propagated K C›
using K unfolding k_def[symmetric] Sk .
have
M_C: ‹trail (?S m) ⊨as CNot (remove1_mset K C)› and
K_C: ‹K ∈# C›
using confl unfolding cdcl⇩W_restart_mset.cdcl⇩W_conflicting_def trail.simps
by (subst (asm)(3) M'; auto simp: k_def[symmetric] Sk Propa)+
have [simp]: ‹k - min (length M') k = 0›
unfolding k_def by auto
have C_N_U: ‹C ∈# N + U'›
using learned kept unfolding cdcl⇩W_restart_mset.cdcl⇩W_learned_clause_alt_def Sk
k_def[symmetric] cdcl⇩W_restart_mset.reasons_in_clauses_def
apply (subst (asm)(4)M')
apply (subst (asm)(10)M')
unfolding K
by (auto simp: K k_def[symmetric] Sk Propa clauses_def)
have ‹cdcl⇩W_restart_mset.propagate (?S m) (?S (Suc m))›
apply (rule cdcl⇩W_restart_mset.propagate_rule[of _ C K])
subgoal by simp
subgoal using C_N_U by (simp add: clauses_def)
subgoal using K_C .
subgoal using M_C .
subgoal using undef_nth_Suc[of m] le K by (simp add: k_def[symmetric] Sk)
subgoal
using le k_le_M' K unfolding k_def[symmetric] Sk
by (auto simp: state_eq_def
state_def Cons_nth_drop_Suc[symmetric])
done
then show ?thesis
by (rule cdcl⇩W_restart_mset.cdcl⇩W.intros)
qed
then show ?thesis
using IH[OF kept'] by simp
qed
qed
qed
lemma cdcl_twl_stgy_get_init_learned_clss_mono:
assumes ‹cdcl_twl_stgy S T›
shows ‹get_init_learned_clss S ⊆# get_init_learned_clss T›
using assms
by induction (auto simp: cdcl_twl_cp.simps cdcl_twl_o.simps)
lemma rtranclp_cdcl_twl_stgy_get_init_learned_clss_mono:
assumes ‹cdcl_twl_stgy⇧*⇧* S T›
shows ‹get_init_learned_clss S ⊆# get_init_learned_clss T›
using assms
by induction (auto dest!: cdcl_twl_stgy_get_init_learned_clss_mono)
lemma cdcl_twl_o_all_learned_diff_learned:
assumes ‹cdcl_twl_o S T›
shows
‹clause `# get_learned_clss S ⊆# clause `# get_learned_clss T ∧
get_init_learned_clss S ⊆# get_init_learned_clss T∧
get_all_init_clss S = get_all_init_clss T›
by (use assms in ‹induction rule: cdcl_twl_o.induct›)
(auto simp: update_clauses.simps size_Suc_Diff1)
lemma cdcl_twl_cp_all_learned_diff_learned:
assumes ‹cdcl_twl_cp S T›
shows
‹clause `# get_learned_clss S = clause `# get_learned_clss T ∧
get_init_learned_clss S = get_init_learned_clss T ∧
get_all_init_clss S = get_all_init_clss T›
apply (use assms in ‹induction rule: cdcl_twl_cp.induct›)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal for D
by (cases D)
(auto simp: update_clauses.simps size_Suc_Diff1 dest!: multi_member_split)
done
lemma cdcl_twl_stgy_all_learned_diff_learned:
assumes ‹cdcl_twl_stgy S T›
shows
‹clause `# get_learned_clss S ⊆# clause `# get_learned_clss T ∧
get_init_learned_clss S ⊆# get_init_learned_clss T∧
get_all_init_clss S = get_all_init_clss T›
by (use assms in ‹induction rule: cdcl_twl_stgy.induct›)
(auto simp: cdcl_twl_cp_all_learned_diff_learned cdcl_twl_o_all_learned_diff_learned)
lemma rtranclp_cdcl_twl_stgy_all_learned_diff_learned:
assumes ‹cdcl_twl_stgy⇧*⇧* S T›
shows
‹clause `# get_learned_clss S ⊆# clause `# get_learned_clss T ∧
get_init_learned_clss S ⊆# get_init_learned_clss T ∧
get_all_init_clss S = get_all_init_clss T›
by (use assms in ‹induction rule: rtranclp_induct›)
(auto dest: cdcl_twl_stgy_all_learned_diff_learned)
lemma rtranclp_cdcl_twl_stgy_all_learned_diff_learned_size:
assumes ‹cdcl_twl_stgy⇧*⇧* S T›
shows
‹size (get_all_learned_clss T) - size (get_all_learned_clss S) ≥
size (get_learned_clss T) - size (get_learned_clss S)›
using rtranclp_cdcl_twl_stgy_all_learned_diff_learned[OF assms]
apply (cases S, cases T)
using size_mset_mono by force+
lemma cdcl_twl_stgy_cdcl⇩W_stgy3:
assumes ‹cdcl_twl_stgy S T› and twl: ‹twl_struct_invs S› and
‹clauses_to_update S = {#}› and
‹literals_to_update S = {#}›
shows ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of T)›
using cdcl_twl_stgy_cdcl⇩W_stgy2[OF assms(1,2)] assms(3-)
by (auto simp: lexn2_conv)
lemma tranclp_cdcl_twl_stgy_cdcl⇩W_stgy:
assumes ST: ‹cdcl_twl_stgy⇧+⇧+ S T› and
twl: ‹twl_struct_invs S› and
‹clauses_to_update S = {#}› and
‹literals_to_update S = {#}›
shows ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy⇧+⇧+ (state⇩W_of S) (state⇩W_of T)›
proof -
obtain S' where
SS': ‹cdcl_twl_stgy S S'› and
S'T: ‹cdcl_twl_stgy⇧*⇧* S' T›
using ST unfolding tranclp_unfold_begin by blast
have 1: ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy (state⇩W_of S) (state⇩W_of S')›
using cdcl_twl_stgy_cdcl⇩W_stgy3[OF SS' assms(2-4)]
by blast
have struct_S': ‹twl_struct_invs S'›
using twl SS' by (blast intro: cdcl_twl_stgy_twl_struct_invs)
have 2: ‹cdcl⇩W_restart_mset.cdcl⇩W_stgy⇧*⇧* (state⇩W_of S') (state⇩W_of T)›
apply (rule rtranclp_cdcl_twl_stgy_cdcl⇩W_stgy)
apply (rule S'T)
by (rule struct_S')
show ?thesis
using 1 2 by auto
qed
definition final_twl_state where
‹final_twl_state S ⟷
no_step cdcl_twl_stgy S ∨ (get_conflict S ≠ None ∧ count_decided (get_trail S) = 0)›
definition conclusive_TWL_run :: ‹'v twl_st ⇒ 'v twl_st nres› where
‹conclusive_TWL_run S = SPEC(λT. cdcl_twl_stgy⇧*⇧* S T ∧ final_twl_state T)›
lemma conflict_of_level_unsatisfiable:
assumes
struct: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv S› and
dec: ‹count_decided (trail S) = 0› and
confl: ‹conflicting S ≠ None› and
‹cdcl⇩W_restart_mset.cdcl⇩W_learned_clauses_entailed_by_init S›
shows ‹unsatisfiable (set_mset (init_clss S))›
proof -
obtain M N U D where S: ‹S = (M, N, U, Some D)›
by (cases S) (use confl in ‹auto simp: cdcl⇩W_restart_mset_state›)
have [simp]: ‹get_all_ann_decomposition M = [([], M)]›
by (rule no_decision_get_all_ann_decomposition)
(use dec in ‹auto simp: count_decided_def filter_empty_conv S cdcl⇩W_restart_mset_state›)
have
N_U: ‹N ⊨psm U› and
M_D: ‹M ⊨as CNot D› and
N_U_M: ‹set_mset N ∪ set_mset U ⊨ps unmark_l M› and
n_d: ‹no_dup M› and
N_U_D: ‹set_mset N ∪ set_mset U ⊨p D›
using assms
by (auto simp: cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def all_decomposition_implies_def
S clauses_def cdcl⇩W_restart_mset.cdcl⇩W_conflicting_def cdcl⇩W_restart_mset_state
cdcl⇩W_restart_mset.cdcl⇩W_learned_clauses_entailed_by_init_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def cdcl⇩W_restart_mset.cdcl⇩W_learned_clause_alt_def)
have ‹set_mset N ∪ set_mset U ⊨ps CNot D›
by (rule true_clss_clss_true_clss_cls_true_clss_clss[OF N_U_M M_D])
then have ‹set_mset N ⊨ps CNot D› ‹set_mset N ⊨p D›
using N_U N_U_D true_clss_clss_left_right by blast+
then have ‹unsatisfiable (set_mset N)›
by (rule true_clss_clss_CNot_true_clss_cls_unsatisfiable)
then show ?thesis
by (auto simp: S clauses_def cdcl⇩W_restart_mset_state dest: satisfiable_decreasing)
qed
lemma conflict_of_level_unsatisfiable2:
assumes
struct: ‹cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv S› and
dec: ‹count_decided (trail S) = 0› and
confl: ‹conflicting S ≠ None›
shows ‹unsatisfiable (set_mset (init_clss S + learned_clss S))›
proof -
obtain M N U D where S: ‹S = (M, N, U, Some D)›
by (cases S) (use confl in ‹auto simp: cdcl⇩W_restart_mset_state›)
have [simp]: ‹get_all_ann_decomposition M = [([], M)]›
by (rule no_decision_get_all_ann_decomposition)
(use dec in ‹auto simp: count_decided_def filter_empty_conv S cdcl⇩W_restart_mset_state›)
have
M_D: ‹M ⊨as CNot D› and
N_U_M: ‹set_mset N ∪ set_mset U ⊨ps unmark_l M› and
n_d: ‹no_dup M› and
N_U_D: ‹set_mset N ∪ set_mset U ⊨p D›
using assms
by (auto simp: cdcl⇩W_restart_mset.cdcl⇩W_all_struct_inv_def all_decomposition_implies_def
S clauses_def cdcl⇩W_restart_mset.cdcl⇩W_conflicting_def cdcl⇩W_restart_mset_state
cdcl⇩W_restart_mset.cdcl⇩W_learned_clauses_entailed_by_init_def
cdcl⇩W_restart_mset.cdcl⇩W_M_level_inv_def cdcl⇩W_restart_mset.cdcl⇩W_learned_clause_alt_def)
have ‹set_mset N ∪ set_mset U ⊨ps CNot D›
by (rule true_clss_clss_true_clss_cls_true_clss_clss[OF N_U_M M_D])
then have ‹set_mset N ∪ set_mset U ⊨ps CNot D› ‹set_mset N ∪ set_mset U ⊨p D›
using N_U_D true_clss_clss_left_right by blast+
then have ‹unsatisfiable (set_mset N ∪ set_mset U)›
by (rule true_clss_clss_CNot_true_clss_cls_unsatisfiable)
then show ?thesis
by (auto simp: S clauses_def cdcl⇩W_restart_mset_state dest: satisfiable_decreasing)
qed
end