Theory Watched_Literals_Watch_List

theory Watched_Literals_Watch_List
  imports Watched_Literals_List Watched_Literals_All_Literals
begin

no_notation funcset (infixr "" 60)

chapter Third Refinement: Remembering watched

section Types

type_synonym clauses_to_update_wl = nat multiset
type_synonym 'v watcher = (nat × 'v literal × bool)
type_synonym 'v watched = 'v watcher list
type_synonym 'v lit_queue_wl = 'v literal multiset

type_synonym 'v twl_st_wl =
  ('v, nat) ann_lits × 'v clauses_l ×
  'v cconflict × 'v clauses × 'v clauses  × 'v clauses × 'v clauses × 'v clauses × 'v clauses ×
  'v clauses × 'v clauses × 'v lit_queue_wl × ('v literal  'v watched)


section Access Functions

fun clauses_to_update_wl :: 'v twl_st_wl  'v literal  nat  clauses_to_update_wl where
  clauses_to_update_wl (_, N, _, _, _, _, _ ,_ , _, _, _, _, W) L i =
      filter_mset (λi. i ∈# dom_m N) (mset (drop i (map fst (W L))))

fun get_trail_wl :: 'v twl_st_wl  ('v, nat) ann_lit list where
  get_trail_wl (M, _, _, _, _, _, _ ,_ , _) = M

fun literals_to_update_wl :: 'v twl_st_wl  'v lit_queue_wl where
  literals_to_update_wl (_, _, _, _, _, _ ,_ , _, _, _, _, Q, _) = Q

fun set_literals_to_update_wl :: 'v lit_queue_wl  'v twl_st_wl  'v twl_st_wl where
  set_literals_to_update_wl Q (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, _, W) =
     (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)

fun get_conflict_wl :: 'v twl_st_wl  'v cconflict where
  get_conflict_wl (_, _, D, _, _, _, _) = D

fun get_clauses_wl :: 'v twl_st_wl  'v clauses_l where
  get_clauses_wl (M, N, D, NE, UE, NEk, UEk, WS, Q) = N

fun get_unit_learned_clss_wl :: 'v twl_st_wl  'v clauses where
  get_unit_learned_clss_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = UE + UEk

fun get_unit_init_clss_wl :: 'v twl_st_wl  'v clauses where
  get_unit_init_clss_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = NE + NEk

fun get_unit_clauses_wl :: 'v twl_st_wl  'v clauses where
  get_unit_clauses_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = NE + NEk + UE + UEk

fun get_kept_unit_clauses_wl :: 'v twl_st_wl  'v clauses where
  get_kept_unit_clauses_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = NEk + UEk

fun get_unkept_learned_clss_wl :: 'v twl_st_wl  'v clauses where
  get_unkept_learned_clss_wl (M, N, D, NE, UE, NEk, UEk, NS, US, WS, Q) = UE

fun get_subsumed_init_clauses_wl :: 'v twl_st_wl  'v clauses where
  get_subsumed_init_clauses_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = NS

fun get_subsumed_learned_clauses_wl :: 'v twl_st_wl  'v clauses where
  get_subsumed_learned_clauses_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = US

abbreviation get_subsumed_clauses_wl  :: 'v twl_st_wl  'v clauses where
  get_subsumed_clauses_wl S  get_subsumed_init_clauses_wl S + get_subsumed_learned_clauses_wl S

fun get_init_clauses0_wl :: 'v twl_st_wl  'v clauses where
  get_init_clauses0_wl (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) = N0

fun get_learned_clauses0_wl :: 'v twl_st_wl  'v clauses where
  get_learned_clauses0_wl (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) = U0

abbreviation get_clauses0_wl  :: 'v twl_st_wl  'v clauses where
  get_clauses0_wl S  get_init_clauses0_wl S + get_learned_clauses0_wl S

definition get_learned_clss_wl :: 'v twl_st_wl  'v clause_l multiset where
  get_learned_clss_wl S = learned_clss_lf (get_clauses_wl S)

abbreviation get_all_learned_clss_wl :: 'v twl_st_wl  'v clause multiset where
  get_all_learned_clss_wl S  mset `# get_learned_clss_wl S + get_unit_learned_clss_wl S +
      get_subsumed_learned_clauses_wl S + get_learned_clauses0_wl S

lemma get_unit_clauses_wl_alt_def:
  get_unit_clauses_wl S = get_unit_init_clss_wl S + get_unit_learned_clss_wl S
  by (cases S) auto

fun get_watched_wl :: 'v twl_st_wl  ('v literal  'v watched) where
  get_watched_wl (_, _, _, _, _, _, _, _, _, _, _, _, W) = W

fun get_unkept_unit_init_clss_wl :: 'v twl_st_wl  'v clauses where
  get_unkept_unit_init_clss_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = NE

fun get_unkept_unit_learned_clss_wl :: 'v twl_st_wl  'v clauses where
  get_unkept_unit_learned_clss_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = UE

fun get_kept_unit_init_clss_wl :: 'v twl_st_wl  'v clauses where
  get_kept_unit_init_clss_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = NEk

fun get_kept_unit_learned_clss_wl :: 'v twl_st_wl  'v clauses where
  get_kept_unit_learned_clss_wl (M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) = UEk

lemma get_unit_init_clss_wl_alt_def:
  get_unit_init_clss_wl T = get_unkept_unit_init_clss_wl T + get_kept_unit_init_clss_wl T
  by (cases T) auto

lemma get_unit_learned_clss_wl_alt_def:
  get_unit_learned_clss_wl T = get_unkept_unit_learned_clss_wl T + get_kept_unit_learned_clss_wl T
  by (cases T) auto

abbreviation get_init_clss_wl :: 'v twl_st_wl  'v clause_l multiset where
  get_init_clss_wl S  init_clss_lf (get_clauses_wl S)

definition all_lits_st :: 'v twl_st_wl  'v literal multiset where
  all_lits_st S  all_lits (get_clauses_wl S) (get_unit_clauses_wl S + get_subsumed_clauses_wl S + get_clauses0_wl S)

definition all_init_lits_of_wl :: 'v twl_st_wl  'v clause where
  all_init_lits_of_wl S'  all_lits_of_mm (mset `# get_init_clss_wl S' + get_unit_init_clss_wl S' +
          get_subsumed_init_clauses_wl S' + get_init_clauses0_wl S')

definition all_learned_lits_of_wl :: 'v twl_st_wl  'v clause where
  all_learned_lits_of_wl S'  all_lits_of_mm (mset `# learned_clss_lf (get_clauses_wl S') + get_unit_learned_clss_wl S' +
          get_subsumed_learned_clauses_wl S' + get_learned_clauses0_wl S')

lemma all_lits_st_alt_def:
  Watched_Literals_Watch_List.all_lits_st S = all_init_lits_of_wl S + all_learned_lits_of_wl S
  apply (auto simp: all_lits_st_def all_init_lits_of_wl_def all_learned_lits_of_wl_def
    ac_simps all_lits_def all_lits_of_mm_union)
   by (metis all_clss_l_ran_m all_lits_of_mm_union get_unit_clauses_wl_alt_def image_mset_union union_assoc) 

lemma all_init_lits_of_wl_all_lits_st:
  set_mset (all_init_lits_of_wl S)  set_mset (all_lits_st S)
  unfolding all_lits_st_def all_init_lits_of_wl_def all_lits_def
  apply (subst (2) all_clss_l_ran_m[symmetric])
  unfolding image_mset_union
  apply (cases S)
  apply (auto simp: all_lits_of_mm_union)
  done

lemma in_all_lits_uminus_iff[simp]: - L ∈# all_lits_st S  L ∈# all_lits_st S
  by (auto simp: all_lits_st_def in_all_lits_of_mm_uminus_iff all_lits_def)

lemma all_lits_ofI[intro]:
  x = get_clauses_wl S  C ∈# dom_m x  n < length (x  C)  x  C ! n ∈# all_lits_st S
  using in_clause_in_all_lits_of_m[of x  C ! n] nth_mem_mset[of n x  C]
  by (auto simp: all_lits_st_def all_lits_def all_lits_of_mm_union ran_m_def all_lits_of_mm_add_mset
    dest!: multi_member_split)


section Watch List Function

definition op_watch_list :: ('v literal  'v watched)  'v literal  nat  'v watcher where
  [simp]: op_watch_list W K i = W K ! i

definition mop_watch_list :: ('v literal  'v watched)  'v literal  nat  'v watcher nres where
  mop_watch_list W K i = do {
      ASSERT(i < length (W K));
      RETURN (W K ! i)
   }

definition op_watch_list_append :: ('v literal  'v watched)  'v literal  'v watcher  ('v literal  'v watched) where
  [simp]: op_watch_list_append W K x = W (K := W K @ [x])

definition mop_watch_list_append :: ('v literal  'v watched)  'v literal  'v watcher  ('v literal  'v watched) nres where
  mop_watch_list_append W K x = do {
      RETURN (W (K := W K @ [x]))
   }

definition op_watch_list_take :: ('v literal  'v watched)  'v literal  nat  ('v literal  'v watched) where
  [simp]: op_watch_list_take W K i = W (K := take i (W K))

definition mop_watch_list_take :: ('v literal  'v watched)  'v literal  nat  ('v literal  'v watched) nres where
  mop_watch_list_take W K i = do {
      ASSERT(i  length (W K));
      RETURN (W (K := take i (W K)))
   }

lemma shows
  op_watch_list:
    i < length (W K)  mop_watch_list W K i  SPEC(λc. (c, op_watch_list W K i)  Id) and
  op_watch_list_append:
    mop_watch_list_append W K x  SPEC(λc. (c, op_watch_list_append W K x)  Id) and
  op_watch_list_take:
    i  length (W K)  mop_watch_list_take W K i  SPEC(λc. (c, op_watch_list_take W K i)  Id)
  by (auto simp: mop_watch_list_def mop_watch_list_append_def mop_watch_list_take_def
   intro!: ASSERT_leI)


section Watch List Invariants

text 
  We cannot just extract the literals of the clauses: we cannot be sure that atoms appear both
  positively and negatively in the clauses. If we could ensure that there are no pure literals, the
  definition of termall_lits_of_mm can be changed to all_lits_of_mm Ls = ∑<^sub># Ls›.

  In this definition termK is the blocking literal.


  We have several different version of the watch-list invariants, either because we need a different
  version or to simplify proofs.

   CDCL: This is the invariant that is the most important.
      binary clauses cannot be deleted
      blocking literals are in the clause
      the watched literals belong to the clause and are at the beginning.
      the set of all literals is the set of all literals (irred+red)

   Inprocessing, deduplicating binary clauses
      binary clauses can be deleted
      blocking literals still are in the clause
      the watched literals belong to the clause and are at the beginning.
      the set of all literals is the set of all irredundant literals (irred)

   Inprocessing, removing true/false literals
      all clauses appear in the watch list
      the set of all literals is the set of all irredundant literals (irred)

(We also have the version for all literals)


   Reduction
      all clauses appear in the watch list
      the set of all literals is the set of all irredundant literals (irred)

   We use the set of irredundant literals because it is easier to handle removing literals --
   deleting a clause does not change the set of all irredundant literals. We then rely on the
   invariants to go back to the set of all literals.


  One additional constraint is that the watch lists do not contain duplicates. This might seem like
  a consequence from the fact that we are correctly watching. However, the invariant talks only
  about non-deleted clauses. Technically we would not need the distinctiveness at this level, but
  during refinement we need it in order to bound the length of the watch lists.


fun correctly_marked_as_binary where
  correctly_marked_as_binary N (i, K, b)  (b  (length (N  i) = 2))

declare correctly_marked_as_binary.simps[simp del]

abbreviation distinct_watched :: 'v watched  bool where
  distinct_watched xs  distinct (map (λ(i, j, k). i) xs)

lemma distinct_watched_alt_def: distinct_watched xs = distinct (map fst xs)
  by (induction xs; auto)


fun correct_watching_except :: nat  nat  'v literal  'v twl_st_wl  bool where
  correct_watching_except i j K (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) 
    (L ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W).
       (L = K 
         distinct_watched (take i (W L) @ drop j (W L)) 
         (((i, K, b)∈#mset (take i (W L) @ drop j (W L)). i ∈# dom_m N  K  set (N  i) 
             K  L  correctly_marked_as_binary N (i, K, b)) 
          ((i, K, b)∈#mset (take i (W L) @ drop j (W L)). b  i ∈# dom_m N) 
         filter_mset (λi. i ∈# dom_m N) (fst `# mset (take i (W L) @ drop j (W L))) = clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))) 
       (L  K 
         distinct_watched (W L) 
         (((i, K, b)∈#mset (W L). i ∈# dom_m N  K  set (N  i)  K  L  correctly_marked_as_binary N (i, K, b)) 
          ((i, K, b)∈#mset (W L). b  i ∈# dom_m N) 
         filter_mset (λi. i ∈# dom_m N) (fst `# mset (W L)) = clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))))

fun correct_watching :: 'v twl_st_wl  bool where
  correct_watching (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) 
    (L ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W).
       distinct_watched (W L) 
       ((i, K, b)∈#mset (W L). i ∈# dom_m N  K  set (N  i)  K  L  correctly_marked_as_binary N (i, K, b)) 
       ((i, K, b)∈#mset (W L).  b  i ∈# dom_m N) 
       filter_mset (λi. i ∈# dom_m N) (fst `# mset (W L)) = clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))

declare correct_watching.simps[simp del]


lemma all_lits_st_simps[simp]:
   all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W(K := WK)) =
   all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
   all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, add_mset K Q, W) =
   all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) ― ‹actually covered below, but still useful for 'unfolding' by hand
  x1 ∈# dom_m x1aa  n < length (x1aa  x1)  n' < length (x1aa  x1) 
     all_lits_st (x1b, x1aa(x1  WB_More_Refinement_List.swap (x1aa  x1) n n'), D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e,
             x2e) =
   all_lits_st
            (x1b, x1aa, D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e,
             x2e)
  all_lits_st (L # M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) =
    all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
  NO_MATCH {#} Q  all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) =
     all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, W)
  NO_MATCH [] M  all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) =
     all_lits_st ([], N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
  NO_MATCH None D  all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) =
  all_lits_st (M, N, None, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
  all_lits_st (set_literals_to_update_wl WS S) = all_lits_st S
  by (cases S; auto simp: all_lits_st_def all_lits_def all_lits_of_mm_union ran_m_clause_upd
    image_mset_remove1_mset_if; fail)+

lemma in_clause_in_all_lits_of_mm[simp]:
  x1 ∈# dom_m x1aa  n < length (x1aa  x1) 
     x1aa  x1 ! n ∈# all_lits_st (x1b, x1aa, D, x1c, x1d, NS, US, N0, U0, x1e,
             x2e)
  by (auto simp: all_lits_st_def all_lits_def all_lits_of_mm_union ran_m_clause_upd
    all_lits_of_mm_add_mset in_clause_in_all_lits_of_m
    image_mset_remove1_mset_if ran_m_def dest!: multi_member_split)

lemma correct_watching_except_correct_watching:
  assumes
    j: j  length (W K) and
    corr: correct_watching_except i j K (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
 shows correct_watching (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W(K := take i (W K)))
proof -
  have
    H1: L i' K' b. L ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) 
       (L = K 
         distinct_watched (take i (W L) @ drop j (W L)) 
         (((i', K', b)∈#mset (take i (W L) @ drop j (W L))  i' ∈# dom_m N 
                K'  set (N  i')  K'  L  correctly_marked_as_binary N (i', K', b)) 
         ((i', K', b)∈#mset (take i (W L) @ drop j (W L))  b  i' ∈# dom_m N) 
         filter_mset (λi. i ∈# dom_m N) (fst `# mset (take i (W L) @ drop j (W L))) =
            clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))) and
    H2: L i K' b. L ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)  (L  K 
         distinct_watched (W L) 
         (((i, K', b)∈#mset (W L)  i ∈# dom_m N  K'  set (N  i)  K'  L 
             (correctly_marked_as_binary N (i, K', b))) 
          ((i, K', b)∈#mset (W L)  b  i ∈# dom_m N) 
         filter_mset (λi. i ∈# dom_m N) (fst `# mset (W L)) =
             clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})))
    using corr unfolding correct_watching_except.simps
    by fast+
  show ?thesis
    unfolding correct_watching.simps
    apply (intro conjI allI impI ballI)
    subgoal for L
      apply (cases L = K)
      subgoal
        using H1[of L] j
        by (auto split: if_splits simp: all_lits_st_def)
      subgoal
        using H2[of L] j
        by (auto split: if_splits simp: all_lits_st_def)
      done
    subgoal for L x
      apply (cases L = K)
      subgoal
        using H1[of L fst x fst (snd x) snd (snd x)] j
        by (auto split: if_splits simp: all_lits_st_def)
      subgoal
        using H2[of L fst x fst (snd x) snd (snd x)]
        by auto
      done
    subgoal for L
      apply (cases L = K)
      subgoal
        using H1[of L _ _] j
        by (auto split: if_splits simp: all_lits_st_def)
      subgoal
        using H2[of L _ _]
        by auto
      done
    subgoal for L
      apply (cases L = K)
      subgoal
        using H1[of L _ _] j
        by (auto split: if_splits)
      subgoal
        using H2[of L _ _]
        by auto
      done
    done
qed

lemma length_ge2I: x  y  x  set xs  y  set xs  length xs  2
  using card_length[of xs] card_mono[of set xs {x, y}]
  by auto

lemma correct_watching_except_alt_def:
  correct_watching_except i j K (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) 
    (L ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W).
       (L = K 
         distinct_watched (take i (W L) @ drop j (W L)) 
         (((i, K, b)∈#mset (take i (W L) @ drop j (W L)). i ∈# dom_m N  K  set (N  i) 
             K  L  L  set (watched_l (N  i))  length (N  i)  2  correctly_marked_as_binary N (i, K, b)) 
          ((i, K, b)∈#mset (take i (W L) @ drop j (W L)). b  i ∈# dom_m N) 
         filter_mset (λi. i ∈# dom_m N) (fst `# mset (take i (W L) @ drop j (W L))) = clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))) 
       (L  K 
         distinct_watched (W L) 
         (((i, K, b)∈#mset (W L). i ∈# dom_m N  K  set (N  i)  K  L  L  set (watched_l(N  i))  length (N  i)  2  correctly_marked_as_binary N (i, K, b)) 
          ((i, K, b)∈#mset (W L). b  i ∈# dom_m N) 
         filter_mset (λi. i ∈# dom_m N) (fst `# mset (W L)) = clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))))
proof -
  have 1: correct_watching_except i j K (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) 
    (L ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W).
       (L = K 
         distinct_watched (take i (W L) @ drop j (W L)) 
         (((i, K, b)∈#mset (take i (W L) @ drop j (W L)). i ∈# dom_m N  K  set (N  i) 
             K  L  L  set (watched_l (N  i))  correctly_marked_as_binary N (i, K, b)) 
          ((i, K, b)∈#mset (take i (W L) @ drop j (W L)). b  i ∈# dom_m N) 
         filter_mset (λi. i ∈# dom_m N) (fst `# mset (take i (W L) @ drop j (W L))) = clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))) 
       (L  K 
         distinct_watched (W L) 
         (((i, K, b)∈#mset (W L). i ∈# dom_m N  K  set (N  i)  K  L  L  set (watched_l (N  i))  correctly_marked_as_binary N (i, K, b)) 
          ((i, K, b)∈#mset (W L). b  i ∈# dom_m N) 
         filter_mset (λi. i ∈# dom_m N) (fst `# mset (W L)) = clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))))
      unfolding correct_watching_except.simps
      apply (intro impI ballI conjI iffI)
      subgoal by auto[]
      defer
      subgoal by fast
      subgoal by fast
      subgoal by fast
      subgoal for K x
        using distinct_mset_dom[of N]
        apply (clarsimp dest!: multi_member_split simp:  ran_m_def clause_to_update_def)
        apply (frule bspec[of set (W K)], assumption)
        apply (auto split: if_splits dest: in_set_takeD union_single_eq_member filter_mset_eq_add_msetD
           simp: in_set_conv_decomp[of _ W K] filter_union_or_split filter_eq_replicate_mset)
       done
      subgoal by fast
      subgoal by fast
      subgoal by fast
      subgoal by fast
      subgoal by fast
      subgoal by fast
      subgoal by fast
      subgoal by fast
      subgoal by fast
      subgoal by fast
      subgoal
        using distinct_mset_dom[of N]
        apply (clarsimp dest!: multi_member_split simp:  ran_m_def clause_to_update_def)
        apply (fastforce split: if_splits dest: in_set_takeD union_single_eq_member
          simp: in_set_conv_decomp[of (_, _, _) take _ _] in_set_conv_decomp[of  (_, _, _) drop _ _])[]
        done
     done
  have H: (K  set (N  i)   K  L  L  set (watched_l (N  i))  P) 
       (K  set (N  i)  K  L  L  set (watched_l (N  i))  length (N  i)  2  P) for i K P L
    using length_ge2I[of K L] by (auto dest: in_set_takeD)
   show ?thesis
      unfolding 1 H[symmetric] ..
qed

definition clause_to_update_wl:: 'v literal  'v twl_st_wl  'v clauses_to_update_lwhere
  clause_to_update_wl L S =
    filter_mset
      (λC::nat. L  set (watched_l (get_clauses_wl S  C)))
      (dom_m (get_clauses_wl S))

(*TODO kill, see alt def below*)
fun watched_by :: 'v twl_st_wl  'v literal  'v watched where
  watched_by (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) L = W L

lemma watched_by_alt_def: watched_by S L = get_watched_wl S L
  by (cases S) auto
(*END*)
definition all_atms :: _  _  'v multiset where
  all_atms N NUE = atm_of `# all_lits N NUE

definition all_atms_st :: 'v twl_st_wl  'v multiset where
  all_atms_st S  all_atms (get_clauses_wl S) (get_unit_clauses_wl S + get_subsumed_clauses_wl S + get_clauses0_wl S)

lemma all_atms_st_alt_def: all_atms_st S = atm_of `# all_lits_st S
  by (auto simp: all_atms_def all_lits_st_def all_atms_st_def)

lemmas all_atms_st_alt_def_sym[simp] =  all_atms_st_alt_def[symmetric]

lemma in_all_lits_minus_iff: -L ∈# all_lits N NUE  L ∈# all_lits N NUE
  unfolding all_lits_def in_all_lits_of_mm_uminus_iff ..

lemma all_lits_of_all_atms_of: K ∈# all_lits N NUE  K ∈# all (all_atms N NUE)
  by (simp add: all_atm_of_all_lits_of_mm all_atms_def all_lits_def)

lemma all_all_atms_all_lits: set_mset (all (all_atms N NE)) = set_mset (all_lits N NE)
  unfolding all_atms_def all_atm_of_all_lits_of_mm all_lits_def ..


definition blits_in_ℒin :: 'v twl_st_wl  bool where
  blits_in_ℒin S = (L ∈# all_lits_st S. (i, K, b)  set (watched_by S L). K ∈# all_lits_st S)


definition literals_are_ℒin :: 'v multiset  'v twl_st_wl  bool where
  literals_are_ℒin 𝒜 S  (is_ℒall 𝒜 (all_lits_st S)  blits_in_ℒin S)


lemma literals_are_in_ℒin_nth:
  fixes C :: nat
  assumes dom: C ∈# dom_m (get_clauses_wl S) and
   literals_are_ℒin 𝒜 S
  shows literals_are_in_ℒin 𝒜 (mset (get_clauses_wl S  C))
proof -
  let ?N = get_clauses_wl S
  have ?N  C ∈# ran_mf ?N
    using dom by (auto simp: ran_m_def)
  then have mset (?N  C) ∈# mset `# (ran_mf ?N)
    by blast
  from all_lits_of_m_subset_all_lits_of_mmD[OF this] show ?thesis
    using assms(2) unfolding is_ℒall_def literals_are_in_ℒin_def literals_are_ℒin_def
    by (auto simp: all_lits_st_def all_lits_of_mm_union all_lits_alt_def)
qed

lemma literals_are_ℒin_set_mset_ℒall[simp]:
  literals_are_ℒin 𝒜 S  set_mset (all (all_atms_st S)) = set_mset (all 𝒜)
  using in_all_lits_of_mm_ain_atms_of_iff
  unfolding set_mset_set_mset_eq_iff is_ℒall_def Ball_def in_ℒall_atm_of_in_atms_of_iff
    in_all_lits_of_mm_ain_atms_of_iff atms_of_ℒall_𝒜in literals_are_ℒin_def all_lits_st_def
    all_lits_def
  by (auto simp: in_all_lits_of_mm_ain_atms_of_iff all_atms_def simp del: all_atms_st_alt_def_sym
    simp: all_lits_def all_atms_st_def)

lemma is_ℒall_all_lits_st_ℒall[simp]:
  is_ℒall 𝒜 (all_lits_st S) 
    set_mset (all (all_atms_st S)) = set_mset (all 𝒜)
  is_ℒall 𝒜 (all_lits N NUE) 
    set_mset (all (all_atms N NUE)) = set_mset (all 𝒜)
  is_ℒall 𝒜 (all_lits N NUE) 
    set_mset (all (atm_of `# all_lits N NUE)) = set_mset (all 𝒜)
  using in_all_lits_of_mm_ain_atms_of_iff
  unfolding set_mset_set_mset_eq_iff is_ℒall_def Ball_def in_ℒall_atm_of_in_atms_of_iff
    in_all_lits_of_mm_ain_atms_of_iff atms_of_ℒall_𝒜in all_lits_st_def all_atms_st_def
  by (auto simp: in_all_lits_of_mm_ain_atms_of_iff all_lits_def all_atms_def)


lemma in_set_all_atms_iff:
  y ∈# all_atms bu bw 
    y  atms_of_mm (mset `# ran_mf bu)  y  atms_of_mm bw
  by (auto simp: all_atms_def all_lits_def in_all_lits_of_mm_ain_atms_of_iff
     atm_of_all_lits_of_mm)


lemma blits_in_ℒin_keep_watch:
  assumes blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g) and
    w:w < length (watched_by (a, b, c, d, e,  NEk, UEk, NS, US, N0, U0, f, g) K)
  shows blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g (K := (g K)[j := g K ! w]))
proof -
  let ?S = (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
  let ?T = (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g (K := (g K)[j := g K ! w]))
  let ?g = g (K := (g K)[j := g K ! w])
  have H: L i K b. L∈# all_lits_st ?S  (i, K, b) set (g L) 
        K ∈# all_lits_st ?S
    using assms
    unfolding blits_in_ℒin_def watched_by.simps
    by blast
  have  L∈# all_lits_st ?S  (i, K', b') set (?g L) 
        K' ∈# all_lits_st ?S for L i K' b'
    using H[of L i K'] H[of L fst (g K ! w) fst (snd (g K ! w))]
      nth_mem[OF w]
    unfolding blits_in_ℒin_def watched_by.simps
    by (cases j < length (g K); cases g K ! w)
      (auto split: if_splits elim!: in_set_upd_cases)
  moreover have all_atms_st ?S = all_atms_st ?T
    by (auto simp: all_lits_def all_atms_def all_atms_st_def)
  ultimately show ?thesis
    unfolding blits_in_ℒin_def watched_by.simps
    by force
qed

lemma all_lits_swap[simp]:
  x1 ∈# dom_m x1aa  n < length (x1aa  x1)  n' < length (x1aa  x1) 
   all_lits
            (x1aa(x1  swap (x1aa  x1) n n'))
            (x1cx1d) =
  all_lits x1aa (x1cx1d)
  using distinct_mset_dom[of x1aa]
  apply (auto simp: ran_m_def all_lits_def dest!: multi_member_split)
  apply (subst image_mset_cong[of _ %x. mset
           (fst (the (fmlookup x1aa x)))])
  apply auto
  done

lemma blits_in_ℒin_propagate:
  x1 ∈# dom_m x1aa  n < length (x1aa  x1)  n' < length (x1aa  x1) 
    blits_in_ℒin (Propagated A x1' # x1b, x1aa
         (x1  swap (x1aa  x1) n n'), D, x1c, x1d,
          NEk, UEk, NS, US, N0, U0, add_mset A' x1e, x2e) 
    blits_in_ℒin (x1b, x1aa, D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e)
  x1 ∈# dom_m x1aa  n < length (x1aa  x1)  n' < length (x1aa  x1) 
    blits_in_ℒin (x1b, x1aa
         (x1  swap (x1aa  x1) n n'), D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e) 
    blits_in_ℒin (x1b, x1aa, D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e)
  blits_in_ℒin
        (Propagated A x1' # x1b, x1aa, D, x1c, x1d,
         NEk, UEk, NS, US, N0, U0, add_mset A' x1e, x2e) 
    blits_in_ℒin (x1b, x1aa, D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e)
  x1' ∈# dom_m x1aa  n < length (x1aa  x1')  n' < length (x1aa  x1') 
    K ∈# all_lits_st (x1b, x1aa, D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e)  blits_in_ℒin
        (x1a, x1aa(x1'  swap (x1aa  x1') n n'), D, x1c, x1d,
         NEk, UEk, NS, US, N0, U0, x1e, x2e
         (x1aa  x1' ! n' :=
            x2e (x1aa  x1' ! n') @ [(x1', K, b')])) 
    blits_in_ℒin (x1a, x1aa, D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e)
  K ∈# all_lits_st (x1b, x1aa, D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e) 
     blits_in_ℒin (x1a, x1aa, D, x1c, x1d,
         NEk, UEk, NS, US, N0, U0, x1e, x2e
         (K' := x2e K' @ [(x1', K, b')])) 
  blits_in_ℒin (x1a, x1aa, D, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e)
  unfolding blits_in_ℒin_def
  by (auto split: if_splits)

lemma blits_in_ℒin_keep_watch':
  assumes K': fst (snd w) ∈# all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g) and
    w: blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
  shows blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g (K := (g K)[j := w]))
proof -
  let ?𝒜 = all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
  let ?g = g (K := (g K)[j := w])
  have H: L i K b'. L∈# ?𝒜  (i, K, b') set (g L)  K ∈# ?𝒜
    using assms
    unfolding blits_in_ℒin_def watched_by.simps
    by blast
  have L∈# ?𝒜  (i, K', b') set (?g L)  K' ∈# ?𝒜 for L i K' b'
    using H[of L i K'] K'
    unfolding blits_in_ℒin_def watched_by.simps
    by (cases j < length (g K))
      (auto split: if_splits elim!: in_set_upd_cases)

  then show ?thesis
    unfolding blits_in_ℒin_def watched_by.simps
    by force
qed


lemma blits_in_ℒin_keep_watch'':
  assumes K': K' ∈# all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
    L' ∈# all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g) and
    w: blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
  shows blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f,
    g (K := (g K)[j := (i, K', b')], L := g L @ [(i', L', b'')]))
  using assms
  unfolding blits_in_ℒin_def
  by (auto split: if_splits elim!: in_set_upd_cases)

lemma clause_to_update_wl_alt_def:
   clause_to_update_wl L (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g) =
     clause_to_update L (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, {#}, {#})
  unfolding clause_to_update_wl_def clause_to_update_def by simp

lemma correct_watching_except_alt_def2:
  correct_watching_except i j K S 
    (L ∈# all_lits_st S.
       (L = K 
         distinct_watched (take i (watched_by S L) @ drop j (watched_by S L)) 
         (((i, K, b)∈#mset (take i (watched_by S L) @ drop j (watched_by S L)). i ∈# dom_m (get_clauses_wl S)  K  set (get_clauses_wl S  i) 
             K  L  L  set (watched_l (get_clauses_wl S  i))  length (get_clauses_wl S  i)  2  correctly_marked_as_binary (get_clauses_wl S) (i, K, b)) 
          ((i, K, b)∈#mset (take i (watched_by S L) @ drop j (watched_by S L)). b  i ∈# dom_m (get_clauses_wl S)) 
         filter_mset (λi. i ∈# dom_m (get_clauses_wl S)) (fst `# mset (take i (watched_by S L) @ drop j (watched_by S L))) = clause_to_update_wl L S)) 
       (L  K 
         distinct_watched (watched_by S L) 
         (((i, K, b)∈#mset (watched_by S L). i ∈# dom_m (get_clauses_wl S)  K  set (get_clauses_wl S  i)  K  L  L  set (watched_l (get_clauses_wl S  i))  length (get_clauses_wl S  i)  2  correctly_marked_as_binary (get_clauses_wl S) (i, K, b)) 
          ((i, K, b)∈#mset (watched_by S L). b  i ∈# dom_m (get_clauses_wl S)) 
         filter_mset (λi. i ∈# dom_m (get_clauses_wl S)) (fst `# mset (watched_by S L)) = clause_to_update_wl L S)))
  by (cases S; hypsubst)
    (simp only: correct_watching_except_alt_def watched_by.simps
      get_clauses_wl.simps clause_to_update_wl_alt_def get_unit_clauses_wl.simps)

fun update_watched :: 'v literal  'v watched  'v twl_st_wl  'v twl_st_wl where
  update_watched L WL (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) =
     (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W(L:= WL))

definition mop_watched_by_at :: 'v twl_st_wl  'v literal  nat  'v watcher nres where
mop_watched_by_at = (λS L w. do {
   ASSERT(L ∈# all_lits_st S);
   ASSERT(w < length (watched_by S L));
  RETURN (watched_by S L ! w)
})


lemma bspec': x  a  xa. P x  P x
  by (rule bspec)

lemma correct_watching_exceptD:
  assumes
    correct_watching_except i j L S and
    L ∈# all_lits_st S and
    w: w < length (watched_by S L) w  j fst (watched_by S L ! w) ∈# dom_m (get_clauses_wl S)
  shows fst (snd (watched_by S L ! w))  set (get_clauses_wl S  (fst (watched_by S L ! w)))
proof -
  have H: x. xset (take i (watched_by S L))  set (drop j (watched_by S L)) 
          case x of (i, K, b)  i ∈# dom_m (get_clauses_wl S)  K  set (get_clauses_wl S  i) 
           K  L
    using assms multi_member_split[OF assms(2)]
    by (cases S; cases watched_by S L ! w)
     (clarsimp, blast)
  have ij. i < length (watched_by S L) 
            watched_by S L ! w = watched_by S L ! i
    by (rule exI[of _ w])
      (use w in auto)
  then show ?thesis
    using H[of watched_by S L ! w] w
    by (cases watched_by S L ! w) (auto simp: in_set_drop_conv_nth)
qed

declare correct_watching_except.simps[simp del]


fun st_l_of_wl :: ('v literal × nat) option  'v twl_st_wl  'v twl_st_l where
  st_l_of_wl None (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) = (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, Q)
| st_l_of_wl (Some (L, j)) (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) =
    (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0,
     (if D  None then {#} else clauses_to_update_wl (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) L j,
      Q))

definition state_wl_l :: ('v literal × nat) option  ('v twl_st_wl × 'v twl_st_l) set where
state_wl_l L = {(T, T'). T' = st_l_of_wl L T}

fun twl_st_of_wl :: ('v literal × nat) option  ('v twl_st_wl × 'v twl_st) set where
  twl_st_of_wl L = state_wl_l L O twl_st_l (map_option fst L)


named_theorems twl_st_wl Conversions simp rules

lemma [twl_st_wl]:
  assumes (S, T)  state_wl_l L
  shows
    get_trail_l T = get_trail_wl S and
    get_clauses_l T = get_clauses_wl S and
    get_conflict_l T = get_conflict_wl S and
    L = None  clauses_to_update_l T = {#}
    L  None  get_conflict_wl S  None  clauses_to_update_l T = {#}
    L  None  get_conflict_wl S = None  clauses_to_update_l T =
       clauses_to_update_wl S (fst (the L)) (snd (the L)) and
    literals_to_update_l T = literals_to_update_wl S
    get_unit_learned_clss_l T = get_unit_learned_clss_wl S
    get_unit_init_clauses_l T = get_unit_init_clss_wl S
    get_unit_learned_clss_l T = get_unit_learned_clss_wl S
    get_unit_clauses_l T = get_unit_clauses_wl S
    get_kept_unit_clauses_l T = get_kept_unit_clauses_wl S
    get_subsumed_init_clauses_l T = get_subsumed_init_clauses_wl S
    get_subsumed_learned_clauses_l T = get_subsumed_learned_clauses_wl S
    get_subsumed_clauses_l T = get_subsumed_clauses_wl S
    get_init_clauses0_l T = get_init_clauses0_wl S
    get_learned_clauses0_l T = get_learned_clauses0_wl S
    get_clauses0_l T = get_clauses0_wl S
    get_init_clss_l T = get_init_clss_wl S
    all_lits_of_st_l T = all_lits_st S
    get_unkept_learned_clss_l T = get_unkept_learned_clss_wl S
    all_lits_of_mm (get_all_clss_l T) = all_lits_st S
  using assms unfolding state_wl_l_def all_clss_lf_ran_m[symmetric]
  apply (cases S; cases T; cases L; auto split: option.splits simp: get_init_clss_l_def
    all_lits_st_def all_lits_def all_lits_of_st_l_def ac_simps; fail)+
  using assms unfolding state_wl_l_def all_clss_lf_ran_m[symmetric] all_lits_st_def all_lits_def
  apply (subst all_clss_lf_ran_m)
  by (cases S; cases T; cases L; auto split: option.splits simp: get_init_clss_l_def ac_simps)

lemma [twl_st_wl]:
  (a, a')  state_wl_l None 
        get_learned_clss_l a' = get_learned_clss_wl a
  unfolding state_wl_l_def by (cases a; cases a')
   (auto simp: get_learned_clss_l_def get_learned_clss_wl_def)

lemma remove_one_lit_from_wq_def:
  remove_one_lit_from_wq L S = set_clauses_to_update_l (clauses_to_update_l S - {#L#}) S
  by (cases S) auto

lemma correct_watching_set_literals_to_update[simp]:
  correct_watching (set_literals_to_update_wl WS T') = correct_watching T'
  by (cases T') (auto simp: correct_watching.simps)

lemma [twl_st_wl]:
  get_clauses_wl (set_literals_to_update_wl W S) = get_clauses_wl S
  get_unit_init_clss_wl (set_literals_to_update_wl W S) = get_unit_init_clss_wl S
  by (cases S; auto; fail)+

lemma get_conflict_wl_set_literals_to_update_wl[twl_st_wl]:
  get_conflict_wl (set_literals_to_update_wl P S) = get_conflict_wl S
  get_unit_clauses_wl (set_literals_to_update_wl P S) = get_unit_clauses_wl S
  get_init_clauses0_wl (set_literals_to_update_wl P S) = get_init_clauses0_wl S
  get_learned_clauses0_wl (set_literals_to_update_wl P S) = get_learned_clauses0_wl S
  get_clauses0_wl (set_literals_to_update_wl P S) = get_clauses0_wl S
  by (cases S; auto; fail)+

definition set_conflict_wl_pre :: nat  'v twl_st_wl  bool where
set_conflict_wl_pre C S 
   (S' b. (S, S')  state_wl_l b  set_conflict_l_pre C S'  blits_in_ℒin S)

definition set_conflict_wl :: nat  'v twl_st_wl  'v twl_st_wl nres where
  set_conflict_wl = (λC (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W). do {
     ASSERT(set_conflict_wl_pre C (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
     RETURN (M, N, Some (mset (N  C)), NE, UE, NEk, UEk, NS, US, N0, U0, {#}, W)
   })

lemma state_wl_l_mark_of_is_decided:
  (x, y)  state_wl_l b 
       get_trail_wl x  [] 
       is_decided (hd (get_trail_l y)) = is_decided (hd (get_trail_wl x))
  by (cases get_trail_wl x; cases get_trail_l y; cases hd (get_trail_wl x);
     cases hd (get_trail_l y); cases b; cases x)
   (auto simp: state_wl_l_def convert_lit.simps st_l_of_wl.simps)

lemma state_wl_l_mark_of_is_proped:
  (x, y)  state_wl_l b 
       get_trail_wl x  [] 
       is_proped (hd (get_trail_l y)) = is_proped (hd (get_trail_wl x))
  by (cases get_trail_wl x; cases get_trail_l y; cases hd (get_trail_wl x);
     cases hd (get_trail_l y); cases b; cases x)
   (auto simp: state_wl_l_def convert_lit.simps)

text We here also update the list of watched clauses termWL.
declare twl_st_wl[simp]


lemma
  assumes
      x2_T: (x2, T)  state_wl_l b and
      struct: twl_struct_invs U and
      T_U: (T, U)  twl_st_l b'
  shows
    literals_are_ℒin_literals_are_ℒin_trail:
      literals_are_ℒin 𝒜in x2  literals_are_in_ℒin_trail 𝒜in (get_trail_wl x2)
     (is _ ?trail) and
    literals_are_ℒin_literals_are_in_ℒin_conflict:
      literals_are_ℒin 𝒜in x2  get_conflict_wl x2  None  literals_are_in_ℒin 𝒜in (the (get_conflict_wl x2)) and
    conflict_not_tautology:
      get_conflict_wl x2  None  ¬tautology (the (get_conflict_wl x2))
proof -
  have
    alien: cdclW_restart_mset.no_strange_atm (stateW_of U) and
    confl: cdclW_restart_mset.cdclW_conflicting (stateW_of U) and
    M_lev: cdclW_restart_mset.cdclW_M_level_inv (stateW_of U) and
    dist: cdclW_restart_mset.distinct_cdclW_state (stateW_of U)
    using struct unfolding twl_struct_invs_def cdclW_restart_mset.cdclW_all_struct_inv_def
      pcdcl_all_struct_invs_def stateW_of_def
   by fast+

  show lits_trail: literals_are_in_ℒin_trail 𝒜in (get_trail_wl x2)
    if literals_are_ℒin 𝒜in x2
    using alien that x2_T T_U unfolding is_ℒall_def
      literals_are_in_ℒin_trail_def cdclW_restart_mset.no_strange_atm_def
      literals_are_ℒin_def all_lits_def all_atms_def all_lits_st_def
    by (subst (asm) all_clss_l_ran_m[symmetric])
      (auto 5 2
        simp del: all_clss_l_ran_m union_filter_mset_complement
        simp: twl_st twl_st_l twl_st_wl all_lits_of_mm_union lits_of_def
        convert_lits_l_def image_image in_all_lits_of_mm_ain_atms_of_iff
        get_unit_clauses_wl_alt_def Un_assoc ac_simps
      simp flip: stateW_of_def)

  {
    assume conf: get_conflict_wl x2  None
    show lits_confl: literals_are_in_ℒin 𝒜in (the (get_conflict_wl x2))
      if literals_are_ℒin 𝒜in x2
      using x2_T T_U alien that conf unfolding is_ℒall_alt_def
       cdclW_restart_mset.no_strange_atm_def literals_are_in_ℒin_alt_def
       literals_are_ℒin_def all_lits_def all_atms_def all_lits_st_def
      apply (subst (asm) all_clss_l_ran_m[symmetric])
      unfolding image_mset_union all_lits_of_mm_union
      by (auto simp add: twl_st all_lits_of_mm_union lits_of_def
         image_image in_all_lits_of_mm_ain_atms_of_iff
        in_all_lits_of_m_ain_atms_of_iff
        get_unit_clauses_wl_alt_def
        simp del: all_clss_l_ran_m
      simp flip: stateW_of_def)

    have M_confl: get_trail_wl x2 ⊨as CNot (the (get_conflict_wl x2))
      using confl conf x2_T T_U unfolding cdclW_restart_mset.cdclW_conflicting_def
      by (auto 5 5 simp: twl_st true_annots_def)
    moreover have n_d: no_dup (get_trail_wl x2)
      using M_lev x2_T T_U unfolding cdclW_restart_mset.cdclW_M_level_inv_def
      by (auto simp: twl_st)
    ultimately show 4: ¬tautology (the (get_conflict_wl x2))
      using n_d M_confl
      by (meson no_dup_consistentD tautology_decomp' true_annots_true_cls_def_iff_negation_in_model)
  }
qed



fun equality_except_conflict_wl :: 'v twl_st_wl  'v twl_st_wl  bool where
equality_except_conflict_wl (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q)
     (M', N', D', NE', UE', NEk', UEk', NS', US', N0', U0', WS', Q') 
  M = M'  N = N'  NE = NE'  UE = UE'  NEk = NEk'  UEk = UEk'  NS = NS'  US = US' 
  N0 = N0'  U0 = U0'  WS = WS'  Q = Q'

fun equality_except_trail_wl :: 'v twl_st_wl  'v twl_st_wl  bool where
equality_except_trail_wl (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q)
     (M', N', D', NE', UE', NEk', UEk', NS', US', N0', U0', WS', Q') 
  N = N'  D = D'  NE = NE'  NEk = NEk'  UEk = UEk'  NS = NS'  US = US'  UE = UE' 
  N0 = N0'  U0 = U0'  WS = WS'  Q = Q'

lemma equality_except_conflict_wl_get_clauses_wl:
    equality_except_conflict_wl S Y  get_clauses_wl S = get_clauses_wl Y and
  equality_except_conflict_wl_get_trail_wl:
    equality_except_conflict_wl S Y  get_trail_wl S = get_trail_wl Y and
  equality_except_trail_wl_get_conflict_wl:
    equality_except_trail_wl S Y  get_conflict_wl S = get_conflict_wl Y and
  equality_except_trail_wl_get_clauses_wl:
    equality_except_trail_wl S Y get_clauses_wl S = get_clauses_wl Y and
  equality_except_trail_wl_get_clauses0_wl:
    equality_except_trail_wl S Y get_clauses0_wl S = get_clauses0_wl Y
 by (cases S; cases Y; solves auto)+


definition unit_prop_body_wl_inv where
unit_prop_body_wl_inv T j i L  (i < length (watched_by T L)  j  i 
    L ∈# all_lits_st T  blits_in_ℒin T 
   (fst (watched_by T L ! i) ∈# dom_m (get_clauses_wl T) 
    (T'. (T, T')  state_wl_l (Some (L, Suc i))  j  i 
    unit_propagation_inner_loop_body_l_inv L (fst (watched_by T L ! i)) T' 
     correct_watching_except (Suc j) (Suc i) L T)))

lemma unit_prop_body_wl_inv_alt_def:
  unit_prop_body_wl_inv T j i L  (i < length (watched_by T L)  j  i 
    L ∈# all_lits_st T  blits_in_ℒin T 
   (fst (watched_by T L ! i) ∈# dom_m (get_clauses_wl T) 
    (T'. (T, T')  state_wl_l (Some (L, Suc i)) 
    unit_propagation_inner_loop_body_l_inv L (fst (watched_by T L ! i)) T' 
     correct_watching_except (Suc j) (Suc i) L T 
    get_conflict_wl T = None 
    length (get_clauses_wl T  fst (watched_by T L ! i))  2)))
  (is ?A = ?B)
proof
  assume ?B
  then show ?A
    unfolding unit_prop_body_wl_inv_def
    by blast
next
  assume ?A
  then show ?B
  proof (cases fst (watched_by T L ! i) ∈# dom_m (get_clauses_wl T))
    case False
    then show ?B
      using ?A unfolding unit_prop_body_wl_inv_def
      by blast
  next
    case True
    then obtain T' where
      i < length (watched_by T L)
      j  i and
      TT': (T, T')  state_wl_l (Some (L, Suc i)) and
      inv: unit_propagation_inner_loop_body_l_inv L (fst (watched_by T L ! i))
         T' and
      L ∈# all_lits_st T
      correct_watching_except (Suc j) (Suc i) L T
      using ?A unfolding unit_prop_body_wl_inv_def
      by blast

    obtain x where
      x: (set_clauses_to_update_l
         (clauses_to_update_l T' + {#fst (watched_by T L ! i)#}) T', x)
        twl_st_l (Some L) and
      struct_invs: twl_struct_invs x and
      twl_stgy_invs x and
      fst (watched_by T L ! i) ∈# dom_m (get_clauses_l T') and
      0 < fst (watched_by T L ! i) and
      0 < length (get_clauses_l T'  fst (watched_by T L ! i)) and
      no_dup (get_trail_l T') and
      (if get_clauses_l T'  fst (watched_by T L ! i) ! 0 = L then 0 else 1)
       < length (get_clauses_l T'  fst (watched_by T L ! i)) and
      1 - (if get_clauses_l T'  fst (watched_by T L ! i) ! 0 = L then 0 else 1)
       < length (get_clauses_l T'  fst (watched_by T L ! i)) and
      L  set (watched_l (get_clauses_l T'  fst (watched_by T L ! i))) and
      confl: get_conflict_l T' = None
      using inv unfolding unit_propagation_inner_loop_body_l_inv_def by blast

    have Multiset.Ball (get_clauses x) struct_wf_twl_cls
      using struct_invs unfolding twl_struct_invs_def twl_st_inv_alt_def by blast
    moreover have twl_clause_of (get_clauses_wl T  fst (watched_by T L ! i)) ∈# get_clauses x
      using TT' x True by auto
    ultimately have 1: length (get_clauses_wl T  fst (watched_by T L ! i))  2
      by auto
    have 2: get_conflict_wl T = None
      using confl TT' x by auto
    show ?B
      using ?A 1 2 unfolding unit_prop_body_wl_inv_def
      by blast
  qed
qed

definition propagate_lit_wl_general :: 'v literal  nat  nat  'v twl_st_wl  'v twl_st_wl nres where
  propagate_lit_wl_general = (λL' C i (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W). do {
      ASSERT(C ∈# dom_m N);
      ASSERT(D = None);
      ASSERT(L' ∈# all_lits_st (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
      ASSERT(i  1);
      M  cons_trail_propagate_l L' C M;
      N  (if length (N  C) > 2 then mop_clauses_swap N C 0 (Suc 0 - i) else RETURN N);
      RETURN (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, add_mset (-L') Q, W) })

definition propagate_lit_wl :: 'v literal  nat  nat  'v twl_st_wl  'v twl_st_wl nres where
  propagate_lit_wl = (λL' C i (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W). do {
      ASSERT(C ∈# dom_m N);
      ASSERT(D = None);
      ASSERT(L' ∈# all_lits_st (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
      ASSERT(i  1);
      M  cons_trail_propagate_l L' C M;
      N  mop_clauses_swap N C 0 (Suc 0 - i);
      RETURN (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, add_mset (-L') Q, W)
   })

definition propagate_lit_wl_bin :: 'v literal  nat  'v twl_st_wl  'v twl_st_wl nres where
  propagate_lit_wl_bin = (λL' C (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W). do {
      ASSERT(D = None);
      ASSERT(C ∈# dom_m N);
      ASSERT(L' ∈# all_lits_st (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
      M  cons_trail_propagate_l L' C M;
      RETURN (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, add_mset (-L') Q, W)})

definition propagate_lit_wl_bin' :: 'v literal  nat  nat  'v twl_st_wl  'v twl_st_wl nres
   where
propagate_lit_wl_bin' L' C i = propagate_lit_wl_bin L' C

definition keep_watch :: 'v literal  nat  nat  'v twl_st_wl  'v twl_st_wl where
  keep_watch = (λL i j (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W).
      (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W(L := (W L)[i := W L ! j])))

definition mop_keep_watch :: 'v literal  nat  nat  'v twl_st_wl  'v twl_st_wl nres where
  mop_keep_watch = (λL i j S. do {
    ASSERT(L ∈# all_lits_st S);
    ASSERT(i < length (watched_by S L));
    ASSERT(j < length (watched_by S L));
    RETURN (keep_watch L i j S)
  })

lemma length_watched_by_keep_watch[twl_st_wl]:
  length (watched_by (keep_watch L i j S) K) = length (watched_by S K)
  by (cases S) (auto simp: keep_watch_def)

lemma watched_by_keep_watch_neq[twl_st_wl, simp]:
  w < length (watched_by S L)  watched_by (keep_watch L j w S) L ! w = watched_by S L ! w
  by (cases S) (auto simp: keep_watch_def)

lemma watched_by_keep_watch_eq[twl_st_wl, simp]:
  j < length (watched_by S L)  watched_by (keep_watch L j w S) L ! j = watched_by S L ! w
  by (cases S) (auto simp: keep_watch_def)


definition update_clause_wl :: 'v literal  'v literal  nat  bool  nat  nat  nat  nat  'v twl_st_wl 
    (nat × nat × 'v twl_st_wl) nres where
  update_clause_wl = (λ(L::'v literal) other_watched C b j w i f (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W). do {
     ASSERT(C ∈# dom_m N  j  w  w < length (W L) 
        correct_watching_except (Suc j) (Suc w) L (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
     ASSERT(L ∈# all_lits_st (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
     ASSERT(other_watched ∈# all_lits_st (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
     K'  mop_clauses_at N C f;
     ASSERT(K' ∈# all_lits_st (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)  L  K');
     N'  mop_clauses_swap N C i f;
     RETURN (j, w+1, (M, N', D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W(K' := W K' @ [(C, L, b)])))
  })


definition update_blit_wl :: 'v literal  nat  bool  nat  nat  'v literal  'v twl_st_wl 
    (nat × nat × 'v twl_st_wl) nres where
  update_blit_wl = (λ(L::'v literal) C b j w K (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W). do {
     ASSERT(L ∈# all_lits_st (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
     ASSERT(K ∈# all_lits_st (M, N,  D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
     ASSERT(j  w);
     ASSERT(w < length (W L));
     ASSERT(C ∈# dom_m N);
     RETURN (j+1, w+1, (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W(L := (W L)[j:=(C, K, b)])))
  })


definition unit_prop_body_wl_find_unwatched_inv where
unit_prop_body_wl_find_unwatched_inv f C S  True

abbreviation remaining_nondom_wl where
remaining_nondom_wl w L S 
  (if get_conflict_wl S = None
          then size (filter_mset (λ(i, _). i ∉# dom_m (get_clauses_wl S)) (mset (drop w (watched_by S L)))) else 0)

definition unit_propagation_inner_loop_wl_loop_inv where
  unit_propagation_inner_loop_wl_loop_inv L = (λ(j, w, S).
    (S'. (S, S')  state_wl_l (Some (L, w))  j w 
       unit_propagation_inner_loop_l_inv L (S', remaining_nondom_wl w L S) 
      correct_watching_except j w L S  w  length (watched_by S L)))

lemma correct_watching_except_correct_watching_except_Suc_Suc_keep_watch:
  assumes
    j_w: j  w and
    w_le: w < length (watched_by S L) and
    corr: correct_watching_except j w L S
  shows correct_watching_except (Suc j) (Suc w) L (keep_watch L j w S)
proof -
  obtain M N D NE UE NEk UEk NS US N0 U0 Q W where S: S = (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) by (cases S)
  let ?ℒ = all_lits_st S
  have
    Hneq: La. La∈#?ℒ 
        (La  L 
	  distinct_watched (W La) 
         ((i, K, b)∈#mset (W La). i ∈# dom_m N  K  set (N  i)  K  La 
             correctly_marked_as_binary N (i, K, b)) 
         ((i, K, b)∈#mset (W La). b  i ∈# dom_m N) 
         {#i ∈# fst `# mset (W La). i ∈# dom_m N#} = clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})) and
    Heq: La. La∈#?ℒ 
        (La = L 
	 distinct_watched (take j (W La) @ drop w (W La)) 
         ((i, K, b)∈#mset (take j (W La) @ drop w (W La)). i ∈# dom_m N  K  set (N  i) 
            K  La  correctly_marked_as_binary N (i, K, b)) 
         ((i, K, b)∈#mset (take j (W La) @ drop w (W La)). b  i ∈# dom_m N) 
         {#i ∈# fst `# mset (take j (W La) @ drop w (W La)). i ∈# dom_m N#} =
         clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))
    using corr unfolding S correct_watching_except.simps
    by fast+

  have eq: mset (take (Suc j) ((W(L := (W L)[j := W L ! w])) La) @ drop (Suc w) ((W(L := (W L)[j := W L ! w])) La)) =
     mset (take j (W La) @ drop w (W La)) if [simp]: La = L for La
    using w_le j_w
    by (auto simp: S take_Suc_conv_app_nth Cons_nth_drop_Suc[symmetric]
        list_update_append)

  have case x of (i, K, b)  i ∈# dom_m N  K  set (N  i)  K  La 
           correctly_marked_as_binary N (i, K, b)
    if
      La ∈# ?ℒ and
      La = L and
      x ∈# mset (take (Suc j) ((W(L := (W L)[j := W L ! w])) La) @
                 drop (Suc w) ((W(L := (W L)[j := W L ! w])) La))
    for La :: 'a literal and x :: nat × 'a literal × bool
    using that Heq[of L]
    apply (subst (asm) eq)
    by (simp_all add: eq)
  moreover have case x of (i, K, b)  b  i ∈# dom_m N
    if
      La ∈# ?ℒ and
      La = L and
      x ∈# mset (take (Suc j) ((W(L := (W L)[j := W L ! w])) La) @
                 drop (Suc w) ((W(L := (W L)[j := W L ! w])) La))
    for La :: 'a literal and x :: nat × 'a literal × bool
    using that Heq[of L]
    by (subst (asm) eq) blast+
  moreover have {#i ∈# fst `#
              mset
               (take (Suc j) ((W(L := (W L)[j := W L ! w])) La) @
                drop (Suc w) ((W(L := (W L)[j := W L ! w])) La)).
       i ∈# dom_m N#} =
      clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
    if
      La ∈# ?ℒ and
      La = L
    for La :: 'a literal
    using that Heq[of L]
    by (subst eq) simp_all
  moreover have case x of (i, K, b)  i ∈# dom_m N  K  set (N  i)  K  La 
        correctly_marked_as_binary N (i, K, b)
    if
      La ∈# ?ℒ and
      La  L and
      x ∈# mset ((W(L := (W L)[j := W L ! w])) La)
    for La :: 'a literal and x :: nat × 'a literal × bool
    using that Hneq[of La]
    by simp
  moreover have case x of (i, K, b)  b  i ∈# dom_m N
    if
      La ∈# ?ℒ and
      La  L and
      x ∈# mset ((W(L := (W L)[j := W L ! w])) La)
    for La :: 'a literal and x :: nat × 'a literal × bool
    using that Hneq[of La]
    by auto
  moreover have {#i ∈# fst `# mset ((W(L := (W L)[j := W L ! w])) La). i ∈# dom_m N#} =
      clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
    if
      La ∈# ?ℒ and
      La  L
    for La :: 'a literal
    using that Hneq[of La]
    by simp
  moreover have distinct_watched ((W(L := (W L)[j := W L ! w])) La)
    if
      La ∈# ?ℒ and
      La  L
    for La :: 'a literal
    using that Hneq[of La]
    by simp
  moreover have distinct_watched (take (Suc j) ((W(L := (W L)[j := W L ! w])) La) @
                drop (Suc w) ((W(L := (W L)[j := W L ! w])) La))
    if
      La ∈# ?ℒ and
      La = L
    for La :: 'a literal
    using that Heq[of La]
    apply (subst distinct_mset_mset_distinct[symmetric])
    apply (subst mset_map)
    apply (subst eq)
    apply (simp add: that)
    apply (subst mset_map[symmetric])
    apply (subst distinct_mset_mset_distinct)
    apply simp
    done
  ultimately show ?thesis
    unfolding S keep_watch_def prod.simps correct_watching_except.simps all_lits_st_simps
    by meson
qed


lemma correct_watching_except_update_blit:
  assumes
    corr: correct_watching_except i j L (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g(L := (g L)[j' := (x1, C, b')])) and
    C': C' ∈# all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
      C'  set (b  x1)
      C'  L and
    corr_watched: correctly_marked_as_binary b (x1, C', b')
  shows correct_watching_except i j L (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g(L := (g L)[j' := (x1, C', b')]))
proof -
  let ?ℒ = all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
  have
    Hdisteq: La i' K' b''. La∈#?ℒ 
        (La = L 
	 distinct_watched (take i ((g(L := (g L)[j' := (x1, C, b')])) La) @ drop j ((g(L := (g L)[j' := (x1, C, b')])) La))) and
    Heq: La i' K' b''. La∈#?ℒ 
        (La = L 
         (((i', K', b'')∈#mset (take i ((g(L := (g L)[j' := (x1, C, b')])) La) @ drop j ((g(L := (g L)[j' := (x1, C, b')])) La)) 
             i' ∈# dom_m b  K'  set (b  i')  K'  La  correctly_marked_as_binary b (i', K', b'')) 
          ((i', K', b'')∈#mset (take i ((g(L := (g L)[j' := (x1, C, b')])) La) @ drop j ((g(L := (g L)[j' := (x1, C, b')])) La)) 
              b''  i' ∈# dom_m b)) 
         {#i ∈# fst `# mset (take i ((g(L := (g L)[j' := (x1, C, b')])) La) @ drop j ((g(L := (g L)[j' := (x1, C, b')])) La)).
          i ∈# dom_m b#} =
         clause_to_update La (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, {#}, {#})) and
    Hdistneq: La i' K' b''. La∈#?ℒ 
        (La  L  distinct_watched (((g(L := (g L)[j' := (x1, C, b')])) La))) and
    Hneq: La i K b''. La∈#?ℒ  La  L 
         distinct_watched (((g(L := (g L)[j' := (x1, C, b')])) La)) 
         ((i, K, b'')∈#mset ((g(L := (g L)[j' := (x1, C, b')])) La) i ∈# dom_m b 
            K  set (b  i)  K  La  correctly_marked_as_binary b (i, K, b'')) 
         ((i, K, b'')∈#mset ((g(L := (g L)[j' := (x1, C, b')])) La) b''  i ∈# dom_m b) 
         {#i ∈# fst `# mset ((g(L := (g L)[j' := (x1, C, b')])) La). i ∈# dom_m b#} =
            clause_to_update La (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, {#}, {#})
    using corr unfolding correct_watching_except.simps all_lits_st_simps
    by fast+
  define g' where g' = g(L := (g L)[j' := (x1, C, b')])
  have g_g': g(L := (g L)[j' := (x1, C', b')]) = g'(L := (g' L)[j' := (x1, C', b')])
    unfolding g'_def by auto

  have H2: fst `# mset ((g'(L := (g' L)[j' := (x1, C', b')])) La) = fst `# mset (g' La) for La
    unfolding g'_def
    by (auto simp flip: mset_map simp: map_update)
  have H3: fst `#
                 mset
                  (take i ((g'(L := (g' L)[j' := (x1, C', b')])) La) @
                   drop j ((g'(L := (g' L)[j' := (x1, C', b')])) La)) =
      fst `#
                 mset
                  (take i (g' La) @
                   drop j (g' La)) for La
    unfolding g'_def
    by (auto simp flip: mset_map drop_map simp: map_update)
  have [simp]:
    fst `# mset ((take i (g' L))[j' := (x1, C', b')]) = fst `# mset (take i (g' L))
    fst `# mset ((drop j ((g' L)[j' := (x1, C', b')]))) = fst `# mset (drop j (g' L))
    ¬j' < j  fst `# mset ((drop j (g' L))[j' - j := (x1, C', b')]) = fst `# mset (drop j (g' L))
    unfolding g'_def
      apply (auto simp flip: mset_map drop_map simp: map_update drop_update_swap; fail)
     apply (auto simp flip: mset_map drop_map simp: map_update drop_update_swap; fail)
    apply (auto simp flip: mset_map drop_map simp: map_update drop_update_swap; fail)
    done
  have j' < length (g' L)  j' < i  (x1, C, b')  set ((take i (g L))[j' := (x1, C, b')])
    using nth_mem[of j' (take i (g L))[j' := (x1, C, b')]] unfolding g'_def
    by auto
  then have H: L ∈#?ℒ  j' < length (g' L) 
       j' < i  b'  x1 ∈# dom_m b
    using C' Heq[of L x1 C b']
    by (cases j' < j) (simp, auto)
  have ¬ j' < j  j' - j < length (g' L) - j 
     (x1, C, b')  set (drop j ((g L)[j' := (x1, C, b')]))
    using nth_mem[of j'-j drop j ((g L)[j' := (x1, C, b')])] unfolding g'_def
    by auto
  then have H': L ∈#?ℒ  ¬ j' < j 
       j' - j < length (g' L) - j  b'  x1 ∈# dom_m b
    using C' Heq[of L x1 C b'] unfolding g'_def
    by auto

  have dist: La∈#?ℒ 
        La = L 
	 distinct_watched (take i ((g'(L := (g' L)[j' := (x1, C', b')])) La) @ drop j ((g'(L := (g' L)[j' := (x1, C', b')])) La))
    for La
    using Hdisteq[of L] unfolding g_g'[symmetric]
    by (cases j' < j)
       (auto simp: map_update drop_update_swap)
  have ℒ_alt: ?ℒ = all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g')
    unfolding g'_def by simp
  have La∈#?ℒ 
        La = L 
	 distinct_watched (take i ((g'(L := (g' L)[j' := (x1, C', b')])) La) @ drop j ((g'(L := (g' L)[j' := (x1, C', b')])) La)) 
         ((i', K, b'')∈#mset (take i ((g'(L := (g' L)[j' := (x1, C', b')])) La) @ drop j ((g'(L := (g' L)[j' := (x1, C', b')])) La)) 
             i' ∈# dom_m b  K  set (b  i')  K  La  correctly_marked_as_binary b (i', K, b'')) 
          ((i', K, b'')∈#mset (take i ((g'(L := (g' L)[j' := (x1, C', b')])) La) @ drop j ((g'(L := (g' L)[j' := (x1, C', b')])) La)) 
            b''  i' ∈# dom_m b) 
         {#i ∈# fst `# mset (take i ((g'(L := (g' L)[j' := (x1, C', b')])) La) @ drop j ((g'(L := (g' L)[j' := (x1, C', b')])) La)).
          i ∈# dom_m b#} =
         clause_to_update La (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, {#}, {#}) for La i' K b''
    using C' Heq[of La i' K] Heq[of La i' K b'] H H' dist[of La] corr_watched unfolding g_g' g'_def[symmetric]
    apply (auto elim!: in_set_upd_cases simp: drop_update_swap simp del: distinct_append)
    apply (cases j' < j; auto elim!: in_set_upd_cases simp: drop_update_swap simp del: distinct_append)+
    done
  moreover have La∈#?ℒ 
       (La  L 
        distinct_watched ((g'(L := (g' L)[j' := (x1, C', b')])) La) 
        ((i, K, ba)∈#mset ((g'(L := (g' L)[j' := (x1, C', b')])) La).
            i ∈# dom_m b 
            K  set (b  i) 
            K  La  correctly_marked_as_binary b (i, K, ba)) 
        ((i, K, ba)∈#mset ((g'(L := (g' L)[j' := (x1, C', b')])) La).
            ba  i ∈# dom_m b) 
        {#i ∈# fst `# mset ((g'(L := (g' L)[j' := (x1, C', b')])) La).
         i ∈# dom_m b#} =
        clause_to_update La (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, {#}, {#}))
     for La
    using Hneq Hdistneq
    unfolding correct_watching_except.simps g_g' g'_def[symmetric]
    by auto
  ultimately show ?thesis
    unfolding correct_watching_except.simps g_g' g'_def[symmetric]
    unfolding H2 H3 all_lits_st_simps ℒ_alt
    by blast
qed


lemma correct_watching_except_correct_watching_except_Suc_notin:
  assumes
    fst (watched_by S L ! w) ∉# dom_m (get_clauses_wl S) and
    j_w: j  w and
    w_le: w < length (watched_by S L) and
    corr: correct_watching_except j w L S
  shows correct_watching_except j (Suc w) L (keep_watch L j w S)
proof -
  obtain M N D NE UE NEk UEk NS US N0 U0 Q W where
    S: S = (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) by (cases S)
  let ?ℒ = all_lits_st S
  have [simp]: fst (W L ! w) ∉# dom_m N
    using assms unfolding S by auto
  have
    Hneq: La. La∈#?ℒ 
        (La  L 
	 distinct_watched (W La) 
         (((i, K, b)∈#mset (W La). i ∈# dom_m N  K  set (N  i)  K  La 
             correctly_marked_as_binary N (i, K, b)) 
          ((i, K, b)∈#mset (W La). b  i ∈# dom_m N)) 
          {#i ∈# fst `# mset (W La). i ∈# dom_m N#} = clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})) and
    Heq: La. La∈#?ℒ 
        (La = L 
	 distinct_watched (take j (W La) @ drop w (W La)) 
         (((i, K, b)∈#mset (take j (W La) @ drop w (W La)). i ∈# dom_m N 
              K  set (N  i)  K  La  correctly_marked_as_binary N (i, K, b)) 
          ((i, K, b)∈#mset (take j (W La) @ drop w (W La)). b  i ∈# dom_m N) 
         {#i ∈# fst `# mset (take j (W La) @ drop w (W La)). i ∈# dom_m N#} =
         clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})))
    using corr unfolding S correct_watching_except.simps
    by fast+

  have eq: mset (take j ((W(L := (W L)[j := W L ! w])) La) @ drop (Suc w) ((W(L := (W L)[j := W L ! w])) La)) =
    remove1_mset (W L ! w) (mset (take j (W La) @ drop w (W La))) if [simp]: La = L for La
    using w_le j_w
    by (auto simp: S take_Suc_conv_app_nth Cons_nth_drop_Suc[symmetric]
        list_update_append)

  have case x of (i, K, b)  i ∈# dom_m N  K  set (N  i)  K  La 
       correctly_marked_as_binary N (i, K, b)
    if
      La ∈# ?ℒ and
      La = L and
      x ∈# mset (take j ((W(L := (W L)[j := W L ! w])) La) @
                 drop (Suc w) ((W(L := (W L)[j := W L ! w])) La))
    for La :: 'a literal and x :: nat × 'a literal × bool
    using that Heq[of L] w_le j_w
    by (subst (asm) eq) (auto dest!: in_diffD)
  moreover have case x of (i, K, b)  b  i ∈# dom_m N
    if
      La ∈# ?ℒ and
      La = L and
      x ∈# mset (take j ((W(L := (W L)[j := W L ! w])) La) @
                 drop (Suc w) ((W(L := (W L)[j := W L ! w])) La))
    for La :: 'a literal and x :: nat × 'a literal × bool
    using that Heq[of L] w_le j_w unfolding S
    by (subst (asm) eq) (fastforce simp: S dest!: in_diffD)+
  moreover have {#i ∈# fst `#
              mset
               (take j ((W(L := (W L)[j := W L ! w])) La) @
                drop (Suc w) ((W(L := (W L)[j := W L ! w])) La)).
       i ∈# dom_m N#} =
      clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
    if
      La ∈# ?ℒ and
      La = L
    for La :: 'a literal
    using that Heq[of L] w_le j_w
    by (subst eq) (auto dest!: in_diffD simp: image_mset_remove1_mset_if)
  moreover have case x of (i, K, b)  i ∈# dom_m N  K  set (N  i)  K  La 
      correctly_marked_as_binary N (i, K, b)
    if
      La ∈# ?ℒ and
      La  L and
      x ∈# mset ((W(L := (W L)[j := W L ! w])) La)
    for La :: 'a literal and x :: nat × 'a literal × bool
    using that Hneq[of La]
    by simp
  moreover have case x of (i, K, b)  b  i ∈# dom_m N
    if
      La ∈# ?ℒ and
      La  L and
      x ∈# mset ((W(L := (W L)[j := W L ! w])) La)
    for La :: 'a literal and x :: nat × 'a literal × bool
    using that Hneq[of La]
    by auto
  moreover have {#i ∈# fst `# mset ((W(L := (W L)[j := W L ! w])) La). i ∈# dom_m N#} =
      clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
    if
      La ∈# ?ℒ and
      La  L
    for La :: 'a literal
    using that Hneq[of La]
    by simp
  moreover have distinct_watched ((W(L := (W L)[j := W L ! w])) La)
    if
      La ∈# ?ℒ and
      La  L
    for La :: 'a literal
    using that Hneq[of La]
    by simp
  moreover have distinct_watched (take j ((W(L := (W L)[j := W L ! w])) La) @
                drop (Suc w) ((W(L := (W L)[j := W L ! w])) La))
    if
      La ∈# ?ℒ and
      La = L
    for La :: 'a literal
    using that Heq[of L] w_le j_w apply -
    apply (subst distinct_mset_mset_distinct[symmetric])
    apply (subst mset_map)
    apply (subst eq)
    apply (solves simp)
    apply (subst (asm) distinct_mset_mset_distinct[symmetric])
    apply (subst (asm) mset_map)
    apply (rule distinct_mset_mono[of _ {#i. (i, j, k) ∈# mset (take j (W L) @ drop w (W L))#}])
    by (auto simp: image_mset_remove1_mset_if split: if_splits)
  ultimately show ?thesis
    unfolding S keep_watch_def prod.simps correct_watching_except.simps
    by simp
qed

lemma correct_watching_except_correct_watching_except_update_clause:
  assumes
    corr: correct_watching_except (Suc j) (Suc w) L
       (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W(L := (W L)[j := W L ! w])) and
    j_w: j  w and
    w_le: w < length (W L) and
    L': L' ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
      L'  set (N  x1)and
    L_L: L ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) and
    L: L  N  x1 ! xa and
    dom: x1 ∈# dom_m N and
    i_xa: i < length (N  x1) xa < length (N  x1) and
    [simp]: W L ! w = (x1, x2, b) and
    N_i: N  x1 ! i = L N  x1 ! (1 -i)  LN  x1 ! xa  L and
    N_xa: N  x1 ! xa  N  x1 ! i N  x1 ! xa  N  x1 ! (Suc 0 - i)and
    i_2: i < 2 and xa  2 and
    L_neq: L'  N  x1 ! xa ―‹The new blocking literal is not the new watched literal.
  shows correct_watching_except j (Suc w) L
          (M, N(x1  swap (N  x1) i xa), D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W
           (L := (W L)[j := (x1, x2, b)],
            N  x1 ! xa := W (N  x1 ! xa) @ [(x1, L', b)]))
proof -
  define W' where W'  W(L := (W L)[j := W L ! w])
  have length (N  x1) > 2
    using i_2 i_xa assms
    by (auto simp: correctly_marked_as_binary.simps)
  let ?ℒ = all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
  have : ?ℒ = all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W')
    using assms unfolding W'_def by simp
  have
    Heq: La i K b. La∈# ?ℒ 
          La = L 
	   distinct_watched (take (Suc j) (W' La) @ drop (Suc w) (W' La)) 
           ((i, K, b)∈#mset (take (Suc j) (W' La) @ drop (Suc w) (W' La)) 
               i ∈# dom_m N  K  set (N  i)  K  La  correctly_marked_as_binary N (i, K, b)) 
           ((i, K, b)∈#mset (take (Suc j) (W' La) @ drop (Suc w) (W' La)) 
               b  i ∈# dom_m N) 
           {#i ∈# fst `#
                   mset
                    (take (Suc j) (W' La) @ drop (Suc w) (W' La)).
            i ∈# dom_m N#} =
           clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}) and
    Hneq: La i K b. La∈#?ℒ 
          La  L 
	   distinct_watched (W' La) 
           ((i, K, b)∈#mset (W' La)  i ∈# dom_m N  K  set (N  i)  K  La 
               correctly_marked_as_binary N (i, K, b)) 
           ((i, K, b)∈#mset (W' La)  b  i ∈# dom_m N) 
           {#i ∈# fst `# mset (W' La). i ∈# dom_m N#} =
           clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}) and
    Hneq2: La. La∈#?ℒ 
          (La  L 
	   distinct_watched (W' La) 
           {#i ∈# fst `# mset (W' La). i ∈# dom_m N#} =
           clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))
    using corr unfolding correct_watching_except.simps W'_def[symmetric] 
    by fast+
  have H1: mset `# ran_mf (N(x1  swap (N  x1) i xa)) = mset `# ran_mf N
    using dom i_xa distinct_mset_dom[of N]
    by (auto simp: ran_m_def dest!: multi_member_split intro!: image_mset_cong2)
  have W_W': W
      (L := (W L)[j := (x1, x2, b)], N  x1 ! xa := W (N  x1 ! xa) @ [(x1, L', b)]) =
     W'(N  x1 ! xa := W (N  x1 ! xa) @ [(x1, L', b)])
    unfolding W'_def
    by auto
  have W_W2: W (N  x1 ! xa) = W' (N  x1 ! xa)
    using L unfolding W'_def by auto
  have H2: set (swap (N  x1) i xa) =  set (N  x1)
    using i_xa by auto
  have [simp]:
    set (fst (the (if x1 = ia then Some (swap (N  x1) i xa, irred N x1) else fmlookup N ia))) =
     set (fst (the (fmlookup N ia))) for ia
    using H2
    by auto
  have H3: i = x1  i ∈# remove1_mset x1 (dom_m N)  i ∈# dom_m N for i
    using dom by (auto dest: multi_member_split)
  have set_N_swap_x1: set (watched_l (swap (N  x1) i xa)) = {N  x1 ! (1 -i), N  x1 ! xa}
    using i_2 i_xa xa  2 N_i
    by (cases N  x1; cases tl (N  x1); cases i; cases i-1; cases xa)
      (auto simp: swap_def split: nat.splits)
  have set_N_x1: set (watched_l (N  x1)) = {N  x1 ! (1 -i), N  x1 ! i}
    using i_2 i_xa xa  2 N_i
    by (cases i) (auto simp: swap_def take_2_if)

  have La_in_notin_swap:  La  set (watched_l (N  x1)) 
       La  set (watched_l (swap (N  x1) i xa))  La = L for La
    using i_2 i_xa xa  2 N_i
    by (auto simp: set_N_x1 set_N_swap_x1)

  have L_notin_swap: L  set (watched_l (swap (N  x1) i xa))
    using i_2 i_xa xa  2 N_i
    by (auto simp: set_N_x1 set_N_swap_x1)
  have N_xa_in_swap: N  x1 ! xa  set (watched_l (swap (N  x1) i xa))
    using i_2 i_xa xa  2 N_i
    by (auto simp: set_N_x1 set_N_swap_x1)
  have H4: (i = x1  K  set (N  x1)  K  La)  (i ∈# remove1_mset x1 (dom_m N)  K  set (N  i)  K  La) 
   (i ∈# dom_m N  K  set (N  i)  K  La) for i P K La
    using dom by (auto dest: multi_member_split)
  have [simp]: x1 ∉# Ab 
       {#C ∈# Ab.
        (x1 = C  Q C) 
        (x1  C  R C)#} =
     {#C ∈# Ab. R C#} for Ab Q R
    by (auto intro: filter_mset_cong)
  have bin:
    correctly_marked_as_binary N (x1, x2, b)
    using Heq[of L fst (W L ! w) fst (snd (W L ! w )) snd (snd (W L ! w))] j_w w_le dom L'
       L_L unfolding all_lits_def
    by (auto simp: take_Suc_conv_app_nth W'_def list_update_append ac_simps)
  have x1_new: x1  fst ` set (W (N  x1 ! xa))
  proof (rule ccontr)
    assume H: "¬ ?thesis"
    have N  x1 ! xa ∈# ?ℒ
      using dom i_xa by auto
    then have {#i ∈# fst `# mset (W (N  x1 ! xa)). i ∈# dom_m N#} =
        clause_to_update (N  x1 ! xa) (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
      using Hneq[of N  x1 ! xa] L unfolding W'_def 
      by simp
    then have x1 ∈# clause_to_update (N  x1 ! xa) (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
      using H dom by (metis (no_types, lifting) mem_Collect_eq set_image_mset
        set_mset_filter set_mset_mset)
    then show False
      using N_xa i_2 i_xa
      by (cases i; cases N  x1 ! xa)
        (auto simp: clause_to_update_def take_2_if split: if_splits)
  qed

  let ?N =  N(x1  swap (N  x1) i xa)
  have L ∈# ?ℒ  La = L 
       x  set (take j (W L))  x  set (drop (Suc w) (W L)) 
       case x of (i, K, b)  i ∈# dom_m N  K  set (N  i)  K  L 
           correctly_marked_as_binary ?N (i, K, b) for La x
    using Heq[of L fst x fst (snd x) snd (snd x)] j_w w_le unfolding [symmetric]
    by (clarsimp simp: take_Suc_conv_app_nth W'_def list_update_append correctly_marked_as_binary.simps
      split: if_splits)
  moreover have L ∈# ?ℒ  La = L 
       x  set (take j (W L))  x  set (drop (Suc w) (W L)) 
       case x of (i, K, b) b  i ∈# dom_m N for La x
    using Heq[of L fst x fst (snd x) snd (snd x)] j_w w_le
    by (auto simp: take_Suc_conv_app_nth W'_def list_update_append correctly_marked_as_binary.simps
      split: if_splits)
  moreover  have L ∈# ?ℒ 
          La = L 
	  distinct_watched (take j (W L) @ drop (Suc w) (W L)) 
          {#i ∈# fst `# mset (take j (W L)). i ∈# dom_m N#} + {#i ∈# fst `# mset (drop (Suc w) (W L)). i ∈# dom_m N#} =
          clause_to_update L (M, N(x1  swap (N  x1) i xa), D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}) for La
    using Heq[of L x1 x2 b] j_w w_le dom L_notin_swap N_xa_in_swap distinct_mset_dom[of N]
    i_xa i_2 assms(12)
    by (auto simp: take_Suc_conv_app_nth W'_def list_update_append set_N_x1 assms(11)
        clause_to_update_def dest!: multi_member_split split: if_splits
        intro: filter_mset_cong2)

  moreover have La ∈# ?ℒ 
       La  L 
       x  set (if La = N  x1 ! xa
                 then W' (N  x1 ! xa) @ [(x1, L', b)]
                 else (W(L := (W L)[j := (x1, x2, b)])) La) 
       case x of
       (i, K, b)  i ∈# dom_m ?N  K  set (?N  i)  K  La  correctly_marked_as_binary ?N (i, K, b) for La x
    using Hneq[of La fst x fst (snd x) snd (snd x)] j_w w_le L' L_neq bin dom unfolding 
    by (fastforce simp: take_Suc_conv_app_nth W'_def list_update_append
      correctly_marked_as_binary.simps split: if_splits)
  moreover have La ∈# ?ℒ 
       La  L 
       x  set (if La = N  x1 ! xa
                 then W' (N  x1 ! xa) @ [(x1, L', b)]
                 else (W(L := (W L)[j := (x1, x2, b)])) La) 
       case x of (i, K, b)  b  i ∈# dom_m N for La x
    using Hneq[of La fst x fst (snd x) snd (snd x)] j_w w_le L' L_neq length (N  x1) > 2
      dom
    by (auto simp: take_Suc_conv_app_nth W'_def list_update_append correctly_marked_as_binary.simps split: if_splits)
  moreover have La ∈# ?ℒ 
       La  L  distinct_watched  ((W
           (L := (W L)[j := (x1, x2, b)],
            N  x1 ! xa := W (N  x1 ! xa) @ [(x1, L', b)])) La) for La x
    using Hneq[of La] j_w w_le L' L_neq length (N  x1) > 2
      dom x1_new
    by (auto simp: take_Suc_conv_app_nth W'_def list_update_append correctly_marked_as_binary.simps split: if_splits)
  moreover {
    have N  x1 ! xa  set (watched_l (N  x1))
      using N_xa
      by (auto simp: set_N_x1 set_N_swap_x1)

    then have  Ab Ac La.
       ?ℒ = add_mset L' (add_mset (N  x1 ! xa) Ac) 
       dom_m N = add_mset x1 Ab 
       N  x1 ! xa  L 
       {#i ∈# fst `# mset (W (N  x1 ! xa)). i = x1  i ∈# Ab#} =
         {#C ∈# Ab. N  x1 ! xa  set (watched_l (N  C))#} 
      using Hneq2[of N  x1 ! xa] L_neq unfolding W_W' W_W2 
      by (auto simp: clause_to_update_def split: if_splits)
    from this[symmetric] have La ∈# ?ℒ 
          La  L 
	  distinct_watched (W' La) 
          (x1 ∈# dom_m N 
           (La = N  x1 ! xa 
            add_mset x1 {#i ∈# fst `# mset (W' (N  x1 ! xa)). i ∈# dom_m N#} =
            clause_to_update (N  x1 ! xa) (M, N(x1  swap (N  x1) i xa), D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})) 
           (La  N  x1 ! xa 
            {#i ∈# fst `# mset (W La). i ∈# dom_m N#} =
            clause_to_update La (M, N(x1  swap (N  x1) i xa), D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))) 
          (x1 ∉# dom_m N 
           (La = N  x1 ! xa 
            {#i ∈# fst `# mset (W' (N  x1 ! xa)). i ∈# dom_m N#} =
            clause_to_update (N  x1 ! xa) (M, N(x1  swap (N  x1) i xa), D, NS, US, NEk, UEk, NE, UE, N0, U0, {#}, {#})) 
           (La  N  x1 ! xa 
            {#i ∈# fst `# mset (W La). i ∈# dom_m N#} =
            clause_to_update La (M, N(x1  swap (N  x1) i xa), D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))) for La
      using Hneq2[of La] j_w w_le L' dom distinct_mset_dom[of N] L_notin_swap N_xa_in_swap L_neq
      apply (auto simp: take_Suc_conv_app_nth W'_def list_update_append clause_to_update_def
        add_mset_eq_add_mset set_N_x1 set_N_swap_x1 assms(11) N_i all_lits_def ac_simps
        eq_commute[of x1]
        eq_commute[of {#y ∈# _. y  _  N  x1 ! xa  set (watched_l (N  y))#}]
        dest!: multi_member_split La_in_notin_swap
        split: if_splits
        intro: image_mset_cong2 intro: filter_mset_cong2)
       by (smt (verit, ccfv_SIG) filter_mset_cong2) 
  }
  ultimately show ?thesis
    using L j_w dom i_xa
    unfolding correct_watching_except.simps H1  W'_def[symmetric] W_W' H2 W_W2 H4 H3 [symmetric]
    by (intro conjI impI ballI)
     (simp_all add: L' W_W' W_W2 H3 H4 drop_map)
qed

definition unit_propagation_inner_loop_wl_loop_pre where
  unit_propagation_inner_loop_wl_loop_pre L = (λ(j, w, S).
     w < length (watched_by S L)  j  w  blits_in_ℒin S 
     unit_propagation_inner_loop_wl_loop_inv L (j, w, S))

definition mop_watched_by :: 'v twl_st_wl  'v literal  nat  _ nres where
mop_watched_by S L w = do {
  ASSERT(w < length (watched_by S L));
  RETURN ((watched_by S L) ! w)
}

definition mop_polarity_wl :: 'v twl_st_wl  'v literal  bool option nres where
  mop_polarity_wl S L = do {
    ASSERT(L ∈# all_lits_st S);
    ASSERT(no_dup (get_trail_wl S));
    RETURN(polarity (get_trail_wl S) L)
  }

lemma mop_polarity_wl_mop_polarity_l:
  (uncurry mop_polarity_wl, uncurry mop_polarity_l)  state_wl_l b ×r Id f Idnres_rel
  unfolding mop_polarity_l_def mop_polarity_wl_def uncurry_def
  by (cases b)
    (intro frefI nres_relI; refine_rcg; auto simp: state_wl_l_def all_lits_def all_lits_of_mm_union
      all_lits_st_def)+

definition is_nondeleted_clause_pre :: nat  'v literal  'v twl_st_wl  bool where
is_nondeleted_clause_pre C L S  C  fst ` set (watched_by S L)  L ∈# all_lits_st S
definition other_watched_wl :: 'v twl_st_wl  'v literal  nat  nat  'v literal nres where
  other_watched_wl S L C i = do {
    ASSERT(get_clauses_wl S  C ! i = L  i < length (get_clauses_wl S  C)  i < 2 
      C ∈# dom_m (get_clauses_wl S)  1-i < length (get_clauses_wl S  C));
    mop_clauses_at (get_clauses_wl S) C (1 - i)
  }

lemma other_watched_wl_other_watched_l:
  (uncurry3 other_watched_wl, uncurry3 other_watched_l)  state_wl_l b ×f Id ×f Id ×f Id f Idnres_rel
  unfolding other_watched_wl_def other_watched_l_def uncurry_def
  by (intro frefI nres_relI)
   (refine_rcg, auto simp: state_wl_l_def)

lemma other_watched_wl_other_watched_l_spec_itself:
  ((S, L, C, i), (S', L', C', i'))  state_wl_l b ×r (Id :: ('v literal × _) set) ×r nat_rel ×r
     nat_rel 
     other_watched_wl S L C i 
       {(L, L'). L = L'  L = get_clauses_wl S  C ! (1 - i)}
        (other_watched_l S' L' C' i')
  using twl_st_wl(2)[of S S' b]
  unfolding other_watched_wl_def other_watched_l_def
    mop_clauses_at_def
  by refine_vcg clarsimp_all

text It was too hard to align the program unto a refinable form directly.
definition unit_propagation_inner_loop_body_wl_int :: 'v literal  nat  nat  'v twl_st_wl 
    (nat × nat × 'v twl_st_wl) nres where
  unit_propagation_inner_loop_body_wl_int L j w S = do {
      ASSERT(unit_propagation_inner_loop_wl_loop_pre L (j, w, S));
      (C, K, b)  mop_watched_by_at S L w;
      S  mop_keep_watch L j w S;
      ASSERT(is_nondeleted_clause_pre C L S);
      val_K  mop_polarity_wl S K;
      if val_K = Some True
      then RETURN (j+1, w+1, S)
      else do { ―‹Now the costly operations:
        if C ∉# dom_m (get_clauses_wl S)
        then RETURN (j, w+1, S)
        else do {
          ASSERT(unit_prop_body_wl_inv S j w L);
          i  pos_of_watched (get_clauses_wl S) C L;
          ASSERT(i  1);
          L'  other_watched_wl S L C i;
          val_L'  mop_polarity_wl S L';
          if val_L' = Some True
          then update_blit_wl L C b j w L' S
          else do {
            f  find_unwatched_l (get_trail_wl S) (get_clauses_wl S) C;
            ASSERT (unit_prop_body_wl_find_unwatched_inv f C S);
            case f of
              None  do {
                if val_L' = Some False
                then do {S  set_conflict_wl C S; RETURN (j+1, w+1, S)}
                else do {S  propagate_lit_wl_general L' C i S; RETURN (j+1, w+1, S)}
              }
            | Some f  do {
                ASSERT(C ∈# dom_m (get_clauses_wl S)  f < length (get_clauses_wl S  C)  f  2);
                K  mop_clauses_at (get_clauses_wl S) C f;
                val_L'  mop_polarity_wl S K;
                if val_L' = Some True
                then update_blit_wl L C b j w K S
                else update_clause_wl L L' C b j w i f S
              }
          }
        }
      }
   }

definition propagate_proper_bin_case where
  propagate_proper_bin_case L L' S C 
    C ∈# dom_m (get_clauses_wl S)  length ((get_clauses_wl S)C) = 2 
    set (get_clauses_wl SC) = {L, L'}  L  L'

definition unit_propagation_inner_loop_body_wl :: 'v literal  nat  nat  'v twl_st_wl 
    (nat × nat × 'v twl_st_wl) nres where
  unit_propagation_inner_loop_body_wl L j w S = do {
      ASSERT(unit_propagation_inner_loop_wl_loop_pre L (j, w, S));
      (C, K, b)  mop_watched_by_at S L w;
      S  mop_keep_watch L j w S;
      ASSERT(is_nondeleted_clause_pre C L S);
      val_K  mop_polarity_wl S K;
      if val_K = Some True
      then RETURN (j+1, w+1, S)
      else do {
        if b then do {
           ASSERT(propagate_proper_bin_case L K S C);
           if val_K = Some False
           then do {S  set_conflict_wl C S;
             RETURN (j+1, w+1, S)}
           else do {
             S  propagate_lit_wl_bin K C S;
             RETURN (j+1, w+1, S)}
        }  ―‹Now the costly operations:
        else if C ∉# dom_m (get_clauses_wl S)
        then RETURN (j, w+1, S)
        else do {
          ASSERT(unit_prop_body_wl_inv S j w L);
          i  pos_of_watched (get_clauses_wl S) C L;
          ASSERT(i  1);
          L'  other_watched_wl S L C i;
          val_L'  mop_polarity_wl S L';
          if val_L' = Some True
          then update_blit_wl L C b j w L' S
          else do {
            f  find_unwatched_l (get_trail_wl S) (get_clauses_wl S) C;
            ASSERT (unit_prop_body_wl_find_unwatched_inv f C S);
            case f of
              None  do {
                if val_L' = Some False
                then do {S  set_conflict_wl C S;
                   RETURN (j+1, w+1, S)}
                else do {S  propagate_lit_wl L' C i S; RETURN (j+1, w+1, S)}
              }
            | Some f  do {
                ASSERT(C ∈# dom_m (get_clauses_wl S)  f < length (get_clauses_wl S  C)  f  2);
                K  mop_clauses_at (get_clauses_wl S) C f;
                val_L'  mop_polarity_wl S K;
                if val_L' = Some True
                then update_blit_wl L C b j w K S
                else update_clause_wl L L' C b j w i f S
              }
          }
        }
      }
   }

lemma [twl_st_wl]: get_clauses_wl (keep_watch L j w S) = get_clauses_wl S
  by (cases S) (auto simp: keep_watch_def)


lemma unit_propagation_inner_loop_body_wl_int_alt_def:
 unit_propagation_inner_loop_body_wl_int L j w S = do {
      ASSERT(unit_propagation_inner_loop_wl_loop_pre L (j, w, S));
      (C, K, b)  mop_watched_by_at S L w;
      let _ = (C ∉# dom_m (get_clauses_wl S));
      S  mop_keep_watch L j w S;
      ASSERT(is_nondeleted_clause_pre C L S);
      let b' = (C ∉# dom_m (get_clauses_wl S));
      if b' then do {
        let K = K;
        val_K  mop_polarity_wl S K;
        if val_K = Some True
        then RETURN (j+1, w+1, S)
        else ―‹Now the costly operations:
          RETURN (j, w+1, S)
      }
      else do {
        let S = S;
        K  SPEC((=) K);
        val_K  mop_polarity_wl S K;
        if val_K = Some True
        then RETURN (j+1, w+1, S)
        else do { ―‹Now the costly operations:
          ASSERT(unit_prop_body_wl_inv S j w L);
          i  pos_of_watched (get_clauses_wl S) C L;
          ASSERT(i  1);
          L'  other_watched_wl S L C i;
          val_L'  mop_polarity_wl S L';
          if val_L' = Some True
          then update_blit_wl L C b j w L' S
          else do {
            f  find_unwatched_l (get_trail_wl S) (get_clauses_wl S) C;
            ASSERT (unit_prop_body_wl_find_unwatched_inv f C S);
            case f of
              None  do {
                if val_L' = Some False
                then do {S  set_conflict_wl C S; RETURN (j+1, w+1, S)}
                else do {S  propagate_lit_wl_general L' C i S; RETURN (j+1, w+1, S)}
              }
            | Some f  do {
                ASSERT(C ∈# dom_m (get_clauses_wl S)  f < length (get_clauses_wl S  C)  f  2);
                K  mop_clauses_at (get_clauses_wl S) C f;
                val_L'  mop_polarity_wl S K;
                if val_L' = Some True
                then update_blit_wl L C b j w K S
                else update_clause_wl L L' C b j w i f S
             }
          }
        }
      }
   }
proof -

  text We first define an intermediate step where both then and else branches are the same.
  have E: unit_propagation_inner_loop_body_wl_int L j w S = do {
      ASSERT(unit_propagation_inner_loop_wl_loop_pre L (j, w, S));
      (C, K, b)  mop_watched_by_at S L w;
      S' mop_keep_watch L j w S;
      ASSERT(is_nondeleted_clause_pre C L S');
      let b' = C ∉# dom_m (get_clauses_wl S');
      if b' then do {
        let K = K;
        val_K  mop_polarity_wl S' K;
        if val_K = Some True
        then RETURN (j+1, w+1, S')
        else do { ―‹Now the costly operations:
          if b'
          then RETURN (j, w+1, S')
          else do {
            ASSERT(unit_prop_body_wl_inv S' j w L);
            i  pos_of_watched (get_clauses_wl S') C L;
            ASSERT(i  1);
            L'  other_watched_wl S' L C i;
            val_L'  mop_polarity_wl S' L';
            if val_L' = Some True
            then update_blit_wl L C b j w L' S'
            else do {
              f  find_unwatched_l (get_trail_wl S') (get_clauses_wl S') C;
              ASSERT (unit_prop_body_wl_find_unwatched_inv f C S');
              case f of
                None  do {
                  if val_L' = Some False
                  then do {S'  set_conflict_wl C S'; RETURN (j+1, w+1, S')}
                  else do {S'  propagate_lit_wl_general L' C i S'; RETURN (j+1, w+1, S')}
                }
              | Some f  do {
                ASSERT(C ∈# dom_m (get_clauses_wl S')  f < length (get_clauses_wl S'  C)  f  2);
                K  mop_clauses_at (get_clauses_wl S') C f;
                val_L'  mop_polarity_wl S' K;
                if val_L' = Some True
                then update_blit_wl L C b j w K S'
                else update_clause_wl L L' C b j w i f S'
                }
            }
          }
        }
      }
      else do {
        K  SPEC((=) K);
        val_K  mop_polarity_wl S' K;
        if val_K = Some True
        then RETURN (j+1, w+1, S')
        else do { ―‹Now the costly operations:
          if b'
          then RETURN (j, w+1, S')
          else do {
            ASSERT(unit_prop_body_wl_inv S' j w L);
            i  pos_of_watched (get_clauses_wl S') C L;
            ASSERT(i  1);
            L'  other_watched_wl S' L C i;
            val_L'  mop_polarity_wl S' L';
            if val_L' = Some True
            then update_blit_wl L C b j w L' S'
            else do {
              f  find_unwatched_l (get_trail_wl S') (get_clauses_wl S') C;
              ASSERT (unit_prop_body_wl_find_unwatched_inv f C S');
              case f of
                None  do {
                  if val_L' = Some False
                  then do {S'  set_conflict_wl C S'; RETURN (j+1, w+1, S')}
                  else do {S'propagate_lit_wl_general L' C i S';RETURN (j+1, w+1, S')}
                }
              | Some f  do {
                ASSERT(C ∈# dom_m (get_clauses_wl S')  f < length (get_clauses_wl S'  C)  f  2);
                K  mop_clauses_at (get_clauses_wl S') C f;
                val_L'  mop_polarity_wl S' K;
                if val_L' = Some True
                then update_blit_wl L C b j w K S'
                else update_clause_wl L L' C b j w i f S'
                }
            }
          }
        }
      }
   }
  (is _ = do {
      ASSERT(unit_propagation_inner_loop_wl_loop_pre L (j, w, S));
      (C, K, b)  mop_watched_by_at S L w;
      S' mop_keep_watch L j w S;
      ASSERT(is_nondeleted_clause_pre C L S');
      let b' = (C ∉# dom_m (get_clauses_wl S'));
      if b' then
        ?P S' C K b b'
      else
        ?Q S' C K b b'
    })
    unfolding unit_propagation_inner_loop_body_wl_int_def if_not_swap bind_to_let_conv
      SPEC_eq_is_RETURN twl_st_wl
    unfolding Let_def if_not_swap bind_to_let_conv
      SPEC_eq_is_RETURN twl_st_wl
    apply (subst if_cancel)
    apply (intro bind_cong case_prod_cong if_cong[OF refl] refl if_cong)
    done
  show ?thesis
    unfolding E
    apply (subst if_replace_cond[of _ ?P _ _ _ _])
    unfolding if_True if_False Let_def
    apply (intro bind_cong_nres case_prod_cong if_cong[OF refl] refl)
    done
qed


section The Functions

subsection Inner Loop

lemma clause_to_update_mapsto_upd_If:
  assumes
    i: i ∈# dom_m N
  shows
  clause_to_update L (M, N(i  C'), C, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q) =
    (if L  set (watched_l C')
     then add_mset i (remove1_mset i (clause_to_update L (M, N, C, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q)))
     else remove1_mset i (clause_to_update L (M, N, C, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q)))
proof -
  define D' where D' = dom_m N - {#i#}
  then have [simp]: dom_m N = add_mset i D'
    using assms by (simp add: mset_set.remove)
  have [simp]: i ∉# D'
    using assms distinct_mset_dom[of N] unfolding D'_def by auto

  have {#C ∈# D'.
     (i = C  L  set (watched_l C')) 
     (i  C  L  set (watched_l (N  C)))#} =
    {#C ∈# D'. L  set (watched_l (N  C))#}
    by (rule filter_mset_cong2) auto
  then show ?thesis
    unfolding clause_to_update_def
    by auto
qed

lemma keep_watch_st_wl[twl_st_wl]:
  get_unit_clauses_wl (keep_watch L j w S) = get_unit_clauses_wl S
  get_init_clauses0_wl (keep_watch L j w S) = get_init_clauses0_wl S
  get_learned_clauses0_wl (keep_watch L j w S) = get_learned_clauses0_wl S
  get_clauses0_wl (keep_watch L j w S) = get_clauses0_wl S
  get_conflict_wl (keep_watch L j w S) = get_conflict_wl S
  get_trail_wl (keep_watch L j w S) = get_trail_wl S
  by (cases S; auto simp: keep_watch_def; fail)+

declare twl_st_wl[simp]

lemma correct_watching_except_correct_watching_except_propagate_lit_wl:
  assumes
    corr: correct_watching_except j w L S and
    i_le: Suc 0 < length (get_clauses_wl S  C) and
    C: C ∈# dom_m (get_clauses_wl S) and undef: undefined_lit (get_trail_wl S) L' and
    confl: get_conflict_wl S = None and
    lit: L' ∈# all_lits_st S and
     i: i  1
  shows propagate_lit_wl_general L' C i S  SPEC(λc. correct_watching_except j w L c)
proof -
  obtain M N D NE UE NEk UEk NS US N0 U0 Q W where S: S = (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) by (cases S)
  let ?ℒ = all_lits_st S
  have
    Hneq: La. La∈#?ℒ 
        La  L 
         ((i, K, b)∈#mset (W La). i ∈# dom_m N  K  set (N  i)  K  La 
            correctly_marked_as_binary N (i, K, b)) 
         ((i, K, b)∈#mset (W La). b  i ∈# dom_m N) 
         {#i ∈# fst `# mset (W La). i ∈# dom_m N#} = clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}) and
    Heq: La. La∈#?ℒ 
        La = L 
         ((i, K, b)∈#mset (take j (W La) @ drop w (W La)). i ∈# dom_m N  K  set (N  i)  K  La 
              correctly_marked_as_binary N (i, K, b)) 
         ((i, K, b)∈#mset (take j (W La) @ drop w (W La)). b  i ∈# dom_m N) 
         {#i ∈# fst `# mset (take j (W La) @ drop w (W La)). i ∈# dom_m N#} =
         clause_to_update La (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
    using corr unfolding S correct_watching_except.simps
    by fast+
  define N' where N'  if length (N  C) > 2 then N(C  swap (N  C) 0 (Suc 0 - i)) else N

  have Suc 0 - i < length (N  C) and 0 < length (N  C)
    using i_le
    by (auto simp: S)
  then have [simp]: mset (swap (N  C) 0 (Suc 0 - i)) = mset (N  C)
    by (auto simp: S)
  have H1[simp]: {#mset (fst x). x ∈# ran_m N'#} =
     {#mset (fst x). x ∈# ran_m N#}
    using C
    by (auto dest!: multi_member_split simp: ran_m_def S N'_def
           intro!: image_mset_cong)

  have H2: mset `# ran_mf N' = mset `# ran_mf N
    using H1 by auto
  have H3: dom_m N' = dom_m N
    using C by (auto simp: S N'_def)
  have H4: set (N'  ia) =
    set (N  ia) for ia
    using i_le
    by (cases C = ia) (auto simp: S N'_def)
  have H5: set (watched_l (N'  ia)) = set (watched_l (N  ia)) for ia
    using i_le
    by (cases C = ia; cases i; cases N  ia; cases tl (N  ia)) (auto simp: N'_def S swap_def)
  have [iff]: correctly_marked_as_binary N C'  correctly_marked_as_binary N' C' for C' ia
    by (cases C')
      (auto simp: correctly_marked_as_binary.simps N'_def)
  have H6: propagate_lit_wl_general L' C i (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) =
     RETURN (Propagated L' C # M, N', D, NE, UE, NEk, UEk, NS, US, N0, U0, add_mset (- L') Q, W)
   using assms Suc 0 - i < length (N  C) undef confl lit
   by (auto simp: mop_clauses_swap_def S propagate_lit_wl_general_def N'_def
     cons_trail_propagate_l_def)
 have : all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) =
   all_lits_st (M, N', D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
   using assms i_le unfolding N'_def by (auto simp: S)
  show ?thesis
    using corr lit unfolding S H6 apply -
    apply (rule RETURN_rule)
    unfolding S prod.simps Let_def correct_watching_except.simps 
      H1 H2 H3 H4 H5 clause_to_update_def get_clauses_l.simps H5 all_lits_st_simps
    by simp
qed


lemma unit_propagation_inner_loop_body_wl_int_alt_def2:
  unit_propagation_inner_loop_body_wl_int L j w S = do {
      ASSERT(unit_propagation_inner_loop_wl_loop_pre L (j, w, S));
      (C, K, b)  mop_watched_by_at S L w;
      S  mop_keep_watch L j w S;
      ASSERT(is_nondeleted_clause_pre C L S);
      val_K  mop_polarity_wl S K;
      if val_K = Some True
      then RETURN (j+1, w+1, S)
      else do { ―‹Now the costly operations:
        if b then
          if C ∉# dom_m (get_clauses_wl S)
          then RETURN (j, w+1, S)
          else do {
            ASSERT(unit_prop_body_wl_inv S j w L);
            i  pos_of_watched (get_clauses_wl S) C L;
            ASSERT(i  1);
            L'  other_watched_wl S L C i;
            val_L'  mop_polarity_wl S L';
            if val_L' = Some True
            then update_blit_wl L C b j w L' S
            else do {
              f  find_unwatched_l (get_trail_wl S) (get_clauses_wl S) C;
              ASSERT (unit_prop_body_wl_find_unwatched_inv f C S);
              case f of
                None  do {
                  if val_L' = Some False
                  then do {S  set_conflict_wl C S; RETURN (j+1, w+1, S)}
                  else do {S  propagate_lit_wl_general L' C i S; RETURN (j+1, w+1, S)}
                }
              | Some f  do {
                  ASSERT(C ∈# dom_m (get_clauses_wl S)  f < length (get_clauses_wl S  C)  f  2);
                  K  mop_clauses_at (get_clauses_wl S) C f;
                  val_L'  mop_polarity_wl S K;
                  if val_L' = Some True
                  then update_blit_wl L C b j w K S
                  else update_clause_wl L L' C b j w i f S
                }
            }
          }
        else
          if C ∉# dom_m (get_clauses_wl S)
          then RETURN (j, w+1, S)
          else do {
            ASSERT(unit_prop_body_wl_inv S j w L);
            i  pos_of_watched (get_clauses_wl S) C L;
            ASSERT(i  1);
            L'  other_watched_wl S L C i;
            val_L'  mop_polarity_wl S L';
            if val_L' = Some True
            then update_blit_wl L C b j w L' S
            else do {
              f  find_unwatched_l (get_trail_wl S) (get_clauses_wl S) C;
              ASSERT (unit_prop_body_wl_find_unwatched_inv f C S);
              case f of
                None  do {
                  if val_L' = Some False
                  then do {S  set_conflict_wl C S; RETURN (j+1, w+1, S)}
                  else do {S  propagate_lit_wl_general L' C i S; RETURN (j+1, w+1, S)}
                }
              | Some f  do {
                  ASSERT(C ∈# dom_m (get_clauses_wl S)  f < length (get_clauses_wl S  C)  f  2);
                  K  mop_clauses_at (get_clauses_wl S) C f;
                  val_L'  mop_polarity_wl S K;
                  if val_L' = Some True
                  then update_blit_wl L C b j w K S
                  else update_clause_wl L L' C b j w i f S
                }
            }
          }
      }
   }
  unfolding unit_propagation_inner_loop_body_wl_int_def if_not_swap bind_to_let_conv
    SPEC_eq_is_RETURN twl_st_wl
  unfolding Let_def if_not_swap bind_to_let_conv
    SPEC_eq_is_RETURN twl_st_wl
  apply (subst if_cancel)
  apply (intro bind_cong_nres case_prod_cong if_cong[OF refl] refl)
  done

lemma unit_propagation_inner_loop_body_wl_alt_def:
  unit_propagation_inner_loop_body_wl (L :: 'v literal) j w S = do {
      ASSERT(unit_propagation_inner_loop_wl_loop_pre L (j, w, S));
      (C, K, b)  mop_watched_by_at S L w;
      S  mop_keep_watch L j w S;
      ASSERT(is_nondeleted_clause_pre C L S);
      val_K  mop_polarity_wl S K;
      if val_K = Some True
      then RETURN (j+1, w+1, S)
      else do {
        if b then do {
          if False
          then RETURN (j, w+1, S)
          else do{
            let i::nat = (if ((get_clauses_wl S)C) ! 0 = L then 0 else 1);
                K = K; val_L' = val_K in
            if False ― ‹termval_L' = Some True
            then RETURN (j, w+1, S)
            else do {
              f  RETURN (None :: nat option);
              case f of
               None  do {
                 ASSERT(propagate_proper_bin_case L K S C);
                 if val_K = Some False
                 then do {S  set_conflict_wl C S; RETURN (j+1, w+1, S)}
                 else do {
                   S  propagate_lit_wl_bin' K C (if ((get_clauses_wl S)C) ! 0 = L then 0 else 1) S;
                   RETURN (j+1, w+1, S)}
               }
             | _  RETURN (j, w+1, S)
            }}
        }  ―‹Now the costly operations:
        else if C ∉# dom_m (get_clauses_wl S)
        then RETURN (j, w+1, S)
        else do {
          ASSERT(unit_prop_body_wl_inv S j w L);
          i  pos_of_watched (get_clauses_wl S) C L;
          ASSERT(i  1);
          L'  other_watched_wl S L C i;
          val_L'  mop_polarity_wl S L';
          if val_L' = Some True
          then update_blit_wl L C b j w L' S
          else do {
            f  find_unwatched_l (get_trail_wl S) (get_clauses_wl S) C;
            ASSERT (unit_prop_body_wl_find_unwatched_inv f C S);
            case f of
              None  do {
                if val_L' = Some False
                then do {S  set_conflict_wl C S; RETURN (j+1, w+1, S)}
                else do {S  propagate_lit_wl L' C i S; RETURN (j+1, w+1, S)}
              }
            | Some f  do {
                 ASSERT(C ∈# dom_m (get_clauses_wl S)  f < length (get_clauses_wl S  C)  f  2);
                 K  mop_clauses_at (get_clauses_wl S) C f;
                val_L' mop_polarity_wl S K;
                if val_L' = Some True
                then update_blit_wl L C b j w K S
                else update_clause_wl L L' C b j w i f S
              }
          }
        }
      }
   }
  unfolding unit_propagation_inner_loop_body_wl_def if_not_swap bind_to_let_conv
    SPEC_eq_is_RETURN twl_st_wl propagate_lit_wl_bin'_def
  unfolding Let_def if_not_swap bind_to_let_conv
    SPEC_eq_is_RETURN twl_st_wl if_False
  apply (intro bind_cong_nres case_prod_cong if_cong[OF refl] refl option.case_cong)
  apply auto
  done

lemma find_unwatched_l_itself:
    (uncurry2 find_unwatched_l, uncurry2 find_unwatched_l)  Id ×f Id ×f Id f Idoption_relnres_rel
   by (auto intro!: frefI nres_relI)

lemma
  fixes S :: 'v twl_st_wl and S' :: 'v twl_st_l and L :: 'v literal and w :: nat
  defines [simp]: C'  fst (watched_by S L ! w)
  defines
    [simp]: T  remove_one_lit_from_wq C' S'

  defines
    [simp]: C''  get_clauses_l S'  C'
  assumes
    S_S': (S, S')  state_wl_l (Some (L, w)) and
    w_le: w < length (watched_by S L) and
    j_w: j  w and
    corr_w: correct_watching_except j w L S and
    inner_loop_inv: unit_propagation_inner_loop_wl_loop_inv L (j, w, S) and
    n: n = size (filter_mset (λ(i, _). i ∉# dom_m (get_clauses_wl S)) (mset (drop w (watched_by S L)))) and
    confl_S: get_conflict_wl S = None
  shows unit_propagation_inner_loop_body_wl_wl_int: unit_propagation_inner_loop_body_wl L j w S 
      Id (unit_propagation_inner_loop_body_wl_int L j w S)
proof -
  have n_d: no_dup (get_trail_wl S)
    using inner_loop_inv unfolding unit_propagation_inner_loop_wl_loop_inv_def prod.case
      unit_propagation_inner_loop_l_inv_def twl_struct_invs_def cdclW_restart_mset.cdclW_all_struct_inv_def
      cdclW_restart_mset.cdclW_M_level_inv_def pcdcl_all_struct_invs_def apply -
    by normalize_goal+ (simp only: twl_st_wl twl_st_l twl_st)

  have pos_of_watched: ((N, C, L), (N', C', L'))  Id  pos_of_watched N C L   nat_rel (pos_of_watched N' C' L')
     for N N' C C' L L'
     by auto
  define SLW where SLW  mop_watched_by_at S L w
  define R_SLW where [simp]: R_SLW = {((C, K, b), (C', K', b')). C = C'  K = K'  b = b' 
               (b  length (get_clauses_wl S  C) = 2  C ∈# dom_m (get_clauses_wl S) 
                  K  set(get_clauses_wl S  C)  L  set(get_clauses_wl S  C)  K  L) 
               (C ∈# dom_m (get_clauses_wl S)  ¬b  length (get_clauses_wl S  C) > 2)}
  define keep where keep  mop_keep_watch L j w S
  have [refine]: unit_propagation_inner_loop_wl_loop_pre L(j, w, S) 
    SLW   R_SLW (mop_watched_by_at S L w)
    using corr_w
    unfolding SLW_def mop_watched_by_at_def R_SLW_def
    apply (refine_rcg, cases S)
    apply  (clarsimp dest!: multi_member_split simp: mop_watched_by_def ASSERT_refine_right
         correctly_marked_as_binary.simps
         correct_watching_except_alt_def simp flip: Cons_nth_drop_Suc
         dest!: )
    apply auto
    done
  have [refine]: unit_propagation_inner_loop_wl_loop_pre L(j, w, S) 
    keep   {(T, T'). (T, T')  Id  get_clauses_wl T = get_clauses_wl S  get_trail_wl T = get_trail_wl S}
            (mop_keep_watch L j w S) (is _  _   ?keep _)
    using corr_w
    unfolding keep_def mop_keep_watch_def
    by (refine_rcg, cases S)
      (clarsimp dest!: multi_member_split simp: mop_watched_by_def ASSERT_refine_right
         correct_watching_except.simps simp flip: Cons_nth_drop_Suc)

  have [refine]: unit_propagation_inner_loop_wl_loop_pre L(j, w, S) 
    (x, x')   R_SLW 
    (T', val_K)  ?keep 
    ¬ x1∉# dom_m (get_clauses_wl val_K) 
    x'= (x1, x2) 
    x= (x2b, x1c) 
    RETURN (if get_clauses_wl T' x2b! 0 = L then 0 else 1 )
       {(K', K). K = K'  K = (if get_clauses_wl T' x2b! 0 = L then 0 else 1)}
      (pos_of_watched (get_clauses_wl val_K) x1 L)
   (is _; _; _; _; _; _  _   ?pos _)
   for x x' x1 x2 x2a x1b x2b x1c x2c T T' val_Ka val_aK val_K
   unfolding pos_of_watched_def R_SLW_def
   by refine_rcg auto

  have bin_mop_clauses_at: unit_propagation_inner_loop_wl_loop_pre L(j, w, S) 
    (x, x')   R_SLW 
    (T', val_K)  ?keep 
    ¬ x1∉# dom_m (get_clauses_wl val_K) 
    x2= (x2a, x1b) 
    x'= (x1, x2) 
    x1c = (x2d, x1d) 
    x= (x2b, x1c)  x1b 
    (if get_clauses_wl T' x2b! 0 = L then 0 else 1, i)  ?pos x2b T' 
    RETURN x2d
       {(K, K'). K = K'  x2d = K'}
    (other_watched_wl val_K L x1 i)
   for x x' x1 x2 x2a x1b x2b x1c x2c T T' val_Ka val_aK val_K x2d x1d i
   unfolding mop_clauses_at_def R_SLW_def other_watched_wl_def
   by refine_rcg (auto simp: length_list_2)

  have bin_mop_polarity_wl: unit_propagation_inner_loop_wl_loop_pre L (j, w, S) 
    unit_propagation_inner_loop_wl_loop_pre L (j, w, S) 
    (Saa, val_K)  ?keep 
    (val_Ka, i) {(v, v'). v = v'  v = polarity (get_trail_wl Saa) x2c} 
    ¬ x2 ∉# dom_m (get_clauses_wl val_K) 
    (if get_clauses_wl Saa  x2b ! 0 = L then 0 else 1, K)
     ?pos x2b Saa 
    (x2c, K') {(K, K'). K = K'  x2c = K'} 
    (x', x1)  R_SLW 
    x1a =(x2a, x1b) 
    x1 =(x2, x1a) 
    x1c =(x2c, Sa) 
    x' =(x2b, x1c) 
    val_Ka  Some True 
    i  Some True 
    Sa 
    x1b 
    RETURN val_Ka
      (Idoption_rel)
       (mop_polarity_wl val_K K')
    for x' x1 x2 x1a x2a x1b x2b x1c x2c Sa Saa val_K val_Ka i K K'
    unfolding mop_polarity_wl_def
    by refine_rcg auto

  have bin_find_unwatched: RETURN None
          ({(b,b'). (b, b')  Idoption_rel  b = None})
           (find_unwatched_l (get_trail_wl Saa) (get_clauses_wl Saa) x1)
      (is _   ?bin_unw _)
    if
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S) and
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S) and
      (x, x')  R_SLW and
      x2 =(x1a, x2a) and
      x' =(x1, x2) and
      x2b =(x1c, x2c) and
      x =(x1b, x2b) and
      (Sa, Saa)  ?keep and
      (val_K, val_Ka)
       {(v, v'). v = v'  v = polarity (get_trail_wl Sa) x1c} and
      val_K  Some True and
      val_Ka  Some True and
      x2c and
      x2a and
      ¬ False and
      ¬ x1 ∉# dom_m (get_clauses_wl Saa) and
      unit_prop_body_wl_inv Saa j w L and
      (if get_clauses_wl Sa  x1b ! 0 = L then 0 else 1, i)
       {(K', K).
          K = K' 
          K = (if get_clauses_wl Sa  x1b ! 0 = L then 0 else 1)} and
      (x1c, L') {(K, K'). K = K'  x1c = K'} and
      (val_K, val_L') bool_reloption_rel and
      ¬ False and
      val_L'  Some True
     for x x' x1 x2 x1a x2a x1b x2b x1c x2c Sa Saa val_K val_Ka i L' val_L'
     using that n_d
     unfolding find_unwatched_l_def
     by (auto simp: RETURN_RES_refine_iff length_list_2)


  have bin_unw_Id: x  ?bin_unw  x  Idoption_rel for x
    by auto
  have prop_bin_prop_gen[THEN fref_to_Down_curry3, refine]:
     (uncurry3 propagate_lit_wl_bin', uncurry3 propagate_lit_wl_general) 
        [λ(((L', C), _), S). length (get_clauses_wl S  C) = 2]f (Id ×f nat_rel  ×f {_. True}  ×f Id)  Idnres_rel
    by (auto simp: propagate_lit_wl_bin'_def propagate_lit_wl_general_def  cons_trail_propagate_l_def le_ASSERT_iff propagate_lit_wl_bin_def
       intro!: frefI nres_relI)

  have prop_prop_gen[THEN fref_to_Down_curry3, refine]: (uncurry3 propagate_lit_wl, uncurry3 propagate_lit_wl_general) 
    [λ(((L', C), _), S). length (get_clauses_wl S  C) > 2]f (Id ×f nat_rel  ×f nat_rel  ×f Id)  Idnres_rel
    by (auto simp: propagate_lit_wl_def propagate_lit_wl_general_def
       intro!: frefI nres_relI)

  have [THEN fref_to_Down_curry2, refine]:
     (uncurry2 mop_clauses_at, uncurry2 mop_clauses_at)  Id ×f Id ×f Id f Idnres_rel
    by (auto intro!: frefI nres_relI)
  have [refine]:
    ((S, L), (S', L'))  Id  mop_polarity_wl S L   {(v, v'). v=v'  v = polarity (get_trail_wl S) L}
       (mop_polarity_wl S' L') for S S' L L'
    by (auto intro!: frefI nres_relI simp: mop_polarity_wl_def ASSERT_refine_right)

  have [THEN fref_to_Down_curry, refine]:
     (uncurry set_conflict_wl, uncurry set_conflict_wl)  Id ×f Id f Idnres_rel
    by (auto intro!: frefI nres_relI)

  have [THEN fref_to_Down_curry3, refine]:
    (uncurry3 other_watched_wl, uncurry3 other_watched_wl)  Id ×f Id ×f Id ×f Id f Idnres_rel
    by (auto intro!: frefI nres_relI)

  show ?thesis
    supply [[goals_limit=1]]
    unfolding unit_propagation_inner_loop_body_wl_int_alt_def2
       unit_propagation_inner_loop_body_wl_alt_def
    apply (subst SLW_def[symmetric])
    apply (subst keep_def[symmetric])
    apply (refine_rcg pos_of_watched
      find_unwatched_l_itself[THEN fref_to_Down_curry2])
    subgoal for SLw SLw' C x2a bin bL C' x2b bin'  bL'
      by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    apply assumption+
    apply (rule bin_mop_clauses_at; assumption)
    apply (rule bin_mop_polarity_wl; assumption)
    subgoal by auto
    subgoal by auto
    apply (rule bin_find_unwatched; assumption)
    apply (rule bin_unw_Id; assumption)
    subgoal
      unfolding propagate_proper_bin_case_def by (auto simp: length_list_2)
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    apply assumption
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    done
qed


lemma nres_add_unrelated:
   P   {(S, S'). R1 S} Q  P   {(S, S'). R2 S S'} Q  P  {(S, S'). R1 S  R2 S S'} Q for P R1 R2 Q
  by (simp add: pw_ref_iff)

lemma nres_add_unrelated2:
   P  SPEC R1  P   {(S, S'). R2 S S'} Q  P  {(S, S'). R1 S  R2 S S'} Q for P R1 R2 Q
  using nres_add_unrelated[of P R1 Q R2]
  by (auto simp add: pw_ref_iff dest: inres_SPEC)

lemma nres_add_unrelated3:
   P   {(S, S'). R1 S} Q  P   {(S, S'). R2 S} Q  P  {(S, S'). R1 S  R2 S} Q for P R1 R2 Q
  by (simp add: pw_ref_iff)

lemma blits_in_ℒin_keep_watch2: w < length (watched_by S L)  blits_in_ℒin S  blits_in_ℒin (keep_watch L j w S)
   apply (cases S; clarsimp simp: blits_in_ℒin_def keep_watch_def)
  using nth_mem[of w (watched_by S L)]
  by (auto elim!: in_set_upd_cases simp: eq_commute[of _ _ ! w] simp del: nth_mem
     dest!: multi_member_split)


lemma unit_propagation_inner_loop_body_l_with_skip_alt_def:
  unit_propagation_inner_loop_body_l_with_skip L (S', n) = do {
      ASSERT (clauses_to_update_l S'  {#}  0 < n);
      ASSERT (unit_propagation_inner_loop_l_inv L (S', n));
      let _ = ();
      b  SPEC (λb. (b  0 < n)  (¬ b  clauses_to_update_l S'  {#}));
      let S' = S';
      let b = b;
      if ¬ b
      then do {
        ASSERT (clauses_to_update_l S'  {#});
        X2  select_from_clauses_to_update S';
        ASSERT (unit_propagation_inner_loop_body_l_inv L (snd X2) (fst X2));
        x  SPEC (λK. K  set (get_clauses_l (fst X2)  snd X2));
        v  mop_polarity_l (fst X2) x;
        if v = Some True then let T = fst X2 in RETURN (T, if get_conflict_l T = None then n else 0)
        else do {
          v  pos_of_watched (get_clauses_l (fst X2)) (snd X2) L;
          va  other_watched_l (fst X2) L (snd X2) v;
          vaa  mop_polarity_l (fst X2) va;
          if vaa = Some True
         then let T = fst X2 in RETURN (T, if get_conflict_l T = None then n else 0)
          else do {
             x  find_unwatched_l (get_trail_l (fst X2)) (get_clauses_l (fst X2)) (snd X2);
             case x of
             None 
               if vaa = Some False
               then do { T  set_conflict_l (snd X2) (fst X2);
                    RETURN (T, if get_conflict_l T = None then n else 0) }
               else do { T  propagate_lit_l va (snd X2) v (fst X2);
                     RETURN (T, if get_conflict_l T = None then n else 0)}
             | Some a  do {
                   x  ASSERT (a < length (get_clauses_l (fst X2)  snd X2));
                   K  mop_clauses_at (get_clauses_l (fst X2)) (snd X2) a;
                   val_K  mop_polarity_l (fst X2) K;
                   if val_K = Some True
                   then let T = fst X2 in RETURN (T, if get_conflict_l T = None then n else 0)
                   else do {
                          T  update_clause_l (snd X2) v a (fst X2);
                          RETURN (T, if get_conflict_l T = None then n else 0)
                        }
                 }
            }
        }
      }
      else RETURN (S', n - 1)
    }
proof -
  have remove_pairs: do {(x2, x2')  (b0 :: _ nres); F x2 x2'} =
        do {X2  b0; F (fst X2) (snd X2)} for T a0 b0 a b c f t F
    by (meson case_prod_unfold)

  have H1: do {T  do {x  a ; b x}; RETURN (f T)} =
        do {x  a; T  b x; RETURN (f T)} for T a0 b0 a b c f t
    by auto
  have H2: do{T  let v = val in g v; (f T :: _ nres)} =
         do{let v = val; T  g v; f T}  for g f T val
    by auto
  have H3: do{T  if b then g else g'; (f T :: _ nres)} =
         (if b then do{T  g; f T} else do{T  g'; f T})  for g g' f T b
    by auto
  have H4: do{T  case x of None  g | Some a  g' a; (f T :: _ nres)} =
         (case x of None  do{T  g; f T} | Some a  do{T  g' a; f T})  for g g' f T b x
    by (cases x) auto
  show ?thesis
    apply (subst Let_def)
    apply (subst Let_def)
    apply (subst Let_def)
    apply (subst Let_def)
    unfolding unit_propagation_inner_loop_body_l_with_skip_def prod.case
      unit_propagation_inner_loop_body_l_def remove_pairs
    unfolding H1 H2 H3 H4 bind_to_let_conv
    by (intro bind_cong[OF refl] if_cong refl option.case_cong) auto
qed

lemma get_subsumed_init_clauses_l_remove_one_lit_from_wq[simp]:
  get_subsumed_init_clauses_l (remove_one_lit_from_wq a S) = get_subsumed_init_clauses_l S
  get_subsumed_clauses_wl (keep_watch L j w S') = get_subsumed_clauses_wl S'
  get_subsumed_clauses_wl (set_literals_to_update_wl x S') = get_subsumed_clauses_wl S'
  by (cases S; cases S'; auto simp: keep_watch_def; fail)+

lemma nth_in_all_lits_stI[simp]:
  x1 ∈# dom_m (get_clauses_wl S)  i < length (get_clauses_wl S  x1) 
  get_clauses_wl S  x1 ! ( i) ∈# all_lits_st S
  using in_clause_in_all_lits_of_m[of get_clauses_wl S  x1 ! i mset(get_clauses_wl S  x1)]
  by (auto simp: all_lits_def all_lits_st_def all_lits_of_mm_union ran_m_def
      all_lits_of_mm_add_mset dest!: multi_member_split)


lemma all_lits_st_keep_watch[simp]: all_lits_st (keep_watch L j w S) = all_lits_st S
  by (cases S) (auto simp: all_lits_st_def)

lemma
  fixes S :: 'v twl_st_wl and S' :: 'v twl_st_l and L :: 'v literal and w :: nat
  defines [simp]: C'  fst (watched_by S L ! w)
  defines
    [simp]: T  remove_one_lit_from_wq C' S'

  defines
    [simp]: C''  get_clauses_l S'  C'
  assumes
    S_S': (S, S')  state_wl_l (Some (L, w)) and
    w_le: w < length (watched_by S L) and
    j_w: j  w and
    corr_w: correct_watching_except j w L S and
    blit_in_lit: blits_in_ℒin S and
    inner_loop_inv: unit_propagation_inner_loop_wl_loop_inv L (j, w, S) and
    n: n = size (filter_mset (λ(i, _). i ∉# dom_m (get_clauses_wl S)) (mset (drop w (watched_by S L)))) and
    confl_S: get_conflict_wl S = None
  shows unit_propagation_inner_loop_body_wl_int_spec: unit_propagation_inner_loop_body_wl_int L j w S 
    {((i, j, T'), (T, n)).
        (T', T)  state_wl_l (Some (L, j)) 
        correct_watching_except i j L T' 
        blits_in_ℒin T' 
        j  length (watched_by T' L) 
        length (watched_by S L) =  length (watched_by T' L) 
        i  j 
        (get_conflict_wl T' = None 
           n = size (filter_mset (λ(i, _). i ∉# dom_m (get_clauses_wl T')) (mset (drop j (watched_by T' L))))) 
        (get_conflict_wl T'  None  n = 0)}
     (unit_propagation_inner_loop_body_l_with_skip L (S', n)) (is ?propa is _   ?unit _)and
    unit_propagation_inner_loop_body_wl_update:
      unit_propagation_inner_loop_body_l_inv L C' T 
         mset `# (ran_mf ((get_clauses_wl S) (C' (swap (get_clauses_wl S  C') 0
                           (1 - (if (get_clauses_wl S)C' ! 0 = L then 0 else 1)))))) =
        mset `# (ran_mf (get_clauses_wl S)) (is _  ?eq)
proof -
  obtain bL where SLw: watched_by S L ! w = (C', bL)
    using C'_def by (cases watched_by S L ! w) auto
  let ?M = get_trail_wl S

  define i :: nat where
    i  (if get_clauses_wl S  C' ! 0 = L then 0 else 1)
  have n_d: no_dup ?M
    using S_S' inner_loop_inv apply -
    unfolding unit_propagation_inner_loop_wl_loop_inv_def
      unit_propagation_inner_loop_l_inv_def twl_struct_invs_def
   unfolding cdclW_restart_mset.cdclW_M_level_inv_def prod.case
     cdclW_restart_mset.cdclW_all_struct_inv_def pcdcl_all_struct_invs_def stateW_of_def
   by normalize_goal+
     (simp only: twl_st_l twl_st_wl twl_st)
  have
    alien_L':
       L ∈# all_lits_st S
       (is ?alien')
    using inner_loop_inv twl_struct_invs_no_alien_in_trail[of _ -L] unfolding unit_propagation_inner_loop_wl_loop_inv_def
     unit_propagation_inner_loop_l_inv_def prod.case apply - apply normalize_goal+
    by (drule twl_struct_invs_no_alien_in_trail[of _ -L])
    (simp_all only: twl_st_wl twl_st_l twl_st multiset.map_comp comp_def clause_twl_clause_of
        ac_simps in_all_lits_uminus_iff)

  have
    clause_ge_0: 0 < length (get_clauses_l T  C') (is ?ge) and
    L_def: defined_lit (get_trail_wl S) L -L  lits_of_l (get_trail_wl S)
      L  lits_of_l (get_trail_wl S) (is ?L_def) and
    i_le: i < length (get_clauses_wl S  C')  (is ?i_le) and
    i_le2: 1-i < length (get_clauses_wl S  C')  (is ?i_le2) and
    C'_dom: C' ∈# dom_m (get_clauses_l T) (is ?C'_dom) and
    L_watched: L  set (watched_l (get_clauses_l T  C')) (is ?L_w) and
    dist_clss: distinct_mset_mset (mset `# ran_mf (get_clauses_wl S)) and
    confl: get_conflict_l T = None (is ?confl) and
    alien_L:
       L ∈# all_lits_st S
       (is ?alien) and
    correctly_marked_as_binary: correctly_marked_as_binary (get_clauses_wl S) (C', bL)
  if
    unit_propagation_inner_loop_body_l_inv L C' T
  proof -
    have unit_propagation_inner_loop_body_l_inv L C' T
      using that unfolding unit_prop_body_wl_inv_def by fast+
    then obtain T' where
      T_T': (set_clauses_to_update_l (clauses_to_update_l T + {#C'#}) T, T')  twl_st_l (Some L) and
      struct_invs: twl_struct_invs T' and
       twl_stgy_invs T' and
      C'_dom: C' ∈# dom_m (get_clauses_l T) and
       0 < C' and
       ge_0: 0 < length (get_clauses_l T  C') and
       no_dup (get_trail_l T) and
       i_le: (if get_clauses_l T  C' ! 0 = L then 0 else 1)
         < length (get_clauses_l T  C') and
       i_le2: 1 - (if get_clauses_l T  C' ! 0 = L then 0 else 1)
         < length (get_clauses_l T  C') and
       L_watched: L  set (watched_l (get_clauses_l T  C')) and
       confl: get_conflict_l T = None
      unfolding unit_propagation_inner_loop_body_l_inv_def by blast
    show ?i_le and ?C'_dom and ?L_w and ?i_le2
      using S_S' i_le C'_dom L_watched i_le2 unfolding i_def by auto
    have
        alien: cdclW_restart_mset.no_strange_atm (stateW_of T') and
        dup: no_duplicate_queued T' and
        lev: cdclW_restart_mset.cdclW_M_level_inv (stateW_of T') and
        dist: cdclW_restart_mset.distinct_cdclW_state (stateW_of T')
      using struct_invs unfolding twl_struct_invs_def cdclW_restart_mset.cdclW_all_struct_inv_def
        pcdcl_all_struct_invs_def stateW_of_def
      by blast+
    have n_d: no_dup (trail (stateW_of T'))
       using lev unfolding cdclW_restart_mset.cdclW_M_level_inv_def by auto
    have 1: C ∈# clauses_to_update T' 
         add_mset (fst C) (literals_to_update T') ⊆#
         uminus `# lit_of `# mset (get_trail T') for C
      using dup unfolding no_duplicate_queued_alt_def
      by blast
    have H: (L, twl_clause_of C'') ∈# clauses_to_update T'
      using twl_st_l(5)[OF T_T']
      by (auto simp: twl_st_l)
    have uL_M: -L  lits_of_l (get_trail T')
      using mset_le_add_mset_decr_left2[OF 1[OF H]]
      by (auto simp: lits_of_def)
    then show defined_lit (get_trail_wl S) L -L  lits_of_l (get_trail_wl S)
      L  lits_of_l (get_trail_wl S)
      using S_S' T_T' n_d by (auto simp: Decided_Propagated_in_iff_in_lits_of_l twl_st
        dest: no_dup_consistentD)
    show L: ?alien
      using alien uL_M T_T' struct_invs twl_st_l(25)[OF T_T'] S_S'
        twl_struct_invs_no_alien_in_trail[of T' -L]
        unit_init_clauses_get_unit_init_clauses_l[OF T_T']
      unfolding cdclW_restart_mset.no_strange_atm_def
      by (auto simp: in_all_lits_of_mm_ain_atms_of_iff)
    then have l_wl_inv: (S, S')  state_wl_l (Some (L, w)) 
         unit_propagation_inner_loop_body_l_inv L (fst (watched_by S L ! w))
          (remove_one_lit_from_wq (fst (watched_by S L ! w)) S') 
         L ∈# all_lits_st S 
         correct_watching_except j w L S 
         w < length (watched_by S L)  get_conflict_wl S = None
      using that assms L unfolding unit_prop_body_wl_inv_def unit_propagation_inner_loop_body_l_inv_def
      by (auto simp: twl_st)


    show ?ge
      by (rule ge_0)
    show distinct_mset_mset (mset `# ran_mf (get_clauses_wl S))
      using dist S_S' twl_st_l(1-8)[OF T_T'] T_T' unfolding cdclW_restart_mset.distinct_cdclW_state_alt_def
      by (auto simp: twl_st)
    show ?confl
      using confl .
    have watched_by S L ! w  set (take j (watched_by S L))  set (drop w (watched_by S L))
      using L alien_L' C'_dom SLw w_le
      by (cases S)
        (auto simp: in_set_drop_conv_nth)
    then show correctly_marked_as_binary (get_clauses_wl S) (C', bL)
      using corr_w alien_L' C'_dom SLw S_S'
      by (cases S; cases watched_by S L ! w)
       (clarsimp simp: correct_watching_except.simps Ball_def all_conj_distrib state_wl_l_def
          simp del: Un_iff
           simp flip: all_lits_alt_def2 all_lits_def
        dest!: multi_member_split[of L])
  qed

  have i_def': i = (if get_clauses_l T  C' ! 0 = L then 0 else 1)
    using S_S' unfolding i_def by auto

  have [simp]: length (watched_by (keep_watch L j w S) L) = length (watched_by S L) for S j w L
    by (cases S) (auto simp: keep_watch_def)
  have keep_watch_state_wl: fst (watched_by S L ! w) ∉# dom_m (get_clauses_wl S) 
     (keep_watch L j w S, S')  state_wl_l (Some (L, Suc w))
    using S_S' w_le j_w by (cases S; cases S')
      (auto simp: state_wl_l_def keep_watch_def Cons_nth_drop_Suc[symmetric]
        drop_map)
  have [simp]: drop (Suc w) (watched_by (keep_watch L j w S) L) = drop (Suc w) (watched_by S L)
    using j_w w_le by (cases S) (auto simp: keep_watch_def)
  have [simp]: get_clauses_wl (keep_watch L j w S) = get_clauses_wl S for L j w S
    by (cases S) (auto simp: keep_watch_def)

  have trail_keep_w: get_trail_wl (keep_watch L j w S) = get_trail_wl S for L j w S
    by (cases S) (auto simp: keep_watch_def)


  let ?pos_of_watched = {(j, j'). j = j'  j < 2  j = i}

  have pos_of_watched:
     ((N, C, K), (N', C', K'))  Id ×r Id ×r Id  N = get_clauses_wl S  C = fst (watched_by S L ! w)  K = L 
     pos_of_watched N C K   ?pos_of_watched (pos_of_watched N' C' K')
     for N N' C C' K K'
     unfolding pos_of_watched_def by refine_rcg (auto simp: i_def)
  have [refine]: mop_watched_by_at S L w   {(Slw, _). Slw = watched_by S L ! w 
           fst (snd (Slw)) ∈# all_lits_st S 
          (fst Slw ∈# dom_m (get_clauses_wl S) 
             fst (snd Slw)  set (get_clauses_l S'  fst (watched_by S L ! w)))} (RETURN ()) (is _   ?Slw _)
    if
      clauses_to_update_l S'  {#}  0 < n and
      unit_propagation_inner_loop_l_inv L (S', n) and
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S)
    using assms that
    unfolding mop_watched_by_at_def unit_propagation_inner_loop_l_inv_def
       prod.case apply -
    apply normalize_goal+
    apply (cases watched_by S L ! w)
    subgoal for T
       using nth_mem[OF w_le] twl_struct_invs_no_alien_in_trail[of T -L] apply -
      by refine_vcg
       (auto simp: in_all_lits_of_mm_uminus_iff
         mset_take_mset_drop_mset' correct_watching_except_alt_def2
          blits_in_ℒin_def in_all_lits_minus_iff
        dest!: multi_member_split[of L all_lits_st S]
        simp flip: Cons_nth_drop_Suc all_lits_def all_lits_alt_def2)
    done
  have [refine]:
     mop_keep_watch L j w S   {(T, T'). T = keep_watch L j w S  T' = S'} (RETURN S')
    using j_w w_le alien_L' unfolding mop_keep_watch_def all_lits_def
    by refine_rcg (auto simp: ac_simps)

  have keep_watch: RETURN Sa
           {(T, (T', C)). (T, T')  state_wl_l (Some (L, Suc w)) 
         C = C'  T' = set_clauses_to_update_l (clauses_to_update_l S' - {#C#}) S'}
        (select_from_clauses_to_update S')
    if
      clauses_to_update_l S'  {#}  0 < n and
      unit_propagation_inner_loop_l_inv L (S', n) and
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S) and
      (x', ())  ?Slw and
      x' =(x1, x2) and
      (x1 ∉# dom_m (get_clauses_wl S), b)  bool_rel and
      (Sa, S') {(T, T'). T = keep_watch L j w S  T' = S'} and
      ¬ x1 ∉# dom_m (get_clauses_wl Sa)
     for x' x1 x2 x1a x2a b Sa
  proof -
    have get_conflict_l S' = None
      using that unfolding unit_propagation_inner_loop_l_inv_def twl_struct_invs_def prod.case
      apply -
      apply normalize_goal+
      by auto
    then show ?thesis
      using S_S' that w_le j_w
      unfolding select_from_clauses_to_update_def keep_watch_def
      by (cases S)
        (auto intro!: RETURN_RES_refine simp: state_wl_l_def drop_map
          Cons_nth_drop_Suc[symmetric] eq_commute[of _ _ ! w])
  qed


  have f: ((M, N, C), (M', N', C'))  Id  find_unwatched_l M N C
        {(found, found'). found = found' 
             (found = None  (Lset (unwatched_l (N  C)). -L  lits_of_l M)) 
             (j. found = Some j  (j < length (N  C)  (undefined_lit M ((N  C)!j)  (N  C)!j  lits_of_l M)  j  2))  distinct (N  C)
           }
            (find_unwatched_l M' N' C')
    (is _  _   ?find _) for M M' N N' C C'
    using S_S' unfolding find_unwatched_l_def
    by refine_rcg (auto intro!: RES_refine)

  have update_blit_wl: update_blit_wl L x1 x2a j w L' Sa
         ?unit
           (RETURN(fst X2, if get_conflict_l (fst X2) = None then n else 0))
    if
      cond: clauses_to_update_l S'  {#}  0 < n and
      loop_inv: unit_propagation_inner_loop_l_inv L (S', n) and
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S) and
      inres (mop_watched_by_at S L w) x' and
      x': (x', ())  ?Slw
        x2 =(x1a, x2a)
        x' =(x1, x2) and
      (x1 ∉# dom_m (get_clauses_wl S), b)  bool_rel and
      inres (mop_keep_watch L j w S) Sa and
      Sa: (Sa, S') {(T, T'). T = keep_watch L j w S  T' = S'} and
      dom: ¬ x1 ∉# dom_m (get_clauses_wl Sa) and
      ¬ b and
      clauses_to_update_l S'  {#} and
      X2: (Sa, X2)
       {(T, T', C).
          (T, T')  state_wl_l (Some(L, Suc w)) 
          C = C' 
          T' =
          set_clauses_to_update_l (remove1_mset C (clauses_to_update_l S'))
           S'} and
      l_inv: unit_propagation_inner_loop_body_l_inv L (snd X2) (fst X2) and
      (K, x)  Id and
      K  Collect ((=) x1a) and
      x {K. K  set (get_clauses_l (fst X2)  snd X2)} and
      (val_K, v)  Id and
      val_K  Some True and
      v  Some True and
      wl_inv: unit_prop_body_wl_inv Sa j w L and
      i: (ia, va) {(j, j'). j = j'  j < 2  j = i} and
      L': (L', vaa)
       {(L, L'). L = L'  L = get_clauses_wl Sa  x1 ! (1 - ia)} and
      val_L' = Some True and
      (val_L', vaaa)  Id and
      vaaa = Some True
   for x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa
  proof -
    have confl: get_conflict_wl S = None and [simp]: Sa = keep_watch L j w S
      using S_S' loop_inv cond Sa unfolding unit_propagation_inner_loop_l_inv_def prod.case apply -
      by normalize_goal+ auto

    then obtain M N NE UE NEk UEk NS US N0 U0 Q W where
      S: S = (M, N, None, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
      by (cases S) (auto simp: twl_st_l)
    have dom': x1 ∈# dom_m (get_clauses_wl (keep_watch L j w S))  True
      using dom by auto
    then have SLW_dom': fst (watched_by Sa L ! w)
        ∈# dom_m (get_clauses_wl Sa) and
      SLw': watched_by S L ! w = (x1, x1a, x2a)
      using w_le x' by (auto simp: eq_commute[of _ _ ! w])
    have bin: correctly_marked_as_binary N (x1, N  x1 ! (Suc 0 - i), x2a)
      using X2 correctly_marked_as_binary l_inv x' dom' SLw
      by (cases bL; cases W L ! w)
       (auto simp: S remove_one_lit_from_wq_def correctly_marked_as_binary.simps)

    obtain x where
      S_x: (Sa, x)  state_wl_l (Some (L, Suc w)) and
      unit_loop_inv:
        unit_propagation_inner_loop_body_l_inv L (fst (watched_by Sa L ! w))
      x and
      L: L ∈# all_lits_st Sa and
      corr_Sa: correct_watching_except (Suc j) (Suc w) L Sa and
      w < length (watched_by Sa L) and
      get_conflict_wl Sa = None
      using wl_inv SLW_dom' unfolding unit_prop_body_wl_inv_alt_def
      by blast
    obtain x' where
      x_x': (set_clauses_to_update_l
        (clauses_to_update_l
          x +
          {#fst (watched_by Sa L ! w)#})
        x,
        x')  twl_st_l (Some L) and
       twl_struct_invs x' and
      ge0: (if get_clauses_l x  fst (watched_by Sa L ! w) ! 0 = L
        then 0 else 1) < length (get_clauses_l x  fst (watched_by Sa L ! w)) and
      ge1i: 1 -
      (if get_clauses_l x  fst (watched_by Sa L ! w) ! 0 = L then 0 else 1)
      < length
          (get_clauses_l x 
          fst (watched_by Sa L ! w)) and
      L_watched: L  set (watched_l
                (get_clauses_l x 
                  fst (watched_by Sa L ! w))) and
      get_conflict_l x = None
      using unit_loop_inv
      unfolding unit_propagation_inner_loop_body_l_inv_def
      by blast

    have unit_T: unit_propagation_inner_loop_body_l_inv L C' T
      using that
      by (auto simp: remove_one_lit_from_wq_def)

    have twl_st_inv x'
      using twl_struct_invs x' unfolding twl_struct_invs_def by fast
    then have x. twl_st_inv
         (x,{#TWL_Clause (mset (watched_l (fst x)))
                (mset (unwatched_l (fst x)))
             . x ∈# init_clss_l N#},
          {#TWL_Clause (mset (watched_l (fst x)))
             (mset (unwatched_l (fst x)))
          . x ∈# learned_clss_l N#},
          None, NE+NEk, UE+UEk, NS, US, N0, U0,
          add_mset
           (L, TWL_Clause
                (mset (watched_l (N  fst ((W L)[j := W L ! w] ! w))))
                (mset (unwatched_l (N  fst ((W L)[j := W L ! w] ! w)))))
           {#(L, TWL_Clause (mset (watched_l (N  x)))
                  (mset (unwatched_l (N  x))))
           . x ∈#{#i ∈# mset
                          (drop (Suc w) (map fst ((W L)[j := W L ! w]))).
                   i ∈# dom_m N#}#},
          Q)
      using x_x' S_x
      apply (cases x)
      apply (auto simp: S twl_st_l_def state_wl_l_def keep_watch_def
        simp del: struct_wf_twl_cls.simps)
      done
    then have Multiset.Ball
       ({#TWL_Clause (mset (watched_l (fst x))) (mset (unwatched_l (fst x)))
        . x ∈# ran_m N#})
       struct_wf_twl_cls
      unfolding twl_st_inv.simps image_mset_union[symmetric] all_clss_l_ran_m
      by blast
    then have distinct_N_x1: distinct (N  x1)
      using dom
      by (auto simp: S ran_m_def mset_take_mset_drop_mset' dest!: multi_member_split)

    have watch_by_S_w: watched_by (keep_watch L j w S) L ! w = (x1, x1a, x2a)
      using j_w w_le SLw x' Sa unfolding i_def
      by (cases S)
        (auto simp: keep_watch_def split: if_splits)
    have i_le: i < length (N  x1)  1-i < length (N  x1)
      using watch_by_S_w ge0 ge1i S_x assms i
      by (auto simp: S)
    have X2: X2 = (set_clauses_to_update_l (remove1_mset x1 (clauses_to_update_l S')) S', x1)
      using SLw X2 S_S' x' unfolding i_def Sa
      by (auto simp add: twl_st_wl eq_commute[of _ _ ! w])
    let ?ℒ = all_lits_st S
    have N_x1_in_L: N  x1 ! (Suc 0 - i) ∈# ?ℒ
      using dom i_le by (auto simp: ran_m_def S all_lits_of_mm_add_mset
        intro!: in_clause_in_all_lits_of_m
        dest!: multi_member_split)
    then have L_i: L = N  x1 ! i
      using j_w w_le L_watched ge0 ge1i SLw S_x i Sa SLw' unfolding i_def
      by (auto simp: take_2_if twl_st_wl S keep_watch_def split: if_splits)

    have ((M, N, None, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W (L := (W L)[j := (x1, L', x2a)])),
       fst X2)  state_wl_l (Some (L, Suc w))
     using S_S' X2 j_w w_le SLw x' dom L' i SLw' dom unfolding Sa
     by (cases W L ! w)
       (auto simp: state_wl_l_def S keep_watch_def drop_map Cons_nth_drop_Suc[symmetric])
    moreover have n = size {#(i, _) ∈# mset (drop (Suc w) (watched_by S L)).
      i ∉# dom_m (get_clauses_wl S)#}
      using dom n w_le
      by (clarsimp_all simp add: SLw' Cons_nth_drop_Suc[symmetric]
          dest!: multi_member_split)
    moreover {
      have Suc 0 - i  i
        using i by (auto simp: i_def split: if_splits)
      then have correct_watching_except (Suc j) (Suc w) L
        (M, N, None, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W(L := (W L)[j := (x1, L', x2a)]))
        using SLw apply -
        apply (rule correct_watching_except_update_blit)
        using N_x1_in_L corr_Sa i_le distinct_N_x1 i_le bin SLw' L' i unfolding S Sa
        by (auto simp: L_i keep_watch_def nth_eq_iff_index_eq S all_lits_def ac_simps)
    }
    ultimately show ?thesis
      using j_w w_le i alien_L' dom unit_T N_x1_in_L L' N_x1_in_L blit_in_lit
      unfolding i[symmetric] T_def[symmetric]
      by (auto simp: S update_blit_wl_def keep_watch_def all_lits_def ac_simps
        intro!: ASSERT_leI blits_in_ℒin_keep_watch'
         blits_in_ℒin_propagate)

  qed

  have set_conflict_wl: set_conflict_wl x1 Sa
          {(U, U'). get_conflict_l U'  None  ((Suc j, Suc w, U), (U', 0))  ?unit}
           (set_conflict_l (snd X2) (fst X2))
    if
      clauses_to_update_l S'  {#}  0 < n and
      unit_propagation_inner_loop_l_inv L (S', n) and
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S) and
      inres (mop_watched_by_at S L w) x' and
      x': (x', ())   ?Slw
        x2 =(x1a, x2a)
        x' =(x1, x2) and
      (x1 ∉# dom_m (get_clauses_wl S), b)  bool_rel and
      inres (mop_keep_watch L j w S) Sa and
      Sa: (Sa, S') {(T, T'). T = keep_watch L j w S  T' = S'} and
      ¬ x1 ∉# dom_m (get_clauses_wl Sa) and
      ¬ b and
      clauses_to_update_l S'  {#} and
      X2: (Sa, X2)
       {(T, T', C).
          (T, T')  state_wl_l (Some(L, Suc w)) 
          C = C' 
          T' =
          set_clauses_to_update_l (remove1_mset C (clauses_to_update_l S'))
           S'} and
      unit_propagation_inner_loop_body_l_inv L (snd X2) (fst X2) and
      (K, x)  Id and
      K  Collect ((=) x1a) and
      x {K. K  set (get_clauses_l (fst X2)  snd X2)} and
      (val_K, v)  Id and
      val_K  Some True and
      v  Some True and
      unit_prop_body_wl_inv Sa j w L and
      (ia, va) {(j, j'). j = j'  j < 2  j = i} and
      (L', vaa)
       {(L, L'). L = L'  L = get_clauses_wl Sa  x1 ! (1 - ia)} and
      (val_L', vaaa)  Id and
      val_L'  Some True and
      vaaa  Some True and
      (f, x'a)  ?find (get_trail_wl Sa) (get_clauses_wl Sa) x1 and
      unit_prop_body_wl_find_unwatched_inv f x1 Sa and
      f = None and
      x'a = None and
      val_L' = Some False and
      vaaa = Some False
   for  x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa f x'a
  proof -
    have [simp]: correct_watching_except a b L (M, N, Some D', NE, UE, NEk, UEk, NS, US, N0, U0, W, oth) 
        correct_watching_except a b L (M, N, None, NE, UE, NEk, UEk, NS, US, N0, U0, W, oth)
         NO_MATCH {#} W  correct_watching_except a b L (M, N, None, NE, UE, NEk, UEk, NS, US, N0, U0, W, oth) 
        correct_watching_except a b L (M, N, None, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, oth)
      for a b M N D' NE UE W oth NS US N0 U0 NEk UEk
      by (simp_all add: correct_watching_except.simps all_lits_st_def
        set_conflict_wl_def prod.case clause_to_update_def get_clauses_l.simps)
     have [simp]: NO_MATCH None d  blits_in_ℒin(x1b, x1aa, d, x1c, x1d, NEk, UEk, NS, US, N0, U0, x2e, g) 
        blits_in_ℒin(x1b, x1aa, None, x1c, x1d, NEk, UEk, NS, US, N0, U0, x2e, g) and
       [simp]: NO_MATCH {#} x2e  blits_in_ℒin(x1b, x1aa, None, x1c, x1d, NEk, UEk,  NS, US, N0, U0, x2e, g) 
        blits_in_ℒin(x1b, x1aa, None, x1c, x1d, NEk, UEk, NS, US, N0, U0, {#}, g) for x1b x1aa x1c x1d x2e g d NS US N0 U0 NEk UEk
        unfolding blits_in_ℒin_def by auto
    have get_conflict_wl (keep_watch L j w S) = None
      using confl_S by (cases S) (auto simp: keep_watch_def)
      then have set_conflict_wl_pre D (keep_watch L j w S) 
        set_conflict_wl D (keep_watch L j w S)  (SPEC(correct_watching_except (Suc j) (Suc w) L)) for D
      using S_S' SLw w_le j_w n that corr_w
        correct_watching_except_correct_watching_except_Suc_Suc_keep_watch[of j w S L]
      by (cases keep_watch L j w S)
        (auto intro!: frefI nres_relI simp: set_conflict_wl_def state_wl_l_def
           keep_watch_state_wl
            corr_w correct_watching_except_correct_watching_except_Suc_notin
        intro!: ASSERT_refine_right intro: ASSERT_leI)

    moreover have N_x1_in_L: get_clauses_wl S  x1 ! (Suc 0 - i) ∈# all_lits_st S
      using that i_le2 by (auto simp: ran_m_def all_lits_of_mm_add_mset all_lits_def
          remove_one_lit_from_wq_def eq_commute[of _ _ ! w] all_lits_st_def
        intro!: in_clause_in_all_lits_of_m nth_mem
        dest!: multi_member_split[of x1])
    ultimately show ?thesis
      using S_S' Sa X2 SLw w_le j_w n corr_w n blit_in_lit x'
      unfolding set_conflict_wl_def set_conflict_l_def
      apply (cases S; cases S'; cases X2; cases watched_by S L ! w)
      apply (refine_rcg)
      subgoal premises p for a ba c d e fa g aa baa ca da ea faa ga x1b x2b x1aa x2aa x1ba x2ba x1c
        x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h x1i x2i x1j x2j x1k x2k unfolding set_conflict_wl_pre_def
        using p apply - by (rule exI[of _ fst X2], rule exI[of _ Some(L, Suc w)])
         (auto simp add: keep_watch_def eq_commute [of _ _ ! w] all_lits_def ac_simps
         intro!: blits_in_ℒin_keep_watch' blits_in_ℒin_propagate
         simp del: )
     using that
     by (simp_all flip: all_lits_alt_def2 all_lits_def
           add: set_conflict_l_def set_conflict_wl_def state_wl_l_def
           keep_watch_state_wl keep_watch_def drop_map Cons_nth_drop_Suc[symmetric]
            correct_watching_except_correct_watching_except_Suc_Suc_keep_watch
         corr_w correct_watching_except_correct_watching_except_Suc_notin
         blits_in_ℒin_keep_watch' blits_in_ℒin_propagate)
  qed

  have propagate_lit_wl_general: propagate_lit_wl_general L' x1 ia Sa
          {(U, U'). ((Suc j, Suc w, U), (U', if get_conflict_l U' = None then n else 0))  ?unit}
           (propagate_lit_l vaa (snd X2) va (fst X2))
    if
      clauses_to_update_l S'  {#}  0 < n and
      unit_propagation_inner_loop_l_inv L (S', n) and
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S) and
      inres (mop_watched_by_at S L w) x' and
      Slw: (x', ())  ?Slw and
      x': x2 =(x1a, x2a)
        x' =(x1, x2) and
      (x1 ∉# dom_m (get_clauses_wl S), b)  bool_rel and
      Sa: (Sa, S') {(T, T'). T = keep_watch L j w S  T' = S'} and
      x1_dom: ¬ x1 ∉# dom_m (get_clauses_wl Sa) and
      ¬ b and
      clauses_to_update_l S'  {#} and
      X2: (Sa, X2)
       {(T, T', C).
          (T, T')  state_wl_l (Some(L, Suc w)) 
          C = C' 
          T' =
          set_clauses_to_update_l (remove1_mset C (clauses_to_update_l S'))
           S'} and
      l_inv: unit_propagation_inner_loop_body_l_inv L (snd X2) (fst X2) and
      (K, x)  Id and
      K  Collect ((=) x1a) and
      x {K. K  set (get_clauses_l (fst X2)  snd X2)} and
      (val_K, v)  Id and
      val_K  Some True and
      v  Some True and
      unit_prop_body_wl_inv Sa j w L and
      i: (ia, va) {(j, j'). j = j'  j < 2  j = i} and
      L': (L', vaa)
       {(L, L'). L = L'  L = get_clauses_wl Sa  x1 ! (1 - ia)} and
      (val_L', vaaa)  Id and
      val_L'  Some True and
      vaaa  Some True and
      (f, x'a)  ?find (get_trail_wl Sa) (get_clauses_wl Sa) x1 and
      unit_prop_body_wl_find_unwatched_inv f x1 Sa and
      f = None and
      x'a = None and
      val_L'  Some False and
      vaaa  Some False
    for x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa f x'a
  proof -
    define i' :: nat where i'  if get_clauses_wl (keep_watch L j w S)  x1 ! 0 = L then 0 else 1
    have x[simp]: watched_by S L ! w = (x1, x1a, x2a) Sa = keep_watch L j w S
      using x' Slw Sa by auto
    have i_alt_def: i' = (if get_clauses_l (fst X2)  snd X2 ! 0 = L then 0 else 1) i = i'
      using X2 S_S' i unfolding i'_def x' i_def by (auto)
    have x1_dom[simp]: x1 ∈# dom_m (get_clauses_wl S)
      using x1_dom X2 by (cases S) (auto simp: keep_watch_def)
    have [simp]: get_clauses_wl S  x1 ! 0  L  get_clauses_wl S  x1 ! Suc 0 = L and
      Suc 0 < length (get_clauses_wl S  x1)
      using l_inv X2 S_S' SLw unfolding unit_propagation_inner_loop_body_l_inv_def
      apply - apply normalize_goal+
      by (cases get_clauses_wl S  x1; cases tl (get_clauses_wl S  x1))
        auto

    have n: n = size {#(i, _) ∈# mset (drop (Suc w) (watched_by S L)).
        i ∉# dom_m (get_clauses_wl S)#}
      using n
      apply (subst (asm) Cons_nth_drop_Suc[symmetric])
      subgoal using w_le by simp
      subgoal using n SLw X2 S_S' unfolding i_def by auto
      done
    have [simp]: get_conflict_l (fst X2) = get_conflict_wl S
      using X2 S_S' by auto

    have eq: a = b  a = get_clauses_wl S  (a, b)  {(a, b). a = b  b = get_clauses_wl S}
      for a b
      by auto
    have ignored: P  SPEC P'  P   {(S, S'). P' S} (RETURN Q) for P P' Q
      by (auto simp: conc_fun_RES RETURN_def)
    have [simp]: i Suc 0
      by (simp add: i'_def i_alt_def(2))
    have get_clauses_wl S  x1 ! (Suc 0 - i) ∈# all_lits_st S
      using i_le x1_dom  Suc 0 < length (get_clauses_wl S  x1)
      by (auto simp: i'_def)
    then have
      (propagate_lit_wl_general (get_clauses_wl S  x1 ! (Suc 0 - i)) x1 i (keep_watch L j w S)
       (state_wl_l (Some (L, Suc w)))
         (propagate_lit_l (get_clauses_l (fst X2)  snd X2 ! (Suc 0 - i)) (snd X2) i (fst X2)))
      using X2 S_S' SLw j_w w_le multi_member_split[OF x1_dom] confl_S x
      by (cases S; cases S')
       (auto simp: state_wl_l_def propagate_lit_wl_general_def keep_watch_def
        propagate_lit_l_def drop_map mop_clauses_swap_def cons_trail_propagate_l_def
          simp flip: all_lits_alt_def2
          intro!: ASSERT_refine_right)
    have correct_watching_except (Suc j) (Suc w) L (keep_watch L j w S)
      by (simp add: corr_w correct_watching_except_correct_watching_except_Suc_Suc_keep_watch j_w w_le)
    then have propagate_lit_wl_general L' x1 ia Sa
         {(S, S'). (correct_watching_except (Suc j) (Suc w) L S)}
          (propagate_lit_l vaa (snd X2) va (fst X2))
      using X2 Sa
      apply (cases X2, cases fst X2, hypsubst)
      unfolding propagate_lit_l_def cons_trail_propagate_l_def nres_monad3
        mop_clauses_swap_def If_bind_distrib apply -
      apply (clarsimp; intro conjI impI ASSERT_refine_right ignored;
         (rule correct_watching_except_correct_watching_except_propagate_lit_wl)?)
      using w_le j_w Suc 0 < length (get_clauses_wl S  x1) confl_S i S_S' L'
      by (cases S; auto simp: keep_watch_def state_wl_l_def simp flip: all_lits_alt_def2; fail)+


    moreover have N_x1_in_L: x1a ∈# all_lits_st S
      using that i_le2 by (auto simp: ran_m_def all_lits_of_mm_add_mset all_lits_def
          remove_one_lit_from_wq_def eq_commute[of _ _ ! w] ac_simps
        intro!: in_clause_in_all_lits_of_m nth_mem
        dest!: multi_member_split[of x1])
    then have propagate_lit_wl_general L' x1 ia Sa
         {(S, S'). (blits_in_ℒin S)}
          (propagate_lit_l vaa (snd X2) va (fst X2))
      using X2 Sa
      apply (cases X2, cases fst X2, hypsubst)
      unfolding propagate_lit_l_def propagate_lit_wl_general_def cons_trail_propagate_l_def nres_monad3
        mop_clauses_swap_def If_bind_distrib apply -
      apply (clarsimp; intro conjI impI ASSERT_refine_right ignored)
      using w_le j_w Suc 0 < length (get_clauses_wl S  x1) confl_S i S_S' L' x x1_dom blit_in_lit
      apply (cases S; auto simp: keep_watch_def state_wl_l_def blits_in_ℒin_keep_watch'
         blits_in_ℒin_propagate simp flip: all_lits_alt_def2 all_lits_def)
      using w_le j_w Suc 0 < length (get_clauses_wl S  x1) confl_S i S_S' L' x x1_dom blit_in_lit
      apply (cases S; auto simp: keep_watch_def state_wl_l_def blits_in_ℒin_keep_watch'
         blits_in_ℒin_propagate simp flip: all_lits_alt_def2 all_lits_def)
      done
    moreover have propagate_lit_wl_general L' x1 ia Sa
     {(U, U').
         ((Suc j, Suc w, U), U', if get_conflict_l U' = None then n else 0)
         {((i, j, T'), T, n).
           (T', T)  state_wl_l (Some (L, Suc w)) 
            j  length (watched_by T' L) 
            length (watched_by S L) = length (watched_by T' L) 
            i  j 
            (get_conflict_wl T' = None 
             n =
             size
              {#(i, _)∈# mset (drop j (watched_by T' L)).
               i ∉# dom_m (get_clauses_wl T')#}) 
            (get_conflict_wl T'  None  n = 0)}}
       (propagate_lit_l vaa (snd X2) va (fst X2))
      using w_le n j_w x1_dom that S_S' confl_S i_le2 x
      unfolding i'_def[symmetric] i_alt_def[symmetric] propagate_lit_wl_general_def
         propagate_lit_l_def If_bind_distrib
      apply (refine_rcg refine_itself2[of cons_trail_propagate_l, THEN fref_to_Down_curry2])
      subgoal by (cases S) (auto simp: keep_watch_def)
      subgoal by (cases S) (auto simp: keep_watch_def)
      subgoal using S_S'  by (cases S) (auto simp: keep_watch_def state_wl_l_def all_lits_st_def
           all_lits_of_st_l_def all_lits_def
         simp flip: all_lits_alt_def2)
      subgoal by fast
      subgoal by (auto simp: state_wl_l_def)
      subgoal by (auto simp: state_wl_l_def)
      subgoal by (auto simp: state_wl_l_def)
      apply (rule mop_clauses_swap_itself_spec2)
      subgoal by (auto simp: state_wl_l_def)
      subgoal
        by (cases X2; cases S; cases S')
         (clarsimp simp: state_wl_l_def mop_clauses_swap_def drop_map
          op_clauses_swap_def keep_watch_def)
      apply (rule eq)
      subgoal by (cases X2; cases S; cases S')
         (auto simp: twl_st_wl simp: op_clauses_swap_def keep_watch_def
             propagate_lit_l_def mop_clauses_swap_def drop_map state_wl_l_def
           intro!: ASSERT_refine_right)
      subgoal by (cases X2; cases S; cases S')
         (clarsimp simp add:  op_clauses_swap_def keep_watch_def
          propagate_lit_l_def mop_clauses_swap_def  state_wl_l_def
          intro!: ASSERT_refine_right)
      done
    ultimately show ?thesis
      apply (rule nres_add_unrelated[OF nres_add_unrelated3, THEN order_trans])
      apply (rule conc_fun_R_mono)
      apply fastforce
      done
  qed

  have update_blit: update_blit_wl L x1 x2a j w Ka Sa
         ?unit
           (RETURN(fst X2, if get_conflict_l (fst X2) = None then n else 0))
    if
      cond: clauses_to_update_l S'  {#}  0 < n and
      loop_inv: unit_propagation_inner_loop_l_inv L (S', n) and
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S) and
      inres (mop_watched_by_at S L w) x' and
      x': (x', ())  ?Slw
        x2 =(x1a, x2a)
        x' =(x1, x2) and
      (x1 ∉# dom_m (get_clauses_wl S), b)  bool_rel and
      inres (mop_keep_watch L j w S) Sa and
      Sa: (Sa, S') {(T, T'). T = keep_watch L j w S  T' = S'} and
      dom: ¬ x1 ∉# dom_m (get_clauses_wl Sa) and
      ¬ b and
      clauses_to_update_l S'  {#} and
     X2: (Sa, X2)
       {(T, T', C).
          (T, T')  state_wl_l (Some(L, Suc w)) 
          C = C' 
          T' =
          set_clauses_to_update_l (remove1_mset C (clauses_to_update_l S'))
           S'} and
      unit_propagation_inner_loop_body_l_inv L (snd X2) (fst X2) and
      (K, x)  Id and
      K  Collect ((=) x1a) and
      x {K. K  set (get_clauses_l (fst X2)  snd X2)} and
      (val_K, v)  Id and
      val_K  Some True and
      v  Some True and
      unit_prop_body_wl_inv Sa j w L and
      (ia, va) {(j, j'). j = j'  j < 2  j = i} and
      (L', vaa)
       {(L, L'). L = L'  L = get_clauses_wl Sa  x1 ! (1 - ia)} and
      (val_L', vaaa)  Id and
      val_L'  Some True and
      vaaa  Some True and
      fx': (f, x'a)  ?find (get_trail_wl Sa) (get_clauses_wl Sa) x1
        f = Some xa
        x'a = Some x'b and
      unit_prop_body_wl_find_unwatched_inv f x1 Sa and
      (xa, x'b)  nat_rel and
      x'b < length (get_clauses_l (fst X2)  snd X2) and
      Ka: (Ka, Kb) {(L, L'). L = L'  L = get_clauses_wl Sa  x1 ! xa} and
      (val_L'a, val_Ka)  Id and
      val_L'a = Some True and
      val_Ka = Some True
   for x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa f x'a xa
       x'b Ka Kb val_L'a val_Ka
  proof -
    have confl: get_conflict_wl S = None
      using S_S' loop_inv cond unfolding unit_propagation_inner_loop_l_inv_def prod.case apply -
      by normalize_goal+ auto

    have unit_T: unit_propagation_inner_loop_body_l_inv L C' T
      using that by (auto simp: remove_one_lit_from_wq_def)

    have correct_watching_except (Suc j) (Suc w) L (keep_watch L j w S)
      by (simp add: corr_w correct_watching_except_correct_watching_except_Suc_Suc_keep_watch
          j_w w_le)
    moreover have correct_watching_except (Suc j) (Suc w) L
       (a, b, None, d, e, NEk, UEk, NS, US, N0, U0, f, ga(L := (ga L)[j := (x1, b  x1 ! xa, x2a)]))
      if
        corr: correct_watching_except (Suc j) (Suc w) L
      (a, b, None, d, e, NEk, UEk, NS, US, N0, U0, f, ga(L := (ga L)[j := (x1, x1a, x2a)])) and
        ga L ! w = (x1, x1a, x2a) and
        S[simp]: S = (a, b, None, d, e, NEk, UEk, NS, US, N0, U0, f, ga) and
        X2 = (set_clauses_to_update_l (remove1_mset x1 (clauses_to_update_l S')) S', x1) and
        (a, b, None, d, e, NEk, UEk, NS, US, N0, U0, 
      {#i ∈# mset (drop (Suc w) (map fst ((ga L)[j := (x1, x1a, x2a)]))). i ∈# dom_m b#}, f) =
      set_clauses_to_update_l (remove1_mset x1 (clauses_to_update_l S')) S'
      for a :: ('v literal, 'v literal,nat) annotated_lit list and
        b :: (nat, 'v literal list ×  bool) fmap and
        d e NS US N0 U0 NEk UEk :: 'v literal multiset multiset and
        f :: 'v literal multiset and
        ga :: 'v literal  (nat × 'v literal × bool) list
    proof -
      have b  x1 ! xa ∈# all_lits_st (a, b, None, d, e, NEk, UEk, NS, US, N0, U0, f, ga)
        using dom fx' Sa by (auto simp: ran_m_def all_lits_of_mm_add_mset x' f twl_st_wl
            dest!: multi_member_split
            simp: all_lits_def
            intro!: in_clause_in_all_lits_of_m)
      moreover have b  x1 ! xa  set (b  x1)
        using dom fx' Sa by (auto simp: ran_m_def all_lits_of_mm_add_mset x' f twl_st_wl
            dest!: multi_member_split
            intro!: in_clause_in_all_lits_of_m)

      moreover have b  x1 ! xa  L
        using X2 L_def[OF unit_T] S_S' SLw fx' x' that Sa
        by (auto simp: polarity_def split: if_splits)
      moreover have correctly_marked_as_binary b (x1, b  x1 ! xa, x2a)
        using correctly_marked_as_binary unit_T dom Sa SLw
        by (auto simp: correctly_marked_as_binary.simps ga L ! w = (x1, x1a, x2a))
      ultimately show ?thesis
        by (rule correct_watching_except_update_blit[OF corr])
    qed
    moreover have in_all: get_clauses_wl S  x1 ! xa ∈# all_lits_st S
      using dom fx' Sa unfolding x'
      by (auto dest!: multi_member_split[of x1]
        simp: ran_m_def all_lits_of_mm_add_mset all_lits_of_mm_union
          in_clause_in_all_lits_of_m all_lits_def)
    moreover have w: watched_by S L ! w = (x1, x1a, x2a)
      using x' by auto
    ultimately have u: update_blit_wl L x1 x2a j w Ka Sa
     SPEC(λ(i, j, T'). correct_watching_except i j L T')
      using X2 confl Sa SLw alien_L' unit_T Ka dom  j_w w_le unfolding T_def[symmetric]
        update_blit_wl_def
      apply (cases S; cases Sa, hypsubst) apply simp
      apply  (auto simp: keep_watch_def state_wl_l_def
           simp flip: all_lits_alt_def2 all_lits_def
           intro!: ASSERT_leI)
      done
    have b: update_blit_wl L x1 x2a j w Ka Sa
     SPEC(λ(i, j, T'). blits_in_ℒin T')
      using X2 confl Sa SLw alien_L' in_all unit_T Ka blit_in_lit j_w w_le dom unfolding T_def[symmetric]
        update_blit_wl_def
      apply (cases S; cases Sa, hypsubst) apply simp
      apply  (auto simp: keep_watch_def state_wl_l_def blits_in_ℒin_keep_watch'
              blits_in_ℒin_propagate all_lits_def ac_simps
           intro!: ASSERT_leI blits_in_ℒin_keep_watch')
      done
    have get_conflict_wl S = None
      using S_S' loop_inv cond unfolding unit_propagation_inner_loop_l_inv_def prod.case apply -
      by normalize_goal+ auto
    moreover have n = size {#(i, _) ∈# mset (drop (Suc w) (watched_by S L)). i ∉# dom_m (get_clauses_wl S)#}
      using n dom X2 w_le S_S' SLw Sa
      by (auto simp: Cons_nth_drop_Suc[symmetric] w)
    ultimately have update_blit_wl L x1 x2a j w Ka Sa
     {((i, j, T'), T, n).
         (T', T)  state_wl_l (Some(L, j)) 
         j  length (watched_by T' L) 
         length (watched_by S L) = length (watched_by T' L) 
         i  j 
         (get_conflict_wl T' = None 
          n =
          size
           {#(i, _)∈# mset (drop j (watched_by T' L)).
            i ∉# dom_m (get_clauses_wl T')#}) 
         (get_conflict_wl T'  None  n = 0)}
       (RETURN(fst X2, if get_conflict_l (fst X2) = None then n else 0))
      using j_w w_le S_S' X2 alien_L' unit_T in_all Sa in_all Ka dom unfolding T_def[symmetric]
      by (cases S)
        (auto simp: update_blit_wl_def keep_watch_def state_wl_l_def drop_map w
           simp flip: all_lits_alt_def2 all_lits_def
           intro!: ASSERT_leI)

    then show ?thesis
      apply (rule nres_add_unrelated2[OF SPEC_rule_conjI[OF u b] _, THEN order_trans])
      apply (rule conc_fun_R_mono)
      apply fastforce
      done
  qed

  have find_is_nat_rel: x  ?find a b c  x  nat_reloption_rel for x a b c
    by auto
  have update_clause_wl: update_clause_wl L L' x1 x2a j w ia xa Sa
         ?unit
           (do{
              T  update_clause_l (snd X2) va x'b (fst X2);
              RETURN(T, if get_conflict_l T = None then n else 0)
            })
    if
      cond: clauses_to_update_l S'  {#}  0 < n and
      loop_inv: unit_propagation_inner_loop_l_inv L (S', n) and
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S) and
      inres (mop_watched_by_at S L w) x' and
      x': (x', ())
        ?Slw and
      [simp]: x2 =(x1a, x2a) x' =(x1, x2) and
      (x1 ∉# dom_m (get_clauses_wl S), b)  bool_rel and
      inres (mop_keep_watch L j w S) Sa and
      Sa: (Sa, S') {(T, T'). T = keep_watch L j w S  T' = S'} and
      dom: ¬ x1 ∉# dom_m (get_clauses_wl Sa) and
      ¬ b and
      clauses_to_update_l S'  {#} and
      X2: (Sa, X2)
       {(T, T', C).
          (T, T')  state_wl_l (Some(L, Suc w)) 
          C = C' 
          T' = set_clauses_to_update_l (remove1_mset C (clauses_to_update_l S')) S'} and
      pre: unit_propagation_inner_loop_body_l_inv L (snd X2) (fst X2) and
      (K, x)  Id and
      K  Collect ((=) x1a) and
      x {K. K  set (get_clauses_l (fst X2)  snd X2)} and
      (val_K, v)  Id and
      val_K  Some True and
      v  Some True and
      wl_inv: unit_prop_body_wl_inv Sa j w L and
      i: (ia, va) {(j, j'). j = j'  j < 2  j = i} and
      L': (L', vaa)
       {(L, L'). L = L'  L = get_clauses_wl Sa  x1 ! (1 - ia)} and
      (val_L', vaaa)  Id and
      val_L'  Some True and
      vaaa  Some True and
      fx': (f, x'a)  ?find (get_trail_wl Sa) (get_clauses_wl Sa) x1 and
      unit_prop_body_wl_find_unwatched_inv f x1 Sa and
      [simp]: f = Some xa x'a = Some x'b and
      (xa, x'b)  nat_rel and
      x'b < length (get_clauses_l (fst X2)  snd X2) and
      Ka: (Ka, Kb) {(L, L'). L = L'  L = get_clauses_wl Sa  x1 ! xa} and
      (val_L'a, val_Ka)  Id and
      val_L'a  Some True and
      val_Ka  Some True
    for x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa f x'a xa
       x'b Ka Kb val_L'a val_Ka
  proof -
    have [simp]: xa = x'b Ka = Kb Kb = get_clauses_wl Sa  x1 ! xa ia = va
       using fx' Ka i by auto
    have update_clause_l_alt_def:
      update_clause_l = (λC i f (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q). do {
           ASSERT (update_clause_l_pre  C i f (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q));
           K  RETURN (op_clauses_at N C f);
           N'  mop_clauses_swap N C i f;
           RETURN (M, N', D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q)
      })
      unfolding update_clause_l_def by (auto intro!: ext)
    have unit_T: unit_propagation_inner_loop_body_l_inv L C' T
      using that by (auto simp: remove_one_lit_from_wq_def)
    have w: watched_by S L ! w = (x1, x1a, x2a)
      using x' by auto
    have alien_K: Ka ∈# all_lits_st Sa
      using Sa dom fx' unfolding all_lits_def
      by (cases S; cases Sa)
        (auto simp: keep_watch_def ran_m_def all_lits_of_mm_add_mset
        intro!: in_clause_in_all_lits_of_m
        dest!: multi_member_split)
    have L_in_watched: L  set (watched_l (get_clauses_wl Sa  x1))
      using corr_w alien_L' Sa X2 w_le w dom by (cases S)
        (clarsimp simp: correct_watching_except_alt_def keep_watch_def
          simp flip: Cons_nth_drop_Suc all_lits_alt_def2 all_lits_def
          dest!: multi_member_split)
    then have LKa: L  Ka
       using fx' Ka by (auto simp: in_set_conv_nth nth_eq_iff_index_eq)
    have L: get_clauses_wl Sa  x1 ! i = L  get_clauses_wl Sa  x1 ! (Suc 0 - i)  L
       get_clauses_wl Sa  x1 ! (Suc 0 - i)  Ka
      using i Sa w L_in_watched fx' Ka unfolding i_def
      by (auto simp: in_set_conv_nth nth_eq_iff_index_eq)
    have corr: correct_watching_except (Suc j) (Suc w) L Sa
      using wl_inv dom w Sa unfolding unit_prop_body_wl_inv_def by auto
    have confl: get_conflict_wl S = None
      using S_S' loop_inv cond unfolding unit_propagation_inner_loop_l_inv_def prod.case apply -
      by normalize_goal+ auto
  have
    alien_L'':
       L' ∈# all_lits_st Sa
    using L' i_le2[OF unit_T] i dom Sa w
    by (auto intro!: nth_in_all_lits_stI)

    show ?thesis
      supply RETURN_as_SPEC_refine[refine2 del]
      unfolding update_clause_l_alt_def update_clause_wl_def
      apply (cases X2; cases fst X2; cases Sa; hypsubst; unfold fst_conv; hypsubst;
        unfold prod.case snd_conv nres_monad3)
      apply (refine_rcg mop_clauses_at_op_clauses_at_spec2
         mop_clauses_swap_itself_spec2)
      subgoal using alien_L' Sa X2 dom by (cases S; auto simp: keep_watch_def)
      subgoal using alien_L' Sa X2 j_w w_le dom by (cases S; auto simp: keep_watch_def)
      subgoal using alien_L' Sa X2 j_w w_le dom by (cases S; auto simp: keep_watch_def)
      subgoal using alien_L' Sa X2 j_w w_le dom corr by (cases S; auto simp: keep_watch_def)
      subgoal using alien_L' Sa X2 by (cases S; auto simp: keep_watch_def
           simp flip: all_lits_alt_def2 all_lits_def)
      subgoal using alien_L'' by auto
      subgoal using dom Sa X2 w S_S' by (cases S; cases S'; auto simp: state_wl_l_def keep_watch_def)
      subgoal using fx' Sa S_S' X2 w dom by (cases S; cases S'; auto simp: state_wl_l_def keep_watch_def)
      subgoal using fx' Sa S_S' X2 w alien_K by (cases S; cases S'; auto simp: state_wl_l_def keep_watch_def all_lits_def)
      subgoal using fx' Sa S_S' X2 w alien_K by (cases S; cases S'; auto simp: state_wl_l_def
        keep_watch_def all_lits_def ac_simps)
      subgoal using LKa Ka by (auto simp: state_wl_l_def keep_watch_def)
      subgoal using fx' Sa S_S' X2 w by (cases S; cases S'; auto simp: state_wl_l_def keep_watch_def)

      subgoal supply [[goals_limit=1]]
        using Sa S_S' X2 w w_le j_w n corr LKa i L_in_watched Ka L fx' blit_in_lit alien_L' L'
           alien_K confl x' unfolding all_lits_def[symmetric] all_lits_alt_def[symmetric]
        apply (cases S; cases S')
        apply (clarsimp simp add: state_wl_l_def  keep_watch_def op_clauses_swap_def  blits_in_ℒin_keep_watch'
              blits_in_ℒin_propagate correct_watching_except_correct_watching_except_update_clause
          simp flip: Cons_nth_drop_Suc
          intro!: blits_in_ℒin_keep_watch'')
        apply (intro conjI correct_watching_except_correct_watching_except_update_clause
          blits_in_ℒin_keep_watch'')
        apply (clarsimp intro!: correct_watching_except_correct_watching_except_update_clause)[]
        apply (auto simp: state_wl_l_def keep_watch_def op_clauses_swap_def  blits_in_ℒin_keep_watch'
              blits_in_ℒin_propagate all_lits_def ac_simps
          simp flip: Cons_nth_drop_Suc  dest: in_set_takeD
          intro!: correct_watching_except_correct_watching_except_update_clause
               blits_in_ℒin_keep_watch'')
         done
      done
  qed

  have in_watched_keepD: x  set (watched_by (keep_watch L j w S) L')  x  set (watched_by S L')
     for x L'
    using j_w w_le by (cases S) (clarsimp simp: keep_watch_def elim!: in_set_upd_cases split: if_splits)

  have is_nondeleted_clause_pre: is_nondeleted_clause_pre x1 L Sa
    if
      clauses_to_update_l S'  {#}  0 < n and
      unit_propagation_inner_loop_l_inv L (S', n) and
      unit_propagation_inner_loop_wl_loop_pre L (j, w, S) and
      inres (mop_watched_by_at S L w) x' and
      x': (x', ())  ?Slw and
      [simp]: x2 = (x1a, x2a)
        x' = (x1, x2) and
      (x1 ∉# dom_m (get_clauses_wl S), b)  bool_rel and
      inres (mop_keep_watch L j w S) Sa and
      Sa: (Sa, S')  {(T, T'). T = keep_watch L j w S  T' = S'}
     for x' x1 x2 x1a x2a b Sa
  proof -
    have watched_by Sa L ! w = x' w < length (watched_by Sa L)
      using that using j_w w_le by auto
    then have  x'  set (watched_by Sa L)
      using j_w w_le by force
    then show ?thesis using j_w w_le x' alien_L' Sa
      by (auto simp: is_nondeleted_clause_pre_def eq_commute[of (_, _) _ ! w]
        simp flip: all_lits_alt_def2)
  qed


  have other_watched_wl_itself_spec2[refine]:
    ((N, C, i, j), (N', C', i', j'))  Id 
       other_watched_wl N C i j   Id (other_watched_wl N' C' i' j')
    for N i j N' i' j' and C C' :: 'v literal
    by (auto intro!: frefI nres_relI ASSERT_refine_right simp: mop_clauses_swap_def
      op_clauses_swap_def)

  show 1: ?propa
    (is _   ?unit _)
    supply trail_keep_w[simp]
    unfolding unit_propagation_inner_loop_body_wl_int_alt_def
      i_def[symmetric] i_def'[symmetric] unit_propagation_inner_loop_body_l_with_skip_alt_def
      unit_propagation_inner_loop_body_l_def
    unfolding i_def[symmetric] SLw prod.case all_lits_alt_def2[symmetric] all_lits_def[symmetric]
    apply (rewrite in if (¬_) then ASSERT _ >>= _ else _ if_not_swap)
    supply RETURN_as_SPEC_refine[refine2 del]
    apply (refine_rcg f (*val f f' find_unwatched_l*)
       mop_polarity_wl_mop_polarity_l[where b = Some(L, Suc w), THEN fref_to_Down_curry]
       mop_clauses_at_itself_spec pos_of_watched
       other_watched_wl_other_watched_l_spec_itself[of _ _ _ _ _ _ _ _  Some(L, Suc w)])
    subgoal using inner_loop_inv w_le j_w blit_in_lit
      unfolding unit_propagation_inner_loop_wl_loop_pre_def by auto
    subgoal
      using j_w w_le S_S'
      by (cases S; cases S')
       (auto simp: n state_wl_l_def unit_propagation_inner_loop_wl_loop_pre_def
           unit_propagation_inner_loop_l_inv_def twl_st_l_def
         simp flip: Cons_nth_drop_Suc)
    subgoal by (rule is_nondeleted_clause_pre)
    subgoal by simp
    subgoal for x' x1 x2 x1a x2a b Sa
      using assms j_w w_le n_d blit_in_lit unfolding unit_prop_body_wl_inv_def mop_polarity_wl_def
         nres_monad3 unit_propagation_inner_loop_wl_loop_pre_def
     using S_S' w_le j_w n confl_S apply -
     apply refine_vcg
     apply (auto simp: keep_watch_state_wl assert_bind_spec_conv Let_def twl_st_wl
        Cons_nth_drop_Suc[symmetric] correct_watching_except_correct_watching_except_Suc_Suc_keep_watch  blits_in_ℒin_def
        corr_w correct_watching_except_correct_watching_except_Suc_notin
        simp flip: all_lits_alt_def2
        split: if_splits dest: multi_member_split dest!: in_watched_keepD)
    done
    apply (rule keep_watch; assumption+)
    subgoal by auto ― ‹selection of K
    subgoal by fast
    subgoal by auto
    subgoal by auto ― ‹If condition
    subgoal for x' x1 x2 x1a x2a b Sa
      using j_w w_le S_S' blit_in_lit
      by (auto simp: n keep_watch_state_wl assert_bind_spec_conv Let_def twl_st_wl
          Cons_nth_drop_Suc[symmetric] correct_watching_except_correct_watching_except_Suc_Suc_keep_watch  blits_in_ℒin_def
          corr_w correct_watching_except_correct_watching_except_Suc_notin
          simp flip: all_lits_def
          split: if_splits dest: multi_member_split dest!: in_watched_keepD)
    subgoal for x' x1 x2 x1a x2a b Sa X2 K x val_K v
      using w_le j_w corr_w alien_L' S_S' blit_in_lit unfolding unit_prop_body_wl_inv_def
           all_lits_def[symmetric] all_lits_alt_def[symmetric] all_lits_alt_def2[symmetric]
       apply (intro conjI impI)
      apply (auto simp: blits_in_ℒin_keep_watch2)[4]
      apply (rule exI[of _ fst X2])
      apply (auto simp: remove_one_lit_from_wq_def
       correct_watching_except_correct_watching_except_Suc_Suc_keep_watch simp flip: all_lits_def)
      done
     subgoal using S_S' by (auto simp: eq_commute[of _ _ ! w])
     subgoal using S_S' by (auto simp: eq_commute[of _ _ ! w])
     subgoal using S_S' by (auto simp: eq_commute[of _ _ ! w])
     subgoal by simp
     subgoal using S_S' by (auto simp: eq_commute[of _ _ ! w])
     subgoal using S_S' by (auto simp: eq_commute[of _ _ ! w])
     subgoal by fast
     subgoal using S_S' by (auto simp: eq_commute[of _ _ ! w])
     subgoal by simp ― ‹polarity
     subgoal for x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa
         by (rule update_blit_wl)
     subgoal using S_S' by (auto simp: eq_commute[of _ _ ! w])
     subgoal unfolding unit_prop_body_wl_find_unwatched_inv_def by fast
     apply (rule find_is_nat_rel; assumption)
     subgoal by auto
     apply (rule set_conflict_wl; assumption)
     subgoal for x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa f x'a
       by force
     apply (rule propagate_lit_wl_general; assumption)
    subgoal
      by auto
    subgoal using S_S' by (auto simp: eq_commute[of _ _ ! w])
    subgoal by fast
    subgoal by auto
    subgoal for x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa f x'a xa
      using S_S' by (cases X2) (auto simp: eq_commute[of _ _ ! w])
    subgoal by fast
    subgoal for x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa f x'a xa
      using S_S' by (cases X2) (auto simp: eq_commute[of _ _ ! w])
    subgoal by simp ― ‹polarity
    subgoal for x' x1 x2 x1a x2a b Sa X2 K x val_K v ia va L' vaa val_L' vaaa f x'a xa
         x'b Ka Kb val_L'a val_Ka
      by (rule update_blit)
    subgoal by (rule update_clause_wl)
    done

  have [simp]: add_mset a (remove1_mset a M) = M  a ∈# M for a M
    by (metis ab_semigroup_add_class.add.commute add.left_neutral multi_self_add_other_not_self
       remove1_mset_eqE union_mset_add_mset_left)

  show ?eq if inv: unit_propagation_inner_loop_body_l_inv L C' T
    using i_le[OF inv] i_le2[OF inv] C'_dom[OF inv] S_S'
    unfolding i_def[symmetric]
    by (auto simp: ran_m_clause_upd image_mset_remove1_mset_if)
qed

lemma
  fixes S :: 'v twl_st_wl and S' :: 'v twl_st_l and L :: 'v literal and w :: nat
  defines [simp]: C'  fst (watched_by S L ! w)
  defines
    [simp]: T  remove_one_lit_from_wq C' S'

  defines
    [simp]: C''  get_clauses_l S'  C'
  assumes
    S_S': (S, S')  state_wl_l (Some (L, w)) and
    w_le: w < length (watched_by S L) and
    j_w: j  w and
    corr_w: correct_watching_except j w L S and
    blit_in_lit: blits_in_ℒin S and
    inner_loop_inv: unit_propagation_inner_loop_wl_loop_inv L (j, w, S) and
    n: n = size (filter_mset (λ(i, _). i ∉# dom_m (get_clauses_wl S)) (mset (drop w (watched_by S L)))) and
    confl_S: get_conflict_wl S = None
  shows unit_propagation_inner_loop_body_wl_spec: unit_propagation_inner_loop_body_wl L j w S 
    {((i, j, T'), (T, n)).
        (T', T)  state_wl_l (Some (L, j)) 
        correct_watching_except i j L T'  blits_in_ℒin T' 
        j  length (watched_by T' L) 
        length (watched_by S L) =  length (watched_by T' L) 
        i  j 
        (get_conflict_wl T' = None 
           n = size (filter_mset (λ(i, _). i ∉# dom_m (get_clauses_wl T')) (mset (drop j (watched_by T' L))))) 
        (get_conflict_wl T'  None  n = 0)}
     (unit_propagation_inner_loop_body_l_with_skip L (S', n))
  apply (rule order_trans)
   apply (rule unit_propagation_inner_loop_body_wl_wl_int[OF S_S' w_le j_w corr_w inner_loop_inv n
       confl_S])
  apply (subst Down_id_eq)
   apply (rule unit_propagation_inner_loop_body_wl_int_spec[OF S_S' w_le j_w corr_w blit_in_lit inner_loop_inv n
       confl_S])
  done




definition unit_propagation_inner_loop_wl_loop
   :: 'v literal  'v twl_st_wl  (nat × nat × 'v twl_st_wl) nres where
  unit_propagation_inner_loop_wl_loop L S0 = do {
    ASSERT(L ∈# all_lits_st S0);
    let n = length (watched_by S0 L);
    WHILETunit_propagation_inner_loop_wl_loop_inv L(λ(j, w, S). w < n  get_conflict_wl S = None)
      (λ(j, w, S). do {
        unit_propagation_inner_loop_body_wl L j w S
      })
      (0, 0, S0)
  }

lemma blits_in_ℒin_cut_watch:
  assumes corr: blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
  shows blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g(L := take j (g L) @ drop w (g L)))
  using assms by (auto simp: blits_in_ℒin_def dest!: in_set_takeD in_set_dropD)

lemma correct_watching_except_correct_watching_cut_watch:
  assumes corr: correct_watching_except j w L (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
  shows correct_watching (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g(L := take j (g L) @ drop w (g L)))
proof -
  let ?ℒ = all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g(L := take j (g L) @ drop w (g L)))
  have : ?ℒ =  all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
    by (auto simp: all_lits_st_def)
  have
    Heq:
      La i K b'. La ∈# ?ℒ 
      (La = L 
       distinct_watched (take j (g La) @ drop w (g La)) 
       ((i, K, b')∈#mset (take j (g La) @ drop w (g La)) 
           i ∈# dom_m b  K  set (b  i)  K  La  correctly_marked_as_binary b (i, K, b')) 
       ((i, K, b')∈#mset (take j (g La) @ drop w (g La)) 
           b'  i ∈# dom_m b) 
       {#i ∈# fst `# mset (take j (g La) @ drop w (g La)). i ∈# dom_m b#} =
       clause_to_update La (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, {#}, {#})) and
    Hneq:
      La i K b'. La∈#?ℒ 
      (La  L 
       distinct_watched (g La) 
       ((i, K, b')∈#mset (g La)  i ∈# dom_m b  K  set (b  i)  K  La
           correctly_marked_as_binary b (i, K, b')) 
        ((i, K, b')∈#mset (g La)  b'  i ∈# dom_m b) 
       {#i ∈# fst `# mset (g La). i ∈# dom_m b#} =
       clause_to_update La (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, {#}, {#}))
    using corr
    unfolding correct_watching.simps correct_watching_except.simps 
    by fast+
  have
    ((i, K, b')∈#mset ((g(L := take j (g L) @ drop w (g L))) La) 
            i ∈# dom_m b  K  set (b  i)  K  La  correctly_marked_as_binary b (i, K, b')) and
    (i, K, b')∈#mset ((g(L := take j (g L) @ drop w (g L))) La) 
            b'  i ∈# dom_m b and
    {#i ∈# fst `# mset ((g(L := take j (g L) @ drop w (g L))) La).
         i ∈# dom_m b#} =
        clause_to_update La (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, {#}, {#})and
    distinct_watched ((g(L := take j (g L) @ drop w (g L))) La)
  if La∈#?ℒ
  for La i K b'
    apply (cases La = L)
    subgoal
      using Heq[of La i K] that by auto
    subgoal
      using Hneq[of La i K] that by auto
    apply (cases La = L)
    subgoal
      using Heq[of La i K] that by auto
    subgoal
      using Hneq[of La i K] that by auto
    apply (cases La = L)
    subgoal
      using Heq[of La i K] that by auto
    subgoal
      using Hneq[of La i K] that by auto
    apply (cases La = L)
    subgoal
      using Heq[of La i K] that by auto
    subgoal
      using Hneq[of La i K] that by auto
    done
  then show ?thesis
    unfolding correct_watching.simps
    by blast
qed

lemma unit_propagation_inner_loop_wl_loop_alt_def:
  unit_propagation_inner_loop_wl_loop L S0 = do {
    ASSERT(L ∈# all_lits_st S0);
    let (_ :: nat) = (if get_conflict_wl S0 = None then remaining_nondom_wl 0 L S0 else 0);
    let n = length (watched_by S0 L);
    WHILETunit_propagation_inner_loop_wl_loop_inv L(λ(j, w, S). w < n  get_conflict_wl S = None)
      (λ(j, w, S). do {
        unit_propagation_inner_loop_body_wl L j w S
      })
      (0, 0, S0)
  }
  
  unfolding unit_propagation_inner_loop_wl_loop_def Let_def by auto

definition cut_watch_list :: nat  nat  'v literal  'v twl_st_wl  'v twl_st_wl nres where
  cut_watch_list j w L =(λ(M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W). do {
      ASSERT(j  w  j  length (W L)  w  length (W L)  L ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
      RETURN (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W(L := take j (W L) @ drop w (W L)))
    })

definition unit_propagation_inner_loop_wl :: 'v literal  'v twl_st_wl  'v twl_st_wl nres where
  unit_propagation_inner_loop_wl L S0 = do {
     (j, w, S)  unit_propagation_inner_loop_wl_loop L S0;
     ASSERT(j  w  w  length (watched_by S L)  L ∈# all_lits_st S);
     cut_watch_list j w L S
  }

lemma correct_watching_correct_watching_except00:
  correct_watching S  correct_watching_except 0 0 L S
  apply (cases S)
  apply (simp only: correct_watching.simps correct_watching_except.simps
    take0 drop0 append.left_neutral)
  by fast

lemma unit_propagation_inner_loop_wl_spec:
  shows (uncurry unit_propagation_inner_loop_wl, uncurry unit_propagation_inner_loop_l) 
    {((L', T'::'v twl_st_wl), (L, T::'v twl_st_l)). L = L'  (T', T)  state_wl_l (Some (L, 0)) 
      correct_watching T'  blits_in_ℒin T'} 
    {(T', T). (T', T)  state_wl_l None  correct_watching T'  blits_in_ℒin T'} nres_rel
     (is ?fg  ?A  ?Bnres_rel is ?fg  ?A  {(T', T). _  ?P T T'}nres_rel)
proof -
  {
    fix L :: 'v literal and S :: 'v twl_st_wl and S' :: 'v twl_st_l
    assume
      corr_w: correct_watching S blits_in_ℒin S and
      SS': (S, S')  state_wl_l (Some (L, 0))
    text To ease the finding the correspondence between the body of the loops, we introduce
      following function:
    let ?R' = {((i, j, T'), (T, n)).
        (T', T)  state_wl_l (Some (L, j)) 
        correct_watching_except i j L T'  blits_in_ℒin T' 
        j  length (watched_by T' L) 
        length (watched_by S L) = length (watched_by T' L) 
        i  j 
        (get_conflict_wl T' = None 
           n = size (filter_mset (λ(i, _). i ∉# dom_m (get_clauses_wl T')) (mset (drop j (watched_by T' L))))) 
        (get_conflict_wl T'  None  n = 0) 
        blits_in_ℒin T'}
    have inv: unit_propagation_inner_loop_wl_loop_inv L iT'
      if
        iT'_Tn: (iT', Tn)  ?R' and
        unit_propagation_inner_loop_l_inv L Tn
        for Tn iT'
    proof -
      obtain i j :: nat and T' where iT': iT' = (i, j, T') by (cases iT')
      obtain T n where Tn[simp]: Tn = (T, n) by (cases Tn)
      have unit_propagation_inner_loop_l_inv L (T, 0::nat)
        if unit_propagation_inner_loop_l_inv L (T, n) and get_conflict_l T  None
        using that iT'_Tn
        unfolding unit_propagation_inner_loop_l_inv_def iT' prod.case
        apply - apply normalize_goal+
        apply (rule_tac x=x in exI)
        by auto
      then show ?thesis
        unfolding unit_propagation_inner_loop_wl_loop_inv_def iT' prod.simps apply -
        apply (rule exI[of _ T])
        using that by (auto simp: iT')
    qed
    have cond: (j < length (watched_by S L)  get_conflict_wl T' = None) =
      (clauses_to_update_l T  {#}  n > 0)
      if
        iT'_T: (ijT', Tn)  ?R' and
        [simp]: ijT' = (i, jT') jT' = (j, T')  Tn = (T, n)
        for ijT' Tn i j T' n T jT'
    proof -
      have [simp]: size {#(i, _) ∈# mset (drop j xs). i ∉# dom_m b#} =
        size {#i ∈# fst `# mset (drop j xs). i ∉# dom_m b#} for xs b
        apply (induction xs arbitrary: j)
        subgoal by auto
        subgoal premises p for a xs j
          using p[of 0] p
          by (cases j) auto
        done
      have [simp]: size (filter_mset (λi. (i ∈# (dom_m b))) (fst `# (mset (drop j (g L))))) +
          size {#i ∈# fst `# mset (drop j (g L)). i ∉# dom_m b#} =
          length (g L) - j for g j b
        apply (subst size_union[symmetric])
        apply (subst multiset_partition[symmetric])
        by auto
      have [simp]: A  {#}  size A > 0 for A
        by (auto dest!: multi_member_split)
      have length (watched_by T' L) = size (clauses_to_update_wl T' L j) + n + j
        if get_conflict_wl T' = None
        using that iT'_T
        by (cases get_conflict_wl T'; cases T')
          (auto simp add: state_wl_l_def drop_map)
      then show ?thesis
        using iT'_T
        by (cases get_conflict_wl T' = None) auto
    qed
    have remaining: RETURN (if get_conflict_wl S = None then remaining_nondom_wl 0 L S else 0)  SPEC (λ_. True)
      by auto

    have [intro]: unit_propagation_inner_loop_wl_loop_inv L
         (j, w, T)  L ∈# all_lits_st T for j w T
       unfolding unit_propagation_inner_loop_wl_loop_inv_def
          unit_propagation_inner_loop_l_inv_def prod.case
       apply normalize_goal+
       apply (drule twl_struct_invs_no_alien_in_trail[of _ -L])
       apply (simp_all only: in_all_lits_of_mm_uminus_iff twl_st_wl twl_st_l twl_st
         all_lits_def multiset.map_comp comp_def clause_twl_clause_of ac_simps
         in_all_lits_uminus_iff)
      done
    have unit_propagation_inner_loop_l_alt_def: unit_propagation_inner_loop_l L S' = do {
        ASSERT(L ∈# all_lits_of_st_l S');
        n  SPEC (λ_::nat. True);
        (S, n)  WHILETunit_propagation_inner_loop_l_inv L(λ(S, n). clauses_to_update_l S  {#}  0 < n)
              (unit_propagation_inner_loop_body_l_with_skip L) (S', n);
        RETURN S} for L S'
      unfolding unit_propagation_inner_loop_l_def by auto
    have unit_propagation_inner_loop_wl_alt_def: unit_propagation_inner_loop_wl L S = do {
      ASSERT(L ∈# all_lits_st S);
      let (n::nat) = (if get_conflict_wl S = None then remaining_nondom_wl 0 L S else 0);
      (j, w, S)  WHILETunit_propagation_inner_loop_wl_loop_inv L(λ(j, w, T). w < length (watched_by S L)   get_conflict_wl T = None)
         (λ(j, x, y). unit_propagation_inner_loop_body_wl L j x y) (0, 0, S);
      ASSERT (j  w  w  length (watched_by S L)  L ∈# all_lits_st S);
      cut_watch_list j w L S}
      unfolding unit_propagation_inner_loop_wl_loop_alt_def unit_propagation_inner_loop_wl_def
      by auto
    have unit_propagation_inner_loop_wl L S 
             {((T'), T). (T', T)  state_wl_l None  ?P T T'}
              (unit_propagation_inner_loop_l L S')
      (is _   ?R _)
      unfolding unit_propagation_inner_loop_l_alt_def uncurry_def
        unit_propagation_inner_loop_wl_alt_def
      apply (refine_vcg WHILEIT_refine_genR[where
            R' = ?R' and
            R = {((i, j, T'), (T, n)). ((i, j, T'), (T, n))  ?R'  i  j 
                length (watched_by S L) =  length (watched_by T' L) 
               (j  length (watched_by T' L)  get_conflict_wl T'  None) 
               unit_propagation_inner_loop_wl_loop_inv L (i, j, T')}]
          remaining)
      subgoal using SS' by (auto simp: all_lits_def ac_simps)
      subgoal using corr_w SS' by (auto simp: correct_watching_correct_watching_except00)
      subgoal by (rule inv)
      subgoal by (rule cond)
      subgoal for n i'w'T' Tn i' w'T' w' T'
        apply (cases Tn)
        apply (rule order_trans)
        apply (rule unit_propagation_inner_loop_body_wl_spec[of _ fst Tn])
        apply (simp only: prod.case in_pair_collect_simp)
        apply normalize_goal+
        apply (auto simp del: twl_st_of_wl.simps intro!: conc_fun_R_mono)
        done
      subgoal by auto
      subgoal by auto
      subgoal by auto
      subgoal by auto
      subgoal for n i'w'T' Tn i' w'T' j L' w' T'
        apply (cases T')
        by (auto simp: state_wl_l_def cut_watch_list_def all_lits_def intro: blits_in_ℒin_cut_watch
          dest!: correct_watching_except_correct_watching_cut_watch)
      done
  }
  note H = this

  show ?thesis
    unfolding fref_param1
    apply (intro frefI nres_relI)
    by (auto simp: intro!: H)
qed


subsection Outer loop

definition select_and_remove_from_literals_to_update_wl :: 'v twl_st_wl  ('v twl_st_wl × 'v literal) nres where
  select_and_remove_from_literals_to_update_wl S = SPEC(λ(S', L). L ∈# literals_to_update_wl S 
     S' = set_literals_to_update_wl (literals_to_update_wl S - {#L#}) S)

definition unit_propagation_outer_loop_wl_inv where
  unit_propagation_outer_loop_wl_inv S 
    (S'. (S, S')  state_wl_l None 
      correct_watching S  blits_in_ℒin S 
      unit_propagation_outer_loop_l_inv S')

definition unit_propagation_outer_loop_wl :: 'v twl_st_wl  'v twl_st_wl nres where
  unit_propagation_outer_loop_wl S0 =
    WHILETunit_propagation_outer_loop_wl_inv(λS. literals_to_update_wl S  {#})
      (λS. do {
        ASSERT(literals_to_update_wl S  {#});
        (S', L)  select_and_remove_from_literals_to_update_wl S;
        ASSERT(L ∈# all_lits_st S');
        unit_propagation_inner_loop_wl L S'
      })
      (S0 :: 'v twl_st_wl)


lemma blits_in_ℒin_simps[simp]: blits_in_ℒin (set_literals_to_update_wl C  xa)  blits_in_ℒin xa
  by (cases xa; auto simp: blits_in_ℒin_def)

lemma unit_propagation_outer_loop_wl_spec:
  (unit_propagation_outer_loop_wl, unit_propagation_outer_loop_l)
  {(T'::'v twl_st_wl, T).
       (T', T)  state_wl_l None 
       correct_watching T'  blits_in_ℒin T'} f
    {(T', T).
       (T', T)  state_wl_l None 
       correct_watching T'  blits_in_ℒin T'}nres_rel
  (is ?u  ?A f ?B nres_rel)
proof -
  have inv: unit_propagation_outer_loop_wl_inv T'
  if
    (T', T)  {(T', T). (T', T)  state_wl_l None  correct_watching T'  blits_in_ℒin T'} and
    unit_propagation_outer_loop_l_inv T
    for T T'
  unfolding unit_propagation_outer_loop_wl_inv_def
  apply (rule exI[of _ T])
  using that by auto

  have select_and_remove_from_literals_to_update_wl:
   select_and_remove_from_literals_to_update_wl S' 
      {((T', L'), (T, L)). L = L'  (T', T)  state_wl_l (Some (L, 0)) 
         T' = set_literals_to_update_wl (literals_to_update_wl S' - {#L#}) S'  L ∈# literals_to_update_wl S' 
         L ∈# all_lits_st S'
       }
       (select_and_remove_from_literals_to_update S)
    if S: (S', S)  state_wl_l None and get_conflict_wl S' = None and
      corr_w: correct_watching S' and
      inv_l: unit_propagation_outer_loop_l_inv S
    for S :: 'v twl_st_l and S' :: 'v twl_st_wl
  proof -
    obtain M N D NE UE NEk UEk NS US N0 U0 W Q where
      S': S' = (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
      by (cases S') auto
    obtain R where
      S_R: (S, R)  twl_st_l None and
      struct_invs: twl_struct_invs R
      using inv_l unfolding unit_propagation_outer_loop_l_inv_def by blast
    have [simp]: (* ‹trail (stateW_of R) = convert_lits_l N M› *)
      init_clss (stateW_of R) = mset `# (init_clss_lf N) + NE + NEk + NS + N0
      atm_of ` lits_of_l (get_trail R) = atm_of ` lits_of_l M
      using S_R S by (auto simp: twl_st S' twl_st_wl simp flip: stateW_of_def)
    have
      no_dup_q: no_duplicate_queued R and
      alien: cdclW_restart_mset.no_strange_atm (stateW_of R)
      using struct_invs that by (auto simp: twl_struct_invs_def
          cdclW_restart_mset.cdclW_all_struct_inv_def
        pcdcl_all_struct_invs_def stateW_of_def)
    then have H1: L ∈# all_lits_of_mm (mset `# ran_mf N + (NE + UE) + (NEk + UEk) + (NS + US) + (N0 + U0))
      if LQ: L ∈# Q for L
    proof -
      have [simp]: (f o g) ` I = f ` g ` I for f g I
        by auto
      obtain K where L = - lit_of K and K ∈# mset (trail (stateW_of R))
        using that no_dup_q LQ S_R S
        mset_le_add_mset_decr_left2[of L remove1_mset L Q Q]
        by (fastforce simp: S' cdclW_restart_mset.no_strange_atm_def cdclW_restart_mset_state
          all_lits_of_mm_def atms_of_ms_def twl_st_l_def state_wl_l_def uminus_lit_swap
          convert_lit.simps
          dest!: multi_member_split[of L Q] mset_subset_eq_insertD in_convert_lits_lD2)
      from imageI[OF this(2), of atm_of o lit_of]
      have atm_of L  atm_of ` lits_of_l (get_trail_wl S') and
        [simp]: atm_of ` lits_of_l (trail (stateW_of R)) = atm_of ` lits_of_l (get_trail_wl S')
        using S_R S S L = - lit_of K
        by (simp_all add: twl_st image_image[symmetric]
            lits_of_def[symmetric])
      then have atm_of L  atm_of ` lits_of_l M
        using S'  by auto
      moreover {
        have atm_of ` lits_of_l M
          (xset_mset (init_clss_lf N). atm_of ` set x) 
           (xset_mset NE. atms_of x) 
           (xset_mset NEk. atms_of x) 
           (xset_mset NS. atms_of x) 
           (xset_mset N0. atms_of x)
          using that alien unfolding cdclW_restart_mset.no_strange_atm_def
          by (auto simp: S' cdclW_restart_mset.no_strange_atm_def cdclW_restart_mset_state
              all_lits_of_mm_def atms_of_ms_def simp flip: stateW_of_def)
        then have atm_of ` lits_of_l M  (xset_mset (init_clss_lf N). atm_of ` set x) 
         (xset_mset NE. atms_of x) 
         (xset_mset NEk. atms_of x) 
         (xset_mset NS. atms_of x) 
         (xset_mset N0. atms_of x)
        unfolding image_Un[symmetric]
          set_append[symmetric]
          append_take_drop_id
          .
        then have atm_of ` lits_of_l M  atms_of_mm (mset `# init_clss_lf N + NE + NEk + NS + N0)
          by (smt UN_Un Un_iff append_take_drop_id atms_of_ms_def atms_of_ms_mset_unfold set_append
              set_image_mset set_mset_mset set_mset_union subset_eq)
      }
      ultimately have atm_of L  atms_of_mm (mset `# ran_mf N + NE + NEk + NS + N0)
        using that
        unfolding all_lits_of_mm_union atms_of_ms_union all_clss_lf_ran_m[symmetric]
          image_mset_union set_mset_union
        by auto
      then show ?thesis
        using that by (auto simp: in_all_lits_of_mm_ain_atms_of_iff)
    qed
    have H: clause_to_update L S = {#i ∈# fst `# mset (W L). i ∈# dom_m N#}
       L ∈# all_lits_st ([], N, None, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, W)
        if L ∈# Q for L
      using corr_w that S H1[OF that] by (auto simp: correct_watching.simps S' clause_to_update_def
        Ball_def ac_simps all_conj_distrib all_lits_st_def all_lits_def
        dest!: multi_member_split)
    show ?thesis
      unfolding select_and_remove_from_literals_to_update_wl_def select_and_remove_from_literals_to_update_def
      apply (rule RES_refine)
      unfolding Bex_def
      by (rule_tac x=(set_clauses_to_update_l (clause_to_update (snd s) S)
              (set_literals_to_update_l
                (remove1_mset (snd s) (literals_to_update_l S)) S), snd s) in exI)
         (use that S' S in auto simp: correct_watching.simps clauses_def state_wl_l_def
          mset_take_mset_drop_mset' cdclW_restart_mset_state all_lits_of_mm_union
          dest: H1 H)
  qed
  have conflict_None: get_conflict_wl T = None
    if
      literals_to_update_wl T  {#} and
      inv1: unit_propagation_outer_loop_wl_inv T
      for T
  proof -
    obtain T' where
      2: (T, T')  state_wl_l None and
      inv2: unit_propagation_outer_loop_l_inv T'
      using inv1 unfolding unit_propagation_outer_loop_wl_inv_def by blast
    obtain T'' where
      3: (T', T'')  twl_st_l None and
      twl_struct_invs T''
      using inv2 unfolding unit_propagation_outer_loop_l_inv_def by blast
    then have get_conflict T''  None 
       clauses_to_update T'' = {#}  literals_to_update T'' = {#}
       unfolding twl_struct_invs_def by fast
    then show ?thesis
      using that 2 3 by (auto simp: twl_st_wl twl_st twl_st_l)
  qed
  show ?thesis
    unfolding unit_propagation_outer_loop_wl_def unit_propagation_outer_loop_l_def
    apply (intro frefI nres_relI)
    apply (refine_rcg select_and_remove_from_literals_to_update_wl
      unit_propagation_inner_loop_wl_spec[unfolded fref_param1, THEN fref_to_Down_curry])
    subgoal by (rule inv)
    subgoal by auto
    subgoal by auto
    subgoal by (rule conflict_None)
    subgoal for T' T by (auto simp: )
    subgoal by (auto simp: twl_st_wl ac_simps)
    subgoal by auto
    done
qed


subsection Decide or Skip

definition find_unassigned_lit_wl :: 'v twl_st_wl  ('v twl_st_wl × 'v literal option) nres where
  find_unassigned_lit_wl = (λS.
     SPEC (λ(T, L). T = S 
         (L  None 
            undefined_lit (get_trail_wl S) (the L) 
            the L ∈# all_lits_st S) 
         (L = None  (L'. undefined_lit (get_trail_wl S) L' 
            L' ∈# all_lits_st S)))
     )

definition decide_wl_or_skip_pre where
decide_wl_or_skip_pre S 
  (S'. (S, S')  state_wl_l None 
   decide_l_or_skip_pre S'  blits_in_ℒin S
  )

definition decide_lit_wl :: 'v literal  'v twl_st_wl  'v twl_st_wl where
  decide_lit_wl = (λL' (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W).
      (Decided L' # M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#- L'#}, W))


definition decide_wl_or_skip :: 'v twl_st_wl  (bool × 'v twl_st_wl) nres where
  decide_wl_or_skip S = (do {
    ASSERT(decide_wl_or_skip_pre S);
    (S, L)  find_unassigned_lit_wl S;
    case L of
      None  RETURN (True, S)
    | Some L  RETURN (False, decide_lit_wl L S)
  })


lemma decide_wl_or_skip_spec:
  (decide_wl_or_skip, decide_l_or_skip)
     {(T':: 'v twl_st_wl, T).
          (T', T)  state_wl_l None 
          correct_watching T'  blits_in_ℒin T' 
          get_conflict_wl T' = None} 
        {((b', T'), (b, T)). b' = b 
         (T', T)  state_wl_l None 
          correct_watching T'  blits_in_ℒin T'}nres_rel
proof -
  have find_unassigned_lit_wl: find_unassigned_lit_wl S'
      {((T, L), L'). T = S'  L = L'}
        (find_unassigned_lit_l S)
    if (S', S)  state_wl_l None
    for S :: 'v twl_st_l and S' :: 'v twl_st_wl
    using that
    by (auto simp: find_unassigned_lit_l_def find_unassigned_lit_wl_def
      intro!: RES_refine)

  have option: (x, x')  Idoption_rel if x = x' for x x'
    using that by (auto)
  show ?thesis
    unfolding decide_wl_or_skip_def decide_l_or_skip_def
    apply (refine_vcg find_unassigned_lit_wl option)
    subgoal unfolding decide_wl_or_skip_pre_def by fast
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal for S S'
      by (cases S) (auto simp: correct_watching.simps clause_to_update_def
          decide_lit_l_def decide_lit_wl_def state_wl_l_def blits_in_ℒin_def)
    done
qed


subsection Skip or Resolve


definition mop_tl_state_wl_pre :: 'v twl_st_wl  bool where
mop_tl_state_wl_pre S 
   (S'. (S, S')  state_wl_l None  correct_watching S  mop_tl_state_l_pre S')

definition tl_state_wl :: 'v twl_st_wl  'v twl_st_wl where
  tl_state_wl = (λ(M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q).
    (tl M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q))

definition mop_tl_state_wl :: 'v twl_st_wl  (bool × 'v twl_st_wl) nres where
  mop_tl_state_wl = (λS. do {
    ASSERT(mop_tl_state_wl_pre S);
    RETURN(False, tl_state_wl S)})

definition resolve_cls_wl' :: 'v twl_st_wl  nat  'v literal  'v clause where
resolve_cls_wl' S C L  =
   remove1_mset L (remove1_mset (-L) (the (get_conflict_wl S) ∪# (mset (get_clauses_wl S  C))))

definition update_confl_tl_wl :: 'v literal  nat  'v twl_st_wl  bool × 'v twl_st_wl where
  update_confl_tl_wl = (λL C (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q).
     let D = resolve_cls_wl' (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q) C L in
        (False, (tl M, N, Some D, NE, UE, NEk, UEk, NS, US, N0, U0, WS, Q)))

definition update_confl_tl_wl_pre :: 'v literal  nat  'v twl_st_wl  bool where
  update_confl_tl_wl_pre L C S 
    (S'. (S, S')  state_wl_l None  update_confl_tl_l_pre L C S' 
     correct_watching S  blits_in_ℒin S)

definition mop_update_confl_tl_wl :: 'v literal  nat  'v twl_st_wl  (bool × 'v twl_st_wl) nres where
  mop_update_confl_tl_wl = (λL C S. do {
     ASSERT(update_confl_tl_wl_pre L C S);
     RETURN (update_confl_tl_wl L C S)
  })

definition mop_hd_trail_wl_pre :: 'v twl_st_wl  bool where
mop_hd_trail_wl_pre S 
   (S'. (S, S')  state_wl_l None  mop_hd_trail_l_pre S' 
     correct_watching S)


definition mop_hd_trail_wl :: 'v twl_st_wl  ('v literal × nat) nres where
  mop_hd_trail_wl S = do{
     ASSERT(mop_hd_trail_wl_pre S);
     SPEC(λ(L, C). Propagated L C = hd (get_trail_wl S))
  }

definition skip_and_resolve_loop_wl_inv :: 'v twl_st_wl  bool  'v twl_st_wl  bool where
  skip_and_resolve_loop_wl_inv S0 brk S 
    (S' S'0. (S, S')  state_wl_l None 
      (S0, S'0)  state_wl_l None 
     skip_and_resolve_loop_inv_l S'0 brk S' 
        correct_watching S)

definition mop_lit_notin_conflict_wl :: 'v literal  'v twl_st_wl  bool nres where
  mop_lit_notin_conflict_wl L S = do {
    ASSERT(get_conflict_wl S  None  -L ∉# the (get_conflict_wl S)  L ∈# all_lits_st S);
    RETURN (L ∉# the (get_conflict_wl S))
  }

definition mop_maximum_level_removed_wl_pre :: 'v literal  'v twl_st_wl  bool where
mop_maximum_level_removed_wl_pre L S 
   (S'. (S, S')  state_wl_l None  mop_maximum_level_removed_l_pre L S' 
      correct_watching S)

definition mop_maximum_level_removed_wl :: 'v literal  'v twl_st_wl  bool nres where
mop_maximum_level_removed_wl L S = do {
   ASSERT (mop_maximum_level_removed_wl_pre L S);
   RETURN (get_maximum_level (get_trail_wl S) (remove1_mset (-L) (the (get_conflict_wl S))) =
      count_decided (get_trail_wl S))
}

definition skip_and_resolve_loop_wl :: 'v twl_st_wl  'v twl_st_wl nres where
  skip_and_resolve_loop_wl S0 =
    do {
      ASSERT(get_conflict_wl S0  None);
      (_, S) 
        WHILETλ(brk, S). skip_and_resolve_loop_wl_inv S0 brk S(λ(brk, S). ¬brk  ¬is_decided (hd (get_trail_wl S)))
        (λ(_, S).
          do {
            (L, C)  mop_hd_trail_wl S;
            b  mop_lit_notin_conflict_wl (-L) S;
            if b then
              mop_tl_state_wl S
            else do {
              b  mop_maximum_level_removed_wl L S;
              if b
              then
                mop_update_confl_tl_wl L C S
              else
                do {RETURN (True, S)}
           }
          }
        )
        (False, S0);
      RETURN S
    }
  

lemma tl_state_wl_tl_state_l:
  (S, S')  state_wl_l None  correct_watching S   blits_in_ℒin S 
    mop_tl_state_wl S  (bool_rel ×f {(S, T). (S, T)  state_wl_l None  correct_watching S  blits_in_ℒin S})
      (mop_tl_state_l S')
  by (cases S) (auto simp: state_wl_l_def tl_state_wl_def tl_state_l_def blits_in_ℒin_def
        mop_tl_state_wl_def mop_tl_state_l_def mop_tl_state_wl_pre_def
        assert_bind_spec_conv correct_watching.simps clause_to_update_def
     intro!: ASSERT_refine_right ASSERT_leI)


lemma mop_update_confl_tl_wl_mop_update_confl_tl_l:
  (S, S')  state_wl_l None  correct_watching S   blits_in_ℒin S 
   ((L, C), (L', C'))  Id 
    mop_update_confl_tl_wl L C S 
     (bool_rel ×f {(S, T). (S, T)  state_wl_l None  correct_watching S  blits_in_ℒin S})
      (mop_update_confl_tl_l L' C' S')
  by (cases S) (auto simp: state_wl_l_def tl_state_wl_def tl_state_l_def blits_in_ℒin_def
        mop_update_confl_tl_wl_def mop_update_confl_tl_l_def
        update_confl_tl_wl_def update_confl_tl_l_def resolve_cls_l'_def
        resolve_cls_wl'_def update_confl_tl_wl_pre_def
        assert_bind_spec_conv correct_watching.simps clause_to_update_def
     intro!: ASSERT_refine_right ASSERT_leI)

lemma mop_hd_trail_wl_mop_hd_trail_l:
  (S, S')  state_wl_l None  correct_watching S   blits_in_ℒin S 
   mop_hd_trail_wl S
      (Id)
        (mop_hd_trail_l S')
  unfolding mop_hd_trail_wl_def mop_hd_trail_l_def
  apply refine_rcg
  subgoal unfolding mop_hd_trail_wl_pre_def by blast
  subgoal by auto
  done

lemma mop_lit_notin_conflict_wl_mop_lit_notin_conflict_l:
  (S, S')  state_wl_l None  correct_watching S   blits_in_ℒin S 
   L = L' 
    mop_lit_notin_conflict_wl L S  bool_rel (mop_lit_notin_conflict_l L' S')
  unfolding mop_lit_notin_conflict_wl_def mop_lit_notin_conflict_l_def
  apply refine_rcg
  subgoal by auto
  subgoal by auto
  subgoal unfolding all_lits_def by (auto simp: ac_simps)
  subgoal by auto
  done

lemma mop_maximum_level_removed_wl_mop_maximum_level_removed_l:
  (S, S')  state_wl_l None  correct_watching S   blits_in_ℒin S 
   L = L' 
    mop_maximum_level_removed_wl L S  bool_rel (mop_maximum_level_removed_l L' S')
  unfolding mop_maximum_level_removed_wl_def mop_maximum_level_removed_l_def
  apply refine_rcg
  subgoal unfolding mop_maximum_level_removed_wl_pre_def by blast
  subgoal by auto
  done


lemma skip_and_resolve_loop_wl_spec:
  (skip_and_resolve_loop_wl, skip_and_resolve_loop_l)
     {(T'::'v twl_st_wl, T).
         (T', T)  state_wl_l None 
          correct_watching T'  blits_in_ℒin T' 
          0 < count_decided (get_trail_wl T')} 
      {(T', T).
         (T', T)  state_wl_l None 
          correct_watching T'  blits_in_ℒin T'}nres_rel
  (is ?s  ?A  ?Bnres_rel)
proof -
  have get_conflict_wl: ((False, S'), False, S)
     Id ×r {(T', T). (T', T)  state_wl_l None  correct_watching T'  blits_in_ℒin T'}
    (is _  ?B)
    if (S', S)  state_wl_l None and correct_watching S' and blits_in_ℒin S'
    for S :: 'v twl_st_l and S' :: 'v twl_st_wl
    using that by (cases S') (auto simp: state_wl_l_def)
  have [simp]: correct_watching  (tl aa, ca, da, ea, fa, NEk, UEk, NS, US, N0, U0, ha, h) 
    correct_watching (aa, ca, None, ea, fa, NEk, UEk, NS, US, N0, U0, ha, h)
    for aa ba ca L da ea fa ha h NS US NEk UEk
    by (auto simp: correct_watching.simps tl_state_wl_def clause_to_update_def)
  have [simp]: NO_MATCH None da  correct_watching  (aa, ca, da, ea, fa, NEk, UEk, NS, US, N0, U0, ha, h) 
    correct_watching (aa, ca, None, ea, fa, NEk, UEk, NS, US, N0, U0, ha, h)
    for aa ba ca L da ea fa ha h NS US N0 U0
    by (auto simp: correct_watching.simps tl_state_wl_def clause_to_update_def)

  have inv: skip_and_resolve_loop_wl_inv S' b' T'
    if
      (S', S)  ?A and
      get_conflict_wl S'  None and
      bt_inv: case bT of (x, xa)  skip_and_resolve_loop_inv_l S x xa and
      (b'T', bT)  ?B and
      b'T': b'T' = (b', T')
    for S' S b'T' bT b' T'
  proof -
    obtain b T where bT: bT = (b, T) by (cases bT)
    show ?thesis
      unfolding skip_and_resolve_loop_wl_inv_def
      apply (rule exI[of _ T])
      apply (rule exI[of _ S])
      using that by (auto simp: bT b'T')
  qed

  show H: ?s  ?A  {(T', T). (T', T)  state_wl_l None  correct_watching T'  blits_in_ℒin T'}nres_rel
    unfolding skip_and_resolve_loop_wl_def skip_and_resolve_loop_l_def
    apply (refine_rcg get_conflict_wl tl_state_wl_tl_state_l
       mop_update_confl_tl_wl_mop_update_confl_tl_l mop_hd_trail_wl_mop_hd_trail_l
       mop_lit_notin_conflict_wl_mop_lit_notin_conflict_l
       mop_maximum_level_removed_wl_mop_maximum_level_removed_l)
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by (rule inv)
    subgoal by auto
    subgoal by auto
    subgoal by (auto intro!: tl_state_wl_tl_state_l)
    subgoal for S' S b'T' bT b' T' by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    done
qed


subsection Backtrack

definition find_decomp_wl :: 'v literal  'v twl_st_wl  'v twl_st_wl nres where
  find_decomp_wl = (λL (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W).
    SPEC(λS. K M2 M1. S = (M1, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) 
         (Decided K # M1, M2)  set (get_all_ann_decomposition M) 
          get_level M K = get_maximum_level M (the D - {#-L#}) + 1))

definition find_lit_of_max_level_wl :: 'v twl_st_wl  'v literal  'v literal nres where
  find_lit_of_max_level_wl =  (λ(M, N, D, NE, UE, NEk, UEk, NS, US, Q, W) L.
    SPEC(λL'. L' ∈# remove1_mset (-L) (the D)  get_level M L' = get_maximum_level M (the D - {#-L#})))


fun extract_shorter_conflict_wl :: 'v twl_st_wl  'v twl_st_wl nres where
  extract_shorter_conflict_wl (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) = SPEC(λS.
     D'. D' ⊆# the D  S = (M, N, Some D', NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) 
     clause `# twl_clause_of `# ran_mf N + NE + NEk + UE + UEk + NS + US ⊨pm D'  -lit_of (hd M) ∈# D')

declare extract_shorter_conflict_wl.simps[simp del]
lemmas extract_shorter_conflict_wl_def = extract_shorter_conflict_wl.simps


definition backtrack_wl_inv where
  backtrack_wl_inv S  (S'. (S, S')  state_wl_l None  backtrack_l_inv S'  correct_watching S  blits_in_ℒin S)
  

text Rougly: we get a fresh index that has not yet been used.
definition get_fresh_index_wl :: 'v clauses_l  _  _  nat nres where
get_fresh_index_wl N NUE W = SPEC(λi. i > 0  i ∉# dom_m N 
   (L ∈# all_lits_of_mm (mset `# ran_mf N + NUE) . i  fst ` set (W L)))

definition (in -) list_of_mset2
  :: 'v literal  'v literal  'v clause  'v clause_l nres
where
  list_of_mset2 L L' D =
    SPEC (λE. mset E = D  E!0 = L  E!1 = L'  length E  2)

definition propagate_bt_wl_pre where
  propagate_bt_wl_pre L L' S 
   (S'. (S, S')  state_wl_l None  propagate_bt_l_pre L L' S')

definition propagate_bt_wl :: 'v literal  'v literal  'v twl_st_wl  'v twl_st_wl nres where
  propagate_bt_wl = (λL L' (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W). do {
    ASSERT(propagate_bt_wl_pre L L' (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
    D''  list_of_mset2 (-L) L' (the D);
    i  get_fresh_index_wl N (NE + UE + NEk + UEk + NS + US + N0 + U0) W;
    let b = (length ([-L, L'] @ (remove1 (-L) (remove1 L' D''))) = 2);
    M  cons_trail_propagate_l (-L) i M;
    RETURN (M,
        fmupd i ([-L, L'] @ (remove1 (-L) (remove1 L' D'')), False) N,
          None, NE, UE, NEk, UEk, NS, US, N0, U0, {#L#}, W(-L:= W (-L) @ [(i, L', b)], L':= W L' @ [(i, -L, b)]))
      })

definition propagate_unit_bt_wl_pre where
  propagate_unit_bt_wl_pre L S 
    (S'. (S, S')  state_wl_l None  propagate_unit_bt_l_pre L S')

definition propagate_unit_bt_wl :: 'v literal  'v twl_st_wl  'v twl_st_wl nres where
  propagate_unit_bt_wl = (λL (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W). do {
    ASSERT(L ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
    ASSERT(propagate_unit_bt_wl_pre L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W));
    M  cons_trail_propagate_l (-L) 0 M;
    RETURN (M, N, None, NE, UE, NEk, add_mset (the D) UEk, NS, US, N0, U0, {#L#}, W)})

definition mop_lit_hd_trail_wl_pre :: 'v twl_st_wl  bool where
mop_lit_hd_trail_wl_pre S 
  (S'. (S, S')  state_wl_l None  mop_lit_hd_trail_l_pre S')

definition mop_lit_hd_trail_wl :: 'v twl_st_wl  ('v literal) nres where
  mop_lit_hd_trail_wl S = do{
     ASSERT(mop_lit_hd_trail_wl_pre S);
     SPEC(λL. L = lit_of (hd (get_trail_wl S)))
  }

definition backtrack_wl :: 'v twl_st_wl  'v twl_st_wl nres where
  backtrack_wl S =
    do {
      ASSERT(backtrack_wl_inv S);
      L  mop_lit_hd_trail_wl S;
      S  extract_shorter_conflict_wl S;
      S  find_decomp_wl L S;

      if size (the (get_conflict_wl S)) > 1
      then do {
        L'  find_lit_of_max_level_wl S L;
        propagate_bt_wl L L' S
      }
      else do {
        propagate_unit_bt_wl L S
     }
  }

lemma all_lits_st_learn_simps:
  assumes
    L1: L1 ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) and
    L2: L2 ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) and
    UW: set_mset (all_lits_of_m (mset UW)) 
       set_mset (all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)) and
    i_dom: i ∉# dom_m N
  shows
    set_mset (all_lits_st (M, fmupd i (L1 # L2 # UW, b') N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)) =
     set_mset (all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W))
  using assms
  by (auto 5 3 simp: all_lits_st_def all_lits_def all_lits_of_mm_union all_lits_of_mm_add_mset
        ran_m_mapsto_upd_notin all_lits_of_m_add_mset in_all_lits_of_mm_ain_atms_of_iff
        in_all_lits_of_m_ain_atms_of_iff atm_of_eq_atm_of)

lemma correct_watching_learn2:
  assumes
    L1: L1 ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) and
    L2: L2 ∈# all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W) and
    UW: set_mset (all_lits_of_m (mset UW)) 
       set_mset (all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)) and
    i_dom: i ∉# dom_m N and
    fresh: L. L∈#all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)  i  fst ` set (W L) and
    [iff]: L1  L2 and
    b: b  length (L1 # L2 # UW) = 2
  shows
  correct_watching (K # M, fmupd i (L1 # L2 # UW, b') N,
    D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W (L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])) 
  correct_watching (M, N, D', NE, UE, NEk, UEk, NS, US, N0, U0, Q', W)
  (is ?l  ?c is correct_watching (_, ?N, _) = _)
proof -
  let ?S = (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q, W)
  let ?T = (M, fmupd i (L1 # L2 # UW, b') N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, W)
  let ?ℒ = all_lits_st ?S
  have ℒ_mono: set_mset ?ℒ = set_mset (all_lits_st ?T)
    using assms(1-3) i_dom
    by (auto 5 3 simp: all_lits_st_def all_lits_def all_lits_of_mm_union all_lits_of_mm_add_mset
        ran_m_mapsto_upd_notin all_lits_of_m_add_mset in_all_lits_of_mm_ain_atms_of_iff
        in_all_lits_of_m_ain_atms_of_iff atm_of_eq_atm_of)

  have ℒ𝒯: set_mset ?ℒ = set_mset (all_lits_st (K # M, fmupd i (L1 # L2 # UW, b') N, D, NE, UE,
       NEk, UEk, NS, US, N0, U0, Q, W (L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])))
      using ℒ_mono by (auto simp: all_lits_st_def)

  have [iff]: L2  L1
    using L1  L2 by (subst eq_commute)
  have [simp]: clause_to_update L1 (M, fmupd i (L1 # L2 # UW, b') N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}) =
         add_mset i (clause_to_update L1 (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})) for L2 UW
    using i_dom
    by (auto simp: clause_to_update_def intro: filter_mset_cong)
  have [simp]: clause_to_update L2 (M, fmupd i (L1 # L2 # UW, b') N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}) =
         add_mset i (clause_to_update L2 (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})) for L1 UW
    using i_dom
    by (auto simp: clause_to_update_def intro: filter_mset_cong)
  have [simp]: x  L1  x  L2 
   clause_to_update x (M, fmupd i (L1 # L2 # UW, b') N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}) =
        clause_to_update x (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}) for x UW
    using i_dom
    by (auto simp: clause_to_update_def intro: filter_mset_cong)

  have H':
     {#ia ∈# fst `# mset (W x). ia = i  ia ∈# dom_m N#} =  {#ia ∈# fst `# mset (W x). ia ∈# dom_m N#}
     if x ∈# ?ℒ for x
    using i_dom fresh[of x] that
    by (auto simp: clause_to_update_def intro!: filter_mset_cong)
  have [simp]:
    clause_to_update L1 (K # M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}) =
       clause_to_update L1 (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
    for L1 N D NE UE M K NS US
    by (auto simp: clause_to_update_def)

  show ?thesis
  proof (rule iffI)
    assume corr: ?l
    have
      H: L ia K' b''. (L∈#?ℒ 
      distinct_watched ((W(L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])) L) 
      ((ia, K', b'')∈#mset ((W(L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])) L) 
          ia ∈# dom_m (fmupd i (L1 # L2 # UW, b') N) 
          K'  set (fmupd i (L1 # L2 # UW, b') N  ia)  K'  L 
          correctly_marked_as_binary (fmupd i (L1 # L2 # UW, b') N) (ia, K', b'') ) 
      ((ia, K', b'')∈#mset ((W(L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])) L) 
          b''  ia ∈# dom_m (fmupd i (L1 # L2 # UW, b') N)) 
      {#ia ∈# fst `#
              mset ((W(L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])) L).
       ia ∈# dom_m (fmupd i (L1 # L2 # UW, b') N)#} =
      clause_to_update L
       (K # M, fmupd i (L1 # L2 # UW, b') N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))
      using corr unfolding correct_watching.simps ℒ𝒯
      by fast+

    have x ∈# ?ℒ 
          distinct_watched (W x) 
         (xa ∈# mset (W x)  (((case xa of (i, K, b'')  i ∈# dom_m N  K  set (N  i)  K  x 
           correctly_marked_as_binary N (i, K, b'')) 
           (case xa of (i, K, b'')  b''  i ∈# dom_m N)))) 
         {#i ∈# fst `# mset (W x). i ∈# dom_m N#} = clause_to_update x (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
      for x xa
      supply correctly_marked_as_binary.simps[simp]
      using H[of x fst xa fst (snd xa) snd (snd xa)] fresh[of x] i_dom
      apply (cases x = L1; cases x = L2)
      subgoal
        by (cases xa)
          (auto dest!: multi_member_split simp: H')
      subgoal
        by (cases xa) (force simp add: H' split: if_splits)
      subgoal
        by (cases xa)
          (force simp add: H' split: if_splits)
      subgoal
        by (cases xa)
          (force simp add: H' split: if_splits)
      done
    then show ?c
      unfolding correct_watching.simps Ball_def
      by (auto 5 5 simp add: all_lits_of_mm_add_mset all_lits_of_m_add_mset
          all_conj_distrib all_lits_of_mm_union clause_to_update_def
         dest: multi_member_split)
  next
    assume corr: ?c
    have ?ℒ = all_lits_st (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, Q', W)
      by auto
    then have
      H: L ia K' b''. (L∈# ?ℒ 
      distinct_watched (W L) 
      ((ia, K', b'')∈#mset (W L) 
          ia ∈# dom_m N 
          K'  set (N  ia)  K'  L  correctly_marked_as_binary N (ia, K', b'')) 
      ((ia, K', b'')∈#mset (W L)  b''  ia ∈# dom_m N) 
      {#ia ∈# fst `# mset (W L). ia ∈# dom_m N#} = clause_to_update L (M, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#}))
      using corr unfolding correct_watching.simps clause_to_update_def get_clauses_l.simps
      by auto
    have x ∈# ?ℒ 
         distinct_watched ((W(L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])) x) 
         (xa ∈# mset ((W(L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])) x) 
               (case xa of (ia, K, b'')  ia ∈# dom_m (fmupd i (L1 # L2 # UW, b') N) 
                 K  set (fmupd i (L1 # L2 # UW, b') N  ia)  K  x 
                    correctly_marked_as_binary (fmupd i (L1 # L2 # UW, b') N) (ia, K, b''))) 
         (xa ∈# mset ((W(L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])) x) 
               (case xa of (ia, K, b'')  b''  ia ∈# dom_m (fmupd i (L1 # L2 # UW, b') N))) 
         {#ia ∈# fst `# mset ((W(L1 := W L1 @ [(i, L2, b)], L2 := W L2 @ [(i, L1, b)])) x). ia ∈# dom_m (fmupd i (L1 # L2 # UW, b') N)#} =
         clause_to_update x (K # M, fmupd i (L1 # L2 # UW, b') N, D, NE, UE, NEk, UEk, NS, US, N0, U0, {#}, {#})
      for x :: 'a literal and xa
      supply correctly_marked_as_binary.simps[simp]
      using H[of x fst xa fst (snd xa) snd (snd xa)] fresh[of x] i_dom b
      apply (cases x = L1; cases x = L2)
      subgoal
        by (cases xa)
          (auto dest!: multi_member_split simp: H')
      subgoal
        using ℒ_mono
        by (cases xa)
          (auto dest!: multi_member_split simp: H')
      subgoal
        by (cases xa)
          (auto dest!: multi_member_split simp: H')
      subgoal
        by (cases xa)
          (auto dest!: multi_member_split simp: H')
      done
  then show ?l
    using ℒ_mono
    unfolding correct_watching.simps Ball_def
    by auto
  qed
qed

lemma all_lits_fmupd_new[simp]:
   i ∉# dom_m NU  all_lits (fmupd i C NU) NUE = all_lits_of_m (mset (fst C)) + all_lits NU NUE
  unfolding all_lits_def by (auto simp: ran_m_mapsto_upd_notin all_lits_of_mm_add_mset)



lemma blits_in_ℒin_keep_watch''':
  assumes
    K': K' ∈# all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
      L' ∈# all_lits_st (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g) and
    w: blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g)
  shows blits_in_ℒin (a, b, c, d, e, NEk, UEk, NS, US, N0, U0, f, g (K := (g K) @[(i, K', b')], L := g L @ [(i', L', b'')]))
  using assms
  unfolding blits_in_ℒin_def
  by (auto split: if_splits elim!: in_set_upd_cases)

lemma backtrack_wl_spec:
  (backtrack_wl, backtrack_l)
     {(T'::'v twl_st_wl, T).
          (T', T)  state_wl_l None 
          correct_watching T'  blits_in_ℒin T' 
          get_conflict_wl T'  None 
          get_conflict_wl T'  Some {#}} 
        {(T', T).
          (T', T)  state_wl_l None 
          correct_watching T'  blits_in_ℒin T'}nres_rel
  (is ?bt  ?A  ?Bnres_rel)
proof -
  have H: A ⊨p C  ¬B ⊨p C  A  B  False for A B C
    by (meson true_clss_cls_subset)
  have extract_shorter_conflict_wl: extract_shorter_conflict_wl S'
      {(U'::'v twl_st_wl, U).
          (U', U)  state_wl_l None  equality_except_conflict_wl U' S' 
          the (get_conflict_wl U') ⊆# the (get_conflict_wl S') 
          get_conflict_wl U'  None  correct_watching U'  blits_in_ℒin U'}
        (extract_shorter_conflict_l S)
    (is _   ?extract _)
    if  (S', S)  ?A
    for S' S
    apply (cases S'; cases S)
    apply clarify
    unfolding extract_shorter_conflict_wl_def extract_shorter_conflict_l_def
    apply (refine_rcg, rule RES_refine)
    using that
    by  (auto simp: extract_shorter_conflict_wl_def extract_shorter_conflict_l_def
        mset_take_mset_drop_mset' state_wl_l_def correct_watching.simps Un_assoc
      clause_to_update_def blits_in_ℒin_def dest!: multi_member_split
      dest: H)
  have find_decomp_wl: find_decomp_wl L T'
      {(U'::'v twl_st_wl, U).
          (U', U)  state_wl_l None  equality_except_trail_wl U' T' 
          (M. get_trail_wl T' = M @ get_trail_wl U') 
          correct_watching U'  blits_in_ℒin U'} (find_decomp L' T)
    (is _   ?find _)
    if (S', S)  ?A L = L' (T', T)  ?extract S'
    for S' S T T' L L'
    using that
    apply (cases T; cases T')
    apply clarify
    unfolding find_decomp_wl_def find_decomp_def prod.case
    apply (rule RES_refine)
    apply (auto 5 5 simp add: state_wl_l_def find_decomp_wl_def find_decomp_def correct_watching.simps
        clause_to_update_def blits_in_ℒin_def dest!: multi_member_split)
    done

  have find_lit_of_max_level_wl: find_lit_of_max_level_wl T' LLK'
      {(L', L). L = L'  L' ∈# the (get_conflict_wl T')  L' ∈# the (get_conflict_wl T') - {#-LLK'#}}
         (find_lit_of_max_level_l T L)
    (is _   ?find_lit _)
    if L = LLK' (T', T)  ?find S'
    for S' S T T' L LLK'
    using that
    apply (cases T; cases T'; cases S')
    apply clarify
    unfolding find_lit_of_max_level_wl_def find_lit_of_max_level_l_def prod.case
    apply (rule RES_refine)
    apply (auto simp add: find_lit_of_max_level_wl_def find_lit_of_max_level_def state_wl_l_def
     dest: in_diffD)
    done
  have empty: literals_to_update_wl S' = {#} if bt: backtrack_wl_inv S' for S'
    using bt apply -
    unfolding backtrack_wl_inv_def backtrack_l_inv_def
    apply normalize_goal+
    apply (auto simp: twl_struct_invs_def)
    done

  have [refine]: (S, S')  {(T', T).
        (T', T)  state_wl_l None 
        correct_watching T' 
        blits_in_ℒin T' 
        get_conflict_wl T'  None 
        get_conflict_wl T'  Some {#}} 
    mop_lit_hd_trail_wl S  Id (mop_lit_hd_trail_l S') for S S'
    unfolding mop_lit_hd_trail_wl_def mop_lit_hd_trail_l_def
    apply refine_rcg
    subgoal unfolding mop_lit_hd_trail_wl_pre_def
      by blast
    subgoal by auto
    done

  have id: (x, x')  Id  f x  Id (f x') for x x' f
    by auto

  have id3: ((x, y, z), (x', y', z'))  Id  f x y z  Id (f x' y' z') for x x' f y y' z z'
    by auto


  have cons_trail_propagate_l: cons_trail_propagate_l L i M0
      {(M, M'). M = M'  M' = Propagated L i # M0}
         (cons_trail_propagate_l L' i' M')
    if L = L' i = i' M0 = M' for L i L' i' M' M0
    using that unfolding cons_trail_propagate_l_def
    by (auto intro!: ASSERT_refine_right)

  have fresh: get_fresh_index_wl N (NE + UE + NEk + UEk + NS + US + N0 + U0) W 
     {(i, i'). i = i'  i ∉# dom_m N 
         (L ∈# all_lits_st (x1, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, x1e, x2e). i  fst ` set (W L))}
       (get_fresh_index N')
    if N = N' propagate_bt_wl_pre L L' (x1, N, D, NE, UE, NEk, UEk, NS, US, N0, U0, x1e, x2e)
    for N N' NUE W NE UE N0 U0 x1e x2e NS US x1 D L L' NEk UEk
    unfolding that get_fresh_index_def get_fresh_index_wl_def all_lits_def all_lits_st_def
    by (auto intro: RES_refine simp: ac_simps)

  have blit: i ∉# dom_m x1a 
   set_mset (all_lits_of_m (mset C))  set_mset (all_lits_st (x1, x1a, Some xd, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e)) 
     blits_in_ℒin
        (x1, x1a, Some xd, x1c, x1d, NEk, UEk, NS, US, N0, U0,  x1e, x2e) 
       blits_in_ℒin
        (x1,
         fmupd i (C, b)  x1a,
         None, x1c, x1d, NEk, UEk, NS, US, N0, U0, {#}, x2e) for x x1a x1 xd x1d x1c x1e x2e i C b NS US N0 U0 NEk UEk
    by (auto simp:  blits_in_ℒin_def dest!: multi_member_split simp: all_lits_st_def)
  have [simp]: mset (unwatched_l (x)) + mset (watched_l (x)) = mset x for x
    by (metis add.commute mset_take_mset_drop_mset')

  have propagate_bt_wl: (Sb, Sc)
     {(U', U).
        (U', U)  state_wl_l None 
        equality_except_trail_wl U' S 
        (M. get_trail_wl S = M @ get_trail_wl U')  
          correct_watching U'  blits_in_ℒin U'} 
    (L, La)  Id 
    (L', L'a)
     {(L', La).
        La = L' 
        L' ∈# the (get_conflict_wl Sb) 
        L' ∈# remove1_mset (- L)
                (the (get_conflict_wl Sb))} 
    propagate_bt_wl L L' Sb
      {(T', T).
          (T', T)  state_wl_l None 
          correct_watching T'  blits_in_ℒin T'}
        (propagate_bt_l La L'a Sc) for a a' L La S Sa Sb Sc L' L'a
    unfolding propagate_bt_wl_def propagate_bt_l_def Let_def list_of_mset2_def list_of_mset_def
    apply (cases Sc; hypsubst)
    unfolding prod.case
    apply (refine_rcg fresh)
    subgoal unfolding propagate_bt_wl_pre_def by fast
    subgoal by (auto simp: propagate_bt_wl_pre_def state_wl_l_def)
    subgoal by (auto simp: state_wl_l_def)
    apply assumption
    apply (rule cons_trail_propagate_l)
    subgoal by (auto simp: state_wl_l_def)
    subgoal by (auto simp: state_wl_l_def)
    subgoal by (auto simp: state_wl_l_def)
    subgoal
      unfolding propagate_bt_l_pre_def propagate_bt_pre_def
      apply normalize_goal+
      apply (simp add: eq_commute[of _ (_, _)])
      apply (auto simp: state_wl_l_def)
      apply (auto 5 3 simp: ran_m_mapsto_upd_notin mset_take_mset_drop_mset'
         uminus_lit_swap blits_in_ℒin_propagate all_lits_of_m_add_mset
        state_wl_l_def twl_st_l_def
        dest: all_lits_of_m_diffD
        simp flip: image_mset_union
        intro!: blit blits_in_ℒin_keep_watch''' correct_watching_learn2[THEN iffD2]
        all_lits_st_learn_simps[THEN arg_cong[of _ _  λx. _  x], THEN iffD2])
      done
    done

  have correct_watching_unit: La ∈# all_lits_st (x1, x1a, x1b, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e) 
        correct_watching (x1, x1a, x1b, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e) 
          correct_watching
           (Propagated (- La) 0 # x1, x1a, None, x1c,
           x1d, NEk, add_mset {#-La#} UEk, NS, US, N0, U0, {#La#}, x2e)
     La ∈# all_lits_st (x1, x1a, x1b, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e) 
        blits_in_ℒin
        (x1, x1a, x1b, x1c, x1d, NEk, UEk, NS, US, N0, U0, x1e, x2e) 
       blits_in_ℒin
        (Propagated (- La) 0 # x1, x1a, None, x1c, x1d, NEk, add_mset {#-La#} UEk, NS, US, N0, U0, {#La#},
         x2e)
     for a b c d e f g x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e
       M Ma La NS US N0 U0 NEk UEk
     unfolding all_lits_st_def all_lits_def
    apply (subst (asm) in_all_lits_of_mm_uminus_iff[symmetric])
    apply (auto simp: correct_watching.simps all_lits_of_mm_add_mset
      all_lits_of_m_add_mset clause_to_update_def blits_in_ℒin_def
      all_lits_st_def all_lits_def
      dest: multi_member_split)
    apply (auto simp: correct_watching.simps all_lits_of_mm_add_mset all_lits_def
        all_lits_of_m_add_mset clause_to_update_def in_all_lits_of_mm_uminus_iff ac_simps
      dest: multi_member_split)
    done

  have propagate_unit_bt_wl: (Sb, Sc)
     {(U', U).
        (U', U)  state_wl_l None 
        equality_except_trail_wl U' S 
        (M. get_trail_wl S = M @ get_trail_wl U')  
          correct_watching U'  blits_in_ℒin U'} 
    (L, La)  Id 
    propagate_unit_bt_wl L Sb
      {(T', T).
          (T', T)  state_wl_l None 
          correct_watching T'  blits_in_ℒin T'}
        (propagate_unit_bt_l La Sc) for a a' L La S Sa Sb Sc L' L'a
    unfolding propagate_unit_bt_wl_def propagate_unit_bt_l_def Let_def list_of_mset_def
    apply (cases Sc; hypsubst)
    unfolding prod.case
    apply (refine_rcg fresh)
    subgoal unfolding in_pair_collect_simp by normalize_goal+ (simp add: eq_commute[of _ (_, _)])
    subgoal using propagate_unit_bt_wl_pre_def by blast
    apply (rule cons_trail_propagate_l)
    subgoal by (auto simp: state_wl_l_def)
    subgoal by (auto simp: state_wl_l_def)
    subgoal by (auto simp: state_wl_l_def)
    subgoal for a b c d e f g x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e
       M Ma
      unfolding propagate_unit_bt_l_pre_def propagate_unit_bt_pre_def
      by normalize_goal+
        (auto simp: all_lits_def ran_m_mapsto_upd_notin all_lits_of_mm_add_mset state_wl_l_def
          twl_st_l_def all_lits_of_m_add_mset mset_take_mset_drop_mset' in_all_lits_of_mm_uminus_iff
          ac_simps
        simp flip: image_mset_union intro!: blit correct_watching_unit dest!: all_lits_of_m_diffD)
    done

  show ?thesis
    unfolding st_l_of_wl.simps get_trail_wl.simps list_of_mset_def
      backtrack_wl_def backtrack_l_def
     apply (refine_vcg find_decomp_wl find_lit_of_max_level_wl extract_shorter_conflict_wl
         propagate_bt_wl propagate_unit_bt_wl;
        remove_dummy_vars)
    subgoal using backtrack_wl_inv_def by blast
    subgoal by auto
    subgoal by auto
    subgoal by auto
    done
qed


subsection Backtrack, Skip, Resolve or Decide

definition cdcl_twl_o_prog_wl_pre where
  cdcl_twl_o_prog_wl_pre S 
     (S'. (S, S')  state_wl_l None 
        correct_watching S 
        cdcl_twl_o_prog_l_pre S')

definition cdcl_twl_o_prog_wl :: 'v twl_st_wl  (bool × 'v twl_st_wl) nres where
  cdcl_twl_o_prog_wl S =
    do {
      ASSERT(cdcl_twl_o_prog_wl_pre S);
      do {
        if get_conflict_wl S = None
        then decide_wl_or_skip S
        else do {
          if count_decided (get_trail_wl S) > 0
          then do {
            T  skip_and_resolve_loop_wl S;
            ASSERT(get_conflict_wl T  None  get_conflict_wl T  Some {#});
            U  backtrack_wl T;
            RETURN (False, U)
          }
          else do {RETURN (True, S)}
        }
      }
    }
  

lemma [simp]: blits_in_ℒin (decide_lit_wl L S)  blits_in_ℒin S
  by (cases S) (auto simp: decide_lit_wl_def blits_in_ℒin_def)

lemma cdcl_twl_o_prog_wl_spec:
  (cdcl_twl_o_prog_wl, cdcl_twl_o_prog_l)  {(S::'v twl_st_wl, S'::'v twl_st_l).
     (S, S')  state_wl_l None 
        correct_watching S  blits_in_ℒin S} f
   {((brk::bool, T::'v twl_st_wl), brk'::bool, T'::'v twl_st_l).
     (T, T')  state_wl_l None 
     brk = brk' 
     correct_watching T  blits_in_ℒin T}nres_rel
   (is ?o  ?A f ?B nres_rel)
proof -

  have [iff]: correct_watching (decide_lit_wl L S)  correct_watching S for L S
    by (cases S; auto simp: decide_lit_wl_def correct_watching.simps clause_to_update_def)
  have [iff]: (decide_lit_wl L S, decide_lit_l L S')  state_wl_l None
    if (S, S')  state_wl_l None
    for L S S'
    using that by (cases S; auto simp: decide_lit_wl_def decide_lit_l_def state_wl_l_def)
  have option_id: x = x'  (x,x')  Idoption_rel for x x' by auto
  have damn_you_are_stupid_isabelle: (a, a')
        {(S, S').
          (S, S')  state_wl_l None 
          correct_watching S  blits_in_ℒin S} 
       get_conflict_l a' = None 
       (a, a')
        {(T', T).
          (T', T)  state_wl_l None 
          correct_watching T' 
          blits_in_ℒin T'  get_conflict_wl T' = None} for a a'
   by auto
  show cdcl_o: ?o  ?A f
   {((brk::bool, T::'v twl_st_wl), brk'::bool, T'::'v twl_st_l).
     (T, T')  state_wl_l None 
     brk = brk' 
     correct_watching T  blits_in_ℒin T}nres_rel
    unfolding cdcl_twl_o_prog_wl_def cdcl_twl_o_prog_l_def
      fref_param1[symmetric]
    apply (refine_vcg skip_and_resolve_loop_wl_spec["to_⇓"] backtrack_wl_spec["to_⇓"]
      decide_wl_or_skip_spec["to_⇓", THEN order_trans]
      option_id)
    subgoal unfolding cdcl_twl_o_prog_wl_pre_def by blast
    subgoal by auto
    apply (rule damn_you_are_stupid_isabelle; assumption)
    subgoal by (auto intro!: conc_fun_R_mono)
    subgoal unfolding decide_wl_or_skip_pre_def by auto
    subgoal by auto
    subgoal by (auto simp: )
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    done
qed


subsection Full Strategy

definition cdcl_twl_stgy_prog_wl_inv :: 'v twl_st_wl  bool × 'v twl_st_wl   bool where
  cdcl_twl_stgy_prog_wl_inv S0  λ(brk, T).
      ( T' S0'.  (T, T')  state_wl_l None 
      (S0, S0')  state_wl_l None 
      cdcl_twl_stgy_prog_l_inv S0' (brk, T'))

definition cdcl_twl_stgy_prog_wl :: 'v twl_st_wl  'v twl_st_wl nres where
  cdcl_twl_stgy_prog_wl S0 =
  do {
    (brk, T)  WHILETcdcl_twl_stgy_prog_wl_inv S0(λ(brk, _). ¬brk)
      (λ(brk, S). do {
        T  unit_propagation_outer_loop_wl S;
        cdcl_twl_o_prog_wl T
      })
      (False, S0);
    RETURN T
  }


theorem cdcl_twl_stgy_prog_wl_spec:
  (cdcl_twl_stgy_prog_wl, cdcl_twl_stgy_prog_l)  {(S::'v twl_st_wl, S').
       (S, S')  state_wl_l None 
       correct_watching S  blits_in_ℒin S} 
    state_wl_l Nonenres_rel
   (is ?o  ?A  ?B nres_rel)
proof -
  have H: ((False, S'), False, S)  {((brk', T'), (brk, T)). (T', T)  state_wl_l None  brk' = brk 
       correct_watching T'  blits_in_ℒin T'}
    if (S', S)  state_wl_l None and
       correct_watching S' blits_in_ℒin S'
    for S' :: 'v twl_st_wl and S :: 'v twl_st_l
    using that by auto

  show ?thesis
    unfolding cdcl_twl_stgy_prog_wl_def cdcl_twl_stgy_prog_l_def
    apply (refine_rcg H unit_propagation_outer_loop_wl_spec[THEN fref_to_Down]
      cdcl_twl_o_prog_wl_spec[THEN fref_to_Down])
    subgoal for S' S by (cases S') auto
    subgoal by auto
    subgoal by auto
    subgoal unfolding cdcl_twl_stgy_prog_wl_inv_def by blast
    subgoal by auto
    subgoal by auto
    subgoal for S' S brk'T' brkT brk' T' by auto
    subgoal by fast
    subgoal by auto
    done
qed

theorem cdcl_twl_stgy_prog_wl_spec':
  (cdcl_twl_stgy_prog_wl, cdcl_twl_stgy_prog_l)  {(S::'v twl_st_wl, S').
       (S, S')  state_wl_l None  correct_watching S  blits_in_ℒin S} 
    {(S::'v twl_st_wl, S').
       (S, S')  state_wl_l None  correct_watching S  blits_in_ℒin S}nres_rel
   (is ?o  ?A  ?B nres_rel)
proof -
  have H: ((False, S'), False, S)  {((brk', T'), (brk, T)). (T', T)  state_wl_l None  brk' = brk 
       correct_watching T'  blits_in_ℒin T'}
    if (S', S)  state_wl_l None and
       correct_watching S' blits_in_ℒin S'
    for S' :: 'v twl_st_wl and S :: 'v twl_st_l
    using that by auto
    thm unit_propagation_outer_loop_wl_spec[THEN fref_to_Down]
  show ?thesis
    unfolding cdcl_twl_stgy_prog_wl_def cdcl_twl_stgy_prog_l_def
    apply (refine_rcg H unit_propagation_outer_loop_wl_spec[THEN fref_to_Down]
      cdcl_twl_o_prog_wl_spec[THEN fref_to_Down])
    subgoal for S' S by (cases S') auto
    subgoal by auto
    subgoal by auto
    subgoal unfolding cdcl_twl_stgy_prog_wl_inv_def by blast
    subgoal by auto
    subgoal by auto
    subgoal for S' S brk'T' brkT brk' T' by auto
    subgoal by fast
    subgoal by auto
    done
qed

definition cdcl_twl_stgy_prog_wl_pre where
  cdcl_twl_stgy_prog_wl_pre S U 
    (T. (S, T)  state_wl_l None  cdcl_twl_stgy_prog_l_pre T U  correct_watching S  blits_in_ℒin S)

lemma cdcl_twl_stgy_prog_wl_spec_final:
  assumes
    cdcl_twl_stgy_prog_wl_pre S S'
  shows
    cdcl_twl_stgy_prog_wl S   (state_wl_l None O twl_st_l None) (conclusive_TWL_norestart_run S')
proof -
  obtain T where T: (S, T)  state_wl_l None cdcl_twl_stgy_prog_l_pre T S' correct_watching S blits_in_ℒin S
    using assms unfolding cdcl_twl_stgy_prog_wl_pre_def by blast
  show ?thesis
    apply (rule order_trans[OF cdcl_twl_stgy_prog_wl_spec["to_⇓", of S T]])
    subgoal using T by auto
    subgoal
      apply (rule order_trans)
      apply (rule ref_two_step')
       apply (rule cdcl_twl_stgy_prog_l_spec_final[of _ S'])
      subgoal using T by fast
      subgoal unfolding conc_fun_chain by auto
      done
    done
qed


definition cdcl_twl_stgy_prog_break_wl :: 'v twl_st_wl  'v twl_st_wl nres where
  cdcl_twl_stgy_prog_break_wl S0 =
  do {
    b  SPEC(λ_. True);
    (b, brk, T)  WHILETλ(_, S). cdcl_twl_stgy_prog_wl_inv S0 S(λ(b, brk, _). b  ¬brk)
      (λ(_, brk, S). do {
        T  unit_propagation_outer_loop_wl S;
        T  cdcl_twl_o_prog_wl T;
        b  SPEC(λ_. True);
        RETURN (b, T)
      })
      (b, False, S0);
    if brk then RETURN T
    else cdcl_twl_stgy_prog_wl T
  }

theorem cdcl_twl_stgy_prog_break_wl_spec':
  (cdcl_twl_stgy_prog_break_wl, cdcl_twl_stgy_prog_break_l)  {(S::'v twl_st_wl, S').
       (S, S')  state_wl_l None  correct_watching S  blits_in_ℒin S} f
    {(S::'v twl_st_wl, S'). (S, S')  state_wl_l None  correct_watching S  blits_in_ℒin S}nres_rel
   (is ?o  ?A f ?B nres_rel)
proof -
  have H: ((b', False, S'), b, False, S)  {((b', brk', T'), (b, brk, T)).
      (T', T)  state_wl_l None  brk' = brk  b' = b 
       correct_watching T'  blits_in_ℒin T'}
    if (S', S)  state_wl_l None and
       correct_watching S' and
       (b', b)  bool_rel blits_in_ℒin S'
    for S' :: 'v twl_st_wl and S :: 'v twl_st_l and b' b :: bool
    using that by auto
  show ?thesis
    unfolding cdcl_twl_stgy_prog_break_wl_def cdcl_twl_stgy_prog_break_l_def fref_param1[symmetric]
    apply (refine_rcg H unit_propagation_outer_loop_wl_spec[THEN fref_to_Down]
      cdcl_twl_o_prog_wl_spec[THEN fref_to_Down]
      cdcl_twl_stgy_prog_wl_spec'[unfolded fref_param1, THEN fref_to_Down])
    subgoal for S' S by (cases S') auto
    subgoal by auto
    subgoal by auto
    subgoal unfolding cdcl_twl_stgy_prog_wl_inv_def by blast
    subgoal by auto
    subgoal by auto
    subgoal for S' S brk'T' brkT brk' T' by auto
    subgoal by fast
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by fast
    subgoal by auto
    done
qed


theorem cdcl_twl_stgy_prog_break_wl_spec:
  (cdcl_twl_stgy_prog_break_wl, cdcl_twl_stgy_prog_break_l)  {(S::'v twl_st_wl, S').
       (S, S')  state_wl_l None 
       correct_watching S  blits_in_ℒin S} f
    state_wl_l Nonenres_rel
   (is ?o  ?A f ?B nres_rel)
  using cdcl_twl_stgy_prog_break_wl_spec'
  apply -
  apply (rule mem_set_trans)
  prefer 2 apply assumption
  apply (match_fun_rel, solves simp)
  apply (match_fun_rel; solves auto)
  done

lemma cdcl_twl_stgy_prog_break_wl_spec_final:
  assumes
    cdcl_twl_stgy_prog_wl_pre S S'
  shows
    cdcl_twl_stgy_prog_break_wl S   (state_wl_l None O twl_st_l None) (conclusive_TWL_norestart_run S')
proof -
  obtain T where T: (S, T)  state_wl_l None cdcl_twl_stgy_prog_l_pre T S' correct_watching S blits_in_ℒin S
    using assms unfolding cdcl_twl_stgy_prog_wl_pre_def by blast
  show ?thesis
    apply (rule order_trans[OF cdcl_twl_stgy_prog_break_wl_spec[unfolded fref_param1[symmetric], "to_⇓", of S T]])
    subgoal using T by auto
    subgoal
      apply (rule order_trans)
      apply (rule ref_two_step')
       apply (rule cdcl_twl_stgy_prog_break_l_spec_final[of _ S'])
      subgoal using T by fast
      subgoal unfolding conc_fun_chain by auto
      done
    done
qed


definition cdcl_twl_stgy_prog_early_wl :: 'v twl_st_wl  (bool × 'v twl_st_wl) nres where
  cdcl_twl_stgy_prog_early_wl S0 =
  do {
    b  SPEC(λ_. True);
    (b, brk, T)  WHILETλ(_, S). cdcl_twl_stgy_prog_wl_inv S0 S(λ(b, brk, _). b  ¬brk)
      (λ(_, brk, S). do {
        T  unit_propagation_outer_loop_wl S;
        T  cdcl_twl_o_prog_wl T;
        b  SPEC(λ_. True);
        RETURN (b, T)
      })
      (b, False, S0);
    RETURN (brk, T)
  }

theorem cdcl_twl_stgy_prog_early_wl_spec':
  (cdcl_twl_stgy_prog_early_wl, cdcl_twl_stgy_prog_early_l)  {(S::'v twl_st_wl, S').
       (S, S')  state_wl_l None  correct_watching S  blits_in_ℒin S} f
    bool_rel ×r {(S::'v twl_st_wl, S'). (S, S')  state_wl_l None  correct_watching S  blits_in_ℒin S}nres_rel
   (is ?o  ?A f ?B nres_rel)
proof -
  have H: ((b', False, S'), b, False, S)  {((b', brk', T'), (b, brk, T)).
      (T', T)  state_wl_l None  brk' = brk  b' = b 
       correct_watching T'  blits_in_ℒin T'}
    if (S', S)  state_wl_l None and
       correct_watching S' and
       (b', b)  bool_rel blits_in_ℒin S'
    for S' :: 'v twl_st_wl and S :: 'v twl_st_l and b' b :: bool
    using that by auto
  show ?thesis
    unfolding cdcl_twl_stgy_prog_early_wl_def cdcl_twl_stgy_prog_early_l_def fref_param1[symmetric]
    apply (refine_rcg H unit_propagation_outer_loop_wl_spec[THEN fref_to_Down]
      cdcl_twl_o_prog_wl_spec[THEN fref_to_Down]
      cdcl_twl_stgy_prog_wl_spec'[unfolded fref_param1, THEN fref_to_Down])
    subgoal for S' S by (cases S') auto
    subgoal by auto
    subgoal by auto
    subgoal unfolding cdcl_twl_stgy_prog_wl_inv_def by blast
    subgoal by auto
    subgoal by auto
    subgoal for S' S brk'T' brkT brk' T' by auto
    subgoal by fast
    subgoal by auto
    subgoal by auto
    done
qed


theorem cdcl_twl_stgy_prog_early_wl_spec:
  (cdcl_twl_stgy_prog_early_wl, cdcl_twl_stgy_prog_early_l)  {(S::'v twl_st_wl, S').
       (S, S')  state_wl_l None 
       correct_watching S  blits_in_ℒin S} f
    bool_rel ×r state_wl_l Nonenres_rel
   (is ?o  ?A f ?B nres_rel)
  using cdcl_twl_stgy_prog_early_wl_spec'
  apply -
  apply (rule mem_set_trans)
  prefer 2 apply assumption
  apply (match_fun_rel, solves simp)
  apply (match_fun_rel; solves auto)
  done

lemma cdcl_twl_stgy_prog_early_wl_spec_final:
  assumes
    cdcl_twl_stgy_prog_wl_pre S S'
  shows
    cdcl_twl_stgy_prog_early_wl S   (bool_rel ×r (state_wl_l None O twl_st_l None)) (partial_conclusive_TWL_norestart_run S')
proof -
  obtain T where T: (S, T)  state_wl_l None cdcl_twl_stgy_prog_l_pre T S' correct_watching S blits_in_ℒin S
    using assms unfolding cdcl_twl_stgy_prog_wl_pre_def by blast
  have H: {((a, b), a', b'). (a, a')  bool_rel  (b, b')  state_wl_l None} O
       (bool_rel ×f twl_st_l None) = bool_rel ×r state_wl_l None O twl_st_l None
    by fastforce
  show ?thesis
    apply (rule order_trans[OF cdcl_twl_stgy_prog_early_wl_spec[unfolded fref_param1[symmetric], "to_⇓", of S T]])
    subgoal using T by auto
    subgoal
      apply (rule order_trans)
      apply (rule ref_two_step')
       apply (rule cdcl_twl_stgy_prog_early_l_spec_final[of _ S'])
      subgoal using T by fast
      subgoal unfolding H conc_fun_chain by auto
      done
    done
qed

subsection Shared for restarts and inprocessing

definition clauses_pointed_to :: 'v literal set  ('v literal  'v watched)  nat set
  where
  clauses_pointed_to 𝒜 W  (((`) fst) ` set ` W ` 𝒜)

lemma clauses_pointed_to_insert[simp]:
  clauses_pointed_to (insert A 𝒜) W =
  fst ` set (W A) 
  clauses_pointed_to 𝒜 W and
  clauses_pointed_to_empty[simp]:
  clauses_pointed_to {} W = {}
  by (auto simp: clauses_pointed_to_def)

lemma clauses_pointed_to_remove1_if:
  Lset (W L). fst L ∉# dom_m aa  xa ∈# dom_m aa 
    xa  clauses_pointed_to (set_mset (remove1_mset L 𝒜))
      (λa. if a = L then [] else W a) 
    xa  clauses_pointed_to (set_mset (remove1_mset L 𝒜)) W
  by (cases L ∈# 𝒜)
    (fastforce simp: clauses_pointed_to_def
    dest!: multi_member_split)+

lemma clauses_pointed_to_remove1_if2:
  Lset (W L). fst L ∉# dom_m aa  xa ∈# dom_m aa 
    xa  clauses_pointed_to (set_mset (𝒜 - {#L, L'#}))
      (λa. if a = L then [] else W a) 
    xa  clauses_pointed_to (set_mset (𝒜 - {#L, L'#})) W
  Lset (W L). fst L ∉# dom_m aa  xa ∈# dom_m aa 
    xa  clauses_pointed_to (set_mset (𝒜 - {#L', L#}))
      (λa. if a = L then [] else W a) 
    xa  clauses_pointed_to (set_mset (𝒜 - {#L', L#})) W
  by (cases L ∈# 𝒜; fastforce simp: clauses_pointed_to_def
    dest!: multi_member_split)+

lemma clauses_pointed_to_remove1_if2_eq:
  Lset (W L). fst L ∉# dom_m aa 
    set_mset (dom_m aa)  clauses_pointed_to (set_mset (𝒜 - {#L, L'#}))
      (λa. if a = L then [] else W a) 
    set_mset (dom_m aa)  clauses_pointed_to (set_mset (𝒜 - {#L, L'#})) W
  Lset (W L). fst L ∉# dom_m aa 
     set_mset (dom_m aa)  clauses_pointed_to (set_mset (𝒜 - {#L', L#}))
      (λa. if a = L then [] else W a) 
     set_mset (dom_m aa)  clauses_pointed_to (set_mset (𝒜 - {#L', L#})) W
  by (auto simp: clauses_pointed_to_remove1_if2)

lemma negs_remove_Neg: A ∉# 𝒜  negs 𝒜 + poss 𝒜 - {#Neg A, Pos A#} =
   negs 𝒜 + poss 𝒜
  by (induction 𝒜) auto
lemma poss_remove_Pos: A ∉# 𝒜  negs 𝒜 + poss 𝒜 - {#Pos A, Neg A#} =
   negs 𝒜 + poss 𝒜
  by (induction 𝒜)  auto

end