Theory IsaSAT_Restart_Heuristics

theory IsaSAT_Restart_Heuristics
imports Watched_Literals_Watch_List_Restart IsaSAT_Rephase IsaSAT_VMTF IsaSAT_Sorting
theory IsaSAT_Restart_Heuristics
imports
  Watched_Literals.WB_Sort Watched_Literals.Watched_Literals_Watch_List_Restart IsaSAT_Rephase
  IsaSAT_Setup IsaSAT_VMTF IsaSAT_Sorting
begin

chapter ‹Restarts›

lemma all_init_atms_alt_def:
  ‹set_mset (all_init_atms N NE) = atms_of_mm (mset `# init_clss_lf N) ∪ atms_of_mm NE›
  unfolding all_init_atms_def all_init_lits_def
  by (auto simp: in_all_lits_of_mm_ain_atms_of_iff
      all_lits_of_mm_def atms_of_ms_def image_UN
      atms_of_def
    dest!: multi_member_split[of ‹(_, _)› ‹ran_m N›]
    dest: multi_member_split atm_of_lit_in_atms_of
    simp del: set_image_mset)

lemma in_set_all_init_atms_iff:
  ‹y ∈# all_init_atms bu bw ⟷
    y ∈ atms_of_mm (mset `# init_clss_lf 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 all_init_atms_alt_def
      simp: in_all_lits_of_mm_ain_atms_of_iff
      all_lits_of_mm_def atms_of_ms_def image_UN
      atms_of_def
    dest!: multi_member_split[of ‹(_, _)› ‹ran_m N›]
    dest: multi_member_split atm_of_lit_in_atms_of
    simp del: set_image_mset)

lemma twl_st_heur_change_subsumed_clauses:
  assumes ‹((M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur,
       vdom, avdom, lcount, opts, old_arena),
     (M, N, D, NE, UE, NS, US, Q, W)) ∈ twl_st_heur›
    ‹set_mset (all_atms N ((NE+UE)+(NS+US))) = set_mset (all_atms N ((NE+UE)+(NS'+US')))›
  shows ‹((M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur,
       vdom, avdom, lcount, opts, old_arena),
     (M, N, D, NE, UE, NS', US', Q, W)) ∈ twl_st_heur›
proof -
  note cong = trail_pol_cong heuristic_rel_cong
      option_lookup_clause_rel_cong D0_cong isa_vmtf_cong phase_saving_cong
      cach_refinement_empty_cong vdom_m_cong isasat_input_nempty_cong
      isasat_input_bounded_cong heuristic_rel_cong
  show ?thesis
    using cong[OF assms(2)] assms(1)
    apply (auto simp add: twl_st_heur_def)
    apply fastforce
    apply force
    done
qed


text ‹
  This is a list of comments (how does it work for glucose and cadical) to prepare the future
  refinement:
  ▸ Reduction
     ▪ every 2000+300*n (rougly since inprocessing changes the real number, cadical)
           (split over initialisation file); don't restart if level < 2 or if the level is less
       than the fast average
     ▪ curRestart * nbclausesbeforereduce;
          curRestart = (conflicts / nbclausesbeforereduce) + 1 (glucose)
  ▸ Killed
     ▪ half of the clauses that can› be deleted (i.e., not used since last restart), not strictly
       LBD, but a probability of being useful.
     ▪ half of the clauses
  ▸ Restarts:
     ▪ EMA-14, aka restart if enough clauses and slow\_glue\_avg * opts.restartmargin > fast\_glue (file ema.cpp)
     ▪ (lbdQueue.getavg() * K) > (sumLBD / conflictsRestarts),
       \<^text>‹conflictsRestarts > LOWER_BOUND_FOR_BLOCKING_RESTART && lbdQueue.isvalid() && trail.size() > R * trailQueue.getavg()›
›
declare all_atms_def[symmetric,simp]


definition twl_st_heur_restart :: ‹(twl_st_wl_heur × nat twl_st_wl) set› where
‹twl_st_heur_restart =
  {((M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur,
       vdom, avdom, lcount, opts, old_arena),
     (M, N, D, NE, UE, NS, US, Q, W)).
    (M', M) ∈ trail_pol (all_init_atms N (NE+NS)) ∧
    valid_arena N' N (set vdom) ∧
    (D', D) ∈ option_lookup_clause_rel (all_init_atms N (NE+NS)) ∧
    (D = None ⟶ j ≤ length M) ∧
    Q = uminus `# lit_of `# mset (drop j (rev M)) ∧
    (W', W) ∈ ⟨Id⟩map_fun_rel (D0 (all_init_atms N (NE+NS))) ∧
    vm ∈ isa_vmtf (all_init_atms N (NE+NS)) M ∧
    no_dup M ∧
    clvls ∈ counts_maximum_level M D ∧
    cach_refinement_empty (all_init_atms N (NE+NS)) cach ∧
    out_learned M D outl ∧
    lcount = size (learned_clss_lf N) ∧
    vdom_m (all_init_atms N (NE+NS))  W N ⊆ set vdom ∧
    mset avdom ⊆# mset vdom ∧
    isasat_input_bounded (all_init_atms N (NE+NS)) ∧
    isasat_input_nempty (all_init_atms N (NE+NS)) ∧
    distinct vdom ∧ old_arena = [] ∧
    heuristic_rel (all_init_atms N (NE+NS)) heur
  }›


abbreviation twl_st_heur'''' where
  ‹twl_st_heur'''' r ≡ {(S, T). (S, T) ∈ twl_st_heur ∧ length (get_clauses_wl_heur S) ≤ r}›

abbreviation twl_st_heur_restart''' where
  ‹twl_st_heur_restart''' r ≡
    {(S, T). (S, T) ∈ twl_st_heur_restart ∧ length (get_clauses_wl_heur S) = r}›

abbreviation twl_st_heur_restart'''' where
  ‹twl_st_heur_restart'''' r ≡
    {(S, T). (S, T) ∈ twl_st_heur_restart ∧ length (get_clauses_wl_heur S) ≤ r}›

definition twl_st_heur_restart_ana :: ‹nat ⇒ (twl_st_wl_heur × nat twl_st_wl) set› where
‹twl_st_heur_restart_ana r =
  {(S, T). (S, T) ∈ twl_st_heur_restart ∧ length (get_clauses_wl_heur S) = r}›

lemma twl_st_heur_restart_anaD: ‹x ∈ twl_st_heur_restart_ana r ⟹ x ∈ twl_st_heur_restart›
  by (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def)

lemma twl_st_heur_restartD:
  ‹x ∈ twl_st_heur_restart ⟹ x ∈ twl_st_heur_restart_ana (length (get_clauses_wl_heur (fst x)))›
  by (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def)

definition clause_score_ordering2 where
  ‹clause_score_ordering2 = (λ(lbd, act) (lbd', act'). lbd < lbd' ∨ (lbd = lbd' ∧ act ≤ act'))›

lemma unbounded_id: ‹unbounded (id :: nat ⇒ nat)›
  by (auto simp: bounded_def) presburger

global_interpretation twl_restart_ops id
  by unfold_locales

global_interpretation twl_restart id
  by standard (rule unbounded_id)

text ‹
  We first fix the function that proves termination. We don't take the ``smallest'' function
  possible (other possibilites that are growing slower include \<^term>‹λ(n::nat). n >> 50›).
  Remark that this scheme is not compatible with Luby (TODO: use Luby restart scheme every once
  in a while like CryptoMinisat?)
›

lemma get_slow_ema_heur_alt_def:
   ‹RETURN o get_slow_ema_heur = (λ(M, N0, D, Q, W, vm, clvls, cach, lbd, outl,
       stats, (fema, sema, (ccount, _)), lcount). RETURN sema)›
  by auto


lemma get_fast_ema_heur_alt_def:
   ‹RETURN o get_fast_ema_heur = (λ(M, N0, D, Q, W, vm, clvls, cach, lbd, outl,
       stats, (fema, sema, ccount), lcount). RETURN fema)›
  by auto

lemma get_learned_count_alt_def:
   ‹RETURN o get_learned_count = (λ(M, N0, D, Q, W, vm, clvls, cach, lbd, outl,
       stats, _, vdom, avdom, lcount, opts). RETURN lcount)›
  by auto

definition (in -) find_local_restart_target_level_int_inv where
  ‹find_local_restart_target_level_int_inv ns cs =
     (λ(brk, i). i ≤ length cs ∧ length cs < uint32_max)›

definition find_local_restart_target_level_int
   :: ‹trail_pol ⇒ isa_vmtf_remove_int ⇒ nat nres›
where
  ‹find_local_restart_target_level_int =
     (λ(M, xs, lvls, reasons, k, cs) ((ns :: nat_vmtf_node list, m :: nat, fst_As::nat, lst_As::nat,
        next_search::nat option), _). do {
     (brk, i) ← WHILETfind_local_restart_target_level_int_inv ns cs
        (λ(brk, i). ¬brk ∧ i < length_uint32_nat cs)
        (λ(brk, i). do {
           ASSERT(i < length cs);
           let t = (cs  ! i);
	   ASSERT(t < length M);
	   let L = atm_of (M ! t);
           ASSERT(L < length ns);
           let brk = stamp (ns ! L) < m;
           RETURN (brk, if brk then i else i+1)
         })
        (False, 0);
    RETURN i
   })›

definition find_local_restart_target_level where
  ‹find_local_restart_target_level M _ = SPEC(λi. i ≤ count_decided M)›

lemma find_local_restart_target_level_alt_def:
  ‹find_local_restart_target_level M vm = do {
      (b, i) ← SPEC(λ(b::bool, i). i ≤ count_decided M);
       RETURN i
    }›
  unfolding find_local_restart_target_level_def by (auto simp: RES_RETURN_RES2 uncurry_def)


lemma find_local_restart_target_level_int_find_local_restart_target_level:
   ‹(uncurry find_local_restart_target_level_int, uncurry find_local_restart_target_level) ∈
     [λ(M, vm). vm ∈ isa_vmtf 𝒜 M]f trail_pol 𝒜 ×r Id → ⟨nat_rel⟩nres_rel›
  unfolding find_local_restart_target_level_int_def find_local_restart_target_level_alt_def
    uncurry_def Let_def
  apply (intro frefI nres_relI)
  apply clarify
  subgoal for a aa ab ac ad b ae af ag ah ba bb ai aj ak al am bc bd
    apply (refine_rcg WHILEIT_rule[where R = ‹measure (λ(brk, i). (If brk 0 1) + length b - i)›]
        assert.ASSERT_leI)
    subgoal by auto
    subgoal
      unfolding find_local_restart_target_level_int_inv_def
      by (auto simp: trail_pol_alt_def control_stack_length_count_dec)
    subgoal by auto
    subgoal by (auto simp: trail_pol_alt_def intro: control_stack_le_length_M)
    subgoal for s x1 x2
      by (subgoal_tac ‹a ! (b ! x2) ∈# ℒall 𝒜›)
        (auto simp: trail_pol_alt_def rev_map lits_of_def rev_nth
            vmtf_def atms_of_def isa_vmtf_def
          intro!: literals_are_in_ℒin_trail_in_lits_of_l)
    subgoal by (auto simp: find_local_restart_target_level_int_inv_def)
    subgoal by (auto simp: trail_pol_alt_def control_stack_length_count_dec
          find_local_restart_target_level_int_inv_def)
    subgoal by auto
    done
  done

definition empty_Q :: ‹twl_st_wl_heur ⇒ twl_st_wl_heur nres› where
  ‹empty_Q = (λ(M, N, D, Q, W, vm, clvls, cach, lbd, outl, stats, (fema, sema, ccount, wasted), vdom,
      lcount). do{
    ASSERT(isa_length_trail_pre M);
    let j = isa_length_trail M;
    RETURN (M, N, D, j, W, vm, clvls, cach, lbd, outl, stats, (fema, sema,
       restart_info_restart_done ccount, wasted), vdom, lcount)
  })›

definition restart_abs_wl_heur_pre  :: ‹twl_st_wl_heur ⇒ bool ⇒ bool› where
  ‹restart_abs_wl_heur_pre S brk  ⟷ (∃T. (S, T) ∈ twl_st_heur ∧ restart_abs_wl_pre T brk)›

text ‹\<^term>‹find_decomp_wl_st_int› is the wrong function here, because unlike in the backtrack case,
  we also have to update the queue of literals to update. This is done in the function \<^term>‹empty_Q›.
›

definition find_local_restart_target_level_st :: ‹twl_st_wl_heur ⇒ nat nres› where
  ‹find_local_restart_target_level_st S = do {
    find_local_restart_target_level_int (get_trail_wl_heur S) (get_vmtf_heur S)
  }›

lemma find_local_restart_target_level_st_alt_def:
  ‹find_local_restart_target_level_st = (λ(M, N, D, Q, W, vm, clvls, cach, lbd, stats). do {
      find_local_restart_target_level_int M vm})›
 apply (intro ext)
  apply (case_tac x)
  by (auto simp: find_local_restart_target_level_st_def)

definition cdcl_twl_local_restart_wl_D_heur
   :: ‹twl_st_wl_heur ⇒ twl_st_wl_heur nres›
where
  ‹cdcl_twl_local_restart_wl_D_heur = (λS. do {
      ASSERT(restart_abs_wl_heur_pre S False);
      lvl ← find_local_restart_target_level_st S;
      if lvl = count_decided_st_heur S
      then RETURN S
      else do {
        S ← find_decomp_wl_st_int lvl S;
        S ← empty_Q S;
        incr_lrestart_stat S
      }
   })›


named_theorems twl_st_heur_restart

lemma [twl_st_heur_restart]:
  assumes ‹(S, T) ∈ twl_st_heur_restart›
  shows ‹(get_trail_wl_heur S, get_trail_wl T) ∈ trail_pol (all_init_atms_st T)›
  using assms by (cases S; cases T)
   (simp only: twl_st_heur_restart_def get_trail_wl_heur.simps get_trail_wl.simps
    mem_Collect_eq prod.case get_clauses_wl.simps get_unit_init_clss_wl.simps
    get_subsumed_init_clauses_wl.simps)

lemma trail_pol_literals_are_in_ℒin_trail:
  ‹(M', M) ∈ trail_pol 𝒜 ⟹ literals_are_in_ℒin_trail 𝒜 M›
  unfolding literals_are_in_ℒin_trail_def trail_pol_def
  by auto

lemma refine_generalise1: "A ≤ B ⟹ do {x ← B; C x} ≤ D ⟹ do {x ← A; C x} ≤ (D:: 'a nres)"
  using Refine_Basic.bind_mono(1) dual_order.trans by blast

lemma refine_generalise2: "A ≤ B ⟹ do {x ← do {x ← B; A' x}; C x} ≤ D ⟹
  do {x ← do {x ← A; A' x}; C x} ≤ (D:: 'a nres)"
  by (simp add: refine_generalise1)

lemma cdcl_twl_local_restart_wl_D_spec_int:
  ‹cdcl_twl_local_restart_wl_spec (M, N, D, NE, UE, NS, US, Q, W) ≥ ( do {
      ASSERT(restart_abs_wl_pre (M, N, D, NE, UE, NS, US, Q, W) False);
      i ← SPEC(λ_. True);
      if i
      then RETURN (M, N, D, NE, UE, NS, {#}, Q, W)
      else do {
        (M, Q') ← SPEC(λ(M', Q'). (∃K M2. (Decided K # M', M2) ∈ set (get_all_ann_decomposition M) ∧
              Q' = {#}) ∨ (M' = M ∧ Q' = Q));
        RETURN (M, N, D, NE, UE, NS, {#}, Q', W)
     }
   })›
proof -
  have If_Res: ‹(if i then (RETURN f) else (RES g)) = (RES (if i then {f} else g))› for i f g
    by auto
  show ?thesis
    unfolding cdcl_twl_local_restart_wl_spec_def prod.case RES_RETURN_RES2 If_Res
    by refine_vcg
      (auto simp: If_Res RES_RETURN_RES2 RES_RES_RETURN_RES uncurry_def
        image_iff split:if_splits)
qed

lemma trail_pol_no_dup: ‹(M, M') ∈ trail_pol 𝒜 ⟹ no_dup M'›
  by (auto simp: trail_pol_def)

lemma heuristic_rel_restart_info_done[intro!, simp]:
  ‹heuristic_rel 𝒜 (fema, sema, ccount, wasted) ⟹
    heuristic_rel 𝒜 ((fema, sema, restart_info_restart_done ccount, wasted))›
  by (auto simp: heuristic_rel_def)

lemma cdcl_twl_local_restart_wl_D_heur_cdcl_twl_local_restart_wl_D_spec:
  ‹(cdcl_twl_local_restart_wl_D_heur, cdcl_twl_local_restart_wl_spec) ∈
    twl_st_heur''' r →f ⟨twl_st_heur''' r⟩nres_rel›
proof -
  have K: ‹( (case S of
               (M, N, D, Q, W, vm, clvls, cach, lbd, outl, stats, (fema, sema,
                ccount), vdom, lcount) ⇒
                 ASSERT (isa_length_trail_pre M) ⤜
                 (λ_. RES {(M, N, D, isa_length_trail M, W, vm, clvls, cach,
                            lbd, outl, stats, (fema, sema,
                            restart_info_restart_done ccount), vdom, lcount)}))) =
        ((ASSERT (case S of (M, N, D, Q, W, vm, clvls, cach, lbd, outl, stats, (fema, sema,
                ccount), vdom, lcount) ⇒ isa_length_trail_pre M)) ⤜
         (λ _. (case S of
               (M, N, D, Q, W, vm, clvls, cach, lbd, outl, stats, (fema, sema,
                ccount), vdom, lcount) ⇒ RES {(M, N, D, isa_length_trail M, W, vm, clvls, cach,
                            lbd, outl, stats, (fema, sema,
                            restart_info_restart_done ccount), vdom, lcount)})))› for S
  by (cases S) auto

  have K2: ‹(case S of
               (a, b) ⇒ RES (Φ a b)) =
        (RES (case S of (a, b) ⇒ Φ a b))› for S
  by (cases S) auto

  have [dest]: ‹av = None› ‹out_learned a av am ⟹ out_learned x1 av am›
    if ‹restart_abs_wl_pre (a, au, av, aw, ax, NS, US, ay, bd) False›
    for a au av aw ax ay bd x1 am NS US
    using that
    unfolding restart_abs_wl_pre_def restart_abs_l_pre_def
      restart_prog_pre_def
    by (auto simp: twl_st_l_def state_wl_l_def out_learned_def)
  have [refine0]:
    ‹find_local_restart_target_level_int (get_trail_wl_heur S) (get_vmtf_heur S) ≤
      ⇓ {(i, b). b = (i = count_decided (get_trail_wl T)) ∧
          i ≤ count_decided (get_trail_wl T)} (SPEC (λ_. True))›
    if ‹(S, T) ∈ twl_st_heur› for S T
    apply (rule find_local_restart_target_level_int_find_local_restart_target_level[THEN
         fref_to_Down_curry, THEN order_trans, of ‹all_atms_st T› ‹get_trail_wl T› ‹get_vmtf_heur S›])
    subgoal using that unfolding twl_st_heur_def by auto
    subgoal using that unfolding twl_st_heur_def by auto
    subgoal by (auto simp: find_local_restart_target_level_def conc_fun_RES)
    done
  have H:
      ‹set_mset (all_atms_st S) =
           set_mset (all_init_atms_st S)› (is ?A)
      ‹set_mset (all_atms_st S) =
           set_mset (all_atms (get_clauses_wl S) (get_unit_clauses_wl S + get_subsumed_init_clauses_wl S))›
           (is ?B)
      ‹get_conflict_wl S = None› (is ?C)
     if pre: ‹restart_abs_wl_pre S False›
       for S
  proof -
    obtain T U where
      ST: ‹(S, T) ∈ state_wl_l None› and
      ‹correct_watching S› and
      ‹blits_in_ℒin S› and
      TU: ‹(T, U) ∈ twl_st_l None› and
      struct: ‹twl_struct_invs U› and
      ‹twl_list_invs T› and
      ‹clauses_to_update_l T = {#}› and
      ‹twl_stgy_invs U› and
      confl: ‹get_conflict U = None›
      using pre unfolding restart_abs_wl_pre_def restart_abs_l_pre_def restart_prog_pre_def apply -
      by blast

   show ?C
     using ST TU confl by auto

   have alien: ‹cdclW_restart_mset.no_strange_atm (stateW_of U)›
     using struct unfolding twl_struct_invs_def cdclW_restart_mset.cdclW_all_struct_inv_def
     by fast+
   then show ?A and ?B
      subgoal
        using ST TU unfolding set_eq_iff in_set_all_atms_iff
          in_set_all_atms_iff in_set_all_init_atms_iff get_unit_clauses_wl_alt_def
        apply (subst all_clss_lf_ran_m[symmetric])
        unfolding image_mset_union
        apply (auto simp: cdclW_restart_mset.no_strange_atm_def twl_st twl_st_l in_set_all_atms_iff
          in_set_all_init_atms_iff)
        done
      subgoal
        using ST TU alien unfolding set_eq_iff in_set_all_atms_iff
          in_set_all_atms_iff in_set_all_init_atms_iff get_unit_clauses_wl_alt_def
        apply (subst all_clss_lf_ran_m[symmetric])
        apply (subst all_clss_lf_ran_m[symmetric])
        unfolding image_mset_union
        by (auto simp: cdclW_restart_mset.no_strange_atm_def twl_st twl_st_l in_set_all_atms_iff
          in_set_all_init_atms_iff)
     done
  qed
  have P: ‹P›
    if
      ST: ‹(((a, aa, ab, ac, ad, b), ae, heur, ah, ai,
	 ((aj, ak, al, am, bb), an, bc), ao, (aq, bd), ar, as,
	 (at', au, av, aw, be), ((ax, ay, az, bf, bg), (bh, bi, bj, bk, bl),
	 (bm, bn)), bo, bp, bq, br, bs),
	bt, bu, bv, bw, bx, NS, US, by, bz)
       ∈ twl_st_heur› and
      ‹restart_abs_wl_pre (bt, bu, bv, bw, bx, NS, US, by, bz) False› and
      ‹restart_abs_wl_heur_pre
	((a, aa, ab, ac, ad, b), ae, heur, ah, ai,
	 ((aj, ak, al, am, bb), an, bc), ao, (aq, bd), ar, as,
	 (at', au, av, aw, be), ((ax, ay, az, bf, bg), (bh, bi, bj, bk, bl),
	 (bm, bn)), bo, bp, bq, br, bs)
	False› and
      lvl: ‹(lvl, i)
       ∈ {(i, b).
	  b = (i = count_decided (get_trail_wl (bt, bu, bv, bw, bx, NS, US, by, bz))) ∧
	  i ≤ count_decided (get_trail_wl (bt, bu, bv, bw, bx, NS, US, by, bz))}› and
      ‹i ∈ {_. True}› and
      ‹lvl ≠
       count_decided_st_heur
	((a, aa, ab, ac, ad, b), ae, heur, ah, ai,
	 ((aj, ak, al, am, bb), an, bc), ao, (aq, bd), ar, as,
	 (at', au, av, aw, be), ((ax, ay, az, bf, bg), (bh, bi, bj, bk, bl),
	 (bm, bn)), bo, bp, bq, br, bs)› and
      i: ‹¬ i› and
    H: ‹(⋀vm0. ((an, bc), vm0) ∈ distinct_atoms_rel (all_atms_st (bt, bu, bv, bw, bx, NS, US, by, bz)) ⟹
           ((aj, ak, al, am, bb), vm0) ∈ vmtf (all_atms_st (bt, bu, bv, bw, bx, NS, US, by, bz)) bt ⟹
      isa_find_decomp_wl_imp (a, aa, ab, ac, ad, b) lvl
        ((aj, ak, al, am, bb), an, bc)
	≤ ⇓ {(a, b). (a,b) ∈ trail_pol (all_atms_st (bt, bu, bv, bw, bx, NS, US, by, bz)) ×f
               (Id ×f distinct_atoms_rel (all_atms_st (bt, bu, bv, bw, bx, NS, US, by, bz)))}
	    (find_decomp_w_ns (all_atms_st (bt, bu, bv, bw, bx, NS, US, by, bz)) bt lvl vm0) ⟹ P)›
    for a aa ab ac ad b ae af ag ba ah ai aj ak al am bb an bc ao aq bd ar as at'
       au av aw be ax ay az bf bg bh bi bj bk bl bm bn bo bp bq br bs bt bu bv
       bw bx "by" bz lvl i x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f
       x1g x2g x1h x2h x1i x2i P NS US heur
  proof -
    let ?𝒜 = ‹all_atms_st (bt, bu, bv, bw, bx, NS, US, by, bz)›
    have
      tr: ‹((a, aa, ab, ac, ad, b), bt) ∈ trail_pol ?𝒜› and
      ‹valid_arena ae bu (set bo)› and
      ‹(heur, bv)
       ∈ option_lookup_clause_rel ?𝒜› and
      ‹by = {#- lit_of x. x ∈# mset (drop ah (rev bt))#}› and
      ‹(ai, bz) ∈ ⟨Id⟩map_fun_rel (D0 ?𝒜)› and
      vm: ‹((aj, ak, al, am, bb), an, bc) ∈ isa_vmtf ?𝒜 bt› and
      ‹no_dup bt› and
      ‹ao ∈ counts_maximum_level bt bv› and
      ‹cach_refinement_empty ?𝒜 (aq, bd)› and
      ‹out_learned bt bv as› and
      ‹bq = size (learned_clss_l bu)› and
      ‹vdom_m ?𝒜 bz bu ⊆ set bo› and
      ‹set bp ⊆ set bo› and
      ‹∀L∈#ℒall ?𝒜. nat_of_lit L ≤ uint32_max› and
      ‹?𝒜 ≠ {#}› and
      bounded: ‹isasat_input_bounded ?𝒜› and
      heur: ‹heuristic_rel ?𝒜 ((ax, ay, az, bf, bg), (bh, bi, bj, bk, bl),
	 (bm, bn))›
      using ST unfolding twl_st_heur_def all_atms_def[symmetric]
      by (auto)

    obtain vm0 where
      vm: ‹((aj, ak, al, am, bb), vm0) ∈ vmtf ?𝒜 bt› and
      vm0: ‹((an, bc), vm0) ∈ distinct_atoms_rel ?𝒜›
      using vm
      by (auto simp: isa_vmtf_def)
    have n_d: ‹no_dup bt›
      using tr by (auto simp: trail_pol_def)
    show ?thesis
      apply (rule H)
      apply (rule vm0)
      apply (rule vm)
      apply (rule isa_find_decomp_wl_imp_find_decomp_wl_imp[THEN fref_to_Down_curry2, THEN order_trans,
        of bt lvl ‹((aj, ak, al, am, bb), vm0)› _ _ _ ‹?𝒜›])
      subgoal using lvl i by auto
      subgoal using vm0 tr by auto
      apply (subst (3) Down_id_eq[symmetric])
      apply (rule order_trans)
      apply (rule ref_two_step')
      apply (rule find_decomp_wl_imp_find_decomp_wl'[THEN fref_to_Down_curry2, of _ bt lvl
        ‹((aj, ak, al, am, bb), vm0)›])
      subgoal
        using that(1-8) vm vm0 bounded n_d tr
	by (auto simp: find_decomp_w_ns_pre_def dest: trail_pol_literals_are_in_ℒin_trail)
      subgoal by auto
        using ST
        by (auto simp: find_decomp_w_ns_def conc_fun_RES twl_st_heur_def)
  qed
  note cong =  trail_pol_cong heuristic_rel_cong
      option_lookup_clause_rel_cong D0_cong isa_vmtf_cong
      cach_refinement_empty_cong vdom_m_cong isasat_input_nempty_cong
      isasat_input_bounded_cong heuristic_rel_cong

  show ?thesis
    supply [[goals_limit=1]]
    unfolding cdcl_twl_local_restart_wl_D_heur_def
    unfolding
      find_decomp_wl_st_int_def find_local_restart_target_level_def incr_lrestart_stat_def
       empty_Q_def find_local_restart_target_level_st_def nres_monad_laws
    apply (intro frefI nres_relI)
    apply clarify
    apply (rule ref_two_step)
     prefer 2
     apply (rule cdcl_twl_local_restart_wl_D_spec_int)
    unfolding bind_to_let_conv Let_def RES_RETURN_RES2 nres_monad_laws
    apply (refine_vcg)
    subgoal unfolding restart_abs_wl_heur_pre_def by blast
    apply assumption
    subgoal by (auto simp: twl_st_heur_def count_decided_st_heur_def trail_pol_def)
    subgoal
      by (drule H(2)) (simp add: twl_st_heur_change_subsumed_clauses)

    apply (rule P)
    apply assumption+
      apply (rule refine_generalise1)
      apply assumption
    subgoal for a aa ab ac ad b ae af ag ba ah ai aj ak al am bb an bc ao ap bd aq ar
       as at au av aw ax ay be az bf bg bh bi bj bk bl bm bn bo bp bq br bs
       bt bu bv bw bx _ _ "by" bz ca cb cc cd ce cf cg ch ci cj ck cl cm cn co cp
       lvl i vm0
      unfolding RETURN_def RES_RES2_RETURN_RES RES_RES13_RETURN_RES find_decomp_w_ns_def conc_fun_RES
        RES_RES13_RETURN_RES K K2
      apply (auto simp: intro_spec_iff intro!: ASSERT_leI isa_length_trail_pre)
      apply (auto simp: isa_length_trail_length_u[THEN fref_to_Down_unRET_Id]
        intro: isa_vmtfI trail_pol_no_dup)
      apply (frule twl_st_heur_change_subsumed_clauses[where US' = ‹{#}› and NS' = cm])
      apply (solves ‹auto dest: H(2)›)[]
      apply (frule H(2))
      apply (frule H(3))
	apply (clarsimp simp: twl_st_heur_def)
	apply (rule_tac x=aja in exI)
	apply (auto simp: isa_length_trail_length_u[THEN fref_to_Down_unRET_Id]
	  intro: isa_vmtfI trail_pol_no_dup)
      apply (rule trail_pol_cong)
      apply assumption
      apply fast
      apply (rule isa_vmtf_cong)
      apply assumption
      apply (fast intro: isa_vmtfI)
      done
    done
qed


definition remove_all_annot_true_clause_imp_wl_D_heur_inv
  :: ‹twl_st_wl_heur ⇒ nat watcher list ⇒ nat × twl_st_wl_heur ⇒ bool›
where
  ‹remove_all_annot_true_clause_imp_wl_D_heur_inv S xs = (λ(i, T).
       ∃S' T'. (S, S') ∈ twl_st_heur_restart ∧ (T, T') ∈ twl_st_heur_restart ∧
         remove_all_annot_true_clause_imp_wl_inv S' (map fst xs) (i, T'))
     ›

definition remove_all_annot_true_clause_one_imp_heur
  :: ‹nat × nat × arena ⇒ (nat × arena) nres›
where
‹remove_all_annot_true_clause_one_imp_heur = (λ(C, j, N). do {
      case arena_status N C of
        DELETED ⇒ RETURN (j, N)
      | IRRED ⇒ RETURN (j, extra_information_mark_to_delete N C)
      | LEARNED ⇒ RETURN (j-1, extra_information_mark_to_delete N C)
  })›

definition remove_all_annot_true_clause_imp_wl_D_pre
  :: ‹nat multiset ⇒ nat literal ⇒ nat twl_st_wl ⇒ bool›
where
  ‹remove_all_annot_true_clause_imp_wl_D_pre 𝒜 L S ⟷ (L ∈# ℒall 𝒜)›

definition remove_all_annot_true_clause_imp_wl_D_heur_pre where
  ‹remove_all_annot_true_clause_imp_wl_D_heur_pre L S ⟷
    (∃S'. (S, S') ∈ twl_st_heur_restart
      ∧ remove_all_annot_true_clause_imp_wl_D_pre (all_init_atms_st S') L S')›


(* TODO: unfold Let when generating code! *)
definition remove_all_annot_true_clause_imp_wl_D_heur
  :: ‹nat literal ⇒ twl_st_wl_heur ⇒ twl_st_wl_heur nres›
where
‹remove_all_annot_true_clause_imp_wl_D_heur = (λL (M, N0, D, Q, W, vm, clvls, cach, lbd, outl,
       stats, heur, vdom, avdom, lcount, opts). do {
    ASSERT(remove_all_annot_true_clause_imp_wl_D_heur_pre L (M, N0, D, Q, W, vm, clvls,
       cach, lbd, outl, stats, heur,
       vdom, avdom, lcount, opts));
    let xs = W!(nat_of_lit L);
    (_, lcount', N) ← WHILETλ(i, j, N).
        remove_all_annot_true_clause_imp_wl_D_heur_inv
           (M, N0, D, Q, W, vm, clvls, cach, lbd, outl, stats,
	  heur, vdom, avdom, lcount, opts) xs
           (i, M, N, D, Q, W, vm, clvls, cach, lbd, outl, stats,
	  heur, vdom, avdom, j, opts)
      (λ(i, j, N). i < length xs)
      (λ(i, j, N). do {
        ASSERT(i < length xs);
        if clause_not_marked_to_delete_heur (M, N, D, Q, W, vm, clvls, cach, lbd, outl, stats,
	  heur, vdom, avdom, lcount, opts) i
        then do {
          (j, N) ← remove_all_annot_true_clause_one_imp_heur (fst (xs!i), j, N);
          ASSERT(remove_all_annot_true_clause_imp_wl_D_heur_inv
             (M, N0, D, Q, W, vm, clvls, cach, lbd, outl, stats,
	       heur, vdom, avdom, lcount, opts) xs
             (i, M, N, D, Q, W, vm, clvls, cach, lbd, outl, stats,
	       heur, vdom, avdom, j, opts));
          RETURN (i+1, j, N)
        }
        else
          RETURN (i+1, j, N)
      })
      (0, lcount, N0);
    RETURN (M, N, D, Q, W, vm, clvls, cach, lbd, outl, stats,
	  heur, vdom, avdom, lcount', opts)
  })›


definition minimum_number_between_restarts :: ‹64 word› where
  ‹minimum_number_between_restarts = 50›

definition five_uint64 :: ‹64 word› where
  ‹five_uint64 = 5›


definition upper_restart_bound_not_reached :: ‹twl_st_wl_heur ⇒ bool› where
  ‹upper_restart_bound_not_reached = (λ(M', N', D', j, W', vm, clvls, cach, lbd, outl,
    (props, decs, confl, restarts, _), heur, vdom, avdom, lcount, opts).
    of_nat lcount < 3000 + 1000 * restarts)›

definition (in -) lower_restart_bound_not_reached :: ‹twl_st_wl_heur ⇒ bool› where
  ‹lower_restart_bound_not_reached = (λ(M', N', D', j, W', vm, clvls, cach, lbd, outl,
        (props, decs, confl, restarts, _), heur,
       vdom, avdom, lcount, opts, old).
     (¬opts_reduce opts ∨ (opts_restart opts ∧ (of_nat lcount < 2000 + 1000 * restarts))))›

definition reorder_vdom_wl :: ‹'v twl_st_wl ⇒ 'v twl_st_wl nres› where
  ‹reorder_vdom_wl S = RETURN S›

definition sort_clauses_by_score :: ‹arena ⇒ nat list ⇒ nat list nres› where
  ‹sort_clauses_by_score arena vdom = do {
      ASSERT(∀i∈set vdom. valid_sort_clause_score_pre_at arena i);
      SPEC(λvdom'. mset vdom = mset vdom')
  }›

definition (in -) quicksort_clauses_by_score :: ‹arena ⇒ nat list ⇒ nat list nres› where
  ‹quicksort_clauses_by_score arena =
    full_quicksort_ref clause_score_ordering2 (clause_score_extract arena)›

lemma quicksort_clauses_by_score_sort:
 ‹(quicksort_clauses_by_score, sort_clauses_by_score) ∈
   Id → Id → ⟨Id⟩nres_rel›
   by (intro fun_relI nres_relI)
    (auto simp: quicksort_clauses_by_score_def sort_clauses_by_score_def
       reorder_list_def  clause_score_extract_def clause_score_ordering2_def
       le_ASSERT_iff
     intro!: insert_sort_reorder_list[THEN fref_to_Down, THEN order_trans])

definition remove_deleted_clauses_from_avdom :: ‹_› where
‹remove_deleted_clauses_from_avdom N avdom0 = do {
  let n = length avdom0;
  (i, j, avdom) ← WHILET λ(i, j, avdom). i ≤ j ∧ j ≤ n ∧ length avdom = length avdom0 ∧
         mset (take i avdom @ drop j avdom) ⊆# mset avdom0
    (λ(i, j, avdom). j < n)
    (λ(i, j, avdom). do {
      ASSERT(j < length avdom);
      if (avdom ! j) ∈# dom_m N then RETURN (i+1, j+1, swap avdom i j)
      else RETURN (i, j+1, avdom)
    })
    (0, 0, avdom0);
  ASSERT(i ≤ length avdom);
  RETURN (take i avdom)
}›

lemma remove_deleted_clauses_from_avdom:
  ‹remove_deleted_clauses_from_avdom N avdom0 ≤ SPEC(λavdom. mset avdom ⊆# mset avdom0)›
  unfolding remove_deleted_clauses_from_avdom_def Let_def
  apply (refine_vcg WHILEIT_rule[where R = ‹measure (λ(i, j, avdom). length avdom - j)›])
  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 for s a b aa ba x1 x2 x1a x2a
     by (cases ‹Suc a ≤ aa›)
      (auto simp: drop_swap_irrelevant swap_only_first_relevant Suc_le_eq take_update_last
        mset_append[symmetric] Cons_nth_drop_Suc simp del: mset_append
      simp flip:  take_Suc_conv_app_nth)
  subgoal by auto
  subgoal by auto
  subgoal by auto
  subgoal by auto
  subgoal for s a b aa ba x1 x2 x1a x2a
     by (cases ‹Suc aa ≤ length x2a›)
       (auto simp: drop_swap_irrelevant swap_only_first_relevant Suc_le_eq take_update_last
         Cons_nth_drop_Suc[symmetric] intro: subset_mset.dual_order.trans
      simp flip:  take_Suc_conv_app_nth)
  subgoal by auto
  subgoal by auto
  subgoal by auto
  done

definition isa_remove_deleted_clauses_from_avdom :: ‹_› where
‹isa_remove_deleted_clauses_from_avdom arena avdom0 = do {
  ASSERT(length avdom0 ≤ length arena);
  let n = length avdom0;
  (i, j, avdom) ← WHILET λ(i, j, _). i ≤ j ∧ j ≤ n
    (λ(i, j, avdom). j < n)
    (λ(i, j, avdom). do {
      ASSERT(j < n);
      ASSERT(arena_is_valid_clause_vdom arena (avdom!j) ∧ j < length avdom ∧ i < length avdom);
      if arena_status arena (avdom ! j) ≠ DELETED then RETURN (i+1, j+1, swap avdom i j)
      else RETURN (i, j+1, avdom)
    }) (0, 0, avdom0);
  ASSERT(i ≤ length avdom);
  RETURN (take i avdom)
}›

lemma isa_remove_deleted_clauses_from_avdom_remove_deleted_clauses_from_avdom:
   ‹valid_arena arena N (set vdom) ⟹ mset avdom0 ⊆# mset vdom ⟹ distinct vdom ⟹
   isa_remove_deleted_clauses_from_avdom arena avdom0 ≤ ⇓Id (remove_deleted_clauses_from_avdom N avdom0)›
  unfolding isa_remove_deleted_clauses_from_avdom_def remove_deleted_clauses_from_avdom_def Let_def
  apply (refine_vcg WHILEIT_refine[where R= ‹Id ×r Id ×r ⟨Id⟩list_rel›])
  subgoal by (auto dest!: valid_arena_vdom_le(2) size_mset_mono simp: distinct_card)
  subgoal by auto
  subgoal by auto
  subgoal by auto
  subgoal by auto
  subgoal by auto
  subgoal for x x' x1 x2 x1a x2a x1b x2b x1c x2c unfolding arena_is_valid_clause_vdom_def
     by (force intro!: exI[of _ N] exI[of _ vdom] dest!: mset_eq_setD dest: mset_le_add_mset simp: Cons_nth_drop_Suc[symmetric])
  subgoal by auto
  subgoal by auto
  subgoal
     by (force simp: arena_lifting arena_dom_status_iff(1) Cons_nth_drop_Suc[symmetric]
       dest!: mset_eq_setD dest: mset_le_add_mset)
  subgoal by auto
  subgoal
     by (force simp: arena_lifting arena_dom_status_iff(1) Cons_nth_drop_Suc[symmetric]
       dest!: mset_eq_setD dest: mset_le_add_mset)
  subgoal by auto
  subgoal by auto
  done

definition (in -) sort_vdom_heur :: ‹twl_st_wl_heur ⇒ twl_st_wl_heur nres› where
  ‹sort_vdom_heur = (λ(M', arena, D', j, W', vm, clvls, cach, lbd, outl, stats, heur,
       vdom, avdom, lcount). do {
    ASSERT(length avdom ≤ length arena);
    avdom ← isa_remove_deleted_clauses_from_avdom arena avdom;
    ASSERT(valid_sort_clause_score_pre arena avdom);
    ASSERT(length avdom ≤ length arena);
    avdom ← sort_clauses_by_score arena avdom;
    RETURN (M', arena, D', j, W', vm, clvls, cach, lbd, outl, stats, heur,
       vdom, avdom, lcount)
    })›


lemma sort_clauses_by_score_reorder:
  ‹valid_arena arena N (set vdom') ⟹ set vdom ⊆ set vdom' ⟹
    sort_clauses_by_score arena vdom ≤ SPEC(λvdom'. mset vdom = mset vdom')›
  unfolding sort_clauses_by_score_def
  apply refine_vcg
  unfolding valid_sort_clause_score_pre_def arena_is_valid_clause_vdom_def
    get_clause_LBD_pre_def arena_is_valid_clause_idx_def arena_act_pre_def
    valid_sort_clause_score_pre_at_def
  apply (auto simp: valid_sort_clause_score_pre_def twl_st_heur_restart_ana_def arena_dom_status_iff(2-)
    arena_dom_status_iff(1)[symmetric] in_set_conv_nth
    arena_act_pre_def get_clause_LBD_pre_def arena_is_valid_clause_idx_def twl_st_heur_restart_def
    intro!: exI[of _ ‹get_clauses_wl y›]  dest!: set_mset_mono mset_subset_eqD)
  using arena_dom_status_iff(1) nth_mem by blast

lemma sort_vdom_heur_reorder_vdom_wl:
  ‹(sort_vdom_heur, reorder_vdom_wl) ∈ twl_st_heur_restart_ana r →f ⟨twl_st_heur_restart_ana r⟩nres_rel›
proof -
  show ?thesis
    unfolding reorder_vdom_wl_def sort_vdom_heur_def
    apply (intro frefI nres_relI)
    apply refine_rcg
    apply (rule ASSERT_leI)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
    apply (rule specify_left)
    apply (rule_tac N1 = ‹get_clauses_wl y› and vdom1 = ‹get_vdom x› in
     order_trans[OF isa_remove_deleted_clauses_from_avdom_remove_deleted_clauses_from_avdom,
      unfolded Down_id_eq, OF _ _ _ remove_deleted_clauses_from_avdom])
    subgoal for x y x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h
       x1i x2i x1j x2l x1m x2m x1n x2n x1o x2o
      by (case_tac y; auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def mem_Collect_eq prod.case)
    subgoal for x y x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h
       x1i x2i x1j x2j x1k x2k x1l x2l x1m x2m
      by (case_tac y; auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def mem_Collect_eq prod.case)
    subgoal for x y x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h
       x1i x2i x1j x2j x1k x2k x1l x2l x1m x2m
      by (case_tac y; auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def mem_Collect_eq prod.case)
    apply (subst assert_bind_spec_conv, intro conjI)
    subgoal for x y
      unfolding valid_sort_clause_score_pre_def arena_is_valid_clause_vdom_def
        get_clause_LBD_pre_def arena_is_valid_clause_idx_def arena_act_pre_def
      by (force simp: valid_sort_clause_score_pre_def twl_st_heur_restart_ana_def arena_dom_status_iff(2-)
        arena_dom_status_iff(1)[symmetric]
        arena_act_pre_def get_clause_LBD_pre_def arena_is_valid_clause_idx_def twl_st_heur_restart_def
        intro!: exI[of _ ‹get_clauses_wl y›]  dest!: set_mset_mono mset_subset_eqD)
    apply (subst assert_bind_spec_conv, intro conjI)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
    subgoal for x y
      apply (rewrite at ‹_ ≤ ⌑› Down_id_eq[symmetric])
      apply (rule bind_refine_spec)
      prefer 2
      apply (rule sort_clauses_by_score_reorder[of _ ‹get_clauses_wl y›  ‹get_vdom x›])
      by (auto 5 3 simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest: mset_eq_setD)
    done
qed

lemma (in -) insort_inner_clauses_by_score_invI:
   ‹valid_sort_clause_score_pre a ba ⟹
       mset ba = mset a2' ⟹
       a1' < length a2' ⟹
       valid_sort_clause_score_pre_at a (a2' ! a1')›
  unfolding valid_sort_clause_score_pre_def all_set_conv_nth valid_sort_clause_score_pre_at_def
  by (metis in_mset_conv_nth)+


lemma sort_clauses_by_score_invI:
  ‹valid_sort_clause_score_pre a b ⟹
       mset b = mset a2' ⟹ valid_sort_clause_score_pre a a2'›
  using mset_eq_setD unfolding valid_sort_clause_score_pre_def by blast

definition partition_main_clause where
  ‹partition_main_clause arena = partition_main clause_score_ordering (clause_score_extract arena)›

definition partition_clause where
  ‹partition_clause arena = partition_between_ref clause_score_ordering  (clause_score_extract arena)›

lemma valid_sort_clause_score_pre_swap:
  ‹valid_sort_clause_score_pre a b ⟹ x < length b ⟹
       ba < length b ⟹ valid_sort_clause_score_pre a (swap b x ba)›
  by (auto simp: valid_sort_clause_score_pre_def)

definition div2 where [simp]: ‹div2 n = n div 2›

definition safe_minus where ‹safe_minus a b = (if b ≥ a then 0 else a - b)›

definition max_restart_decision_lvl :: nat where
  ‹max_restart_decision_lvl = 300›

definition max_restart_decision_lvl_code :: ‹32 word› where
  ‹max_restart_decision_lvl_code = 300›

definition GC_EVERY :: ‹64 word› where
  ‹GC_EVERY = 15› ―‹hard-coded limit›

fun (in -) get_reductions_count :: ‹twl_st_wl_heur ⇒ 64 word› where
  ‹get_reductions_count (_, _, _, _, _, _, _,_,_,_,
       (_, _, _, lres, _, _), _)
      = lres›

definition get_restart_phase :: ‹twl_st_wl_heur ⇒ 64 word› where
  ‹get_restart_phase = (λ(_, _, _, _, _, _, _, _, _, _, _, heur, _).
     current_restart_phase heur)›

definition GC_required_heur :: "twl_st_wl_heur ⇒ nat ⇒ bool nres" where
  ‹GC_required_heur S n = do {
    n ← RETURN (full_arena_length_st S);
    wasted ← RETURN (wasted_bytes_st S);
    RETURN (3*wasted > ((of_nat n)>>2))
 }›


definition FLAG_no_restart :: ‹8 word› where
  ‹FLAG_no_restart = 0›

definition FLAG_restart :: ‹8 word› where
  ‹FLAG_restart = 1›

definition FLAG_GC_restart :: ‹8 word› where
  ‹FLAG_GC_restart = 2›

definition restart_flag_rel :: ‹(8 word × restart_type) set› where
  ‹restart_flag_rel = {(FLAG_no_restart, NO_RESTART), (FLAG_restart, RESTART), (FLAG_GC_restart, GC)}›

definition restart_required_heur :: "twl_st_wl_heur ⇒ nat ⇒ 8 word nres" where
  ‹restart_required_heur S n = do {
    let opt_red = opts_reduction_st S;
    let opt_res = opts_restart_st S;
    let curr_phase = get_restart_phase S;
    let lcount = get_learned_count S;
    let can_res = (lcount > n);

    if ¬can_res ∨ ¬opt_res ∨ ¬opt_red then RETURN FLAG_no_restart
    else if curr_phase = QUIET_PHASE
    then do {
      GC_required ← GC_required_heur S n;
      let upper = upper_restart_bound_not_reached S;
      if (opt_res ∨ opt_red) ∧ ¬upper
      then RETURN FLAG_GC_restart
      else RETURN FLAG_no_restart
    }
    else do {
      let sema = ema_get_value (get_slow_ema_heur S);
      let limit = (shiftr (11 * sema) (4::nat));
      let fema = ema_get_value (get_fast_ema_heur S);
      let ccount = get_conflict_count_since_last_restart_heur S;
      let min_reached = (ccount > minimum_number_between_restarts);
      let level = count_decided_st_heur S;
      let should_not_reduce = (¬opt_red ∨ upper_restart_bound_not_reached S);
      let should_reduce = ((opt_res ∨ opt_red) ∧
         (should_not_reduce ⟶ limit > fema) ∧ min_reached ∧ can_res ∧
        level > 2 ∧ ⌦‹This comment from Marijn Heule seems not to help:
           \<^term>‹level < max_restart_decision_lvl››
        of_nat level > (shiftr fema 32));
      GC_required ← GC_required_heur S n;
      if should_reduce
      then if GC_required
        then RETURN FLAG_GC_restart
        else RETURN FLAG_restart
      else RETURN FLAG_no_restart
     }
   }›


lemma (in -) get_reduction_count_alt_def:
   ‹RETURN o get_reductions_count = (λ(M, N0, D, Q, W, vm, clvls, cach, lbd, outl,
       (_, _, _, lres, _, _), heur, lcount). RETURN lres)›
  by auto


definition mark_to_delete_clauses_wl_D_heur_pre :: ‹twl_st_wl_heur ⇒ bool› where
  ‹mark_to_delete_clauses_wl_D_heur_pre S ⟷
    (∃S'. (S, S') ∈ twl_st_heur_restart ∧ mark_to_delete_clauses_wl_pre S')›

lemma mark_to_delete_clauses_wl_post_alt_def:
  ‹mark_to_delete_clauses_wl_post S0 S ⟷
    (∃T0 T.
        (S0, T0) ∈ state_wl_l None ∧
        (S, T) ∈ state_wl_l None ∧
        blits_in_ℒin S0 ∧
        blits_in_ℒin S ∧
        (∃U0 U. (T0, U0) ∈ twl_st_l None ∧
               (T, U) ∈ twl_st_l None ∧
               remove_one_annot_true_clause** T0 T ∧
               twl_list_invs T0 ∧
               twl_struct_invs U0 ∧
               twl_list_invs T ∧
               twl_struct_invs U ∧
               get_conflict_l T0 = None ∧
	       clauses_to_update_l T0 = {#}) ∧
        correct_watching S0 ∧ correct_watching S)›
  unfolding mark_to_delete_clauses_wl_post_def mark_to_delete_clauses_l_post_def
    mark_to_delete_clauses_l_pre_def
  apply (rule iffI; normalize_goal+)
  subgoal for T0 T U0
    apply (rule exI[of _ T0])
    apply (rule exI[of _ T])
    apply (intro conjI)
    apply auto[4]
    apply (rule exI[of _ U0])
    apply auto
    using rtranclp_remove_one_annot_true_clause_cdcl_twl_restart_l2[of T0 T U0]
      rtranclp_cdcl_twl_restart_l_list_invs[of T0]
    apply (auto dest: )
     using rtranclp_cdcl_twl_restart_l_list_invs by blast
  subgoal for T0 T U0 U
    apply (rule exI[of _ T0])
    apply (rule exI[of _ T])
    apply (intro conjI)
    apply auto[3]
    apply (rule exI[of _ U0])
    apply auto
    done
  done

lemma mark_to_delete_clauses_wl_D_heur_pre_alt_def:
    ‹mark_to_delete_clauses_wl_D_heur_pre S ⟷
      (∃S'. (S, S') ∈ twl_st_heur ∧ mark_to_delete_clauses_wl_pre S')› (is ?A) and
    mark_to_delete_clauses_wl_D_heur_pre_twl_st_heur:
      ‹mark_to_delete_clauses_wl_pre T ⟹
        (S, T) ∈ twl_st_heur ⟷ (S, T) ∈ twl_st_heur_restart› (is ‹_ ⟹ _ ?B›) and
    mark_to_delete_clauses_wl_post_twl_st_heur:
      ‹mark_to_delete_clauses_wl_post T0 T ⟹
        (S, T) ∈ twl_st_heur ⟷ (S, T) ∈ twl_st_heur_restart› (is ‹_ ⟹ _ ?C›)
proof -
  note cong = trail_pol_cong heuristic_rel_cong
      option_lookup_clause_rel_cong D0_cong isa_vmtf_cong phase_saving_cong
      cach_refinement_empty_cong vdom_m_cong isasat_input_nempty_cong
      isasat_input_bounded_cong

  show ?A
    supply [[goals_limit=1]]
    unfolding mark_to_delete_clauses_wl_D_heur_pre_def mark_to_delete_clauses_wl_pre_def
      mark_to_delete_clauses_l_pre_def
    apply (rule iffI)
    apply normalize_goal+
    subgoal for T U V
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T U V]
        cong[of ‹all_init_atms_st T› ‹all_atms_st T›]
	vdom_m_cong[of ‹all_init_atms_st T› ‹all_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
      apply -
      apply (rule exI[of _ T])
      apply (intro conjI) defer
      apply (rule exI[of _ U])
      apply (intro conjI) defer
      apply (rule exI[of _ V])
      apply (simp_all del: isasat_input_nempty_def isasat_input_bounded_def)
      apply (cases S; cases T)
      by (simp add: twl_st_heur_def twl_st_heur_restart_def del: isasat_input_nempty_def)
    apply normalize_goal+
    subgoal for T U V
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T U V]
        cong[of ‹all_atms_st T› ‹all_init_atms_st T›]
	vdom_m_cong[of ‹all_atms_st T› ‹all_init_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
      apply -
      apply (rule exI[of _ T])
      apply (intro conjI) defer
      apply (rule exI[of _ U])
      apply (intro conjI) defer
      apply (rule exI[of _ V])
      apply (simp_all del: isasat_input_nempty_def isasat_input_bounded_def)
      apply (cases S; cases T)
      by (simp add: twl_st_heur_def twl_st_heur_restart_def del: isasat_input_nempty_def)
    done
  show ‹mark_to_delete_clauses_wl_pre T ⟹ ?B›
    supply [[goals_limit=1]]
    unfolding mark_to_delete_clauses_wl_D_heur_pre_def mark_to_delete_clauses_wl_pre_def
      mark_to_delete_clauses_l_pre_def mark_to_delete_clauses_wl_pre_def
    apply normalize_goal+
    apply (rule iffI)
    subgoal for U V
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T U V]
        cong[of ‹all_atms_st T› ‹all_init_atms_st T›]
	vdom_m_cong[of ‹all_atms_st T› ‹all_init_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
      apply -
      apply (simp_all del: isasat_input_nempty_def isasat_input_bounded_def)
      apply (cases S; cases T)
      by (simp add: twl_st_heur_def twl_st_heur_restart_def del: isasat_input_nempty_def)
    subgoal for U V
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T U V]
        cong[of ‹all_init_atms_st T› ‹all_atms_st T›]
	vdom_m_cong[of ‹all_init_atms_st T› ‹all_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
      apply -
      apply (cases S; cases T)
      by (simp add: twl_st_heur_def twl_st_heur_restart_def del: isasat_input_nempty_def)
    done
  show  ‹mark_to_delete_clauses_wl_post T0 T ⟹ ?C›
    supply [[goals_limit=1]]
    unfolding mark_to_delete_clauses_wl_post_alt_def
    apply normalize_goal+
    apply (rule iffI)
    subgoal for U0 U V0 V
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T U V]
        cong[of ‹all_atms_st T› ‹all_init_atms_st T›]
	vdom_m_cong[of ‹all_atms_st T› ‹all_init_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
      apply -
      apply (simp_all del: isasat_input_nempty_def isasat_input_bounded_def)
      apply (cases S; cases T)
      apply (simp add: twl_st_heur_def twl_st_heur_restart_def del: isasat_input_nempty_def)
      done
    subgoal for U0 U V0 V
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T U V]
        cong[of ‹all_init_atms_st T› ‹all_atms_st T›]
	vdom_m_cong[of ‹all_init_atms_st T› ‹all_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
      apply -
      apply (cases S; cases T)
      by (simp add: twl_st_heur_def twl_st_heur_restart_def del: isasat_input_nempty_def)
    done

qed

lemma mark_garbage_heur_wl:
  assumes
    ‹(S, T) ∈ twl_st_heur_restart› and
    ‹C ∈# dom_m (get_clauses_wl T)› and
    ‹¬ irred (get_clauses_wl T) C› and ‹i < length (get_avdom S)›
  shows ‹(mark_garbage_heur C i S, mark_garbage_wl C T) ∈ twl_st_heur_restart›
  using assms
  apply (cases S; cases T)
   apply (simp add: twl_st_heur_restart_def mark_garbage_heur_def mark_garbage_wl_def)
  apply (auto simp: twl_st_heur_restart_def mark_garbage_heur_def mark_garbage_wl_def
         learned_clss_l_l_fmdrop size_remove1_mset_If
     simp: all_init_atms_def all_init_lits_def mset_butlast_remove1_mset
     simp del: all_init_atms_def[symmetric]
     intro: valid_arena_extra_information_mark_to_delete'
      dest!: in_set_butlastD in_vdom_m_fmdropD
      elim!: in_set_upd_cases)
  done


lemma mark_garbage_heur_wl_ana:
  assumes
    ‹(S, T) ∈ twl_st_heur_restart_ana r› and
    ‹C ∈# dom_m (get_clauses_wl T)› and
    ‹¬ irred (get_clauses_wl T) C› and ‹i < length (get_avdom S)›
  shows ‹(mark_garbage_heur C i S, mark_garbage_wl C T) ∈ twl_st_heur_restart_ana r›
  using assms
  apply (cases S; cases T)
   apply (simp add: twl_st_heur_restart_ana_def mark_garbage_heur_def mark_garbage_wl_def)
  apply (auto simp: twl_st_heur_restart_def mark_garbage_heur_def mark_garbage_wl_def
         learned_clss_l_l_fmdrop size_remove1_mset_If init_clss_l_fmdrop_irrelev
     simp: all_init_atms_def all_init_lits_def
     simp del: all_init_atms_def[symmetric]
     intro: valid_arena_extra_information_mark_to_delete'
      dest!: in_set_butlastD in_vdom_m_fmdropD
      elim!: in_set_upd_cases)
  done

lemma mark_unused_st_heur_ana:
  assumes
    ‹(S, T) ∈ twl_st_heur_restart_ana r› and
    ‹C ∈# dom_m (get_clauses_wl T)›
  shows ‹(mark_unused_st_heur C S, T) ∈ twl_st_heur_restart_ana r›
  using assms
  apply (cases S; cases T)
   apply (simp add: twl_st_heur_restart_ana_def mark_unused_st_heur_def)
  apply (auto simp: twl_st_heur_restart_def mark_garbage_heur_def mark_garbage_wl_def
         learned_clss_l_l_fmdrop size_remove1_mset_If
     simp: all_init_atms_def all_init_lits_def
     simp del: all_init_atms_def[symmetric]
     intro!: valid_arena_mark_unused valid_arena_arena_decr_act
     dest!: in_set_butlastD in_vdom_m_fmdropD
     elim!: in_set_upd_cases)
  done

lemma twl_st_heur_restart_valid_arena[twl_st_heur_restart]:
  assumes
    ‹(S, T) ∈ twl_st_heur_restart›
  shows ‹valid_arena (get_clauses_wl_heur S) (get_clauses_wl T) (set (get_vdom S))›
  using assms by (auto simp: twl_st_heur_restart_def)

lemma twl_st_heur_restart_get_avdom_nth_get_vdom[twl_st_heur_restart]:
  assumes
    ‹(S, T) ∈ twl_st_heur_restart› ‹i < length (get_avdom S)›
  shows ‹get_avdom S ! i ∈ set (get_vdom S)›
  using assms by (auto 5 3 simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: set_mset_mono)

lemma [twl_st_heur_restart]:
  assumes
    ‹(S, T) ∈ twl_st_heur_restart› and
    ‹C ∈ set (get_avdom S)›
  shows ‹clause_not_marked_to_delete_heur S C ⟷
         (C ∈# dom_m (get_clauses_wl T))› and
    ‹C ∈# dom_m (get_clauses_wl T) ⟹ arena_lit (get_clauses_wl_heur S) C = get_clauses_wl T ∝ C ! 0›and
    ‹C ∈# dom_m (get_clauses_wl T) ⟹ arena_status (get_clauses_wl_heur S) C = LEARNED ⟷ ¬irred (get_clauses_wl T) C›
    ‹C ∈# dom_m (get_clauses_wl T) ⟹ arena_length (get_clauses_wl_heur S) C = length (get_clauses_wl T ∝ C)›
proof -
  show ‹clause_not_marked_to_delete_heur S C ⟷ (C ∈# dom_m (get_clauses_wl T))›
    using assms
    by (cases S; cases T)
      (auto simp add: twl_st_heur_restart_def clause_not_marked_to_delete_heur_def
          arena_dom_status_iff(1)
        split: prod.splits)
  assume C: ‹C ∈# dom_m (get_clauses_wl T)›
  show ‹arena_lit (get_clauses_wl_heur S) C = get_clauses_wl T ∝ C ! 0›
    using assms C
    by (cases S; cases T)
      (auto simp add: twl_st_heur_restart_def clause_not_marked_to_delete_heur_def
          arena_lifting
        split: prod.splits)
  show ‹arena_status (get_clauses_wl_heur S) C = LEARNED ⟷ ¬irred (get_clauses_wl T) C›
    using assms C
    by (cases S; cases T)
      (auto simp add: twl_st_heur_restart_def clause_not_marked_to_delete_heur_def
          arena_lifting
        split: prod.splits)
  show ‹arena_length (get_clauses_wl_heur S) C = length (get_clauses_wl T ∝ C)›
    using assms C
    by (cases S; cases T)
      (auto simp add: twl_st_heur_restart_def clause_not_marked_to_delete_heur_def
          arena_lifting
        split: prod.splits)
qed

definition number_clss_to_keep :: ‹twl_st_wl_heur ⇒ nat nres› where
  ‹number_clss_to_keep = (λ(M', N', D', j, W', vm, clvls, cach, lbd, outl,
      (props, decs, confl, restarts, _), heur,
       vdom, avdom, lcount).
    RES UNIV)›

definition number_clss_to_keep_impl :: ‹twl_st_wl_heur ⇒ nat nres› where
  ‹number_clss_to_keep_impl = (λ(M', N', D', j, W', vm, clvls, cach, lbd, outl,
      (props, decs, confl, restarts, _), heur,
       vdom, avdom, lcount).
    let n = unat (1000 + 150 * restarts) in RETURN (if n ≥ sint64_max then sint64_max else n))›

lemma number_clss_to_keep_impl_number_clss_to_keep:
  ‹(number_clss_to_keep_impl, number_clss_to_keep) ∈ Id →f ⟨nat_rel⟩nres_rel›
  by (auto simp: number_clss_to_keep_impl_def number_clss_to_keep_def Let_def intro!: frefI nres_relI)

(* TODO Missing : The sorting function + definition of l should depend on the number of initial
  clauses *)
definition (in -) MINIMUM_DELETION_LBD :: nat where
  ‹MINIMUM_DELETION_LBD = 3›

lemma in_set_delete_index_and_swapD:
  ‹x ∈ set (delete_index_and_swap xs i) ⟹ x ∈ set xs›
  apply (cases ‹i < length xs›)
  apply (auto dest!: in_set_butlastD)
  by (metis List.last_in_set in_set_upd_cases list.size(3) not_less_zero)


lemma delete_index_vdom_heur_twl_st_heur_restart:
  ‹(S, T) ∈ twl_st_heur_restart ⟹ i < length (get_avdom S) ⟹
    (delete_index_vdom_heur i S, T) ∈ twl_st_heur_restart›
  by (auto simp: twl_st_heur_restart_def delete_index_vdom_heur_def
    dest: in_set_delete_index_and_swapD)


lemma delete_index_vdom_heur_twl_st_heur_restart_ana:
  ‹(S, T) ∈ twl_st_heur_restart_ana r ⟹ i < length (get_avdom S) ⟹
    (delete_index_vdom_heur i S, T) ∈ twl_st_heur_restart_ana r›
  by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def delete_index_vdom_heur_def
    dest: in_set_delete_index_and_swapD)

definition mark_clauses_as_unused_wl_D_heur
  :: ‹nat ⇒ twl_st_wl_heur ⇒ twl_st_wl_heur nres›
where
‹mark_clauses_as_unused_wl_D_heur  = (λi S. do {
    (_, T) ← WHILET
      (λ(i, S). i < length (get_avdom S))
      (λ(i, T). do {
        ASSERT(i < length (get_avdom T));
        ASSERT(length (get_avdom T) ≤ length (get_avdom S));
        ASSERT(access_vdom_at_pre T i);
        let C = get_avdom T ! i;
        ASSERT(clause_not_marked_to_delete_heur_pre (T, C));
        if ¬clause_not_marked_to_delete_heur T C then RETURN (i, delete_index_vdom_heur i T)
        else do {
          ASSERT(arena_act_pre (get_clauses_wl_heur T) C);
          RETURN (i+1, (mark_unused_st_heur C T))
        }
      })
      (i, S);
    RETURN T
  })›

lemma avdom_delete_index_vdom_heur[simp]:
  ‹get_avdom (delete_index_vdom_heur i S) =
     delete_index_and_swap (get_avdom S) i›
  by (cases S) (auto simp: delete_index_vdom_heur_def)

lemma incr_wasted_st:
  assumes
    ‹(S, T) ∈ twl_st_heur_restart_ana r›
  shows ‹(incr_wasted_st C S, T) ∈ twl_st_heur_restart_ana r›
  using assms
  apply (cases S; cases T)
   apply (simp add: twl_st_heur_restart_ana_def incr_wasted_st_def)
  apply (auto simp: twl_st_heur_restart_def mark_garbage_heur_def mark_garbage_wl_def
         learned_clss_l_l_fmdrop size_remove1_mset_If
     simp: all_init_atms_def all_init_lits_def heuristic_rel_def
     simp del: all_init_atms_def[symmetric]
     intro!: valid_arena_mark_unused valid_arena_arena_decr_act
     dest!: in_set_butlastD in_vdom_m_fmdropD
     elim!: in_set_upd_cases)
  done

lemma incr_wasted_st_twl_st[simp]:
  ‹get_avdom (incr_wasted_st w T) = get_avdom T›
  ‹get_vdom (incr_wasted_st w T) = get_vdom T›
  ‹get_trail_wl_heur (incr_wasted_st w T) = get_trail_wl_heur T›
  ‹get_clauses_wl_heur (incr_wasted_st C T) = get_clauses_wl_heur T›
  ‹get_conflict_wl_heur (incr_wasted_st C T) = get_conflict_wl_heur T›
  ‹get_learned_count (incr_wasted_st C T) = get_learned_count T›
  ‹get_conflict_count_heur (incr_wasted_st C T) = get_conflict_count_heur T›
  by (cases T; auto simp: incr_wasted_st_def)+

lemma mark_clauses_as_unused_wl_D_heur:
  assumes ‹(S, T) ∈ twl_st_heur_restart_ana r›
  shows ‹mark_clauses_as_unused_wl_D_heur i S ≤ ⇓ (twl_st_heur_restart_ana r) (SPEC ( (=) T))›
proof -
  have 1: ‹ ⇓ (twl_st_heur_restart_ana r) (SPEC ((=) T)) = do {
      (i, T) ← SPEC (λ(i::nat, T'). (T', T) ∈ twl_st_heur_restart_ana r);
      RETURN T
    }›
    by (auto simp: RES_RETURN_RES2 uncurry_def conc_fun_RES)
  show ?thesis
    unfolding mark_clauses_as_unused_wl_D_heur_def 1 mop_arena_length_st_def
    apply (rule Refine_Basic.bind_mono)
    subgoal
      apply (refine_vcg
         WHILET_rule[where R = ‹measure (λ(i, T). length (get_avdom T) - i)› and
	   I = ‹λ(_, S'). (S', T) ∈ twl_st_heur_restart_ana r ∧ length (get_avdom S') ≤ length(get_avdom S)›])
      subgoal by auto
      subgoal using assms by auto
      subgoal by auto
      subgoal by auto
      subgoal by auto
      subgoal unfolding access_vdom_at_pre_def by auto
      subgoal for st a S'
        unfolding clause_not_marked_to_delete_heur_pre_def
	  arena_is_valid_clause_vdom_def
        by (auto 7 3 simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: set_mset_mono
          intro!: exI[of _ ‹get_clauses_wl T›]  exI[of _ ‹set (get_vdom S')›])
      subgoal
        by (auto intro: delete_index_vdom_heur_twl_st_heur_restart_ana)
      subgoal by auto
      subgoal by auto
      subgoal
        unfolding arena_is_valid_clause_idx_def
	  arena_is_valid_clause_vdom_def arena_act_pre_def
       by (fastforce simp: twl_st_heur_restart_def twl_st_heur_restart
            dest!: twl_st_heur_restart_anaD)
      subgoal for s a b
        apply (auto intro!: mark_unused_st_heur_ana)
        unfolding arena_act_pre_def arena_is_valid_clause_idx_def
          arena_is_valid_clause_idx_def
          arena_is_valid_clause_vdom_def arena_act_pre_def
        by (fastforce simp: twl_st_heur_restart_def twl_st_heur_restart
            intro!: mark_unused_st_heur_ana
            dest!: twl_st_heur_restart_anaD)
      subgoal
        unfolding twl_st_heur_restart_ana_def
        by (auto simp: twl_st_heur_restart_def)
      subgoal
        by (auto intro!: mark_unused_st_heur_ana incr_wasted_st simp: twl_st_heur_restart
          dest: twl_st_heur_restart_anaD)
      subgoal by auto
      done
      subgoal by auto
      done
qed

definition mark_to_delete_clauses_wl_D_heur
  :: ‹twl_st_wl_heur ⇒ twl_st_wl_heur nres›
where
‹mark_to_delete_clauses_wl_D_heur  = (λS0. do {
    ASSERT(mark_to_delete_clauses_wl_D_heur_pre S0);
    S ← sort_vdom_heur S0;
    l ← number_clss_to_keep S;
    ASSERT(length (get_avdom S) ≤ length (get_clauses_wl_heur S0));
    (i, T) ← WHILETλ_. True
      (λ(i, S). i < length (get_avdom S))
      (λ(i, T). do {
        ASSERT(i < length (get_avdom T));
        ASSERT(access_vdom_at_pre T i);
        let C = get_avdom T ! i;
        ASSERT(clause_not_marked_to_delete_heur_pre (T, C));
        b ← mop_clause_not_marked_to_delete_heur T C;
        if ¬b then RETURN (i, delete_index_vdom_heur i T)
        else do {
          ASSERT(access_lit_in_clauses_heur_pre ((T, C), 0));
          ASSERT(length (get_clauses_wl_heur T) ≤ length (get_clauses_wl_heur S0));
          ASSERT(length (get_avdom T) ≤ length (get_clauses_wl_heur T));
          L ← mop_access_lit_in_clauses_heur T C 0;
          D ← get_the_propagation_reason_pol (get_trail_wl_heur T) L;
          lbd ← mop_arena_lbd (get_clauses_wl_heur T) C;
          length ← mop_arena_length (get_clauses_wl_heur T) C;
          status ← mop_arena_status (get_clauses_wl_heur T) C;
          used ← mop_marked_as_used (get_clauses_wl_heur T) C;
          let can_del = (D ≠ Some C) ∧
	     lbd > MINIMUM_DELETION_LBD ∧
             status = LEARNED ∧
             length ≠ 2 ∧
	     ¬used;
          if can_del
          then
            do {
              wasted ← mop_arena_length_st T C;
              T ← mop_mark_garbage_heur C i (incr_wasted_st (of_nat wasted) T);
              RETURN (i, T)
            }
          else do {
	    T ← mop_mark_unused_st_heur C T;
            RETURN (i+1, T)
	  }
        }
      })
      (l, S);
    ASSERT(length (get_avdom T) ≤ length (get_clauses_wl_heur S0));
    T ← mark_clauses_as_unused_wl_D_heur i T;
    incr_restart_stat T
  })›

lemma twl_st_heur_restart_same_annotD:
  ‹(S, T) ∈ twl_st_heur_restart ⟹ Propagated L C ∈ set (get_trail_wl T) ⟹
     Propagated L C' ∈ set (get_trail_wl T) ⟹ C = C'›
  ‹(S, T) ∈ twl_st_heur_restart ⟹ Propagated L C ∈ set (get_trail_wl T) ⟹
     Decided L ∈ set (get_trail_wl T) ⟹ False›
  by (auto simp: twl_st_heur_restart_def dest: no_dup_no_propa_and_dec
    no_dup_same_annotD)

lemma all_mono:
  ‹set_mset 𝒜 ⊆ set_mset ℬ ⟹ L  ∈# ℒall 𝒜 ⟹ L  ∈# ℒall ℬ›
  by (auto simp: all_def)

lemma all_lits_of_mm_mono2:
  ‹x ∈# (all_lits_of_mm A) ⟹ set_mset A ⊆ set_mset B ⟹ x ∈# (all_lits_of_mm B)›
  by (auto simp: all_lits_of_mm_def)


lemma all_init_all:
  ‹L ∈# ℒall (all_init_atms_st x1a) ⟹ L ∈# ℒall (all_atms_st x1a)›
  apply (rule all_mono)
  defer
  apply assumption
  by (cases x1a)
    (auto simp: all_init_atms_def all_lits_def all_init_lits_def
        all_atm_of_all_lits_of_mm all_atms_def intro: all_lits_of_mm_mono2 intro!: imageI
      simp del: all_init_atms_def[symmetric]
      simp flip: image_mset_union)

lemma get_vdom_mark_garbage[simp]:
  ‹get_vdom (mark_garbage_heur C i S) = get_vdom S›
  ‹get_avdom (mark_garbage_heur C i S) = delete_index_and_swap (get_avdom S) i›
  by (cases S; auto simp: mark_garbage_heur_def; fail)+

lemma mark_to_delete_clauses_wl_D_heur_alt_def:
    ‹mark_to_delete_clauses_wl_D_heur  = (λS0. do {
          ASSERT (mark_to_delete_clauses_wl_D_heur_pre S0);
          S ← sort_vdom_heur S0;
          _ ← RETURN (get_avdom S);
          l ← number_clss_to_keep S;
          ASSERT
               (length (get_avdom S) ≤ length (get_clauses_wl_heur S0));
          (i, T) ←
            WHILETλ_. True (λ(i, S). i < length (get_avdom S))
             (λ(i, T). do {
                   ASSERT (i < length (get_avdom T));
                   ASSERT (access_vdom_at_pre T i);
                   ASSERT
                        (clause_not_marked_to_delete_heur_pre
                          (T, get_avdom T ! i));
                   b ← mop_clause_not_marked_to_delete_heur T
                        (get_avdom T ! i);
                   if ¬b then RETURN (i, delete_index_vdom_heur i T)
                   else do {
                          ASSERT
                               (access_lit_in_clauses_heur_pre
                                 ((T, get_avdom T ! i), 0));
                          ASSERT
                               (length (get_clauses_wl_heur T)
                                ≤ length (get_clauses_wl_heur S0));
                          ASSERT
                               (length (get_avdom T)
                                ≤ length (get_clauses_wl_heur T));
                          L ← mop_access_lit_in_clauses_heur T
                               (get_avdom T ! i) 0;
                          D ← get_the_propagation_reason_pol
                               (get_trail_wl_heur T) L;
                          ASSERT
                               (get_clause_LBD_pre (get_clauses_wl_heur T)
                                 (get_avdom T ! i));
                          ASSERT
                               (arena_is_valid_clause_idx
                                 (get_clauses_wl_heur T) (get_avdom T ! i));
                          ASSERT
                               (arena_is_valid_clause_vdom
                                 (get_clauses_wl_heur T) (get_avdom T ! i));
                          ASSERT
                               (marked_as_used_pre
                                 (get_clauses_wl_heur T) (get_avdom T ! i));
                          let can_del = (D ≠ Some (get_avdom T ! i) ∧
                             MINIMUM_DELETION_LBD
                             < arena_lbd (get_clauses_wl_heur T)
                                (get_avdom T ! i) ∧
                             arena_status (get_clauses_wl_heur T)
                              (get_avdom T ! i) =
                             LEARNED ∧
                             arena_length (get_clauses_wl_heur T)
                              (get_avdom T ! i) ≠
                             2 ∧
                             ¬ marked_as_used (get_clauses_wl_heur T)
                                (get_avdom T ! i));
                          if can_del
                          then do {
                                wasted ← mop_arena_length_st T (get_avdom T ! i);
                                 ASSERT(mark_garbage_pre
                                   (get_clauses_wl_heur T, get_avdom T ! i) ∧
                                   1 ≤ get_learned_count T ∧ i < length (get_avdom T));
                                 RETURN
                                  (i, mark_garbage_heur (get_avdom T ! i) i (incr_wasted_st (of_nat wasted) T))
                               }
                          else do {
                                 ASSERT(arena_act_pre (get_clauses_wl_heur T) (get_avdom T ! i));
                                 RETURN
                                  (i + 1,
                                   mark_unused_st_heur (get_avdom T ! i) T)
                               }
                        }
                 })
             (l, S);
          ASSERT
               (length (get_avdom T) ≤ length (get_clauses_wl_heur S0));
          mark_clauses_as_unused_wl_D_heur i T ⤜ incr_restart_stat
        })›
    unfolding mark_to_delete_clauses_wl_D_heur_def
      mop_arena_lbd_def mop_arena_status_def mop_arena_length_def
      mop_marked_as_used_def bind_to_let_conv Let_def
      nres_monad3 mop_mark_garbage_heur_def mop_mark_unused_st_heur_def
      incr_wasted_st_twl_st
    by (auto intro!: ext simp: get_clauses_wl_heur.simps)

lemma mark_to_delete_clauses_wl_D_heur_mark_to_delete_clauses_wl_D:
  ‹(mark_to_delete_clauses_wl_D_heur, mark_to_delete_clauses_wl) ∈
     twl_st_heur_restart_ana r →f ⟨twl_st_heur_restart_ana r⟩nres_rel›
proof -
  have mark_to_delete_clauses_wl_D_alt_def:
    ‹mark_to_delete_clauses_wl  = (λS0. do {
      ASSERT(mark_to_delete_clauses_wl_pre S0);
      S ← reorder_vdom_wl S0;
      xs ← collect_valid_indices_wl S;
      l ← SPEC(λ_::nat. True);
      (_, S, _) ← WHILETmark_to_delete_clauses_wl_inv S xs
        (λ(i, T, xs). i < length xs)
        (λ(i, T, xs). do {
          b ← RETURN (xs!i ∈# dom_m (get_clauses_wl T));
          if ¬b then RETURN (i, T, delete_index_and_swap xs i)
          else do {
            ASSERT(0 < length (get_clauses_wl T∝(xs!i)));
	    ASSERT (get_clauses_wl T ∝ (xs ! i) ! 0 ∈# ℒall (all_init_atms_st T));
            K ← RETURN (get_clauses_wl T ∝ (xs ! i) ! 0);
            b ← RETURN (); ― ‹propagation reason›
            can_del ← SPEC(λb. b ⟶
              (Propagated (get_clauses_wl T∝(xs!i)!0) (xs!i) ∉ set (get_trail_wl T)) ∧
               ¬irred (get_clauses_wl T) (xs!i) ∧ length (get_clauses_wl T ∝ (xs!i)) ≠ 2);
            ASSERT(i < length xs);
            if can_del
            then
              RETURN (i, mark_garbage_wl (xs!i) T, delete_index_and_swap xs i)
            else
              RETURN (i+1, T, xs)
          }
        })
        (l, S, xs);
      remove_all_learned_subsumed_clauses_wl S
    })›
    unfolding mark_to_delete_clauses_wl_def reorder_vdom_wl_def bind_to_let_conv Let_def
    by (force intro!: ext)
  have mono: ‹g = g' ⟹ do {f; g} = do {f; g'}›
     ‹(⋀x. h x = h' x) ⟹ do {x ← f; h x} = do {x ← f; h' x}› for f f' :: ‹_ nres› and g g' and h h'
    by auto

  have [refine0]: ‹RETURN (get_avdom x) ≤ ⇓ {(xs, xs'). xs = xs' ∧ xs = get_avdom x} (collect_valid_indices_wl y)›
    if
      ‹(x, y) ∈ twl_st_heur_restart_ana r› and
      ‹mark_to_delete_clauses_wl_D_heur_pre x›
    for x y
  proof -
    show ?thesis by (auto simp: collect_valid_indices_wl_def simp: RETURN_RES_refine_iff)
  qed
  have init_rel[refine0]: ‹(x, y) ∈ twl_st_heur_restart_ana r ⟹
       (l, la) ∈ nat_rel ⟹
       ((l, x), la, y) ∈ nat_rel ×f {(S, T). (S, T) ∈ twl_st_heur_restart_ana r ∧ get_avdom S = get_avdom x}›
    for x y l la
    by auto

  define reason_rel where
    ‹reason_rel K x1a ≡ {(C, _ :: unit).
          (C ≠ None) = (Propagated K (the C) ∈ set (get_trail_wl x1a)) ∧
          (C = None) = (Decided K ∈ set (get_trail_wl x1a) ∨
             K ∉ lits_of_l (get_trail_wl x1a)) ∧
         (∀C1. (Propagated K C1 ∈ set (get_trail_wl x1a) ⟶ C1 = the C))}› for K :: ‹nat literal› and x1a
  have get_the_propagation_reason:
    ‹get_the_propagation_reason_pol (get_trail_wl_heur x2b) L
        ≤ SPEC (λD. (D, ()) ∈ reason_rel K x1a)›
  if
    ‹(x, y) ∈ twl_st_heur_restart_ana r› and
    ‹mark_to_delete_clauses_wl_pre y› and
    ‹mark_to_delete_clauses_wl_D_heur_pre x› and
    ‹(S, Sa)
     ∈ {(U, V).
        (U, V) ∈ twl_st_heur_restart_ana r ∧
        V = y ∧
        (mark_to_delete_clauses_wl_pre y ⟶
         mark_to_delete_clauses_wl_pre V) ∧
        (mark_to_delete_clauses_wl_D_heur_pre x ⟶
         mark_to_delete_clauses_wl_D_heur_pre U)}› and
    ‹(ys, xs) ∈ {(xs, xs'). xs = xs' ∧ xs = get_avdom S}› and
    ‹(l, la) ∈ nat_rel› and
    ‹la ∈ {_. True}› and
    xa_x': ‹(xa, x')
     ∈ nat_rel ×f {(Sa, T, xs). (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
    ‹case xa of (i, S) ⇒ i < length (get_avdom S)› and
    ‹case x' of (i, T, xs) ⇒ i < length xs› and
    ‹x1b < length (get_avdom x2b)› and
    ‹access_vdom_at_pre x2b x1b› and
    dom: ‹(b, ba)
       ∈ {(b, b').
          (b, b') ∈ bool_rel ∧
          b = (x2a ! x1 ∈# dom_m (get_clauses_wl x1a))}›
      ‹¬ ¬ b›
      ‹¬ ¬ ba› and
    ‹0 < length (get_clauses_wl x1a ∝ (x2a ! x1))› and
    ‹access_lit_in_clauses_heur_pre ((x2b, get_avdom x2b ! x1b), 0)› and
    st:
      ‹x2 = (x1a, x2a)›
      ‹x' = (x1, x2)›
      ‹xa = (x1b, x2b)› and
    L: ‹get_clauses_wl x1a ∝ (x2a ! x1) ! 0 ∈# ℒall (all_init_atms_st x1a)› and
    L': ‹(L, K)
    ∈ {(L, L').
       (L, L') ∈ nat_lit_lit_rel ∧
       L' = get_clauses_wl x1a ∝ (x2a ! x1) ! 0}›
    for x y S Sa xs' xs l la xa x' x1 x2 x1a x2a x1' x2' x3 x1b ys x2b L K b ba
  proof -
    have L: ‹arena_lit (get_clauses_wl_heur x2b) (x2a ! x1) ∈# ℒall (all_init_atms_st x1a)›
      using L that by (auto simp: twl_st_heur_restart st arena_lifting dest: all_init_all twl_st_heur_restart_anaD)

    show ?thesis
      apply (rule order.trans)
      apply (rule get_the_propagation_reason_pol[THEN fref_to_Down_curry,
        of ‹all_init_atms_st x1a› ‹get_trail_wl x1a›
	  ‹arena_lit (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b + 0)›])
      subgoal
        using xa_x' L L' by (auto simp add: twl_st_heur_restart_def st)
      subgoal
        using xa_x' L' dom by (auto simp add: twl_st_heur_restart_ana_def twl_st_heur_restart_def st arena_lifting)
      using that unfolding get_the_propagation_reason_def reason_rel_def apply -
      by (auto simp: twl_st_heur_restart lits_of_def get_the_propagation_reason_def
          conc_fun_RES
        dest!: twl_st_heur_restart_anaD dest: twl_st_heur_restart_same_annotD imageI[of _ _ lit_of])
  qed
  have ‹((M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur, vdom, avdom, lcount),
           S')
          ∈ twl_st_heur_restart ⟹
    ((M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur, vdom, avdom', lcount),
           S')
          ∈ twl_st_heur_restart›
    if ‹mset avdom' ⊆# mset avdom›
    for M' N' D' j W' vm clvls cach lbd outl stats fast_ema slow_ema
      ccount vdom lcount S' avdom' avdom heur
    using that unfolding twl_st_heur_restart_def
    by auto
  then have mark_to_delete_clauses_wl_D_heur_pre_vdom':
    ‹mark_to_delete_clauses_wl_D_heur_pre (M', N', D', j, W', vm, clvls, cach, lbd, outl, stats,
       heur, vdom, avdom', lcount) ⟹
      mark_to_delete_clauses_wl_D_heur_pre (M', N', D', j, W', vm, clvls, cach, lbd, outl, stats,
        heur, vdom, avdom, lcount)›
    if ‹mset avdom ⊆# mset avdom'›
    for M' N' D' j W' vm clvls cach lbd outl stats fast_ema slow_ema avdom avdom'
      ccount vdom lcount heur
    using that
    unfolding mark_to_delete_clauses_wl_D_heur_pre_def
    by metis
  have [refine0]:
    ‹sort_vdom_heur S ≤ ⇓ {(U, V). (U, V) ∈ twl_st_heur_restart_ana r ∧ V = T ∧
         (mark_to_delete_clauses_wl_pre T ⟶ mark_to_delete_clauses_wl_pre V) ∧
         (mark_to_delete_clauses_wl_D_heur_pre S ⟶ mark_to_delete_clauses_wl_D_heur_pre U)}
         (reorder_vdom_wl T)›
    if ‹(S, T) ∈ twl_st_heur_restart_ana r› for S T
    using that unfolding reorder_vdom_wl_def sort_vdom_heur_def
    apply (refine_rcg ASSERT_leI)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
    apply (rule specify_left)
    apply (rule_tac N1 = ‹get_clauses_wl T› and vdom1 = ‹(get_vdom S)› in
     order_trans[OF isa_remove_deleted_clauses_from_avdom_remove_deleted_clauses_from_avdom,
      unfolded Down_id_eq, OF _ _ _ remove_deleted_clauses_from_avdom])
    subgoal for x y x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h
       x1i x2i x1j x2j x1k x2k x1l x2l
      by (case_tac T; auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def mem_Collect_eq prod.case)
    subgoal for x y x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h
       x1i x2i x1j x2j x1k x2k x1l x2l
      by (case_tac T; auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def mem_Collect_eq prod.case)
    subgoal for x y x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h
       x1i x2i x1j x2j x1k x2k x1l x2l
      by (case_tac T; auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def mem_Collect_eq prod.case)
    apply (subst assert_bind_spec_conv, intro conjI)
    subgoal for x y
      unfolding valid_sort_clause_score_pre_def arena_is_valid_clause_vdom_def
        get_clause_LBD_pre_def arena_is_valid_clause_idx_def arena_act_pre_def
      by (force simp: valid_sort_clause_score_pre_def twl_st_heur_restart_ana_def arena_dom_status_iff
        arena_act_pre_def get_clause_LBD_pre_def arena_is_valid_clause_idx_def twl_st_heur_restart_def
         intro!: exI[of _ ‹get_clauses_wl T›] dest!: set_mset_mono mset_subset_eqD)
    apply (subst assert_bind_spec_conv, intro conjI)
    subgoal
     by (auto simp: twl_st_heur_restart_ana_def valid_arena_vdom_subset twl_st_heur_restart_def
        dest!: size_mset_mono valid_arena_vdom_subset)
    subgoal
      apply (rewrite at ‹_ ≤ ⌑› Down_id_eq[symmetric])
      apply (rule bind_refine_spec)
      prefer 2
      apply (rule sort_clauses_by_score_reorder[of _ ‹get_clauses_wl T›  ‹get_vdom S›])
      by (auto 5 3 simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest: mset_eq_setD
        simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def
         intro: mark_to_delete_clauses_wl_D_heur_pre_vdom'
         dest: mset_eq_setD)
    done
  have already_deleted:
    ‹((x1b, delete_index_vdom_heur x1b x2b), x1, x1a,
       delete_index_and_swap x2a x1)
      ∈ nat_rel ×f {(Sa, T, xs). (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}›
    if
      ‹(x, y) ∈ twl_st_heur_restart_ana r› and
      ‹mark_to_delete_clauses_wl_D_heur_pre x› and
      ‹(S, Sa)
     ∈ {(U, V).
        (U, V) ∈ twl_st_heur_restart_ana r ∧
        V = y ∧
        (mark_to_delete_clauses_wl_pre y ⟶
         mark_to_delete_clauses_wl_pre V) ∧
        (mark_to_delete_clauses_wl_D_heur_pre x ⟶
         mark_to_delete_clauses_wl_D_heur_pre U)}› and
      ‹(l, la) ∈ nat_rel› and
      ‹la ∈ {_. True}› and
      xx: ‹(xa, x')
     ∈ nat_rel ×f {(Sa, T, xs). (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
      ‹case xa of (i, S) ⇒ i < length (get_avdom S)› and
      ‹case x' of (i, T, xs) ⇒ i < length xs› and
      st:
      ‹x2 = (x1a, x2a)›
      ‹x' = (x1, x2)›
      ‹xa = (x1b, x2b)› and
      le: ‹x1b < length (get_avdom x2b)› and
      ‹access_vdom_at_pre x2b x1b› and
      ‹(b, ba) ∈ {(b, b'). (b, b') ∈ bool_rel ∧ b = (x2a ! x1 ∈# dom_m (get_clauses_wl x1a))}› and
      ‹¬ba›
    for x y S xs l la xa x' xz x1 x2 x1a x2a x2b x2c x2d ys x1b Sa ba b
  proof -
    show ?thesis
      using xx le unfolding st
      by (auto simp: twl_st_heur_restart_ana_def delete_index_vdom_heur_def
          twl_st_heur_restart_def mark_garbage_heur_def mark_garbage_wl_def
          learned_clss_l_l_fmdrop size_remove1_mset_If
          intro: valid_arena_extra_information_mark_to_delete'
          dest!: in_set_butlastD in_vdom_m_fmdropD
          elim!: in_set_upd_cases)
  qed
  have get_learned_count_ge: ‹Suc 0 ≤ get_learned_count x2b›
    if
      xy: ‹(x, y) ∈ twl_st_heur_restart_ana r› and
      ‹(xa, x')
       ∈ nat_rel ×f
         {(Sa, T, xs).
          (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
      ‹x2 = (x1a, x2a)› and
      ‹x' = (x1, x2)› and
      ‹xa = (x1b, x2b)› and
      dom: ‹(b, ba)
         ∈ {(b, b').
            (b, b') ∈ bool_rel ∧
            b = (x2a ! x1 ∈# dom_m (get_clauses_wl x1a))}›
        ‹¬ ¬ b›
        ‹¬ ¬ ba› and
      ‹MINIMUM_DELETION_LBD
    < arena_lbd (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b) ∧
    arena_status (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b) = LEARNED ∧
    arena_length (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b) ≠ 2 ∧
    ¬ marked_as_used (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b)› and
      ‹can_del› for x y S Sa uu xs l la xa x' x1 x2 x1a x2a x1b x2b D can_del b ba
  proof -
    have ‹¬irred (get_clauses_wl x1a) (x2a ! x1)› and ‹(x2b, x1a) ∈ twl_st_heur_restart_ana r›
      using that by (auto simp: twl_st_heur_restart arena_lifting
        dest!: twl_st_heur_restart_anaD twl_st_heur_restart_valid_arena)
    then show ?thesis
     using dom by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def ran_m_def
       dest!: multi_member_split)
  qed
  have mop_clause_not_marked_to_delete_heur:
    ‹mop_clause_not_marked_to_delete_heur x2b (get_avdom x2b ! x1b)
        ≤ SPEC
           (λc. (c, x2a ! x1 ∈# dom_m (get_clauses_wl x1a))
                ∈ {(b, b'). (b,b') ∈ bool_rel ∧ (b ⟷ x2a ! x1 ∈# dom_m (get_clauses_wl x1a))})›
    if
      ‹(xa, x')
       ∈ nat_rel ×f
         {(Sa, T, xs).
          (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
      ‹case xa of (i, S) ⇒ i < length (get_avdom S)› and
      ‹case x' of (i, T, xs) ⇒ i < length xs› and
      ‹mark_to_delete_clauses_wl_inv Sa xs x'› and
      ‹x2 = (x1a, x2a)› and
      ‹x' = (x1, x2)› and
      ‹xa = (x1b, x2b)› and
      ‹clause_not_marked_to_delete_heur_pre (x2b, get_avdom x2b ! x1b)›
    for x y S Sa uu xs l la xa x' x1 x2 x1a x2a x1b x2b
    unfolding mop_clause_not_marked_to_delete_heur_def
    apply refine_vcg
    subgoal
      using that by blast
    subgoal
      using that by (auto simp: twl_st_heur_restart arena_lifting dest!: twl_st_heur_restart_anaD)
    done


  have init:
    ‹(u, xs) ∈ {(xs, xs'). xs = xs' ∧ xs = get_avdom S} ⟹
    (l, la) ∈ nat_rel ⟹
    (S, Sa) ∈ twl_st_heur_restart_ana r ⟹
    ((l, S), la, Sa, xs) ∈  nat_rel ×f
       {(Sa, (T, xs)). (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}›
       for x y S Sa xs l la u
    by auto
  have mop_access_lit_in_clauses_heur:
    ‹mop_access_lit_in_clauses_heur x2b (get_avdom x2b ! x1b) 0
        ≤ SPEC
           (λc. (c, get_clauses_wl x1a ∝ (x2a ! x1) ! 0)
                ∈ {(L, L'). (L, L') ∈ nat_lit_lit_rel ∧ L' = get_clauses_wl x1a ∝ (x2a ! x1) ! 0})›
    if
      ‹(x, y) ∈ twl_st_heur_restart_ana r› and
      ‹mark_to_delete_clauses_wl_pre y› and
      ‹mark_to_delete_clauses_wl_D_heur_pre x› and
      ‹(S, Sa)
       ∈ {(U, V).
          (U, V) ∈ twl_st_heur_restart_ana r ∧
          V = y ∧
          (mark_to_delete_clauses_wl_pre y ⟶
           mark_to_delete_clauses_wl_pre V) ∧
          (mark_to_delete_clauses_wl_D_heur_pre x ⟶
           mark_to_delete_clauses_wl_D_heur_pre U)}› and
      ‹(uu, xs) ∈ {(xs, xs'). xs = xs' ∧ xs = get_avdom S}› and
      ‹(l, la) ∈ nat_rel› and
      ‹la ∈ {uu. True}› and
      ‹length (get_avdom S) ≤ length (get_clauses_wl_heur x)› and
      ‹(xa, x')
       ∈ nat_rel ×f
         {(Sa, T, xs).
          (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
      ‹case xa of (i, S) ⇒ i < length (get_avdom S)› and
      ‹case x' of (i, T, xs) ⇒ i < length xs› and
      ‹mark_to_delete_clauses_wl_inv Sa xs x'› and
      ‹x2 = (x1a, x2a)› and
      ‹x' = (x1, x2)› and
      ‹xa = (x1b, x2b)› and
      ‹x1b < length (get_avdom x2b)› and
      ‹access_vdom_at_pre x2b x1b› and
      ‹clause_not_marked_to_delete_heur_pre (x2b, get_avdom x2b ! x1b)› and
      ‹(b, ba)
       ∈ {(b, b').
          (b, b') ∈ bool_rel ∧
          b = (x2a ! x1 ∈# dom_m (get_clauses_wl x1a))}› and
      ‹¬ ¬ b› and
      ‹¬ ¬ ba› and
      ‹0 < length (get_clauses_wl x1a ∝ (x2a ! x1))› and
      ‹get_clauses_wl x1a ∝ (x2a ! x1) ! 0
       ∈# ℒall (all_init_atms_st x1a)› and
      pre: ‹access_lit_in_clauses_heur_pre ((x2b, get_avdom x2b ! x1b), 0)›
     for x y S Sa uu xs l la xa x' x1 x2 x1a x2a x1b x2b b ba
  unfolding mop_access_lit_in_clauses_heur_def mop_arena_lit2_def
  apply refine_vcg
  subgoal using pre unfolding access_lit_in_clauses_heur_pre_def by simp
  subgoal using that by (auto dest!: twl_st_heur_restart_anaD twl_st_heur_restart_valid_arena simp: arena_lifting)
  done

  have incr_restart_stat: ‹incr_restart_stat T
    ≤ ⇓ (twl_st_heur_restart_ana r) (remove_all_learned_subsumed_clauses_wl S)›
    if ‹(T, S) ∈ twl_st_heur_restart_ana r› for S T i
    using that
    by (cases S; cases T)
      (auto simp: conc_fun_RES incr_restart_stat_def
        twl_st_heur_restart_ana_def twl_st_heur_restart_def
        remove_all_learned_subsumed_clauses_wl_def
        RES_RETURN_RES)

  have [refine0]: ‹mark_clauses_as_unused_wl_D_heur i T⤜ incr_restart_stat
    ≤ ⇓ (twl_st_heur_restart_ana r)
       (remove_all_learned_subsumed_clauses_wl S)›
    if ‹(T, S) ∈ twl_st_heur_restart_ana r› for S T i
    apply (cases S)
    apply (rule bind_refine_res[where R = Id, simplified])
    defer
    apply (rule mark_clauses_as_unused_wl_D_heur[unfolded conc_fun_RES, OF that, of i])
    apply (rule incr_restart_stat[THEN order_trans, of _ S])
    by auto

  show ?thesis
    supply sort_vdom_heur_def[simp] twl_st_heur_restart_anaD[dest] [[goals_limit=1]]
    unfolding mark_to_delete_clauses_wl_D_heur_alt_def mark_to_delete_clauses_wl_D_alt_def
      access_lit_in_clauses_heur_def
    apply (intro frefI nres_relI)
    apply (refine_vcg sort_vdom_heur_reorder_vdom_wl[THEN fref_to_Down])
    subgoal
      unfolding mark_to_delete_clauses_wl_D_heur_pre_def by fast
    subgoal by auto
    subgoal by auto
    subgoal for x y S T unfolding number_clss_to_keep_def by (cases S) (auto)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def
       dest!: valid_arena_vdom_subset size_mset_mono)
    apply (rule init; solves auto)
    subgoal by auto
    subgoal by auto
    subgoal by (auto simp: access_vdom_at_pre_def)
    subgoal for x y S xs l la xa x' xz x1 x2 x1a x2a x2b x2c x2d
      unfolding clause_not_marked_to_delete_heur_pre_def arena_is_valid_clause_vdom_def
        prod.simps
      by (rule exI[of _ ‹get_clauses_wl x2a›], rule exI[of _ ‹set (get_vdom x2d)›])
         (auto simp: twl_st_heur_restart dest: twl_st_heur_restart_get_avdom_nth_get_vdom)
    apply (rule mop_clause_not_marked_to_delete_heur; assumption)
    subgoal for x y S Sa uu xs l la xa x' x1 x2 x1a x2a x1b x2b
      by (auto simp: twl_st_heur_restart)
    subgoal
      by (rule already_deleted)
    subgoal for x y _ _ _ _ _ xs l la xa x' x1 x2 x1a x2a
      unfolding access_lit_in_clauses_heur_pre_def prod.simps arena_lit_pre_def
        arena_is_valid_clause_idx_and_access_def
      by (rule bex_leI[of _ ‹get_avdom x2a ! x1a›], simp, rule exI[of _ ‹get_clauses_wl x1›])
        (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
    subgoal premises p using p(7-) by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
     apply (rule mop_access_lit_in_clauses_heur; assumption)
    apply (rule get_the_propagation_reason; assumption)
    subgoal for x y S Sa _ xs l la xa x' x1 x2 x1a x2a x1b x2b
      unfolding prod.simps
        get_clause_LBD_pre_def arena_is_valid_clause_idx_def
      by (rule exI[of _ ‹get_clauses_wl x1a›], rule exI[of _ ‹set (get_vdom x2b)›])
        (auto simp: twl_st_heur_restart dest: twl_st_heur_restart_valid_arena)
    subgoal for x y S Sa _ xs l la xa x' x1 x2 x1a x2a x1b x2b
      unfolding prod.simps
        arena_is_valid_clause_vdom_def arena_is_valid_clause_idx_def
      by (rule exI[of _ ‹get_clauses_wl x1a›], rule exI[of _ ‹set (get_vdom x2b)›])
        (auto simp: twl_st_heur_restart dest: twl_st_heur_restart_valid_arena
	  twl_st_heur_restart_get_avdom_nth_get_vdom)
    subgoal for x y S Sa _ xs l la xa x' x1 x2 x1a x2a x1b x2b
      unfolding prod.simps
        arena_is_valid_clause_vdom_def arena_is_valid_clause_idx_def
      by (rule exI[of _ ‹get_clauses_wl x1a›], rule exI[of _ ‹set (get_vdom x2b)›])
        (auto simp: twl_st_heur_restart arena_dom_status_iff
          dest: twl_st_heur_restart_valid_arena twl_st_heur_restart_get_avdom_nth_get_vdom)
    subgoal
      unfolding marked_as_used_pre_def
      by (auto simp: twl_st_heur_restart reason_rel_def)
    subgoal
      unfolding marked_as_used_pre_def
      by (auto simp: twl_st_heur_restart reason_rel_def)
    subgoal
      by (auto simp: twl_st_heur_restart)
    subgoal
      by (auto dest!: twl_st_heur_restart_anaD twl_st_heur_restart_valid_arena simp: arena_lifting)
    subgoal by fast
    subgoal for x y S Sa _ xs l la xa x' x1 x2 x1a x2a x1b x2b
      unfolding mop_arena_length_st_def
      apply (rule mop_arena_length[THEN fref_to_Down_curry, THEN order_trans,
        of ‹get_clauses_wl x1a› ‹get_avdom x2b ! x1b› _ _ ‹set (get_vdom x2b)›])
      subgoal
        by auto
      subgoal
        by (auto simp: twl_st_heur_restart_valid_arena)
      subgoal
        apply (auto intro!: incr_wasted_st_twl_st ASSERT_leI)
        subgoal
          unfolding prod.simps mark_garbage_pre_def
            arena_is_valid_clause_vdom_def arena_is_valid_clause_idx_def
          by (rule exI[of _ ‹get_clauses_wl x1a›], rule exI[of _ ‹set (get_vdom x2b)›])
            (auto simp: twl_st_heur_restart dest: twl_st_heur_restart_valid_arena)
        subgoal
           apply (rule get_learned_count_ge; assumption?; fast?)
           apply auto
           done
        subgoal
          by (use arena_lifting(24)[of ‹get_clauses_wl_heur x2b› _ _  ‹get_avdom x2b ! x1›] in
            ‹auto intro!: incr_wasted_st mark_garbage_heur_wl_ana
            dest: twl_st_heur_restart_valid_arena twl_st_heur_restart_anaD›)
        done
     done
   subgoal for x y
      unfolding valid_sort_clause_score_pre_def arena_is_valid_clause_vdom_def
        get_clause_LBD_pre_def arena_is_valid_clause_idx_def arena_act_pre_def
      by (force simp: valid_sort_clause_score_pre_def twl_st_heur_restart_ana_def arena_dom_status_iff
        arena_act_pre_def get_clause_LBD_pre_def arena_is_valid_clause_idx_def twl_st_heur_restart_def
         intro!: exI[of _ ‹get_clauses_wl T›] dest!: set_mset_mono mset_subset_eqD)
    subgoal
      by (auto intro!: mark_unused_st_heur_ana)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
    subgoal
      by auto
    done
qed

definition cdcl_twl_full_restart_wl_prog_heur where
‹cdcl_twl_full_restart_wl_prog_heur S = do {
  _ ← ASSERT (mark_to_delete_clauses_wl_D_heur_pre S);
  T ← mark_to_delete_clauses_wl_D_heur S;
  RETURN T
}›

lemma cdcl_twl_full_restart_wl_prog_heur_cdcl_twl_full_restart_wl_prog_D:
  ‹(cdcl_twl_full_restart_wl_prog_heur, cdcl_twl_full_restart_wl_prog) ∈
     twl_st_heur''' r →f ⟨twl_st_heur''' r⟩nres_rel›
  unfolding cdcl_twl_full_restart_wl_prog_heur_def cdcl_twl_full_restart_wl_prog_def
  apply (intro frefI nres_relI)
  apply (refine_vcg
    mark_to_delete_clauses_wl_D_heur_mark_to_delete_clauses_wl_D[THEN fref_to_Down])
  subgoal
    unfolding mark_to_delete_clauses_wl_D_heur_pre_alt_def
    by fast
  apply (rule twl_st_heur_restartD)
  subgoal
    by (subst mark_to_delete_clauses_wl_D_heur_pre_twl_st_heur[symmetric]) auto
  subgoal
    by (auto simp: mark_to_delete_clauses_wl_post_twl_st_heur twl_st_heur_restart_anaD)
     (auto simp: twl_st_heur_restart_ana_def)
  done

definition cdcl_twl_restart_wl_heur where
‹cdcl_twl_restart_wl_heur S = do {
    let b = lower_restart_bound_not_reached S;
    if b then cdcl_twl_local_restart_wl_D_heur S
    else cdcl_twl_full_restart_wl_prog_heur S
  }›


lemma cdcl_twl_restart_wl_heur_cdcl_twl_restart_wl_D_prog:
  ‹(cdcl_twl_restart_wl_heur, cdcl_twl_restart_wl_prog) ∈
    twl_st_heur''' r →f ⟨twl_st_heur''' r⟩nres_rel›
  unfolding cdcl_twl_restart_wl_prog_def cdcl_twl_restart_wl_heur_def
  apply (intro frefI nres_relI)
  apply (refine_rcg
    cdcl_twl_local_restart_wl_D_heur_cdcl_twl_local_restart_wl_D_spec[THEN fref_to_Down]
    cdcl_twl_full_restart_wl_prog_heur_cdcl_twl_full_restart_wl_prog_D[THEN fref_to_Down])
  subgoal by auto
  subgoal by auto
  done


definition isasat_replace_annot_in_trail
  :: ‹nat literal ⇒ nat ⇒ twl_st_wl_heur ⇒ twl_st_wl_heur nres›
where
  ‹isasat_replace_annot_in_trail L C = (λ((M, val, lvls, reason, k), oth). do {
      ASSERT(atm_of L < length reason);
      RETURN ((M, val, lvls, reason[atm_of L := 0], k), oth)
    })›

lemma all_atm_of_all_init_lits_of_mm:
  ‹set_mset (ℒall (atm_of `# all_init_lits N NUE)) = set_mset (all_init_lits N NUE)›
  by (auto simp: all_init_lits_def all_atm_of_all_lits_of_mm)

lemma trail_pol_replace_annot_in_trail_spec:
  assumes
    ‹atm_of x2 < length x1e› and
    x2: ‹atm_of x2 ∈# all_init_atms_st (ys @ Propagated x2 C # zs, x2n')› and
    ‹(((x1b, x1c, x1d, x1e, x2d), x2n),
        (ys @ Propagated x2 C # zs, x2n'))
       ∈ twl_st_heur_restart_ana r›
  shows
    ‹(((x1b, x1c, x1d, x1e[atm_of x2 := 0], x2d), x2n),
        (ys @ Propagated x2 0 # zs, x2n'))
       ∈ twl_st_heur_restart_ana r›
proof -
  let ?S = ‹(ys @ Propagated x2 C # zs, x2n')›
  let ?𝒜 = ‹all_init_atms_st ?S›
  have pol: ‹((x1b, x1c, x1d, x1e, x2d), ys @ Propagated x2 C # zs)
         ∈ trail_pol (all_init_atms_st ?S)›
    using assms(3) unfolding twl_st_heur_restart_ana_def twl_st_heur_restart_def
    by auto
  obtain x y where
    x2d: ‹x2d = (count_decided (ys @ Propagated x2 C # zs), y)› and
    reasons: ‹((map lit_of (rev (ys @ Propagated x2 C # zs)), x1e),
      ys @ Propagated x2 C # zs)
     ∈ ann_lits_split_reasons ?𝒜› and
    pol: ‹∀L∈#ℒall ?𝒜.        nat_of_lit L < length x1c ∧
        x1c ! nat_of_lit L = polarity (ys @ Propagated x2 C # zs) L› and
    n_d: ‹no_dup (ys @ Propagated x2 C # zs)› and
    lvls: ‹∀L∈#ℒall ?𝒜. atm_of L < length x1d ∧
        x1d ! atm_of L = get_level (ys @ Propagated x2 C # zs) L› and
    ‹undefined_lit ys (lit_of (Propagated x2 C))› and
    ‹undefined_lit zs (lit_of (Propagated x2 C))› and
    inA:‹∀L∈set (ys @ Propagated x2 C # zs). lit_of L ∈# ℒall ?𝒜› and
    cs: ‹control_stack y (ys @ Propagated x2 C # zs)› and
    ‹literals_are_in_ℒin_trail ?𝒜 (ys @ Propagated x2 C # zs)› and
    ‹length (ys @ Propagated x2 C # zs) < uint32_max› and
    ‹length (ys @ Propagated x2 C # zs) ≤ uint32_max div 2 + 1› and
    ‹count_decided (ys @ Propagated x2 C # zs) < uint32_max› and
    ‹length (map lit_of (rev (ys @ Propagated x2 C # zs))) =
     length (ys @ Propagated x2 C # zs)› and
    bounded: ‹isasat_input_bounded ?𝒜› and
    x1b: ‹x1b = map lit_of (rev (ys @ Propagated x2 C # zs))›
    using pol unfolding trail_pol_alt_def
    by blast
  have lev_eq: ‹get_level (ys @ Propagated x2 C # zs) =
    get_level (ys @ Propagated x2 0 # zs)›
    ‹count_decided (ys @ Propagated x2 C # zs) =
      count_decided (ys @ Propagated x2 0 # zs)›
    by (auto simp: get_level_cons_if get_level_append_if)
  have [simp]: ‹atm_of x2 < length x1e›
    using reasons x2 in_ℒall_atm_of_𝒜in
    by (auto simp: ann_lits_split_reasons_def all_all_init_atms all_init_atms_def
        all_atm_of_all_init_lits_of_mm
      simp del: all_init_atms_def[symmetric]
      dest: multi_member_split)

  have ‹((x1b, x1e[atm_of x2 := 0]), ys @ Propagated x2 0 # zs)
       ∈ ann_lits_split_reasons ?𝒜›
    using reasons n_d undefined_notin
    by (auto simp: ann_lits_split_reasons_def x1b
      DECISION_REASON_def atm_of_eq_atm_of)
  moreover have n_d': ‹no_dup (ys @ Propagated x2 0 # zs)›
    using n_d by auto
  moreover have ‹∀L∈#ℒall ?𝒜.
     nat_of_lit L < length x1c ∧
        x1c ! nat_of_lit L = polarity (ys @ Propagated x2 0 # zs) L›
    using pol by (auto simp: polarity_def)
  moreover have ‹∀L∈#ℒall ?𝒜.
    atm_of L < length x1d ∧
           x1d ! atm_of L = get_level (ys @ Propagated x2 0 # zs) L›
    using lvls by (auto simp: get_level_append_if get_level_cons_if)
  moreover have ‹control_stack y (ys @ Propagated x2 0 # zs)›
    using cs apply -
    apply (subst control_stack_alt_def[symmetric, OF n_d'])
    apply (subst (asm) control_stack_alt_def[symmetric, OF n_d])
    unfolding control_stack'_def lev_eq
    apply normalize_goal
    apply (intro ballI conjI)
    apply (solves auto)
    unfolding set_append list.set(2) Un_iff insert_iff
    apply (rule disjE, assumption)
    subgoal for L
      by (drule_tac x = L in bspec)
        (auto simp: nth_append nth_Cons split: nat.splits)
    apply (rule disjE[of ‹_ = _›], assumption)
    subgoal for L
      by (auto simp: nth_append nth_Cons split: nat.splits)
    subgoal for L
      by (drule_tac x = L in bspec)
        (auto simp: nth_append nth_Cons split: nat.splits)

    done
  ultimately have
    ‹((x1b, x1c, x1d, x1e[atm_of x2 := 0], x2d), ys @ Propagated x2 0 # zs)
         ∈ trail_pol ?𝒜›
    using n_d x2 inA bounded
    unfolding trail_pol_def x2d
    by simp
  moreover { fix aaa ca
    have ‹vmtf_ℒall (all_init_atms aaa ca) (ys @ Propagated x2 C # zs) =
       vmtf_ℒall (all_init_atms aaa ca) (ys @ Propagated x2 0 # zs)›
       by (auto simp: vmtf_ℒall_def)
    then have ‹isa_vmtf (all_init_atms aaa ca) (ys @ Propagated x2 C # zs) =
      isa_vmtf (all_init_atms aaa ca) (ys @ Propagated x2 0 # zs)›
      by (auto simp: isa_vmtf_def vmtf_def
	image_iff)
  }
  moreover { fix D
    have ‹get_level (ys @ Propagated x2 C # zs) = get_level (ys @ Propagated x2 0 # zs)›
      by (auto simp: get_level_append_if get_level_cons_if)
    then have ‹counts_maximum_level (ys @ Propagated x2 C # zs) D =
      counts_maximum_level (ys @ Propagated x2 0 # zs) D› and
      ‹out_learned (ys @ Propagated x2 C # zs) = out_learned (ys @ Propagated x2 0 # zs)›
      by (auto simp: counts_maximum_level_def card_max_lvl_def
        out_learned_def intro!: ext)
  }
  ultimately show ?thesis
    using assms(3) unfolding twl_st_heur_restart_ana_def
    by (cases x2n; cases x2n')
      (auto simp: twl_st_heur_restart_def
        simp flip: mset_map drop_map)
qed

lemmas trail_pol_replace_annot_in_trail_spec2 =
  trail_pol_replace_annot_in_trail_spec[of ‹- _›, simplified]

lemma all_ball_all:
  ‹(∀L ∈# ℒall (all_atms N NUE). P L) = (∀L ∈# all_lits N NUE. P L)›
  ‹(∀L ∈# ℒall (all_init_atms N NUE). P L) = (∀L ∈# all_init_lits N NUE. P L)›
  by (simp_all add: all_all_atms_all_lits  all_all_init_atms)

lemma twl_st_heur_restart_ana_US_empty:
  ‹NO_MATCH {#} US ⟹ (S, M, N, D, NE, UE, NS, US, W, Q) ∈ twl_st_heur_restart_ana r ⟷
   (S, M, N, D, NE, UE, NS, {#}, W, Q)
       ∈ twl_st_heur_restart_ana r›
   by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def)

fun equality_except_trail_empty_US_wl :: ‹'v twl_st_wl ⇒ 'v twl_st_wl ⇒ bool› where
‹equality_except_trail_empty_US_wl (M, N, D, NE, UE, NS, US, WS, Q)
     (M', N', D', NE', UE', NS', US', WS', Q') ⟷
    N = N' ∧ D = D' ∧ NE = NE' ∧ NS = NS' ∧ US = {#} ∧ UE = UE' ∧ 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_empty_US_wl_get_conflict_wl:
    ‹equality_except_trail_empty_US_wl S Y ⟹ get_conflict_wl S = get_conflict_wl Y› and
  equality_except_trail_empty_US_wl_get_clauses_wl:
    ‹equality_except_trail_empty_US_wl S Y⟹ get_clauses_wl S = get_clauses_wl Y›
 by (cases S; cases Y; solves auto)+

lemma isasat_replace_annot_in_trail_replace_annot_in_trail_spec:
  ‹(((L, C), S), ((L', C'), S')) ∈ Id ×f Id ×f twl_st_heur_restart_ana r ⟹
  isasat_replace_annot_in_trail L C S ≤
    ⇓{(U, U'). (U, U') ∈ twl_st_heur_restart_ana r ∧
       get_clauses_wl_heur U = get_clauses_wl_heur S ∧
       get_vdom U = get_vdom S ∧
       equality_except_trail_empty_US_wl U' S'}
    (replace_annot_wl L' C' S')›
  unfolding isasat_replace_annot_in_trail_def replace_annot_wl_def
    uncurry_def
  apply refine_rcg
  subgoal
    by (auto simp: trail_pol_alt_def ann_lits_split_reasons_def all_ball_all
      twl_st_heur_restart_def twl_st_heur_restart_ana_def replace_annot_wl_pre_def)
  subgoal for x y x1 x1a x2 x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f
      x2f x1g x2g x1h x1i
      x2h x1j x2i x1k x2j x1l
    unfolding replace_annot_wl_pre_def replace_annot_l_pre_def
    apply (clarify dest!: split_list[of ‹Propagated _ _›])
    apply (rule RETURN_SPEC_refine)
    apply (rule_tac x = ‹(ys @ Propagated L 0 # zs, x1, x2, x1b,
        x1c, x1d, {#}, x1f, x2f)› in exI)
    apply (intro conjI)
    prefer 2
    apply (rule_tac x = ‹ys @ Propagated L 0 # zs› in exI)
    apply (intro conjI)
    apply blast
    by (cases x1l; auto intro!: trail_pol_replace_annot_in_trail_spec
        trail_pol_replace_annot_in_trail_spec2
      simp: atm_of_eq_atm_of all_init_atms_def replace_annot_wl_pre_def
        all_ball_all replace_annot_l_pre_def state_wl_l_def
        twl_st_heur_restart_ana_US_empty
      simp del: all_init_atms_def[symmetric])+
  done

definition remove_one_annot_true_clause_one_imp_wl_D_heur
  :: ‹nat ⇒ twl_st_wl_heur ⇒ (nat × twl_st_wl_heur) nres›
where
‹remove_one_annot_true_clause_one_imp_wl_D_heur = (λi S. do {
      (L, C) ← do {
        L ← isa_trail_nth (get_trail_wl_heur S) i;
	C ← get_the_propagation_reason_pol (get_trail_wl_heur S) L;
	RETURN (L, C)};
      ASSERT(C ≠ None ∧ i + 1 ≤ Suc (uint32_max div 2));
      if the C = 0 then RETURN (i+1, S)
      else do {
        ASSERT(C ≠ None);
        S ← isasat_replace_annot_in_trail L (the C) S;
	ASSERT(mark_garbage_pre (get_clauses_wl_heur S, the C) ∧ arena_is_valid_clause_vdom (get_clauses_wl_heur S) (the C));
        S ← mark_garbage_heur2 (the C) S;
        ― ‹\<^text>‹S ← remove_all_annot_true_clause_imp_wl_D_heur L S;››
        RETURN (i+1, S)
      }
  })›

definition cdcl_twl_full_restart_wl_D_GC_prog_heur_post :: ‹twl_st_wl_heur ⇒ twl_st_wl_heur ⇒ bool› where
‹cdcl_twl_full_restart_wl_D_GC_prog_heur_post S T ⟷
  (∃S' T'. (S, S') ∈ twl_st_heur_restart ∧ (T, T') ∈ twl_st_heur_restart ∧
    cdcl_twl_full_restart_wl_GC_prog_post S' T')›

definition remove_one_annot_true_clause_imp_wl_D_heur_inv
  :: ‹twl_st_wl_heur ⇒ (nat × twl_st_wl_heur) ⇒ bool› where
  ‹remove_one_annot_true_clause_imp_wl_D_heur_inv S = (λ(i, T).
    (∃S' T'. (S, S') ∈ twl_st_heur_restart ∧ (T, T') ∈ twl_st_heur_restart ∧
     remove_one_annot_true_clause_imp_wl_inv S' (i, T')))›

definition remove_one_annot_true_clause_imp_wl_D_heur :: ‹twl_st_wl_heur ⇒ twl_st_wl_heur nres›
where
‹remove_one_annot_true_clause_imp_wl_D_heur = (λS. do {
    ASSERT((isa_length_trail_pre o get_trail_wl_heur) S);
    k ← (if count_decided_st_heur S = 0
      then RETURN (isa_length_trail (get_trail_wl_heur S))
      else get_pos_of_level_in_trail_imp (get_trail_wl_heur S) 0);
    (_, S) ← WHILETremove_one_annot_true_clause_imp_wl_D_heur_inv S
      (λ(i, S). i < k)
      (λ(i, S). remove_one_annot_true_clause_one_imp_wl_D_heur i S)
      (0, S);
    RETURN S
  })›


lemma get_pos_of_level_in_trail_le_decomp:
  assumes
    ‹(S, T) ∈ twl_st_heur_restart›
  shows ‹get_pos_of_level_in_trail (get_trail_wl T) 0
         ≤ SPEC
            (λk. ∃M1. (∃M2 K.
                          (Decided K # M1, M2)
                          ∈ set (get_all_ann_decomposition (get_trail_wl T))) ∧
                      count_decided M1 = 0 ∧ k = length M1)›
  unfolding get_pos_of_level_in_trail_def
proof (rule SPEC_rule)
  fix x
  assume H: ‹x < length (get_trail_wl T) ∧
        is_decided (rev (get_trail_wl T) ! x) ∧
        get_level (get_trail_wl T) (lit_of (rev (get_trail_wl T) ! x)) = 0 + 1›
  let ?M1 = ‹rev (take x (rev (get_trail_wl T)))›
  let ?K =  ‹Decided ((lit_of(rev (get_trail_wl T) ! x)))›
  let ?M2 = ‹rev (drop  (Suc x) (rev (get_trail_wl T)))›
  have T: ‹(get_trail_wl T) = ?M2 @ ?K # ?M1› and
     K: ‹Decided (lit_of ?K) = ?K›
    apply (subst append_take_drop_id[symmetric, of _ ‹length (get_trail_wl T) - Suc x›])
    apply (subst Cons_nth_drop_Suc[symmetric])
    using H
    apply (auto simp: rev_take rev_drop rev_nth)
    apply (cases ‹rev (get_trail_wl T) ! x›)
    apply (auto simp: rev_take rev_drop rev_nth)
    done
  have n_d: ‹no_dup (get_trail_wl T)›
    using assms(1)
    by (auto simp: twl_st_heur_restart_def)
  obtain M2 where
    ‹(?K # ?M1, M2) ∈ set (get_all_ann_decomposition (get_trail_wl T))›
    using get_all_ann_decomposition_ex[of ‹lit_of ?K› ?M1 ?M2]
    apply (subst (asm) K)
    apply (subst (asm) K)
    apply (subst (asm) T[symmetric])
    by blast
  moreover have ‹count_decided ?M1 = 0›
    using n_d H
    by (subst (asm)(1) T, subst (asm)(11)T, subst T) auto
  moreover have ‹x = length ?M1›
    using n_d H by auto
  ultimately show ‹∃M1. (∃M2 K. (Decided K # M1, M2)
                 ∈ set (get_all_ann_decomposition (get_trail_wl T))) ∧
             count_decided M1 = 0 ∧ x = length M1 ›
    by blast
qed

lemma twl_st_heur_restart_isa_length_trail_get_trail_wl:
  ‹(S, T) ∈ twl_st_heur_restart_ana r ⟹ isa_length_trail (get_trail_wl_heur S) = length (get_trail_wl T)›
  unfolding isa_length_trail_def twl_st_heur_restart_ana_def twl_st_heur_restart_def trail_pol_def
  by (cases S) (auto dest: ann_lits_split_reasons_map_lit_of)

lemma twl_st_heur_restart_count_decided_st_alt_def:
  fixes S :: twl_st_wl_heur
  shows ‹(S, T) ∈ twl_st_heur_restart_ana r ⟹ count_decided_st_heur S = count_decided (get_trail_wl T)›
  unfolding count_decided_st_def twl_st_heur_restart_ana_def trail_pol_def twl_st_heur_restart_def
  by (cases S) (auto simp: count_decided_st_heur_def)

lemma twl_st_heur_restart_trailD:
  ‹(S, T) ∈ twl_st_heur_restart_ana r ⟹
    (get_trail_wl_heur S, get_trail_wl T) ∈ trail_pol (all_init_atms_st T)›
  by (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def)

lemma no_dup_nth_proped_dec_notin:
  ‹no_dup M ⟹ k < length M ⟹ M ! k = Propagated L C ⟹ Decided L ∉ set M›
  apply (auto dest!: split_list simp: nth_append nth_Cons defined_lit_def in_set_conv_nth
    split: if_splits nat.splits)
  by (metis no_dup_no_propa_and_dec nth_mem)

lemma remove_all_annot_true_clause_imp_wl_inv_length_cong:
  ‹remove_all_annot_true_clause_imp_wl_inv S xs T ⟹
    length xs = length ys ⟹ remove_all_annot_true_clause_imp_wl_inv S ys T›
  by (auto simp: remove_all_annot_true_clause_imp_wl_inv_def
    remove_all_annot_true_clause_imp_inv_def)

lemma get_literal_and_reason:
  assumes
    ‹((k, S), k', T) ∈ nat_rel ×f twl_st_heur_restart_ana r› and
    ‹remove_one_annot_true_clause_one_imp_wl_pre k' T› and
    proped: ‹is_proped (rev (get_trail_wl T) ! k')›
  shows ‹do {
           L ← isa_trail_nth (get_trail_wl_heur S) k;
           C ← get_the_propagation_reason_pol (get_trail_wl_heur S) L;
           RETURN (L, C)
         } ≤ ⇓ {((L, C), L', C'). L = L' ∧ C' = the C ∧ C ≠ None}
              (SPEC (λp. rev (get_trail_wl T) ! k' = Propagated (fst p) (snd p)))›
proof -
  have n_d: ‹no_dup (get_trail_wl T)› and
   res: ‹((k, S), k', T) ∈ nat_rel ×f twl_st_heur_restart›
    using assms by (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def)
  from no_dup_nth_proped_dec_notin[OF this(1), of ‹length (get_trail_wl T) - Suc k'›]
  have dec_notin: ‹Decided (lit_of (rev (fst T) ! k')) ∉ set (fst T)›
    using proped assms(2) by (cases T; cases ‹rev (get_trail_wl T) ! k'›)
     (auto simp: twl_st_heur_restart_def state_wl_l_def
      remove_one_annot_true_clause_one_imp_wl_pre_def twl_st_l_def
      remove_one_annot_true_clause_one_imp_pre_def rev_nth
      dest: no_dup_nth_proped_dec_notin)
  have k': ‹k' < length (get_trail_wl T)› and [simp]: ‹fst T = get_trail_wl T›
    using proped assms(2)
    by (cases T; auto simp: twl_st_heur_restart_def state_wl_l_def
      remove_one_annot_true_clause_one_imp_wl_pre_def twl_st_l_def
      remove_one_annot_true_clause_one_imp_pre_def; fail)+
  define k'' where ‹k'' ≡ length (get_trail_wl T) - Suc k'›
  have k'': ‹k'' < length (get_trail_wl T)›
    using k' by (auto simp: k''_def)
  have ‹rev (get_trail_wl T) ! k' = get_trail_wl T ! k''›
    using k' k'' by (auto simp: k''_def nth_rev)
  then have ‹rev_trail_nth (fst T) k' ∈# ℒall (all_init_atms_st T)›
    using k'' assms nth_mem[OF k']
    by (auto simp: twl_st_heur_restart_ana_def rev_trail_nth_def
      trail_pol_alt_def twl_st_heur_restart_def)
  then have 1: ‹(SPEC (λp. rev (get_trail_wl T) ! k' = Propagated (fst p) (snd p))) =
    do {
      L ← RETURN (rev_trail_nth (fst T) k');
      ASSERT(L ∈# ℒall (all_init_atms_st T));
      C ← (get_the_propagation_reason (fst T) L);
      ASSERT(C ≠ None);
      RETURN (L, the C)
    }›
    using proped dec_notin k' nth_mem[OF k''] no_dup_same_annotD[OF n_d]
    apply (subst order_class.eq_iff)
    apply (rule conjI)
    subgoal
      unfolding get_the_propagation_reason_def
      by (cases ‹rev (get_trail_wl T) ! k'›)
        (auto simp: RES_RES_RETURN_RES rev_trail_nth_def
            get_the_propagation_reason_def lits_of_def rev_nth
  	    RES_RETURN_RES
          dest: split_list
	  simp flip: k''_def
	  intro!: le_SPEC_bindI[of _ ‹Some (mark_of (get_trail_wl T ! k''))›])
    subgoal
      apply (cases ‹rev (get_trail_wl T) ! k'›) (*TODO proof*)
      apply  (auto simp: RES_RES_RETURN_RES rev_trail_nth_def
          get_the_propagation_reason_def lits_of_def rev_nth
	  RES_RETURN_RES
        simp flip: k''_def
        dest: split_list
        intro!: exI[of _ ‹Some (mark_of (rev (fst T) ! k'))›])
	  apply (subst RES_ASSERT_moveout)
	  apply (auto simp: RES_RETURN_RES
        dest: split_list)
	done
    done

  show ?thesis
    supply RETURN_as_SPEC_refine[refine2 del]
    apply (subst 1)
    apply (refine_rcg
      isa_trail_nth_rev_trail_nth[THEN fref_to_Down_curry, unfolded comp_def,
        of _ _ _ _ ‹all_init_atms_st T›]
      get_the_propagation_reason_pol[THEN fref_to_Down_curry, unfolded comp_def,
        of ‹all_init_atms_st T›])
    subgoal using k' by auto
    subgoal using assms by (cases S; auto dest: twl_st_heur_restart_trailD)
    subgoal by auto
    subgoal for K K'
      using assms by (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def)
    subgoal
      by auto
    done
qed


lemma red_in_dom_number_of_learned_ge1: ‹C' ∈# dom_m baa ⟹ ¬ irred baa C' ⟹ Suc 0 ≤ size (learned_clss_l baa)›
  by (auto simp: ran_m_def dest!: multi_member_split)

lemma mark_garbage_heur2_remove_and_add_cls_l:
  ‹(S, T) ∈ twl_st_heur_restart_ana r ⟹ (C, C') ∈ Id ⟹
    mark_garbage_heur2 C S
       ≤ ⇓ (twl_st_heur_restart_ana r) (remove_and_add_cls_wl C' T)›
  unfolding mark_garbage_heur2_def remove_and_add_cls_wl_def Let_def
  apply (cases S; cases T)
  apply refine_rcg
  subgoal
    by  (auto simp: twl_st_heur_restart_def arena_lifting
      valid_arena_extra_information_mark_to_delete'
      all_init_atms_fmdrop_add_mset_unit learned_clss_l_l_fmdrop
      learned_clss_l_l_fmdrop_irrelev twl_st_heur_restart_ana_def ASSERT_refine_left
      size_Diff_singleton red_in_dom_number_of_learned_ge1 intro!: ASSERT_leI
    dest: in_vdom_m_fmdropD)
  subgoal
    by  (auto simp: twl_st_heur_restart_def arena_lifting
      valid_arena_extra_information_mark_to_delete'
      all_init_atms_fmdrop_add_mset_unit learned_clss_l_l_fmdrop
      learned_clss_l_l_fmdrop_irrelev twl_st_heur_restart_ana_def
      size_Diff_singleton red_in_dom_number_of_learned_ge1
    dest: in_vdom_m_fmdropD)
  done

lemma remove_one_annot_true_clause_one_imp_wl_pre_fst_le_uint32:
  assumes ‹(x, y) ∈ nat_rel ×f twl_st_heur_restart_ana r› and
    ‹remove_one_annot_true_clause_one_imp_wl_pre (fst y) (snd y)›
  shows ‹fst x + 1 ≤ Suc (uint32_max div 2)›
proof -
  have [simp]: ‹fst y = fst x›
    using assms by (cases x, cases y) auto
  have ‹fst x < length (get_trail_wl (snd y))›
    using assms apply -
    unfolding
     remove_one_annot_true_clause_one_imp_wl_pre_def
     remove_one_annot_true_clause_one_imp_pre_def
   by normalize_goal+ auto
  moreover have ‹(get_trail_wl_heur (snd x), get_trail_wl (snd y)) ∈ trail_pol (all_init_atms_st (snd y))›
    using assms
    by (cases x, cases y) (simp add: twl_st_heur_restart_ana_def
      twl_st_heur_restart_def)
  ultimately show ‹?thesis›
    by (auto simp add: trail_pol_alt_def)
qed

lemma remove_one_annot_true_clause_one_imp_wl_D_heur_remove_one_annot_true_clause_one_imp_wl_D:
  ‹(uncurry remove_one_annot_true_clause_one_imp_wl_D_heur,
    uncurry remove_one_annot_true_clause_one_imp_wl) ∈
    nat_rel ×f twl_st_heur_restart_ana r →f ⟨nat_rel ×f twl_st_heur_restart_ana r⟩nres_rel›
  unfolding remove_one_annot_true_clause_one_imp_wl_D_heur_def
    remove_one_annot_true_clause_one_imp_wl_def case_prod_beta uncurry_def
  apply (intro frefI nres_relI)
  subgoal for x y
  apply (refine_rcg get_literal_and_reason[where r=r]
    isasat_replace_annot_in_trail_replace_annot_in_trail_spec
      [where r=r]
    mark_garbage_heur2_remove_and_add_cls_l[where r=r])
  subgoal by auto
  subgoal unfolding remove_one_annot_true_clause_one_imp_wl_pre_def
    by auto
  subgoal
    by (rule remove_one_annot_true_clause_one_imp_wl_pre_fst_le_uint32)
  subgoal for p pa
    by (cases pa)
      (auto simp: all_init_atms_def simp del: all_init_atms_def[symmetric])
  subgoal
    by (cases x, cases y)
     (fastforce simp: twl_st_heur_restart_def
       trail_pol_alt_def)+
  subgoal by auto
  subgoal for p pa
    by (cases pa; cases p; cases x; cases y)
      (auto simp: all_init_atms_def simp del: all_init_atms_def[symmetric])

  subgoal for p pa S Sa
    unfolding mark_garbage_pre_def
      arena_is_valid_clause_idx_def
      prod.case
    apply (rule_tac x = ‹get_clauses_wl Sa› in exI)
    apply (rule_tac x = ‹set (get_vdom S)› in exI)
    apply (case_tac S, case_tac Sa; cases y)
    apply (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def)
    done
  subgoal for p pa S Sa
    unfolding arena_is_valid_clause_vdom_def
    apply (rule_tac x = ‹get_clauses_wl Sa› in exI)
    apply (rule_tac x = ‹set (get_vdom S)› in exI)
    apply (case_tac S, case_tac Sa; cases y)
    apply (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def)
    done
  subgoal
    by auto
  subgoal
    by auto
  subgoal
    by (cases x, cases y) fastforce
  done
  done

(*
lemma remove_one_annot_true_clause_imp_wl_D_heur_remove_one_annot_true_clause_imp_wl_D:
  ‹(remove_one_annot_true_clause_imp_wl_D_heur, remove_one_annot_true_clause_imp_wl_D) ∈
    twl_st_heur_restart →f ⟨twl_st_heur_restart⟩nres_rel›
  unfolding remove_one_annot_true_clause_imp_wl_D_heur_def
    remove_one_annot_true_clause_imp_wl_D_def
  apply (intro frefI nres_relI)
  apply (refine_vcg WHILEIT_refine[where R = ‹nat_rel ×r {(S, T). (S, T) ∈ twl_st_heur_restart ∧
     literals_are_ℒin' (all_init_atms_st T) T}›])
  subgoal by (auto simp: twl_st_heur_restart_count_decided_st_alt_def
    twl_st_heur_restart_isa_length_trail_get_trail_wl)
  subgoal for S T
    apply (rule order_trans)
      apply (rule get_pos_of_level_in_trail_imp_get_pos_of_level_in_trail_CS[THEN fref_to_Down_curry,
        of ‹get_trail_wl T› ‹0::nat› ‹get_trail_wl_heur S› _ ‹all_init_atms_st T›])
    apply (auto simp: get_pos_of_level_in_trail_pre_def twl_st_heur_restart_count_decided_st_alt_def
      dest: twl_st_heur_restart_trailD
      intro!: get_pos_of_level_in_trail_le_decomp)
    done
  subgoal
    by (auto simp: remove_one_annot_true_clause_imp_wl_D_inv_def)
  subgoal for x y k ka xa x'
    unfolding remove_one_annot_true_clause_imp_wl_D_heur_inv_def
    apply (subst case_prod_beta)
    apply (subst (asm)(11) surjective_pairing)
    apply (subst (asm)(9) surjective_pairing)
    unfolding prod_rel_iff
    apply (rule_tac x=y in exI)
    apply (rule_tac x= ‹snd x'› in exI)
    apply auto
    done
  subgoal by auto
  subgoal sor ry
  subgoal by auto
  done


    *)
definition find_decomp_wl0 :: ‹'v twl_st_wl ⇒ 'v twl_st_wl ⇒ bool› where
  ‹find_decomp_wl0 = (λ(M, N, D, NE, UE, NS, US, Q, W) (M', N', D', NE', UE', NS', US', Q', W').
	 (∃K M2. (Decided K # M', M2) ∈ set (get_all_ann_decomposition M) ∧
	    count_decided M' = 0) ∧
	  (N', D', NE', UE', NS, US, Q', W') = (N, D, NE, UE, NS', US', Q, W))›

definition empty_Q_wl  :: ‹'v twl_st_wl ⇒ 'v twl_st_wl› where
‹empty_Q_wl = (λ(M', N, D, NE, UE, NS, US, _, W). (M', N, D, NE, UE, NS, {#}, {#}, W))›

definition empty_US_wl  :: ‹'v twl_st_wl ⇒ 'v twl_st_wl› where
‹empty_US_wl = (λ(M', N, D, NE, UE, NS, US, Q, W). (M', N, D, NE, UE, NS, {#}, Q, W))›

lemma cdcl_twl_local_restart_wl_spec0_alt_def:
  ‹cdcl_twl_local_restart_wl_spec0 = (λS. do {
    ASSERT(restart_abs_wl_pre2 S False);
    if count_decided (get_trail_wl S) > 0
    then do {
      T ← SPEC(find_decomp_wl0 S);
      RETURN (empty_Q_wl T)
    } else RETURN (empty_US_wl S)})›
  by (intro ext; case_tac S)
   (auto 5 3 simp: cdcl_twl_local_restart_wl_spec0_def
	RES_RETURN_RES2 image_iff RES_RETURN_RES empty_US_wl_def
	find_decomp_wl0_def empty_Q_wl_def uncurry_def
       intro!: bind_cong[OF refl]
      dest: get_all_ann_decomposition_exists_prepend)

lemma cdcl_twl_local_restart_wl_spec0:
  assumes Sy:  ‹(S, y) ∈ twl_st_heur_restart_ana r› and
    ‹get_conflict_wl y = None›
  shows ‹do {
      if count_decided_st_heur S > 0
      then do {
        S ← find_decomp_wl_st_int 0 S;
        empty_Q S
      } else RETURN S
    }
         ≤ ⇓ (twl_st_heur_restart_ana r) (cdcl_twl_local_restart_wl_spec0 y)›
proof -
  define upd :: ‹_ ⇒ _ ⇒ twl_st_wl_heur ⇒ twl_st_wl_heur› where
    ‹upd M' vm = (λ (_, N, D, Q, W, _, clvls, cach, lbd, stats).
       (M', N, D, Q, W, vm, clvls, cach, lbd, stats))›
     for M' :: trail_pol and vm

  have find_decomp_wl_st_int_alt_def:
    ‹find_decomp_wl_st_int = (λhighest S. do{
      (M', vm) ← isa_find_decomp_wl_imp (get_trail_wl_heur S) highest (get_vmtf_heur S);
      RETURN (upd M' vm S)
    })›
    unfolding upd_def find_decomp_wl_st_int_def
    by (auto intro!: ext)

  have [refine0]: ‹do {
	  (M', vm) ←
	    isa_find_decomp_wl_imp (get_trail_wl_heur S) 0 (get_vmtf_heur S);
	  RETURN (upd M' vm S)
	} ≤ ⇓ {((M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, (fast_ema,
         slow_ema, ccount, wasted),
       vdom, avdom, lcount, opts),
     T).
       ((M', N', D', isa_length_trail M', W', vm, clvls, cach, lbd, outl, stats, (fast_ema,
         slow_ema, restart_info_restart_done ccount, wasted), vdom, avdom, lcount, opts),
	  (empty_Q_wl T)) ∈ twl_st_heur_restart_ana r ∧
	  isa_length_trail_pre M'} (SPEC (find_decomp_wl0 y))›
     (is ‹_ ≤ ⇓ ?A _›)
    if
      ‹0 < count_decided_st_heur S› and
      ‹0 < count_decided (get_trail_wl y)›
  proof -
    have A:
      ‹?A = {((M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, (fast_ema, slow_ema,
	  ccount, wasted),
       vdom, avdom, lcount, opts),
     T).
       ((M', N', D', length (get_trail_wl T), W', vm, clvls, cach, lbd, outl, stats, (fast_ema,
         slow_ema, restart_info_restart_done ccount, wasted), vdom, avdom, lcount, opts),
	  (empty_Q_wl T)) ∈ twl_st_heur_restart_ana r ∧
	  isa_length_trail_pre M'}›
	  supply[[goals_limit=1]]
      apply (rule ; rule)
      subgoal for x
        apply clarify
	apply (frule twl_st_heur_restart_isa_length_trail_get_trail_wl)
        by (auto simp:  trail_pol_def empty_Q_wl_def
            isa_length_trail_def
          dest!: ann_lits_split_reasons_map_lit_of)
      subgoal for x
        apply clarify
	apply (frule twl_st_heur_restart_isa_length_trail_get_trail_wl)
        by (auto simp:  trail_pol_def empty_Q_wl_def
            isa_length_trail_def
          dest!: ann_lits_split_reasons_map_lit_of)
      done

    let ?𝒜 = ‹all_init_atms_st y›
    have ‹get_vmtf_heur S ∈ isa_vmtf ?𝒜 (get_trail_wl y)›and
      n_d: ‹no_dup (get_trail_wl y)›
      using Sy
      by (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def)
    then obtain vm' where
      vm': ‹(get_vmtf_heur S, vm') ∈ Id ×f distinct_atoms_rel ?𝒜› and
      vm: ‹vm' ∈ vmtf (all_init_atms_st y) (get_trail_wl y)›
      unfolding isa_vmtf_def
      by force

    have find_decomp_w_ns_pre:
      ‹find_decomp_w_ns_pre (all_init_atms_st y) ((get_trail_wl y, 0), vm')›
      using that assms vm' vm unfolding find_decomp_w_ns_pre_def
      by (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def
        dest: trail_pol_literals_are_in_ℒin_trail)
    have 1: ‹isa_find_decomp_wl_imp (get_trail_wl_heur S) 0 (get_vmtf_heur S) ≤
       ⇓ ({(M, M'). (M, M') ∈ trail_pol ?𝒜 ∧ count_decided M' = 0} ×f (Id ×f distinct_atoms_rel ?𝒜))
         (find_decomp_w_ns ?𝒜 (get_trail_wl y) 0 vm')›
      apply (rule  order_trans)
      apply (rule isa_find_decomp_wl_imp_find_decomp_wl_imp[THEN fref_to_Down_curry2,
        of ‹get_trail_wl y› 0 vm' _ _ _ ?𝒜])
      subgoal using that by auto
      subgoal
        using Sy vm'
	by (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def)
      apply (rule order_trans, rule ref_two_step')
      apply (rule find_decomp_wl_imp_find_decomp_wl'[THEN fref_to_Down_curry2,
        of ?𝒜 ‹get_trail_wl y› 0 vm'])
      subgoal by (rule find_decomp_w_ns_pre)
      subgoal by auto
      subgoal
        using n_d
        by (fastforce simp: find_decomp_w_ns_def conc_fun_RES Image_iff)
      done
    show ?thesis
      supply [[goals_limit=1]] unfolding A
      apply (rule bind_refine_res[OF _ 1[unfolded find_decomp_w_ns_def conc_fun_RES]])
      apply (case_tac y, cases S)
      apply clarify
      apply (rule RETURN_SPEC_refine)
      using assms
      by (auto simp: upd_def find_decomp_wl0_def
        intro!: RETURN_SPEC_refine simp: twl_st_heur_restart_def out_learned_def
	    empty_Q_wl_def twl_st_heur_restart_ana_def
	  intro: isa_vmtfI isa_length_trail_pre dest: no_dup_appendD)
  qed

  have Sy': ‹(S, empty_US_wl y) ∈ twl_st_heur_restart_ana r›
    using Sy by (cases y; cases S; auto simp: twl_st_heur_restart_ana_def
       empty_US_wl_def twl_st_heur_restart_def)
  show ?thesis
    unfolding find_decomp_wl_st_int_alt_def
      cdcl_twl_local_restart_wl_spec0_alt_def
    apply refine_vcg
    subgoal
      using Sy by (auto simp: twl_st_heur_restart_count_decided_st_alt_def)
    subgoal
      unfolding empty_Q_def empty_Q_wl_def
      by refine_vcg
        (simp_all add: twl_st_heur_restart_isa_length_trail_get_trail_wl)
    subgoal
      using Sy' .
    done
qed

lemma no_get_all_ann_decomposition_count_dec0:
  ‹(∀M1. (∀M2 K. (Decided K # M1, M2) ∉ set (get_all_ann_decomposition M))) ⟷
  count_decided M = 0›
  apply (induction M rule: ann_lit_list_induct)
  subgoal by auto
  subgoal for L M
    by auto
  subgoal for L C M
    by (cases ‹get_all_ann_decomposition M›) fastforce+
  done

lemma get_pos_of_level_in_trail_decomp_iff:
  assumes ‹no_dup M›
  shows ‹((∃M1 M2 K.
                (Decided K # M1, M2)
                ∈ set (get_all_ann_decomposition M) ∧
                count_decided M1 = 0 ∧ k = length M1)) ⟷
    k < length M ∧ count_decided M > 0 ∧ is_decided (rev M ! k) ∧ get_level M (lit_of (rev M ! k)) = 1›
  (is ‹?A ⟷ ?B›)
proof
  assume ?A
  then obtain K M1 M2 where
    decomp: ‹(Decided K # M1, M2) ∈ set (get_all_ann_decomposition M)› and
    [simp]: ‹count_decided M1 = 0› and
    k_M1: ‹length M1 = k›
    by auto
  then have ‹k < length M›
    by auto
  moreover have ‹rev M ! k = Decided K›
    using decomp
    by (auto dest!: get_all_ann_decomposition_exists_prepend
      simp: nth_append
      simp flip: k_M1)
  moreover have ‹get_level M (lit_of (rev M ! k)) = 1›
    using assms decomp
    by (auto dest!: get_all_ann_decomposition_exists_prepend
      simp: get_level_append_if nth_append
      simp flip: k_M1)
  ultimately show ?B
    using decomp by auto
next
  assume ?B
  define K where ‹K = lit_of (rev M ! k)›
  obtain M1 M2 where
    M: ‹M = M2 @ Decided K # M1› and
    k_M1: ‹length M1 = k›
    apply (subst (asm) append_take_drop_id[of ‹length M - Suc k›, symmetric])
    apply (subst (asm) Cons_nth_drop_Suc[symmetric])
    unfolding K_def
    subgoal using ‹?B› by auto
    subgoal using ‹?B› K_def by (cases ‹rev M ! k›) (auto simp: rev_nth)
    done
  moreover have ‹count_decided M1 = 0›
    using assms ‹?B› unfolding M
    by (auto simp: nth_append k_M1)
  ultimately show ?A
    using get_all_ann_decomposition_ex[of K M1 M2]
    unfolding M
    by auto
qed

lemma remove_all_learned_subsumed_clauses_wl_id:
  ‹(x2a, x2) ∈ twl_st_heur_restart_ana r ⟹
   RETURN x2a
    ≤ ⇓ (twl_st_heur_restart_ana r)
       (remove_all_learned_subsumed_clauses_wl x2)›
   by (cases x2a; cases x2)
    (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def
     remove_all_learned_subsumed_clauses_wl_def)

lemma remove_one_annot_true_clause_imp_wl_D_heur_remove_one_annot_true_clause_imp_wl_D:
  ‹(remove_one_annot_true_clause_imp_wl_D_heur, remove_one_annot_true_clause_imp_wl) ∈
    twl_st_heur_restart_ana r →f ⟨twl_st_heur_restart_ana r⟩nres_rel›
  unfolding remove_one_annot_true_clause_imp_wl_def
    remove_one_annot_true_clause_imp_wl_D_heur_def
  apply (intro frefI nres_relI)
  apply (refine_vcg
    WHILEIT_refine[where R = ‹nat_rel ×r twl_st_heur_restart_ana r›]
    remove_one_annot_true_clause_one_imp_wl_D_heur_remove_one_annot_true_clause_one_imp_wl_D[THEN
      fref_to_Down_curry])
  subgoal by (auto simp: trail_pol_alt_def isa_length_trail_pre_def
    twl_st_heur_restart_def twl_st_heur_restart_ana_def)
  subgoal by (auto simp: twl_st_heur_restart_isa_length_trail_get_trail_wl
    twl_st_heur_restart_count_decided_st_alt_def)
  subgoal for x y
    apply (rule order_trans)
    apply (rule get_pos_of_level_in_trail_imp_get_pos_of_level_in_trail_CS[THEN fref_to_Down_curry,
        of ‹get_trail_wl y› 0 _ _ ‹all_init_atms_st y›])
    subgoal by (auto simp: get_pos_of_level_in_trail_pre_def
      twl_st_heur_restart_count_decided_st_alt_def)
    subgoal by (auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def)
    subgoal
      apply (subst get_pos_of_level_in_trail_decomp_iff)
      apply (solves ‹auto simp: twl_st_heur_restart_def twl_st_heur_restart_ana_def›)
      apply (auto simp: get_pos_of_level_in_trail_def
        twl_st_heur_restart_count_decided_st_alt_def)
      done
    done
    subgoal by auto
    subgoal for x y k k' T T'
      apply (subst (asm)(12) surjective_pairing)
      apply (subst (asm)(10) surjective_pairing)
      unfolding remove_one_annot_true_clause_imp_wl_D_heur_inv_def
        prod_rel_iff
      apply (subst (10) surjective_pairing, subst prod.case)
      by (fastforce intro: twl_st_heur_restart_anaD simp: twl_st_heur_restart_anaD)
    subgoal by auto
    subgoal by auto
    subgoal by (auto intro!: remove_all_learned_subsumed_clauses_wl_id)
  done


lemma mark_to_delete_clauses_wl_D_heur_mark_to_delete_clauses_wl2_D:
  ‹(mark_to_delete_clauses_wl_D_heur, mark_to_delete_clauses_wl2) ∈
     twl_st_heur_restart_ana r →f ⟨twl_st_heur_restart_ana r⟩nres_rel›
proof -
  have mark_to_delete_clauses_wl2_D_alt_def:
    ‹mark_to_delete_clauses_wl2  = (λS0. do {
      ASSERT(mark_to_delete_clauses_wl_pre S0);
      S ← reorder_vdom_wl S0;
      xs ← collect_valid_indices_wl S;
      l ← SPEC(λ_::nat. True);
      (_, S, _) ← WHILETmark_to_delete_clauses_wl2_inv S xs
        (λ(i, T, xs). i < length xs)
        (λ(i, T, xs). do {
          b ← RETURN (xs!i ∈# dom_m (get_clauses_wl T));
          if ¬b then RETURN (i, T, delete_index_and_swap xs i)
          else do {
            ASSERT(0 < length (get_clauses_wl T∝(xs!i)));
	    ASSERT (get_clauses_wl T ∝ (xs ! i) ! 0 ∈# ℒall (all_init_atms_st T));
            K ← RETURN (get_clauses_wl T ∝ (xs ! i) ! 0);
            b ← RETURN (); ― ‹propagation reason›
            can_del ← SPEC(λb. b ⟶
              (Propagated (get_clauses_wl T∝(xs!i)!0) (xs!i) ∉ set (get_trail_wl T)) ∧
               ¬irred (get_clauses_wl T) (xs!i) ∧ length (get_clauses_wl T ∝ (xs!i)) ≠ 2);
            ASSERT(i < length xs);
            if can_del
            then
              RETURN (i, mark_garbage_wl (xs!i) T, delete_index_and_swap xs i)
            else
              RETURN (i+1, T, xs)
          }
        })
        (l, S, xs);
      remove_all_learned_subsumed_clauses_wl S
    })›
    unfolding mark_to_delete_clauses_wl2_def reorder_vdom_wl_def bind_to_let_conv Let_def
    by (force intro!: ext)
  have mono: ‹g = g' ⟹ do {f; g} = do {f; g'}›
     ‹(⋀x. h x = h' x) ⟹ do {x ← f; h x} = do {x ← f; h' x}› for f f' :: ‹_ nres› and g g' and h h'
    by auto

  have [refine0]: ‹RETURN (get_avdom x) ≤ ⇓ {(xs, xs'). xs = xs' ∧ xs = get_avdom x} (collect_valid_indices_wl y)›
    if
      ‹(x, y) ∈ twl_st_heur_restart_ana r› and
      ‹mark_to_delete_clauses_wl_D_heur_pre x›
    for x y
  proof -
    show ?thesis by (auto simp: collect_valid_indices_wl_def simp: RETURN_RES_refine_iff)
  qed
  have init_rel[refine0]: ‹(x, y) ∈ twl_st_heur_restart_ana r ⟹
       (l, la) ∈ nat_rel ⟹
       ((l, x), la, y) ∈ nat_rel ×f {(S, T). (S, T) ∈ twl_st_heur_restart_ana r ∧ get_avdom S = get_avdom x}›
    for x y l la
    by auto

  define reason_rel where
    ‹reason_rel K x1a ≡ {(C, _ :: unit).
          (C ≠ None) = (Propagated K (the C) ∈ set (get_trail_wl x1a)) ∧
          (C = None) = (Decided K ∈ set (get_trail_wl x1a) ∨
             K ∉ lits_of_l (get_trail_wl x1a)) ∧
         (∀C1. (Propagated K C1 ∈ set (get_trail_wl x1a) ⟶ C1 = the C))}› for K :: ‹nat literal› and x1a
  have get_the_propagation_reason:
    ‹get_the_propagation_reason_pol (get_trail_wl_heur x2b) L
        ≤ SPEC (λD. (D, ()) ∈ reason_rel K x1a)›
  if
    ‹(x, y) ∈ twl_st_heur_restart_ana r› and
    ‹mark_to_delete_clauses_wl_pre y› and
    ‹mark_to_delete_clauses_wl_D_heur_pre x› and
    ‹(S, Sa)
     ∈ {(U, V).
        (U, V) ∈ twl_st_heur_restart_ana r ∧
        V = y ∧
        (mark_to_delete_clauses_wl_pre y ⟶
         mark_to_delete_clauses_wl_pre V) ∧
        (mark_to_delete_clauses_wl_D_heur_pre x ⟶
         mark_to_delete_clauses_wl_D_heur_pre U)}› and
    ‹(ys, xs) ∈ {(xs, xs'). xs = xs' ∧ xs = get_avdom S}› and
    ‹(l, la) ∈ nat_rel› and
    ‹la ∈ {_. True}› and
    xa_x': ‹(xa, x')
     ∈ nat_rel ×f {(Sa, T, xs). (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
    ‹case xa of (i, S) ⇒ i < length (get_avdom S)› and
    ‹case x' of (i, T, xs) ⇒ i < length xs› and
    ‹x1b < length (get_avdom x2b)› and
    ‹access_vdom_at_pre x2b x1b› and
    dom: ‹(b, ba)
       ∈ {(b, b').
          (b, b') ∈ bool_rel ∧
          b = (x2a ! x1 ∈# dom_m (get_clauses_wl x1a))}›
      ‹¬ ¬ b›
      ‹¬ ¬ ba› and
    ‹0 < length (get_clauses_wl x1a ∝ (x2a ! x1))› and
    ‹access_lit_in_clauses_heur_pre ((x2b, get_avdom x2b ! x1b), 0)› and
    st:
      ‹x2 = (x1a, x2a)›
      ‹x' = (x1, x2)›
      ‹xa = (x1b, x2b)› and
    L: ‹get_clauses_wl x1a ∝ (x2a ! x1) ! 0 ∈# ℒall (all_init_atms_st x1a)› and
    L': ‹(L, K)
    ∈ {(L, L').
       (L, L') ∈ nat_lit_lit_rel ∧
       L' = get_clauses_wl x1a ∝ (x2a ! x1) ! 0}›
    for x y S Sa xs' xs l la xa x' x1 x2 x1a x2a x1' x2' x3 x1b ys x2b L K b ba
  proof -
    have L: ‹arena_lit (get_clauses_wl_heur x2b) (x2a ! x1) ∈# ℒall (all_init_atms_st x1a)›
      using L that by (auto simp: twl_st_heur_restart st arena_lifting dest: all_init_all twl_st_heur_restart_anaD)

    show ?thesis
      apply (rule order.trans)
      apply (rule get_the_propagation_reason_pol[THEN fref_to_Down_curry,
        of ‹all_init_atms_st x1a› ‹get_trail_wl x1a›
	  ‹arena_lit (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b + 0)›])
      subgoal
        using xa_x' L L' by (auto simp add: twl_st_heur_restart_def st)
      subgoal
        using xa_x' L' dom by (auto simp add: twl_st_heur_restart_ana_def twl_st_heur_restart_def st arena_lifting)
      using that unfolding get_the_propagation_reason_def reason_rel_def apply -
      by (auto simp: twl_st_heur_restart lits_of_def get_the_propagation_reason_def
          conc_fun_RES
        dest!: twl_st_heur_restart_anaD dest: twl_st_heur_restart_same_annotD imageI[of _ _ lit_of])
  qed
  have ‹((M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur, vdom, avdom, lcount),
           S')
          ∈ twl_st_heur_restart ⟹
    ((M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur, vdom, avdom', lcount),
           S')
          ∈ twl_st_heur_restart›
    if ‹mset avdom' ⊆# mset avdom›
    for M' N' D' j W' vm clvls cach lbd outl stats fast_ema slow_ema
      ccount vdom lcount S' avdom' avdom heur
    using that unfolding twl_st_heur_restart_def
    by auto
  then have mark_to_delete_clauses_wl_D_heur_pre_vdom':
    ‹mark_to_delete_clauses_wl_D_heur_pre (M', N', D', j, W', vm, clvls, cach, lbd, outl, stats,
       heur, vdom, avdom', lcount) ⟹
      mark_to_delete_clauses_wl_D_heur_pre (M', N', D', j, W', vm, clvls, cach, lbd, outl, stats,
        heur, vdom, avdom, lcount)›
    if ‹mset avdom ⊆# mset avdom'›
    for M' N' D' j W' vm clvls cach lbd outl stats fast_ema slow_ema avdom avdom'
      ccount vdom lcount heur
    using that
    unfolding mark_to_delete_clauses_wl_D_heur_pre_def
    by metis
  have [refine0]:
    ‹sort_vdom_heur S ≤ ⇓ {(U, V). (U, V) ∈ twl_st_heur_restart_ana r ∧ V = T ∧
         (mark_to_delete_clauses_wl_pre T ⟶ mark_to_delete_clauses_wl_pre V) ∧
         (mark_to_delete_clauses_wl_D_heur_pre S ⟶ mark_to_delete_clauses_wl_D_heur_pre U)}
         (reorder_vdom_wl T)›
    if ‹(S, T) ∈ twl_st_heur_restart_ana r› for S T
    using that unfolding reorder_vdom_wl_def sort_vdom_heur_def
    apply (refine_rcg ASSERT_leI)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
    apply (rule specify_left)
    apply (rule_tac N1 = ‹get_clauses_wl T› and vdom1 = ‹(get_vdom S)› in
     order_trans[OF isa_remove_deleted_clauses_from_avdom_remove_deleted_clauses_from_avdom,
      unfolded Down_id_eq, OF _ _ _ remove_deleted_clauses_from_avdom])
    subgoal for x y x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h
       x1i x2i x1j x2j x1k x2k x1l x2l
      by (case_tac T; auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def mem_Collect_eq prod.case)
    subgoal for x y x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h
       x1i x2i x1j x2j x1k x2k x1l x2l
      by (case_tac T; auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def mem_Collect_eq prod.case)
    subgoal for x y x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e x1f x2f x1g x2g x1h x2h
       x1i x2i x1j x2j x1k x2k x1l x2l
      by (case_tac T; auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def mem_Collect_eq prod.case)
    apply (subst assert_bind_spec_conv, intro conjI)
    subgoal for x y
      unfolding valid_sort_clause_score_pre_def arena_is_valid_clause_vdom_def
        get_clause_LBD_pre_def arena_is_valid_clause_idx_def arena_act_pre_def
      by (force simp: valid_sort_clause_score_pre_def twl_st_heur_restart_ana_def arena_dom_status_iff
        arena_act_pre_def get_clause_LBD_pre_def arena_is_valid_clause_idx_def twl_st_heur_restart_def
         intro!: exI[of _ ‹get_clauses_wl T›] dest!: set_mset_mono mset_subset_eqD)
    apply (subst assert_bind_spec_conv, intro conjI)
    subgoal
     by (auto simp: twl_st_heur_restart_ana_def valid_arena_vdom_subset twl_st_heur_restart_def
        dest!: size_mset_mono valid_arena_vdom_subset)
    subgoal
      apply (rewrite at ‹_ ≤ ⌑› Down_id_eq[symmetric])
      apply (rule bind_refine_spec)
      prefer 2
      apply (rule sort_clauses_by_score_reorder[of _ ‹get_clauses_wl T› ‹get_vdom S›])
      by (auto 5 3 simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest: mset_eq_setD
         intro: mark_to_delete_clauses_wl_D_heur_pre_vdom'
         dest: mset_eq_setD)
    done
  have already_deleted:
    ‹((x1b, delete_index_vdom_heur x1b x2b), x1, x1a,
       delete_index_and_swap x2a x1)
      ∈ nat_rel ×f {(Sa, T, xs). (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}›
    if
      ‹(x, y) ∈ twl_st_heur_restart_ana r› and
      ‹mark_to_delete_clauses_wl_D_heur_pre x› and
      ‹(S, Sa)
     ∈ {(U, V).
        (U, V) ∈ twl_st_heur_restart_ana r ∧
        V = y ∧
        (mark_to_delete_clauses_wl_pre y ⟶
         mark_to_delete_clauses_wl_pre V) ∧
        (mark_to_delete_clauses_wl_D_heur_pre x ⟶
         mark_to_delete_clauses_wl_D_heur_pre U)}› and
      ‹(l, la) ∈ nat_rel› and
      ‹la ∈ {_. True}› and
      xx: ‹(xa, x')
     ∈ nat_rel ×f {(Sa, T, xs). (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
      ‹case xa of (i, S) ⇒ i < length (get_avdom S)› and
      ‹case x' of (i, T, xs) ⇒ i < length xs› and
      st:
      ‹x2 = (x1a, x2a)›
      ‹x' = (x1, x2)›
      ‹xa = (x1b, x2b)› and
      le: ‹x1b < length (get_avdom x2b)› and
      ‹access_vdom_at_pre x2b x1b› and
      ‹(b, ba) ∈ {(b, b'). (b, b') ∈ bool_rel ∧ b = (x2a ! x1 ∈# dom_m (get_clauses_wl x1a))}› and
      ‹¬ba›
    for x y S xs l la xa x' xz x1 x2 x1a x2a x2b x2c x2d ys x1b Sa ba b
  proof -
    show ?thesis
      using xx le unfolding st
      by (auto simp: twl_st_heur_restart_ana_def delete_index_vdom_heur_def
          twl_st_heur_restart_def mark_garbage_heur_def mark_garbage_wl_def
          learned_clss_l_l_fmdrop size_remove1_mset_If
          intro: valid_arena_extra_information_mark_to_delete'
          dest!: in_set_butlastD in_vdom_m_fmdropD
          elim!: in_set_upd_cases)
  qed
  have get_learned_count_ge: ‹Suc 0 ≤ get_learned_count x2b›
    if
      xy: ‹(x, y) ∈ twl_st_heur_restart_ana r› and
      ‹(xa, x')
       ∈ nat_rel ×f
         {(Sa, T, xs).
          (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
      ‹x2 = (x1a, x2a)› and
      ‹x' = (x1, x2)› and
      ‹xa = (x1b, x2b)› and
      dom: ‹(b, ba)
         ∈ {(b, b').
            (b, b') ∈ bool_rel ∧
            b = (x2a ! x1 ∈# dom_m (get_clauses_wl x1a))}›
        ‹¬ ¬ b›
        ‹¬ ¬ ba› and
      ‹MINIMUM_DELETION_LBD
    < arena_lbd (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b) ∧
    arena_status (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b) = LEARNED ∧
    arena_length (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b) ≠ 2 ∧
    ¬ marked_as_used (get_clauses_wl_heur x2b) (get_avdom x2b ! x1b)› and
      ‹can_del› for x y S Sa uu xs l la xa x' x1 x2 x1a x2a x1b x2b D can_del b ba
  proof -
    have ‹¬irred (get_clauses_wl x1a) (x2a ! x1)› and ‹(x2b, x1a) ∈ twl_st_heur_restart_ana r›
      using that by (auto simp: twl_st_heur_restart arena_lifting
        dest!: twl_st_heur_restart_anaD twl_st_heur_restart_valid_arena)
    then show ?thesis
     using dom by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def ran_m_def
       dest!: multi_member_split)
  qed
  have mop_clause_not_marked_to_delete_heur:
    ‹mop_clause_not_marked_to_delete_heur x2b (get_avdom x2b ! x1b)
        ≤ SPEC
           (λc. (c, x2a ! x1 ∈# dom_m (get_clauses_wl x1a))
                ∈ {(b, b'). (b,b') ∈ bool_rel ∧ (b ⟷ x2a ! x1 ∈# dom_m (get_clauses_wl x1a))})›
    if
      ‹(xa, x')
       ∈ nat_rel ×f
         {(Sa, T, xs).
          (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
      ‹case xa of (i, S) ⇒ i < length (get_avdom S)› and
      ‹case x' of (i, T, xs) ⇒ i < length xs› and
      ‹mark_to_delete_clauses_wl2_inv Sa xs x'› and
      ‹x2 = (x1a, x2a)› and
      ‹x' = (x1, x2)› and
      ‹xa = (x1b, x2b)› and
      ‹clause_not_marked_to_delete_heur_pre (x2b, get_avdom x2b ! x1b)›
    for x y S Sa uu xs l la xa x' x1 x2 x1a x2a x1b x2b
    unfolding mop_clause_not_marked_to_delete_heur_def
    apply refine_vcg
    subgoal
      using that by blast
    subgoal
      using that by (auto simp: twl_st_heur_restart arena_lifting dest!: twl_st_heur_restart_anaD)
    done


  have init:
    ‹(u, xs) ∈ {(xs, xs'). xs = xs' ∧ xs = get_avdom S} ⟹
    (l, la) ∈ nat_rel ⟹
    (S, Sa) ∈ twl_st_heur_restart_ana r ⟹
    ((l, S), la, Sa, xs) ∈  nat_rel ×f
       {(Sa, (T, xs)). (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}›
       for x y S Sa xs l la u
    by auto
  have mop_access_lit_in_clauses_heur:
    ‹mop_access_lit_in_clauses_heur x2b (get_avdom x2b ! x1b) 0
        ≤ SPEC
           (λc. (c, get_clauses_wl x1a ∝ (x2a ! x1) ! 0)
                ∈ {(L, L'). (L, L') ∈ nat_lit_lit_rel ∧ L' = get_clauses_wl x1a ∝ (x2a ! x1) ! 0})›
    if
      ‹(x, y) ∈ twl_st_heur_restart_ana r› and
      ‹mark_to_delete_clauses_wl_pre y› and
      ‹mark_to_delete_clauses_wl_D_heur_pre x› and
      ‹(S, Sa)
       ∈ {(U, V).
          (U, V) ∈ twl_st_heur_restart_ana r ∧
          V = y ∧
          (mark_to_delete_clauses_wl_pre y ⟶
           mark_to_delete_clauses_wl_pre V) ∧
          (mark_to_delete_clauses_wl_D_heur_pre x ⟶
           mark_to_delete_clauses_wl_D_heur_pre U)}› and
      ‹(uu, xs) ∈ {(xs, xs'). xs = xs' ∧ xs = get_avdom S}› and
      ‹(l, la) ∈ nat_rel› and
      ‹la ∈ {uu. True}› and
      ‹length (get_avdom S) ≤ length (get_clauses_wl_heur x)› and
      ‹(xa, x')
       ∈ nat_rel ×f
         {(Sa, T, xs).
          (Sa, T) ∈ twl_st_heur_restart_ana r ∧ xs = get_avdom Sa}› and
      ‹case xa of (i, S) ⇒ i < length (get_avdom S)› and
      ‹case x' of (i, T, xs) ⇒ i < length xs› and
      ‹mark_to_delete_clauses_wl2_inv Sa xs x'› and
      ‹x2 = (x1a, x2a)› and
      ‹x' = (x1, x2)› and
      ‹xa = (x1b, x2b)› and
      ‹x1b < length (get_avdom x2b)› and
      ‹access_vdom_at_pre x2b x1b› and
      ‹clause_not_marked_to_delete_heur_pre (x2b, get_avdom x2b ! x1b)› and
      ‹(b, ba)
       ∈ {(b, b').
          (b, b') ∈ bool_rel ∧
          b = (x2a ! x1 ∈# dom_m (get_clauses_wl x1a))}› and
      ‹¬ ¬ b› and
      ‹¬ ¬ ba› and
      ‹0 < length (get_clauses_wl x1a ∝ (x2a ! x1))› and
      ‹get_clauses_wl x1a ∝ (x2a ! x1) ! 0
       ∈# ℒall (all_init_atms_st x1a)› and
      pre: ‹access_lit_in_clauses_heur_pre ((x2b, get_avdom x2b ! x1b), 0)›
     for x y S Sa uu xs l la xa x' x1 x2 x1a x2a x1b x2b b ba
  unfolding mop_access_lit_in_clauses_heur_def mop_arena_lit2_def
  apply refine_vcg
  subgoal using pre unfolding access_lit_in_clauses_heur_pre_def by simp
  subgoal using that by (auto dest!: twl_st_heur_restart_anaD twl_st_heur_restart_valid_arena simp: arena_lifting)
  done

  have incr_restart_stat: ‹incr_restart_stat T
    ≤ ⇓ (twl_st_heur_restart_ana r) (remove_all_learned_subsumed_clauses_wl S)›
    if ‹(T, S) ∈ twl_st_heur_restart_ana r› for S T i
    using that
    by (cases S; cases T)
      (auto simp: conc_fun_RES incr_restart_stat_def
        twl_st_heur_restart_ana_def twl_st_heur_restart_def
        remove_all_learned_subsumed_clauses_wl_def
        RES_RETURN_RES)

  have [refine0]: ‹mark_clauses_as_unused_wl_D_heur i T⤜ incr_restart_stat
    ≤ ⇓ (twl_st_heur_restart_ana r)
       (remove_all_learned_subsumed_clauses_wl S)›
    if ‹(T, S) ∈ twl_st_heur_restart_ana r› for S T i
    apply (cases S)
    apply (rule bind_refine_res[where R = Id, simplified])
    defer
    apply (rule mark_clauses_as_unused_wl_D_heur[unfolded conc_fun_RES, OF that, of i])
    apply (rule incr_restart_stat[THEN order_trans, of _ S])
    by auto

  show ?thesis
    supply sort_vdom_heur_def[simp] twl_st_heur_restart_anaD[dest] [[goals_limit=1]]
    unfolding mark_to_delete_clauses_wl_D_heur_alt_def mark_to_delete_clauses_wl2_D_alt_def
      access_lit_in_clauses_heur_def
    apply (intro frefI nres_relI)
    apply (refine_vcg sort_vdom_heur_reorder_vdom_wl[THEN fref_to_Down])
    subgoal
      unfolding mark_to_delete_clauses_wl_D_heur_pre_def by fast
    subgoal by auto
    subgoal by auto
    subgoal for x y S T unfolding number_clss_to_keep_def by (cases S) (auto)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def
       dest!: valid_arena_vdom_subset size_mset_mono)
    apply (rule init; solves auto)
    subgoal by auto
    subgoal by auto
    subgoal by (auto simp: access_vdom_at_pre_def)
    subgoal for x y S xs l la xa x' xz x1 x2 x1a x2a x2b x2c x2d
      unfolding clause_not_marked_to_delete_heur_pre_def arena_is_valid_clause_vdom_def
        prod.simps
      by (rule exI[of _ ‹get_clauses_wl x2a›], rule exI[of _ ‹set (get_vdom x2d)›])
         (auto simp: twl_st_heur_restart dest: twl_st_heur_restart_get_avdom_nth_get_vdom)
    apply (rule mop_clause_not_marked_to_delete_heur; assumption)
    subgoal for x y S Sa uu xs l la xa x' x1 x2 x1a x2a x1b x2b
      by (auto simp: twl_st_heur_restart)
    subgoal
      by (rule already_deleted)
    subgoal for x y _ _ _ _ _ xs l la xa x' x1 x2 x1a x2a
      unfolding access_lit_in_clauses_heur_pre_def prod.simps arena_lit_pre_def
        arena_is_valid_clause_idx_and_access_def
      by (rule bex_leI[of _ ‹get_avdom x2a ! x1a›], simp, rule exI[of _ ‹get_clauses_wl x1›])
        (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
    subgoal premises p using p(7-) by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
     apply (rule mop_access_lit_in_clauses_heur; assumption)
    apply (rule get_the_propagation_reason; assumption)
    subgoal for x y S Sa _ xs l la xa x' x1 x2 x1a x2a x1b x2b
      unfolding prod.simps
        get_clause_LBD_pre_def arena_is_valid_clause_idx_def
      by (rule exI[of _ ‹get_clauses_wl x1a›], rule exI[of _ ‹set (get_vdom x2b)›])
        (auto simp: twl_st_heur_restart dest: twl_st_heur_restart_valid_arena)
    subgoal for x y S Sa _ xs l la xa x' x1 x2 x1a x2a x1b x2b
      unfolding prod.simps
        arena_is_valid_clause_vdom_def arena_is_valid_clause_idx_def
      by (rule exI[of _ ‹get_clauses_wl x1a›], rule exI[of _ ‹set (get_vdom x2b)›])
        (auto simp: twl_st_heur_restart dest: twl_st_heur_restart_valid_arena
	  twl_st_heur_restart_get_avdom_nth_get_vdom)
    subgoal for x y S Sa _ xs l la xa x' x1 x2 x1a x2a x1b x2b
      unfolding prod.simps
        arena_is_valid_clause_vdom_def arena_is_valid_clause_idx_def
      by (rule exI[of _ ‹get_clauses_wl x1a›], rule exI[of _ ‹set (get_vdom x2b)›])
        (auto simp: twl_st_heur_restart arena_dom_status_iff
          dest: twl_st_heur_restart_valid_arena twl_st_heur_restart_get_avdom_nth_get_vdom)
    subgoal
      unfolding marked_as_used_pre_def
      by (auto simp: twl_st_heur_restart reason_rel_def)
    subgoal
      by (auto simp: twl_st_heur_restart reason_rel_def)
    subgoal
      by (auto simp: twl_st_heur_restart)
    subgoal
      by (auto dest!: twl_st_heur_restart_anaD twl_st_heur_restart_valid_arena simp: arena_lifting)
    subgoal by fast
    subgoal for x y S Sa _ xs l la xa x' x1 x2 x1a x2a x1b x2b
      unfolding mop_arena_length_st_def
      apply (rule mop_arena_length[THEN fref_to_Down_curry, THEN order_trans,
        of ‹get_clauses_wl x1a› ‹get_avdom x2b ! x1b› _ _ ‹set (get_vdom x2b)›])
      subgoal
        by auto
      subgoal
        by (auto simp: twl_st_heur_restart_valid_arena)
      subgoal
        apply (auto intro!: incr_wasted_st_twl_st ASSERT_leI)
        subgoal
          unfolding prod.simps mark_garbage_pre_def
            arena_is_valid_clause_vdom_def arena_is_valid_clause_idx_def
          by (rule exI[of _ ‹get_clauses_wl x1a›], rule exI[of _ ‹set (get_vdom x2b)›])
            (auto simp: twl_st_heur_restart dest: twl_st_heur_restart_valid_arena)
        subgoal
           apply (rule get_learned_count_ge; assumption?; fast?)
           apply auto
           done
        subgoal
          by (use arena_lifting(24)[of ‹get_clauses_wl_heur x2b› _ _  ‹get_avdom x2b ! x1›] in
            ‹auto intro!: incr_wasted_st mark_garbage_heur_wl_ana
            dest: twl_st_heur_restart_valid_arena twl_st_heur_restart_anaD›)
        done
     done
   subgoal for x y
      unfolding valid_sort_clause_score_pre_def arena_is_valid_clause_vdom_def
        get_clause_LBD_pre_def arena_is_valid_clause_idx_def arena_act_pre_def
      by (force simp: valid_sort_clause_score_pre_def twl_st_heur_restart_ana_def arena_dom_status_iff
        arena_act_pre_def get_clause_LBD_pre_def arena_is_valid_clause_idx_def twl_st_heur_restart_def
         intro!: exI[of _ ‹get_clauses_wl T›] dest!: set_mset_mono mset_subset_eqD)
    subgoal
      by (auto intro!: mark_unused_st_heur_ana)
    subgoal by (auto simp: twl_st_heur_restart_ana_def twl_st_heur_restart_def dest!: valid_arena_vdom_subset size_mset_mono)
    subgoal
      by auto
    done
qed


definition iterate_over_VMTF where
  ‹iterate_over_VMTF ≡ (λf (I :: 'a ⇒ bool) (ns :: (nat, nat) vmtf_node list, n) x. do {
      (_, x) ← WHILETλ(n, x). I x
        (λ(n, _). n ≠ None)
        (λ(n, x). do {
          ASSERT(n ≠ None);
          let A = the n;
          ASSERT(A < length ns);
          ASSERT(A ≤ uint32_max div 2);
          x ← f A x;
          RETURN (get_next ((ns ! A)), x)
        })
        (n, x);
      RETURN x
    })›

definition iterate_over_ℒall where
  ‹iterate_over_ℒall = (λf 𝒜0 I x. do {
    𝒜 ← SPEC(λ𝒜. set_mset 𝒜 = set_mset 𝒜0 ∧ distinct_mset 𝒜);
    (_, x) ← WHILETλ(_, x). I x
      (λ(ℬ, _). ℬ ≠ {#})
      (λ(ℬ, x). do {
        ASSERT(ℬ ≠ {#});
        A ← SPEC (λA. A ∈# ℬ);
        x ← f A x;
        RETURN (remove1_mset A ℬ, x)
      })
      (𝒜, x);
    RETURN x
  })›

lemma iterate_over_VMTF_iterate_over_ℒall:
  fixes x :: 'a
  assumes vmtf: ‹((ns, m, fst_As, lst_As, next_search), to_remove) ∈ vmtf 𝒜 M› and
    nempty: ‹𝒜 ≠ {#}› ‹isasat_input_bounded 𝒜›
  shows ‹iterate_over_VMTF f I (ns, Some fst_As) x ≤ ⇓ Id (iterate_over_ℒall f 𝒜 I x)›
proof -
  obtain xs' ys' where
    vmtf_ns: ‹vmtf_ns (ys' @ xs') m ns› and
    ‹fst_As = hd (ys' @ xs')› and
    ‹lst_As = last (ys' @ xs')› and
    vmtf_ℒ: ‹vmtf_ℒall 𝒜 M ((set xs', set ys'), to_remove)› and
    fst_As: ‹fst_As = hd (ys' @ xs')› and
    le: ‹∀L∈atms_of (ℒall 𝒜). L < length ns›
    using vmtf unfolding vmtf_def
    by blast
  define zs where ‹zs = ys' @ xs'›
  define is_lasts where
    ‹is_lasts ℬ n m ⟷ set_mset ℬ = set (drop m zs) ∧ set_mset ℬ ⊆ set_mset 𝒜 ∧
        distinct_mset ℬ ∧
        card (set_mset ℬ) ≤ length zs ∧
        card (set_mset ℬ) + m = length zs ∧
        (n = option_hd (drop m zs)) ∧
        m ≤ length zs› for  and n :: ‹nat option› and m
  have card_𝒜: ‹card (set_mset 𝒜) = length zs›
    ‹set_mset 𝒜 = set zs› and
    nempty': ‹zs ≠ []› and
    dist_zs: ‹distinct zs›
    using vmtf_ℒ vmtf_ns_distinct[OF vmtf_ns] nempty
    unfolding vmtf_ℒall_def eq_commute[of _ ‹atms_of _›] zs_def
    by (auto simp: atms_of_ℒall_𝒜in card_Un_disjoint distinct_card)
  have hd_zs_le: ‹hd zs < length ns›
    using vmtf_ns_le_length[OF vmtf_ns, of ‹hd zs›] nempty'
    unfolding zs_def[symmetric]
    by auto
  have [refine0]: ‹
       (the x1a, A) ∈ nat_rel ⟹
       x = x2b ⟹
       f (the x1a) x2b ≤ ⇓ Id (f A x)› for x1a A x x2b
      by auto
  define iterate_over_VMTF2 where
    ‹iterate_over_VMTF2 ≡ (λf (I :: 'a ⇒ bool) (vm :: (nat, nat) vmtf_node list, n) x. do {
      let _ = remdups_mset 𝒜;
      (_, _, x) ← WHILETλ(n,m, x). I x
        (λ(n, _, _). n ≠ None)
        (λ(n, m, x). do {
          ASSERT(n ≠ None);
          let A = the n;
          ASSERT(A < length ns);
          ASSERT(A ≤ uint32_max div 2);
          x ← f A x;
          RETURN (get_next ((ns ! A)), Suc m, x)
        })
        (n, 0, x);
      RETURN x
    })›
  have iterate_over_VMTF2_alt_def:
    ‹iterate_over_VMTF2 ≡ (λf (I :: 'a ⇒ bool) (vm :: (nat, nat) vmtf_node list, n) x. do {
      (_, _, x) ← WHILETλ(n,m, x). I x
        (λ(n, _, _). n ≠ None)
        (λ(n, m, x). do {
          ASSERT(n ≠ None);
          let A = the n;
          ASSERT(A < length ns);
          ASSERT(A ≤ uint32_max div 2);
          x ← f A x;
          RETURN (get_next ((ns ! A)), Suc m, x)
        })
        (n, 0, x);
      RETURN x
    })›
    unfolding iterate_over_VMTF2_def by force
  have nempty_iff: ‹(x1 ≠ None) = (x1b ≠ {#})›
  if
    ‹(remdups_mset 𝒜, 𝒜') ∈ Id› and
    H: ‹(x, x') ∈ {((n, m, x), 𝒜', y). is_lasts 𝒜' n m ∧ x = y}› and
    ‹case x of (n, m, xa) ⇒ I xa› and
    ‹case x' of (uu_, x) ⇒ I x› and
    st[simp]:
      ‹x2 = (x1a, x2a)›
      ‹x = (x1, x2)›
      ‹x' = (x1b, xb)›
    for 𝒜' x x' x1 x2 x1a x2a x1b xb
  proof
    show ‹x1b ≠ {#}› if ‹x1 ≠ None›
      using that H
      by (auto simp: is_lasts_def)
    show ‹x1 ≠ None› if  ‹x1b ≠ {#}›
      using that H
      by (auto simp: is_lasts_def)
  qed
  have IH: ‹((get_next (ns ! the x1a), Suc x1b, xa), remove1_mset A x1, xb)
        ∈ {((n, m, x), 𝒜', y). is_lasts 𝒜' n m ∧ x = y}›
     if
      ‹(remdups_mset 𝒜, 𝒜') ∈ Id› and
      H: ‹(x, x') ∈ {((n, m, x), 𝒜', y). is_lasts 𝒜' n m ∧ x = y}› and
      ‹case x of (n, uu_, uua_) ⇒ n ≠ None› and
      nempty: ‹case x' of (ℬ, uu_) ⇒ ℬ ≠ {#}› and
      ‹case x of (n, m, xa) ⇒ I xa› and
      ‹case x' of (uu_, x) ⇒ I x› and
      st:
        ‹x' = (x1, x2)›
        ‹x2a = (x1b, x2b)›
        ‹x = (x1a, x2a)›
        ‹(xa, xb) ∈ Id› and
      ‹x1 ≠ {#}› and
      ‹x1a ≠ None› and
      A: ‹(the x1a, A) ∈ nat_rel› and
      ‹the x1a < length ns›
      for 𝒜' x x' x1 x2 x1a x2a x1b x2b A xa xb
  proof -
    have [simp]: ‹distinct_mset x1› ‹x1b < length zs›
      using H A nempty
      apply (auto simp: st is_lasts_def simp flip: Cons_nth_drop_Suc)
      apply (cases ‹x1b = length zs›)
      apply auto
      done
    then have [simp]: ‹zs ! x1b ∉ set (drop (Suc x1b) zs)›
      by (auto simp: in_set_drop_conv_nth nth_eq_iff_index_eq dist_zs)
    have [simp]: ‹length zs - Suc x1b + x1b = length zs ⟷ False›
      using ‹x1b < length zs› by presburger
    have ‹vmtf_ns (take x1b zs @ zs ! x1b # drop (Suc x1b) zs) m ns›
      using vmtf_ns
      by (auto simp: Cons_nth_drop_Suc simp flip: zs_def)
    from vmtf_ns_last_mid_get_next_option_hd[OF this]
    show ?thesis
      using H A st
      by (auto simp: st is_lasts_def dist_zs distinct_card distinct_mset_set_mset_remove1_mset
           simp flip: Cons_nth_drop_Suc)
  qed
  have WTF[simp]: ‹length zs - Suc 0 = length zs ⟷ zs = []›
    by (cases zs) auto
  have zs2: ‹set (xs' @ ys') = set zs›
    by (auto simp: zs_def)
  have is_lasts_le:  ‹is_lasts x1 (Some A) x1b ⟹ A < length ns› for x2 xb x1b x1 A
    using vmtf_ℒ le nth_mem[of ‹x1b› zs] unfolding is_lasts_def prod.case vmtf_ℒall_def
      set_append[symmetric]zs_def[symmetric] zs2
    by (auto simp: eq_commute[of ‹set zs› ‹atms_of (ℒall 𝒜)›] hd_drop_conv_nth
      simp del: nth_mem)
  have le_uint32_max: ‹the x1a ≤ uint32_max div 2›
    if
      ‹(remdups_mset 𝒜, 𝒜') ∈ Id› and
      ‹(x, x') ∈ {((n, m, x), 𝒜', y). is_lasts 𝒜' n m ∧ x = y}› and
      ‹case x of (n, uu_, uua_) ⇒ n ≠ None› and
      ‹case x' of (ℬ, uu_) ⇒ ℬ ≠ {#}› and
      ‹case x of (n, m, xa) ⇒ I xa› and
      ‹case x' of (uu_, x) ⇒ I x› and
      ‹x' = (x1, x2)› and
      ‹x2a = (x1b, xb)› and
      ‹x = (x1a, x2a)› and
      ‹x1 ≠ {#}› and
      ‹x1a ≠ None› and
      ‹(the x1a, A) ∈ nat_rel› and
      ‹the x1a < length ns›
    for 𝒜' x x' x1 x2 x1a x2a x1b xb A
  proof -
    have ‹the x1a ∈# 𝒜›
      using that by (auto simp: is_lasts_def)
    then show ?thesis
      using nempty by (auto dest!: multi_member_split simp: all_add_mset)
  qed
  have  ‹iterate_over_VMTF2 f I (ns, Some fst_As) x ≤ ⇓ Id (iterate_over_ℒall f 𝒜 I x)›
    unfolding iterate_over_VMTF2_def iterate_over_ℒall_def prod.case
    apply (refine_vcg WHILEIT_refine[where R = ‹{((n :: nat option, m::nat, x::'a), (𝒜' :: nat multiset, y)).
        is_lasts 𝒜' n m ∧ x = y}›])
    subgoal by simp
    subgoal by simp
    subgoal
      using card_𝒜 fst_As nempty nempty' hd_conv_nth[OF nempty'] hd_zs_le unfolding zs_def[symmetric]
        is_lasts_def
      by (simp_all add:  eq_commute[of ‹remdups_mset _›])
    subgoal by auto
    subgoal for 𝒜' x x' x1 x2 x1a x2a x1b xb
      by (rule nempty_iff)
    subgoal by auto
    subgoal for 𝒜' x x' x1 x2 x1a x2a x1b xb
      by (simp add: is_lasts_def in_set_dropI)
    subgoal for 𝒜' x x' x1 x2 x1a x2a x1b xb
      by (auto simp: is_lasts_le)
    subgoal by (rule le_uint32_max)
    subgoal by auto
    subgoal for 𝒜' x x' x1 x2 x1a x2a x1b x2b A xa xb
      by (rule IH)
    subgoal by auto
    done
  moreover have ‹iterate_over_VMTF f I (ns, Some fst_As) x  ≤ ⇓ Id (iterate_over_VMTF2 f I (ns, Some fst_As) x)›
    unfolding iterate_over_VMTF2_alt_def iterate_over_VMTF_def prod.case
    by (refine_vcg WHILEIT_refine[where R = ‹{((n :: nat option, x::'a), (n' :: nat option, m'::nat, x'::'a)).
        n = n' ∧ x = x'}›]) auto
  ultimately show ?thesis
    by simp
qed


definition arena_is_packed :: ‹arena ⇒ nat clauses_l ⇒ bool› where
‹arena_is_packed arena N ⟷ length arena = (∑C ∈# dom_m N. length (N ∝ C) + header_size (N ∝ C))›

lemma arena_is_packed_empty[simp]: ‹arena_is_packed [] fmempty›
  by (auto simp: arena_is_packed_def)

(*TODO Move *)
lemma sum_mset_cong:
  ‹(⋀A. A ∈# M ⟹ f A = g A) ⟹ (∑ A ∈# M. f A) = (∑ A ∈# M. g A)›
  by (induction M) auto
lemma arena_is_packed_append:
  assumes ‹arena_is_packed (arena) N› and
    [simp]: ‹length C = length (fst C') + header_size (fst C')› and
    [simp]: ‹a ∉# dom_m N›
  shows ‹arena_is_packed (arena @ C) (fmupd a C' N)›
proof -
  show ?thesis
    using assms(1) by (auto simp: arena_is_packed_def
     intro!: sum_mset_cong)
qed
(*END Move*)

lemma arena_is_packed_append_valid:
  assumes
    in_dom: ‹fst C ∈# dom_m x1a› and
    valid0: ‹valid_arena x1c x1a vdom0› and
    valid: ‹valid_arena x1d x2a (set x2d)› and
    packed: ‹arena_is_packed x1d x2a› and
    n: ‹n = header_size  (x1a ∝ (fst C))›
  shows ‹arena_is_packed
          (x1d @
           Misc.slice (fst C - n)
            (fst C + arena_length x1c (fst C)) x1c)
          (fmupd (length x1d + n) (the (fmlookup x1a (fst C))) x2a)›
proof -
  have [simp]: ‹length x1d + n ∉# dom_m x2a›
  using valid by (auto dest: arena_lifting(2) valid_arena_in_vdom_le_arena
    simp: arena_is_valid_clause_vdom_def header_size_def)
  have [simp]: ‹arena_length x1c (fst C) = length (x1a ∝ (fst C))› ‹fst C ≥ n›
      ‹fst C - n < length x1c›  ‹fst C < length x1c›
    using valid0 valid in_dom by (auto simp: arena_lifting n less_imp_diff_less)
  have [simp]: ‹length
     (Misc.slice (fst C - n)
       (fst C + length (x1a ∝ (fst C))) x1c) =
     length (x1a ∝ fst C) + header_size (x1a ∝ fst C)›
     using valid in_dom arena_lifting(10)[OF valid0]
     by (fastforce simp: slice_len_min_If min_def arena_lifting(4) simp flip: n)
  show ?thesis
    by (rule arena_is_packed_append[OF packed]) auto
qed

definition move_is_packed :: ‹arena ⇒ _ ⇒ arena ⇒ _ ⇒ bool› where
‹move_is_packed arenao No arena N ⟷
   ((∑C∈#dom_m No. length (No ∝ C) + header_size (No ∝ C)) +
   (∑C∈#dom_m N. length (N ∝ C) + header_size (N ∝ C)) ≤ length arenao)›

definition isasat_GC_clauses_prog_copy_wl_entry
  :: ‹arena ⇒ (nat watcher) list list ⇒ nat literal ⇒
         (arena × _ × _) ⇒ (arena × (arena × _ × _)) nres›
where
‹isasat_GC_clauses_prog_copy_wl_entry = (λN0 W A (N', vdm, avdm). do {
    ASSERT(nat_of_lit A < length W);
    ASSERT(length (W ! nat_of_lit A) ≤ length N0);
    let le = length (W ! nat_of_lit A);
    (i, N, N', vdm, avdm) ← WHILET
      (λ(i, N, N', vdm, avdm). i < le)
      (λ(i, N, (N', vdm, avdm)). do {
        ASSERT(i < length (W ! nat_of_lit A));
        let C = fst (W ! nat_of_lit A ! i);
        ASSERT(arena_is_valid_clause_vdom N C);
        let st = arena_status N C;
        if st ≠ DELETED then do {
          ASSERT(arena_is_valid_clause_idx N C);
          ASSERT(length N' + (if arena_length N C > 4 then 5 else 4) + arena_length N C ≤ length N0);
          ASSERT(length N = length N0);
          ASSERT(length vdm < length N0);
          ASSERT(length avdm < length N0);
          let D = length N' + (if arena_length N C > 4 then 5 else 4);
          N' ← fm_mv_clause_to_new_arena C N N';
          ASSERT(mark_garbage_pre (N, C));
	  RETURN (i+1, extra_information_mark_to_delete N C, N', vdm @ [D],
             (if st = LEARNED then avdm @ [D] else avdm))
        } else RETURN (i+1, N, (N', vdm, avdm))
      }) (0, N0, (N', vdm, avdm));
    RETURN (N, (N', vdm, avdm))
  })›

definition isasat_GC_entry :: ‹_› where
‹isasat_GC_entry 𝒜 vdom0 arena_old W'  = {((arenao, (arena, vdom, avdom)), (No, N)). valid_arena arenao No vdom0 ∧ valid_arena arena N (set vdom) ∧ vdom_m 𝒜 W' No ⊆ vdom0 ∧ dom_m N = mset vdom ∧ distinct vdom ∧
    arena_is_packed arena N ∧ mset avdom ⊆# mset vdom ∧ length arenao = length arena_old ∧
    move_is_packed arenao No arena N}›

definition isasat_GC_refl :: ‹_› where
‹isasat_GC_refl 𝒜 vdom0 arena_old = {((arenao, (arena, vdom, avdom), W), (No, N, W')). valid_arena arenao No vdom0 ∧ valid_arena arena N (set vdom) ∧
     (W, W') ∈ ⟨Id⟩map_fun_rel (D0 𝒜) ∧ vdom_m 𝒜 W' No ⊆ vdom0 ∧ dom_m N = mset vdom ∧ distinct vdom ∧
    arena_is_packed arena N ∧ mset avdom ⊆# mset vdom ∧ length arenao = length arena_old ∧
    (∀L ∈# ℒall 𝒜. length (W' L) ≤ length arenao) ∧move_is_packed arenao No arena N}›

lemma move_is_packed_empty[simp]: ‹valid_arena arena N vdom ⟹ move_is_packed arena N [] fmempty›
  by (auto simp: move_is_packed_def valid_arena_ge_length_clauses)

lemma move_is_packed_append:
  assumes
    dom: ‹C ∈# dom_m x1a› and
    E: ‹length E = length (x1a ∝ C) + header_size (x1a ∝ C)› ‹(fst E') = (x1a ∝ C)›
     ‹n = header_size (x1a ∝ C)› and
    valid: ‹valid_arena x1d x2a D'› and
    packed: ‹move_is_packed x1c x1a x1d x2a›
  shows ‹move_is_packed (extra_information_mark_to_delete x1c C)
          (fmdrop C x1a)
          (x1d @ E)
          (fmupd (length x1d + n) E' x2a)›
proof -
  have [simp]: ‹(∑x∈#remove1_mset C
          (dom_m
            x1a). length
                   (fst (the (if x = C then None
                              else fmlookup x1a x))) +
                  header_size
                   (fst (the (if x = C then None
                              else fmlookup x1a x)))) =
     (∑x∈#remove1_mset C
          (dom_m
            x1a). length
                   (x1a ∝ x) +
                  header_size
                   (x1a ∝ x))›
   by (rule sum_mset_cong)
    (use distinct_mset_dom[of x1a] in ‹auto dest!: simp: distinct_mset_remove1_All›)
  have [simp]: ‹(length x1d + header_size (x1a ∝ C)) ∉# (dom_m x2a)›
    using valid arena_lifting(2) by blast
  have [simp]: ‹(∑x∈#(dom_m x2a). length
                    (fst (the (if length x1d + header_size (x1a ∝ C) = x
                               then Some E'
                               else fmlookup x2a x))) +
                   header_size
                    (fst (the (if length x1d + header_size (x1a ∝ C) = x
                               then Some E'
                               else fmlookup x2a x)))) =
    (∑x∈#dom_m x2a. length
                    (x2a ∝ x) +
                   header_size
                    (x2a ∝ x))›
   by (rule sum_mset_cong)
    (use distinct_mset_dom[of x2a] in ‹auto dest!: simp: distinct_mset_remove1_All›)
  show ?thesis
    using packed dom E
    by (auto simp: move_is_packed_def split: if_splits dest!: multi_member_split)
qed

definition arena_header_size :: ‹arena ⇒ nat ⇒ nat› where
‹arena_header_size arena C = (if arena_length arena C > 4 then 5 else 4)›

lemma valid_arena_header_size:
  ‹valid_arena arena N vdom ⟹ C ∈# dom_m N ⟹ arena_header_size arena C = header_size (N ∝ C)›

  by (auto simp: arena_header_size_def header_size_def arena_lifting)
lemma isasat_GC_clauses_prog_copy_wl_entry:
  assumes ‹valid_arena arena N vdom0› and
    ‹valid_arena arena' N' (set vdom)› and
    vdom: ‹vdom_m 𝒜 W N ⊆ vdom0› and
    L: ‹atm_of A ∈# 𝒜› and
    L'_L: ‹(A', A) ∈ nat_lit_lit_rel› and
    W: ‹(W' , W) ∈ ⟨Id⟩map_fun_rel (D0 𝒜)› and
    ‹dom_m N' = mset vdom› ‹distinct vdom› and
   ‹arena_is_packed arena' N'› and
    avdom: ‹mset avdom ⊆# mset vdom› and
    r: ‹length arena = r› and
    le: ‹∀L ∈# ℒall 𝒜. length (W L) ≤ length arena› and
    packed: ‹move_is_packed arena N arena' N'›
  shows ‹isasat_GC_clauses_prog_copy_wl_entry arena W' A' (arena', vdom, avdom)
     ≤ ⇓ (isasat_GC_entry 𝒜 vdom0 arena W)
         (cdcl_GC_clauses_prog_copy_wl_entry N (W A) A N')›
     (is ‹_ ≤ ⇓ (?R) _›)
proof -
  have A: ‹A' = A› and K[simp]: ‹W' ! nat_of_lit A = W A›
    using L'_L L W apply auto
    by (cases A) (auto simp: map_fun_rel_def all_add_mset dest!: multi_member_split)
  have A_le: ‹nat_of_lit A < length W'›
    using W L by (cases A; auto simp: map_fun_rel_def all_add_mset dest!: multi_member_split)
  have length_slice: ‹C ∈# dom_m x1a ⟹ valid_arena x1c x1a vdom' ⟹
      length
     (Misc.slice (C - header_size (x1a ∝ C))
       (C + arena_length x1c C) x1c) =
    arena_length x1c C + header_size (x1a ∝ C)› for x1c x1a C vdom'
     using arena_lifting(1-4,10)[of x1c x1a vdom' C]
    by (auto simp: header_size_def slice_len_min_If min_def split: if_splits)
  show ?thesis
    unfolding isasat_GC_clauses_prog_copy_wl_entry_def cdcl_GC_clauses_prog_copy_wl_entry_def prod.case A
    arena_header_size_def[symmetric]
    apply (refine_vcg ASSERT_leI WHILET_refine[where R = ‹nat_rel ×r ?R›])
    subgoal using A_le by (auto simp: isasat_GC_entry_def)
    subgoal using le L K by (cases A) (auto dest!: multi_member_split simp: all_add_mset)
    subgoal using assms by (auto simp: isasat_GC_entry_def)
    subgoal using W L by auto
    subgoal by auto
    subgoal for x x' x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d
      using vdom L
      unfolding arena_is_valid_clause_vdom_def K isasat_GC_entry_def
      by (cases A)
        (force dest!: multi_member_split simp: vdom_m_def all_add_mset)+
    subgoal
      using vdom L
      unfolding arena_is_valid_clause_vdom_def K isasat_GC_entry_def
      by (subst arena_dom_status_iff)
        (cases A ; auto dest!: multi_member_split simp: arena_lifting arena_dom_status_iff
            vdom_m_def all_add_mset; fail)+
   subgoal
     unfolding arena_is_valid_clause_idx_def isasat_GC_entry_def
     by auto
   subgoal unfolding isasat_GC_entry_def move_is_packed_def arena_is_packed_def
       by (auto simp: valid_arena_header_size arena_lifting dest!: multi_member_split)
   subgoal using r by (auto simp: isasat_GC_entry_def)
   subgoal by (auto dest: valid_arena_header_size simp: arena_lifting dest!: valid_arena_vdom_subset multi_member_split simp: arena_header_size_def isasat_GC_entry_def
    split: if_splits)
    subgoal by (auto simp: isasat_GC_entry_def dest!: size_mset_mono)
   subgoal
     by (force simp: isasat_GC_entry_def dest: arena_lifting(2))
   subgoal by (auto simp: arena_header_size_def)
   subgoal for x x' x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d D
     by (rule order_trans[OF fm_mv_clause_to_new_arena])
       (auto intro: valid_arena_extra_information_mark_to_delete'
         simp: arena_lifting remove_1_mset_id_iff_notin
            mark_garbage_pre_def isasat_GC_entry_def min_def
            valid_arena_header_size
       dest: in_vdom_m_fmdropD arena_lifting(2)
       intro!: arena_is_packed_append_valid subset_mset_trans_add_mset
        move_is_packed_append length_slice)
   subgoal
     by auto
   subgoal
     by auto
   done
 qed

definition isasat_GC_clauses_prog_single_wl
  :: ‹arena ⇒  (arena × _ × _) ⇒ (nat watcher) list list ⇒ nat ⇒
        (arena × (arena × _ × _) × (nat watcher) list list) nres›
where
‹isasat_GC_clauses_prog_single_wl = (λN0 N' WS A. do {
    let L = Pos A; ⌦‹use phase saving instead›
    ASSERT(nat_of_lit L < length WS);
    ASSERT(nat_of_lit (-L) < length WS);
    (N, (N', vdom, avdom)) ← isasat_GC_clauses_prog_copy_wl_entry N0 WS L N';
    let WS = WS[nat_of_lit L := []];
    ASSERT(length N = length N0);
    (N, N') ← isasat_GC_clauses_prog_copy_wl_entry N WS (-L) (N', vdom, avdom);
    let WS = WS[nat_of_lit (-L) := []];
    RETURN (N, N', WS)
  })›


lemma isasat_GC_clauses_prog_single_wl:
  assumes
    ‹(X, X') ∈ isasat_GC_refl 𝒜 vdom0 arena0› and
    X: ‹X = (arena, (arena', vdom, avdom), W)› ‹X' = (N, N', W')› and
    L: ‹A ∈# 𝒜› and
    st: ‹(A, A') ∈ Id› and st': ‹narena = (arena', vdom, avdom)› and
    ae: ‹length arena0 = length arena› and
    le_all: ‹∀L ∈# ℒall 𝒜. length (W' L) ≤ length arena›
  shows ‹isasat_GC_clauses_prog_single_wl arena narena  W A
     ≤ ⇓ (isasat_GC_refl 𝒜 vdom0 arena0)
         (cdcl_GC_clauses_prog_single_wl N W' A' N')›
     (is ‹_ ≤ ⇓ ?R _›)
proof -
  have H:
    ‹valid_arena arena N vdom0›
    ‹valid_arena arena' N' (set vdom)› and
    vdom: ‹vdom_m 𝒜 W' N ⊆ vdom0› and
    L: ‹A ∈# 𝒜› and
    eq: ‹A' = A› and
    WW': ‹(W, W') ∈ ⟨Id⟩map_fun_rel (D0 𝒜)› and
    vdom_dom: ‹dom_m N' = mset vdom› and
    dist: ‹distinct vdom› and
    packed: ‹arena_is_packed arena' N'› and
    avdom: ‹mset avdom ⊆# mset vdom› and
    packed2: ‹move_is_packed arena N arena' N'› and
    incl: ‹vdom_m 𝒜 W' N ⊆ vdom0›
    using assms X st by (auto simp: isasat_GC_refl_def)

  have vdom2: ‹vdom_m 𝒜 W' x1 ⊆ vdom0 ⟹ vdom_m 𝒜 (W'(L := [])) x1 ⊆ vdom0› for x1 L
    by (force simp: vdom_m_def dest!: multi_member_split)
  have vdom_m_upd: ‹x ∈ vdom_m 𝒜 (W(Pos A := [], Neg A := [])) N ⟹ x ∈ vdom_m 𝒜 W N› for x W A N
    by (auto simp: image_iff vdom_m_def dest: multi_member_split)
  have vdom_m3: ‹x ∈ vdom_m 𝒜 W a ⟹ dom_m a ⊆# dom_m b ⟹ dom_m b ⊆# dom_m c ⟹x ∈ vdom_m 𝒜 W c› for a b c W x
    unfolding vdom_m_def by auto
  have W: ‹(W[2 * A := [], Suc (2 * A) := []], W'(Pos A := [], Neg A := []))
       ∈ ⟨Id⟩map_fun_rel (D0 𝒜)› for A
    using WW' unfolding map_fun_rel_def
    apply clarify
    apply (intro conjI)
    apply auto[]
    apply (drule multi_member_split)
    apply (case_tac L)
    apply (auto dest!: multi_member_split)
    done
  have le: ‹nat_of_lit (Pos A) < length W› ‹nat_of_lit (Neg A) < length W›
    using WW' L by (auto dest!: multi_member_split simp: map_fun_rel_def all_add_mset)
  have [refine0]: ‹RETURN (Pos A) ≤ ⇓ Id (RES {Pos A, Neg A})› by auto
  have vdom_upD:‹ x ∈ vdom_m 𝒜 (W'(Pos A := [], Neg A := [])) xd ⟹ x ∈  vdom_m 𝒜 (λa. if a = Pos A then [] else W' a) xd›
    for W' a A x xd
    by (auto simp: vdom_m_def)
  show ?thesis
    unfolding isasat_GC_clauses_prog_single_wl_def
      cdcl_GC_clauses_prog_single_wl_def eq st' isasat_GC_refl_def
    apply (refine_vcg
      isasat_GC_clauses_prog_copy_wl_entry[where r= ‹length arena› and 𝒜 = 𝒜])
    subgoal using le by auto
    subgoal using le by auto
    apply (rule H(1); fail)
    apply (rule H(2); fail)
    subgoal using incl by auto
    subgoal using L by auto
    subgoal using WW' by auto
    subgoal using vdom_dom by blast
    subgoal using dist by blast
    subgoal using packed by blast
    subgoal using avdom by blast
    subgoal by blast
    subgoal using le_all by auto
    subgoal using packed2 by auto
    subgoal using ae by (auto simp: isasat_GC_entry_def)
    apply (solves ‹auto simp: isasat_GC_entry_def›)
    apply (solves ‹auto simp: isasat_GC_entry_def›)
    apply (rule vdom2; auto)
    supply isasat_GC_entry_def[simp]
    subgoal using WW' by (auto simp: map_fun_rel_def dest!: multi_member_split simp: all_add_mset)
    subgoal using L by auto
    subgoal using L by auto
    subgoal using WW' by (auto simp: map_fun_rel_def dest!: multi_member_split simp: all_add_mset)
    subgoal using WW' by (auto simp: map_fun_rel_def dest!: multi_member_split simp: all_add_mset)
    subgoal using WW' le_all by (auto simp: map_fun_rel_def dest!: multi_member_split simp: all_add_mset)
    subgoal using WW' le_all by (auto simp: map_fun_rel_def dest!: multi_member_split simp: all_add_mset)
    subgoal using WW' le_all by (auto simp: map_fun_rel_def dest!: multi_member_split simp: all_add_mset)
    subgoal using WW' le_all by (auto simp: map_fun_rel_def dest!: multi_member_split simp: all_add_mset)
    subgoal using WW' le_all by (auto simp: map_fun_rel_def dest!: multi_member_split simp: all_add_mset)
    subgoal using W ae le_all vdom by (auto simp: dest!: vdom_upD)
    done
qed

definition isasat_GC_clauses_prog_wl2 where
  ‹isasat_GC_clauses_prog_wl2 ≡ (λ(ns :: (nat, nat) vmtf_node list, n) x0. do {
      (_, x) ← WHILETλ(n, x). length (fst x) = length (fst x0)
        (λ(n, _). n ≠ None)
        (λ(n, x). do {
          ASSERT(n ≠ None);
          let A = the n;
          ASSERT(A < length ns);
          ASSERT(A ≤ uint32_max div 2);
          x ← (λ(arenao, arena, W). isasat_GC_clauses_prog_single_wl arenao arena W A) x;
          RETURN (get_next ((ns ! A)), x)
        })
        (n, x0);
      RETURN x
    })›

definition cdcl_GC_clauses_prog_wl2  where
  ‹cdcl_GC_clauses_prog_wl2 = (λN0 𝒜0 WS. do {
    𝒜 ← SPEC(λ𝒜. set_mset 𝒜 = set_mset 𝒜0);
    (_, (N, N', WS)) ← WHILETcdcl_GC_clauses_prog_wl_inv 𝒜 N0
      (λ(ℬ, _). ℬ ≠ {#})
      (λ(ℬ, (N, N', WS)). do {
        ASSERT(ℬ ≠ {#});
        A ← SPEC (λA. A ∈# ℬ);
        (N, N', WS) ← cdcl_GC_clauses_prog_single_wl N WS A N';
        RETURN (remove1_mset A ℬ, (N, N', WS))
      })
      (𝒜, (N0, fmempty, WS));
    RETURN (N, N', WS)
  })›


lemma WHILEIT_refine_with_invariant_and_break:
  assumes R0: "I' x' ⟹ (x,x')∈R"
  assumes IREF: "⋀x x'. ⟦ (x,x')∈R; I' x' ⟧ ⟹ I x"
  assumes COND_REF: "⋀x x'. ⟦ (x,x')∈R; I x; I' x' ⟧ ⟹ b x = b' x'"
  assumes STEP_REF:
    "⋀x x'. ⟦ (x,x')∈R; b x; b' x'; I x; I' x' ⟧ ⟹ f x ≤ ⇓R (f' x')"
  shows "WHILEIT I b f x ≤⇓{(x, x'). (x, x') ∈ R ∧ I x ∧  I' x' ∧ ¬b' x'} (WHILEIT I' b' f' x')"
  (is ‹_ ≤ ⇓?R' _›)
    apply (subst (2)WHILEIT_add_post_condition)
    apply (refine_vcg WHILEIT_refine_genR[where R'=R and R = ?R'])
    subgoal by (auto intro: assms)[]
    subgoal by (auto intro: assms)[]
    subgoal using COND_REF by (auto)
    subgoal by (auto intro: assms)[]
    subgoal by (auto intro: assms)[]
    done

lemma cdcl_GC_clauses_prog_wl_inv_cong_empty:
  ‹set_mset 𝒜 = set_mset ℬ ⟹
  cdcl_GC_clauses_prog_wl_inv 𝒜 N ({#}, x) ⟹ cdcl_GC_clauses_prog_wl_inv ℬ N ({#}, x)›
  by (auto simp: cdcl_GC_clauses_prog_wl_inv_def)

lemma isasat_GC_clauses_prog_wl2:
  assumes ‹valid_arena arenao No vdom0› and
    ‹valid_arena arena N (set vdom)› and
    vdom: ‹vdom_m 𝒜 W' No ⊆ vdom0› and
    vmtf: ‹((ns, m, n, lst_As1, next_search1), to_remove1) ∈ vmtf 𝒜 M› and
    nempty: ‹𝒜 ≠ {#}› and
    W_W': ‹(W, W') ∈ ⟨Id⟩map_fun_rel (D0 𝒜)› and
    bounded: ‹isasat_input_bounded 𝒜› and old: ‹old_arena = []› and
    le_all: ‹∀L ∈# ℒall 𝒜. length (W' L) ≤ length arenao
 shows
    ‹isasat_GC_clauses_prog_wl2 (ns, Some n) (arenao, (old_arena, [], []), W)
        ≤ ⇓ ({((arenao', (arena, vdom, avdom), W), (No', N, W')). valid_arena arenao' No' vdom0 ∧
                valid_arena arena N (set vdom) ∧
       (W, W') ∈ ⟨Id⟩map_fun_rel (D0 𝒜) ∧ vdom_m 𝒜 W' No' ⊆ vdom0 ∧
        cdcl_GC_clauses_prog_wl_inv 𝒜 No ({#}, No', N, W') ∧ dom_m N = mset vdom ∧ distinct vdom ∧
        arena_is_packed arena N ∧ mset avdom ⊆# mset vdom ∧ length arenao' = length arenao})
         (cdcl_GC_clauses_prog_wl2 No 𝒜 W')›
proof -
  define f where
    ‹f A ≡ (λ(arenao, arena, W). isasat_GC_clauses_prog_single_wl arenao arena W A)› for A :: nat
  let ?R = ‹{((𝒜', arenao', (arena, vdom), W), (𝒜'', No', N, W')). 𝒜' = 𝒜'' ∧
      ((arenao', (arena, vdom), W), (No', N, W')) ∈ isasat_GC_refl 𝒜 vdom0 arenao ∧
      length arenao' = length arenao}›
  have H: ‹(X, X') ∈ ?R ⟹ X = (x1, x2) ⟹ x2 = (x3, x4) ⟹ x4 = (x5, x6) ⟹
     X' = (x1', x2') ⟹ x2' = (x3', x4') ⟹ x4' = (x5', x6') ⟹
     ((x3, (fst x5, fst (snd x5), snd (snd x5)), x6), (x3', x5', x6')) ∈ isasat_GC_refl 𝒜 vdom0 arenao
    for X X' A B x1 x1' x2 x2' x3 x3' x4 x4' x5 x5' x6 x6' x0 x0' x x'
     supply [[show_types]]
    by auto
  have isasat_GC_clauses_prog_wl_alt_def:
    ‹isasat_GC_clauses_prog_wl2 n x0 = iterate_over_VMTF f (λx. length (fst x) = length (fst x0)) n x0›
    for n x0
    unfolding f_def isasat_GC_clauses_prog_wl2_def iterate_over_VMTF_def by (cases n) (auto intro!: ext)
  show ?thesis
    unfolding isasat_GC_clauses_prog_wl_alt_def prod.case f_def[symmetric] old
    apply (rule order_trans[OF iterate_over_VMTF_iterate_over_ℒall[OF vmtf nempty bounded]])
    unfolding Down_id_eq iterate_over_ℒall_def cdcl_GC_clauses_prog_wl2_def f_def
    apply (refine_vcg WHILEIT_refine_with_invariant_and_break[where R = ?R]
            isasat_GC_clauses_prog_single_wl)
    subgoal by fast
    subgoal using assms by (auto simp: valid_arena_empty isasat_GC_refl_def)
    subgoal by auto
    subgoal by auto
    subgoal by auto
    subgoal by auto
    apply (rule H; assumption; fail)
    apply (rule refl)+
    subgoal by (auto simp add: cdcl_GC_clauses_prog_wl_inv_def)
    subgoal by auto
    subgoal by auto
    subgoal using le_all by (auto simp: isasat_GC_refl_def split: prod.splits)
    subgoal by (auto simp: isasat_GC_refl_def)
    subgoal by (auto simp: isasat_GC_refl_def
      dest: cdcl_GC_clauses_prog_wl_inv_cong_empty)
    done
qed


lemma cdcl_GC_clauses_prog_wl_alt_def:
  ‹cdcl_GC_clauses_prog_wl = (λ(M, N0, D, NE, UE, NS, US, Q, WS). do {
    ASSERT(cdcl_GC_clauses_pre_wl (M, N0, D, NE, UE, NS, US, Q, WS));
    (N, N', WS) ← cdcl_GC_clauses_prog_wl2 N0 (all_init_atms N0 (NE+NS)) WS;
    RETURN (M, N', D, NE, UE, NS, US, Q, WS)
     })›
 proof -
   have [refine0]: ‹(x1c, x1) ∈ Id ⟹ RES (set_mset x1c)
       ≤ ⇓ Id  (RES (set_mset x1))› for x1 x1c
     by auto
   have [refine0]: ‹(xa, x') ∈ Id ⟹
       x2a = (x1b, x2b) ⟹
       x2 = (x1a, x2a) ⟹
       x' = (x1, x2) ⟹
       x2d = (x1e, x2e) ⟹
       x2c = (x1d, x2d) ⟹
       xa = (x1c, x2c) ⟹
       (A, Aa) ∈ Id ⟹
       cdcl_GC_clauses_prog_single_wl x1d x2e A x1e
       ≤ ⇓ Id
          (cdcl_GC_clauses_prog_single_wl x1a x2b Aa x1b)›
      for 𝒜 x xa x' x1 x2 x1a x2a x1b x2b x1c x2c x1d x2d x1e x2e A aaa Aa
      by auto
   show ?thesis
     unfolding cdcl_GC_clauses_prog_wl_def cdcl_GC_clauses_prog_wl2_def
       while.imonad3

     apply (intro ext)
     apply (clarsimp simp add: while.imonad3)
     apply (subst order_class.eq_iff[of ‹(_ :: _ nres)›])
     apply (intro conjI)
     subgoal
       by (rewrite at ‹_ ≤ ⌑› Down_id_eq[symmetric]) (refine_rcg WHILEIT_refine[where R = Id], auto)
     subgoal
       by (rewrite at ‹_ ≤ ⌑› Down_id_eq[symmetric]) (refine_rcg WHILEIT_refine[where R = Id], auto)
     done
qed

definition isasat_GC_clauses_prog_wl :: ‹twl_st_wl_heur ⇒ twl_st_wl_heur nres› where
  ‹isasat_GC_clauses_prog_wl = (λ(M', N', D', j, W', ((ns, st, fst_As, lst_As, nxt), to_remove), clvls, cach, lbd, outl, stats,
    heur,  vdom, avdom, lcount, opts, old_arena). do {
    ASSERT(old_arena = []);
    (N, (N', vdom, avdom), WS) ← isasat_GC_clauses_prog_wl2 (ns, Some fst_As) (N', (old_arena, take 0 vdom, take 0 avdom), W');
    RETURN (M', N', D', j, WS, ((ns, st, fst_As, lst_As, nxt), to_remove), clvls, cach, lbd, outl, incr_GC stats, set_zero_wasted heur,
       vdom, avdom, lcount, opts, take 0 N)
  })›

lemma length_watched_le'':
  assumes
    xb_x'a: ‹(x1a, x1) ∈ twl_st_heur_restart› and
    prop_inv: ‹correct_watching'' x1›
  shows ‹∀x2 ∈# ℒall (all_init_atms_st x1). length (watched_by x1 x2) ≤ length (get_clauses_wl_heur x1a)›
proof
  fix x2
  assume x2: ‹x2 ∈# ℒall (all_init_atms_st x1)›
  have ‹correct_watching'' x1›
    using prop_inv unfolding unit_propagation_outer_loop_wl_inv_def
      unit_propagation_outer_loop_wl_inv_def
    by auto
  then have dist: ‹distinct_watched (watched_by x1 x2)›
    using x2
    by (cases x1; auto simp: all_all_init_atms correct_watching''.simps
      simp flip: all_init_lits_def all_init_lits_alt_def)
  then have dist: ‹distinct_watched (watched_by x1 x2)›
    using xb_x'a
    by (cases x1; auto simp: all_atm_of_all_lits_of_mm correct_watching.simps)
  have dist_vdom: ‹distinct (get_vdom x1a)›
    using xb_x'a
    by (cases x1)
      (auto simp: twl_st_heur_restart_def)
  have x2: ‹x2 ∈# ℒall (all_init_atms_st x1)›
    using x2 xb_x'a unfolding all_init_atms_def all_init_lits_def
    by auto

  have
      valid: ‹valid_arena (get_clauses_wl_heur x1a) (get_clauses_wl x1) (set (get_vdom x1a))›
    using xb_x'a unfolding all_atms_def all_lits_def
    by (cases x1)
     (auto simp: twl_st_heur_restart_def)

  have ‹vdom_m (all_init_atms_st x1) (get_watched_wl x1) (get_clauses_wl x1) ⊆ set (get_vdom x1a)›
    using xb_x'a
    by (cases x1)
      (auto simp: twl_st_heur_restart_def all_atms_def[symmetric])
  then have subset: ‹set (map fst (watched_by x1 x2)) ⊆ set (get_vdom x1a)›
    using x2 unfolding vdom_m_def
    by (cases x1)
      (force simp: twl_st_heur_restart_def simp flip: all_init_atms_def
        dest!: multi_member_split)
  have watched_incl: ‹mset (map fst (watched_by x1 x2)) ⊆# mset (get_vdom x1a)›
    by (rule distinct_subseteq_iff[THEN iffD1])
      (use dist[unfolded distinct_watched_alt_def] dist_vdom subset in
         ‹simp_all flip: distinct_mset_mset_distinct›)
  have vdom_incl: ‹set (get_vdom x1a) ⊆ {4..< length (get_clauses_wl_heur x1a)}›
    using valid_arena_in_vdom_le_arena[OF valid] arena_dom_status_iff[OF valid] by auto

  have ‹length (get_vdom x1a) ≤ length (get_clauses_wl_heur x1a)›
    by (subst distinct_card[OF dist_vdom, symmetric])
      (use card_mono[OF _ vdom_incl] in auto)
  then show ‹length (watched_by x1 x2) ≤ length (get_clauses_wl_heur x1a)›
    using size_mset_mono[OF watched_incl] xb_x'a
    by (auto intro!: order_trans[of ‹length (watched_by x1 x2)› ‹length (get_vdom x1a)›])
qed

lemma isasat_GC_clauses_prog_wl:
  ‹(isasat_GC_clauses_prog_wl, cdcl_GC_clauses_prog_wl) ∈
   twl_st_heur_restart →f
   ⟨{(S, T). (S, T) ∈ twl_st_heur_restart ∧ arena_is_packed (get_clauses_wl_heur S) (get_clauses_wl T)}⟩nres_rel›
  (is ‹_ ∈ ?T →f _›)
proof-
  have [refine0]: ‹⋀x1 x1a x1b x1c x1d x1e x2e x1f x1g x1h x1i x1j x1m x1n x1o x1p x2n x2o x1q
       x1r x1s x1t x1u x1v x1w x1x x1y x1z x1aa x1ab x2ab NS US.
       ((x1f, x1g, x1h, x1i, x1j, ((x1m, x1n, x1o, x1p, x2n), x2o), x1q,
         x1s, x1t, x1w, x1x, x1y, x1z, x1aa, x1ab, x2ab),
        x1, x1a, x1b, x1c, x1d, NS, US, x1e, x2e)
       ∈ ?T ⟹
       valid_arena x1g x1a (set x1z)›
     unfolding twl_st_heur_restart_def
     by auto
  have [refine0]: ‹⋀x1 x1a x1b x1c x1d x1e x2e x1f x1g x1h x1i x1j x1m x1n x1o x1p x2n x2o x1q
       x1r x1s x1t x1u x1v x1w x1x x1y x1z x1aa x1ab x2ab NS US.
       ((x1f, x1g, x1h, x1i, x1j, ((x1m, x1n, x1o, x1p, x2n), x2o), x1q,
         x1s, x1t, x1w, x1x, x1y, x1z, x1aa, x1ab, x2ab),
        x1, x1a, x1b, x1c, x1d, NS, US, x1e, x2e)
       ∈ ?T ⟹
       isasat_input_bounded (all_init_atms x1a (x1c + NS))›
     unfolding twl_st_heur_restart_def
     by auto
  have [refine0]: ‹⋀x1 x1a x1b x1c x1d x1e x2e x1f x1g x1h x1i x1j x1m x1n x1o x1p x2n x2o x1q
       x1r x1s x1t x1u x1v x1w x1x x1y x1z x1aa x1ab x2ab NS US.
       ((x1f, x1g, x1h, x1i, x1j, ((x1m, x1n, x1o, x1p, x2n), x2o), x1q,
         x1s, x1t, x1w, x1x, x1y, x1z, x1aa, x1ab, x2ab),
        x1, x1a, x1b, x1c, x1d, NS, US, x1e, x2e)
       ∈ ?T ⟹
       vdom_m (all_init_atms x1a (x1c+NS)) x2e x1a ⊆ set x1z›
     unfolding twl_st_heur_restart_def
     by auto
  have [refine0]: ‹⋀x1 x1a x1b x1c x1d x1e x2e x1f x1g x1h x1i x1j x1m x1n x1o x1p x2n x2o x1q
       x1r x1s x1t x1u x1v x1w x1x x1y x1z x1aa x1ab x2ab NS US.
       ((x1f, x1g, x1h, x1i, x1j, ((x1m, x1n, x1o, x1p, x2n), x2o), x1q, x1r,
         x1s, x1t, x1w, x1x, x1y, x1z, x1aa, x1ab, x2ab),
        x1, x1a, x1b, x1c, x1d, NS, US, x1e, x2e)
       ∈ ?T ⟹
       all_init_atms x1a (x1c+NS) ≠ {#}›
     unfolding twl_st_heur_restart_def
     by auto
  have [refine0]: ‹⋀x1 x1a x1b x1c x1d x1e x2e x1f x1g x1h x1i x1j x1m x1n x1o x1p x2n x2o x1q
       x1r x1s x1t x1u x1v x1w x1x x1y x1z x1aa x1ab x2ab NS US.
       ((x1f, x1g, x1h, x1i, x1j, ((x1m, x1n, x1o, x1p, x2n), x2o), x1q,
         x1s, x1t, x1w, x1x, x1y, x1z, x1aa, x1ab, x2ab),
        x1, x1a, x1b, x1c, x1d, NS, US, x1e, x2e)
       ∈ ?T ⟹
       ((x1m, x1n, x1o, x1p, x2n), set (fst x2o)) ∈ vmtf (all_init_atms x1a (x1c+NS)) x1›
       ‹⋀x1 x1a x1b x1c x1d x1e x2e x1f x1g x1h x1i x1j x1m x1n x1o x1p x2n x2o x1q
       x1r x1s x1t x1u x1v x1w x1x x1y x1z x1aa x1ab x2ab NS US.
       ((x1f, x1g, x1h, x1i, x1j, ((x1m, x1n, x1o, x1p, x2n), x2o), x1q,
         x1s, x1t, x1w, x1x, x1y, x1z, x1aa, x1ab, x2ab),
        x1, x1a, x1b, x1c, x1d, NS, US, x1e, x2e)
       ∈ ?T ⟹ (x1j, x2e) ∈ ⟨Id⟩map_fun_rel (D0 (all_init_atms x1a (x1c+NS)))›
     unfolding twl_st_heur_restart_def isa_vmtf_def distinct_atoms_rel_def distinct_hash_atoms_rel_def
     by auto
  have H: ‹vdom_m (all_init_atms x1a x1c) x2ad x1ad ⊆ set x2af›
    if
       empty: ‹∀A∈#all_init_atms x1a x1c. x2ad (Pos A) = [] ∧ x2ad (Neg A) = []› and
      rem: ‹GC_remap** (x1a, Map.empty, fmempty) (fmempty, m, x1ad)› and
      ‹dom_m x1ad = mset x2af›
    for m :: ‹nat ⇒ nat option› and y :: ‹nat literal multiset› and x :: ‹nat› and
      x1 x1a x1b x1c x1d x1e x2e x1f x1g x1h x1i x1j x1m x1n x1o x1p x2n x2o x1q
         x1r x1s x1t x1u x1v x1w x1x x1y x1z x1aa x1ab x2ab x1ac x1ad x2ad x1ae
         x1ag x2af x2ag
  proof -
    have ‹xa ∈# ℒall (all_init_atms x1a x1c) ⟹ x2ad xa = []› for xa
      using empty by (cases xa) (auto simp: in_ℒall_atm_of_𝒜in)
    then show ?thesis
      using ‹dom_m x1ad = mset x2af›
      by (auto simp: vdom_m_def)
  qed
  have H': ‹mset x2ag ⊆# mset x1ah ⟹ x ∈ set x2ag ⟹ x ∈ set x1ah› for x2ag x1ah x
    by (auto dest: mset_eq_setD)
  show ?thesis
    supply [[goals_limit=1]]
    unfolding isasat_GC_clauses_prog_wl_def cdcl_GC_clauses_prog_wl_alt_def take_0
    apply (intro frefI nres_relI)
    apply (refine_vcg isasat_GC_clauses_prog_wl2[where 𝒜 = ‹all_init_atms _ _›]; remove_dummy_vars)
    subgoal
      by (clarsimp simp add: twl_st_heur_restart_def
        cdcl_GC_clauses_prog_wl_inv_def H H'
        rtranclp_GC_remap_all_init_atms
        rtranclp_GC_remap_learned_clss_l)
    subgoal
      unfolding cdcl_GC_clauses_pre_wl_def
      by (drule length_watched_le'')
        (clarsimp_all simp add: twl_st_heur_restart_def
          cdcl_GC_clauses_prog_wl_inv_def H H'
          rtranclp_GC_remap_all_init_atms
         rtranclp_GC_remap_learned_clss_l)
    subgoal
      by (clarsimp simp add: twl_st_heur_restart_def
        cdcl_GC_clauses_prog_wl_inv_def H H'
        rtranclp_GC_remap_all_init_atms
        rtranclp_GC_remap_learned_clss_l)
    done
qed

definition cdcl_remap_st :: ‹'v twl_st_wl ⇒ 'v twl_st_wl nres› where
‹cdcl_remap_st = (λ(M, N0, D, NE, UE, NS, US, Q, WS).
  SPEC (λ(M', N', D', NE', UE', NS', US', Q', WS').
         (M', D', NE', UE', NS', US', Q') = (M, D, NE, UE, NS, US, Q) ∧
         (∃m. GC_remap** (N0, (λ_. None), fmempty) (fmempty, m, N')) ∧
         0 ∉# dom_m N'))›

definition rewatch_spec :: ‹nat twl_st_wl ⇒ nat twl_st_wl nres› where
‹rewatch_spec = (λ(M, N, D, NE, UE, NS, US, Q, WS).
  SPEC (λ(M', N', D', NE', UE', NS', US', Q', WS').
     (M', N', D', NE', UE', NS', US', Q') = (M, N, D, NE, UE, NS, {#}, Q) ∧
     correct_watching' (M, N', D, NE, UE, NS', US, Q', WS') ∧
     literals_are_ℒin' (M, N', D, NE, UE, NS', US, Q', WS')))›

lemma blits_in_ℒin'_restart_wl_spec0':
  ‹literals_are_ℒin' (a, aq, ab, ac, ad, ae, af, Q, b) ⟹
       literals_are_ℒin' (a, aq, ab, ac, ad, ae, af, {#}, b)›
  by (auto simp: literals_are_ℒin'_empty blits_in_ℒin'_restart_wl_spec0)

lemma cdcl_GC_clauses_wl_D_alt_def:
  ‹cdcl_GC_clauses_wl = (λS. do {
    ASSERT(cdcl_GC_clauses_pre_wl S);
    let b = True;
    if b then do {
      S ← cdcl_remap_st S;
      S ← rewatch_spec S;
      RETURN S
    }
    else remove_all_learned_subsumed_clauses_wl S})›
  supply [[goals_limit=1]]
  unfolding cdcl_GC_clauses_wl_def
  by (fastforce intro!: ext simp: RES_RES_RETURN_RES2 cdcl_remap_st_def
      RES_RES9_RETURN_RES uncurry_def image_iff cdcl_remap_st_def
      RES_RETURN_RES_RES2 RES_RETURN_RES RES_RES2_RETURN_RES rewatch_spec_def
      rewatch_spec_def remove_all_learned_subsumed_clauses_wl_def
      literals_are_ℒin'_empty blits_in_ℒin'_restart_wl_spec0'
    intro!: bind_cong_nres intro: literals_are_ℒin'_empty(4))

definition isasat_GC_clauses_pre_wl_D :: ‹twl_st_wl_heur ⇒ bool› where
‹isasat_GC_clauses_pre_wl_D S ⟷ (
  ∃T. (S, T) ∈ twl_st_heur_restart ∧ cdcl_GC_clauses_pre_wl T
  )›


definition isasat_GC_clauses_wl_D :: ‹twl_st_wl_heur ⇒ twl_st_wl_heur nres› where
‹isasat_GC_clauses_wl_D = (λS. do {
  ASSERT(isasat_GC_clauses_pre_wl_D S);
  let b = True;
  if b then do {
    T ← isasat_GC_clauses_prog_wl S;
    ASSERT(length (get_clauses_wl_heur T) ≤ length (get_clauses_wl_heur S));
    ASSERT(∀i ∈ set (get_vdom T). i < length (get_clauses_wl_heur S));
    U ← rewatch_heur_st T;
    RETURN U
  }
  else RETURN S})›


lemma cdcl_GC_clauses_prog_wl2_st:
  assumes ‹(T, S) ∈ state_wl_l None›
  ‹correct_watching'' T ∧ cdcl_GC_clauses_pre S ∧
   set_mset (dom_m (get_clauses_wl T)) ⊆ clauses_pointed_to
      (Neg ` set_mset (all_init_atms_st T) ∪
       Pos ` set_mset (all_init_atms_st T))
       (get_watched_wl T) ∧
    literals_are_ℒin' T› and
    ‹get_clauses_wl T = N0'›
  shows
    ‹cdcl_GC_clauses_prog_wl T ≤
       ⇓ {((M', N'', D', NE', UE', NS', US', Q', WS'), (N, N')).
       (M', D', NE', UE', NS', US', Q') = (get_trail_wl T, get_conflict_wl T, get_unit_init_clss_wl T,
           get_unit_learned_clss_wl T, get_subsumed_init_clauses_wl T, get_subsumed_learned_clauses_wl T,
           literals_to_update_wl T) ∧ N'' = N ∧
           (∀L∈#all_init_lits_st T. WS' L = []) ∧
           all_init_lits_st T = all_init_lits N (NE'+NS') ∧
           (∃m. GC_remap** (get_clauses_wl T, Map.empty, fmempty)
               (fmempty, m, N))}
      (SPEC(λ(N'::(nat, 'a literal list × bool) fmap, m).
         GC_remap** (N0', (λ_. None), fmempty) (fmempty, m, N') ∧
	  0 ∉# dom_m N'))›
   using cdcl_GC_clauses_prog_wl2[of ‹get_trail_wl T› ‹get_clauses_wl T› ‹get_conflict_wl T›
     ‹get_unit_init_clss_wl T› ‹get_unit_learned_clss_wl T› ‹get_subsumed_init_clauses_wl T›
     ‹get_subsumed_learned_clauses_wl T› ‹literals_to_update_wl T›
     ‹get_watched_wl T› S N0'] assms
   by (cases T) auto

lemma correct_watching''_clauses_pointed_to:
  assumes
    xa_xb: ‹(xa, xb) ∈ state_wl_l None› and
    corr: ‹correct_watching'' xa› and
    pre: ‹cdcl_GC_clauses_pre xb› and
    L: ‹literals_are_ℒin' xa›
  shows ‹set_mset (dom_m (get_clauses_wl xa))
         ⊆ clauses_pointed_to
            (Neg `
             set_mset
              (all_init_atms_st xa) ∪
             Pos `
             set_mset
              (all_init_atms_st xa))
            (get_watched_wl xa)›
        (is ‹_ ⊆ ?A›)
proof
  let ?𝒜 = ‹all_init_atms (get_clauses_wl xa) (get_unit_init_clss_wl xa)›
  fix C
  assume C: ‹C ∈# dom_m (get_clauses_wl xa)›
  obtain M N D NE UE NS US Q W where
    xa: ‹xa = (M, N, D, NE, UE, NS, US, Q, W)›
    by (cases xa)
  obtain x where
    xb_x: ‹(xb, x) ∈ twl_st_l None› and
    ‹twl_list_invs xb› and
    struct_invs: ‹twl_struct_invs x› and
    ‹get_conflict_l xb = None› and
    ‹clauses_to_update_l xb = {#}› and
    ‹count_decided (get_trail_l xb) = 0› and
    ‹∀L∈set (get_trail_l xb). mark_of L = 0›
    using pre unfolding cdcl_GC_clauses_pre_def by fast
  have ‹twl_st_inv x›
    using xb_x C struct_invs
    by (auto simp: twl_struct_invs_def
      cdclW_restart_mset.cdclW_all_struct_inv_def)
  then have le0: ‹get_clauses_wl xa ∝ C ≠ []›
    using xb_x C xa_xb
    by (cases x; cases ‹irred N C›)
      (auto simp: twl_struct_invs_def twl_st_inv.simps
        twl_st_l_def state_wl_l_def xa ran_m_def conj_disj_distribR
        Collect_disj_eq Collect_conv_if
      dest!: multi_member_split)
  then have le: ‹N ∝ C ! 0 ∈ set (watched_l (N ∝ C))›
    by (cases ‹N ∝ C›) (auto simp: xa)
  have eq: ‹set_mset (ℒall (all_init_atms N NE)) =
        set_mset (all_lits_of_mm (mset `# init_clss_lf N + NE))›
     by (auto simp del: all_init_atms_def[symmetric]
        simp: all_init_atms_def xa all_atm_of_all_lits_of_mm[symmetric]
          all_init_lits_def)

  have H: ‹get_clauses_wl xa ∝ C ! 0 ∈# all_init_lits_st xa›
    using L C le0 apply -
    unfolding all_init_atms_def[symmetric] all_init_lits_def[symmetric]
    apply (subst literals_are_ℒin'_literals_are_ℒin_iff(4)[OF xa_xb xb_x struct_invs])
    apply (cases ‹N ∝ C›; auto simp: literals_are_ℒin_def all_lits_def ran_m_def eq
          all_lits_of_mm_add_mset is_ℒall_def xa all_lits_of_m_add_mset
          all_all_atms_all_lits
        dest!: multi_member_split)
    done

  moreover {
    have ‹{#i ∈# fst `# mset (W (N ∝ C ! 0)). i ∈# dom_m N#} =
           add_mset C {#Ca ∈# remove1_mset C (dom_m N). N ∝ C ! 0 ∈ set (watched_l (N ∝ Ca))#}›
      using corr H C le unfolding xa
      by (auto simp: clauses_pointed_to_def correct_watching''.simps xa
        simp flip: all_init_atms_def all_init_lits_def all_init_atms_alt_def
          all_init_lits_alt_def
        simp: clause_to_update_def
        simp del: all_init_atms_def[symmetric]
        dest!: multi_member_split)
    from arg_cong[OF this, of set_mset] have ‹C ∈ fst ` set (W (N ∝ C ! 0))›
      using corr H C le unfolding xa
      by (auto simp: clauses_pointed_to_def correct_watching''.simps xa
        simp: all_init_atms_def all_init_lits_def clause_to_update_def
        simp del: all_init_atms_def[symmetric]
        dest!: multi_member_split) }
  ultimately show ‹C ∈ ?A›
    by (cases ‹N ∝ C ! 0›)
      (auto simp: clauses_pointed_to_def correct_watching''.simps xa
        simp flip: all_init_lits_def all_init_atms_alt_def
          all_init_lits_alt_def
        simp: clause_to_update_def all_init_atms_def
        simp del: all_init_atms_def[symmetric]
      dest!: multi_member_split)
qed

abbreviation isasat_GC_clauses_rel where
  ‹isasat_GC_clauses_rel y ≡ {(S, T). (S, T) ∈ twl_st_heur_restart ∧
           (∀L∈#all_init_lits_st y. get_watched_wl T L = [])∧
           get_trail_wl T = get_trail_wl y ∧
           get_conflict_wl T = get_conflict_wl y ∧
           get_unit_init_clss_wl T = get_unit_init_clss_wl y ∧
           get_unit_learned_clss_wl T = get_unit_learned_clss_wl y ∧
           get_subsumed_init_clauses_wl T = get_subsumed_init_clauses_wl y ∧
           get_subsumed_learned_clauses_wl T = get_subsumed_learned_clauses_wl y ∧
           (∃m. GC_remap** (get_clauses_wl y, (λ_. None), fmempty) (fmempty, m, get_clauses_wl T)) ∧
           arena_is_packed (get_clauses_wl_heur S) (get_clauses_wl T)}›

lemma ref_two_step'': ‹R ⊆ R' ⟹ A ≤ B ⟹ ⇓ R A ≤  ⇓ R' B›
  by (simp add: "weaken_⇓" ref_two_step')

lemma isasat_GC_clauses_prog_wl_cdcl_remap_st:
  assumes
    ‹(x, y) ∈ twl_st_heur_restart''' r› and
    ‹cdcl_GC_clauses_pre_wl y›
  shows ‹isasat_GC_clauses_prog_wl x ≤ ⇓ (isasat_GC_clauses_rel y) (cdcl_remap_st y)›
proof -
  have xy: ‹(x, y) ∈ twl_st_heur_restart›
    using assms(1) by fast
  have H: ‹isasat_GC_clauses_rel y =
    {(S, T). (S, T) ∈ twl_st_heur_restart ∧ arena_is_packed (get_clauses_wl_heur S) (get_clauses_wl T)} O
    {(S, T). S = T ∧ (∀L∈#all_init_lits_st y. get_watched_wl T L = [])∧
           get_trail_wl T = get_trail_wl y ∧
           get_conflict_wl T = get_conflict_wl y ∧
           get_unit_init_clss_wl T = get_unit_init_clss_wl y ∧
           get_unit_learned_clss_wl T = get_unit_learned_clss_wl y ∧
           get_subsumed_init_clauses_wl T = get_subsumed_init_clauses_wl y ∧
           get_subsumed_learned_clauses_wl T = get_subsumed_learned_clauses_wl y ∧
           (∃m. GC_remap** (get_clauses_wl y, (λ_. None), fmempty) (fmempty, m, get_clauses_wl T))}›
    by blast
  show ?thesis
    using assms apply -
    apply (rule order_trans[OF isasat_GC_clauses_prog_wl[THEN fref_to_Down]])
    subgoal by fast
    apply (rule xy)
    unfolding conc_fun_chain[symmetric] H
    apply (rule ref_two_step')
    unfolding cdcl_GC_clauses_pre_wl_D_def cdcl_GC_clauses_pre_wl_def
    apply normalize_goal+
    apply (rule order_trans[OF cdcl_GC_clauses_prog_wl2_st])
    apply assumption
    subgoal for xa
      using assms(2) by (simp add: correct_watching''_clauses_pointed_to
        cdcl_GC_clauses_pre_wl_def)
    apply (rule refl)
    subgoal by (auto simp: cdcl_remap_st_def conc_fun_RES split: prod.splits)
    done
qed

fun correct_watching''' :: ‹_ ⇒ 'v twl_st_wl ⇒ bool› where
  ‹correct_watching''' 𝒜 (M, N, D, NE, UE, NS, US, Q, W) ⟷
    (∀L ∈# all_lits_of_mm 𝒜.
       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)) ∧
        fst `# mset (W L) = clause_to_update L (M, N, D, NE, UE, NS, US, {#}, {#}))›

declare correct_watching'''.simps[simp del]

lemma correct_watching'''_add_clause:
  assumes
    corr: ‹correct_watching''' 𝒜 ((a, aa, CD, ac, ad, NS, US, Q, b))› and
    leC: ‹2 ≤ length C› and
    i_notin[simp]: ‹i ∉# dom_m aa› and
    dist[iff]: ‹C ! 0 ≠ C ! Suc 0›
  shows ‹correct_watching''' 𝒜
          ((a, fmupd i (C, red) aa, CD, ac, ad, NS, US, Q, b
            (C ! 0 := b (C ! 0) @ [(i, C ! Suc 0, length C = 2)],
             C ! Suc 0 := b (C ! Suc 0) @ [(i, C ! 0, length C = 2)])))›
proof -
  have [iff]: ‹C ! Suc 0 ≠ C ! 0›
    using  ‹C ! 0 ≠ C ! Suc 0› by argo
  have [iff]: ‹C ! Suc 0 ∈# all_lits_of_m (mset C)› ‹C ! 0 ∈# all_lits_of_m (mset C)›
    ‹C ! Suc 0 ∈ set C› ‹ C ! 0 ∈ set C› ‹C ! 0 ∈ set (watched_l C)› ‹C ! Suc 0 ∈ set (watched_l C)›
    using leC by (force intro!: in_clause_in_all_lits_of_m nth_mem simp: in_set_conv_iff
        intro: exI[of _ 0] exI[of _ ‹Suc 0›])+
  have [dest!]: ‹⋀L. L ≠ C ! 0 ⟹ L ≠ C ! Suc 0 ⟹ L ∈ set (watched_l C) ⟹ False›
     by (cases C; cases ‹tl C›; auto)+
  have i: ‹i ∉ fst ` set (b L)› if ‹L∈#all_lits_of_mm 𝒜›for L
    using corr i_notin that unfolding correct_watching'''.simps
    by force
  have [iff]: ‹(i,c, d) ∉ set (b L)› if ‹L∈#all_lits_of_mm 𝒜› for L c d
    using i[of L, OF that] by (auto simp: image_iff)
  then show ?thesis
    using corr
    by (force simp: correct_watching'''.simps ran_m_mapsto_upd_notin
        all_lits_of_mm_add_mset all_lits_of_mm_union clause_to_update_mapsto_upd_notin correctly_marked_as_binary.simps
        split: if_splits)
qed


lemma rewatch_correctness:
  assumes empty: ‹⋀L. L ∈# all_lits_of_mm 𝒜 ⟹ W L = []› and
    H[dest]: ‹⋀x. x ∈# dom_m N ⟹ distinct (N ∝ x) ∧ length (N ∝ x) ≥ 2› and
    incl: ‹set_mset (all_lits_of_mm (mset `# ran_mf N)) ⊆ set_mset (all_lits_of_mm 𝒜)›
  shows
    ‹rewatch N W ≤ SPEC(λW. correct_watching''' 𝒜 (M, N, C, NE, UE, NS, US, Q, W))›
proof -
  define I where
    ‹I ≡ λ(a :: nat list) (b :: nat list) W.
        correct_watching''' 𝒜 ((M, fmrestrict_set (set a) N, C, NE, UE, NS, US, Q, W))›
  have I0: ‹set_mset (dom_m N) ⊆ set x ∧ distinct x ⟹ I [] x W› for x
    using empty unfolding I_def by (auto simp: correct_watching'''.simps
       all_blits_are_in_problem_init.simps clause_to_update_def
       all_lits_of_mm_union)
  have le: ‹length (σ L) < size (dom_m N)›
     if ‹correct_watching''' 𝒜 (M, fmrestrict_set (set l1) N, C, NE, UE, NS, US, Q, σ)› and
      ‹set_mset (dom_m N) ⊆ set x ∧ distinct x› and
     ‹x = l1 @ xa # l2› ‹xa ∈# dom_m N› ‹L ∈ set (N ∝ xa)›
     for L l1 σ xa l2 x
  proof -
    have 1: ‹card (set l1) ≤ length l1›
      by (auto simp: card_length)
    have ‹L ∈# all_lits_of_mm 𝒜›
      using that incl in_clause_in_all_lits_of_m[of L ‹mset (N ∝ xa)›]
      by (auto simp: correct_watching'''.simps dom_m_fmrestrict_set' ran_m_def
          all_lits_of_mm_add_mset all_lits_of_m_add_mset atm_of_all_lits_of_m
          in_all_lits_of_mm_ain_atms_of_iff
        dest!: multi_member_split)
    then have ‹distinct_watched (σ L)› and ‹fst ` set (σ L) ⊆ set l1 ∩ set_mset (dom_m N)›
      using that incl
      by (auto simp: correct_watching'''.simps dom_m_fmrestrict_set' dest!: multi_member_split)
    then have ‹length (map fst (σ L)) ≤ card (set l1 ∩ set_mset (dom_m N))›
      using 1 by (subst distinct_card[symmetric])
       (auto simp: distinct_watched_alt_def intro!: card_mono intro: order_trans)
    also have ‹... < card (set_mset (dom_m N))›
      using that by (auto intro!: psubset_card_mono)
    also have ‹... = size (dom_m N)›
      by (simp add: distinct_mset_dom distinct_mset_size_eq_card)
    finally show ?thesis by simp
  qed
  show ?thesis
    unfolding rewatch_def
    apply (refine_vcg
      nfoldli_rule[where I = ‹I›])
    subgoal by (rule I0)
    subgoal using assms unfolding I_def by auto
    subgoal for x xa l1 l2 σ  using H[of xa] unfolding I_def apply -
      by (rule, subst (asm)nth_eq_iff_index_eq)
        linarith+
    subgoal for x xa l1 l2 σ unfolding I_def by (rule le) (auto intro!: nth_mem)
    subgoal for x xa l1 l2 σ unfolding I_def by (drule le[where L = ‹N ∝ xa ! 1›]) (auto simp: I_def dest!: le)
    subgoal for x xa l1 l2 σ
      unfolding I_def
      by (cases ‹the (fmlookup N xa)›)
       (auto intro!: correct_watching'''_add_clause simp: dom_m_fmrestrict_set')
    subgoal
      unfolding I_def
      by auto
    subgoal by auto
    subgoal unfolding I_def
      by (auto simp: fmlookup_restrict_set_id')
    done
qed

inductive_cases GC_remapE: ‹GC_remap (a, aa, b) (ab, ac, ba)›
lemma rtranclp_GC_remap_ran_m_remap:
  ‹GC_remap** (old, m, new) (old', m', new')  ⟹ C ∈# dom_m old ⟹ C ∉# dom_m old' ⟹
         m' C ≠ None ∧
         fmlookup new' (the (m' C)) = fmlookup old C›
  apply (induction rule: rtranclp_induct[of r ‹(_, _, _)› ‹(_, _, _)›, split_format(complete), of for r])
  subgoal by auto
  subgoal for a aa b ab ac ba
    apply (cases ‹C ∉# dom_m a›)
    apply (auto dest: GC_remap_ran_m_remap GC_remap_ran_m_no_rewrite_map
       GC_remap_ran_m_no_rewrite)
    apply (metis GC_remap_ran_m_no_rewrite_fmap GC_remap_ran_m_no_rewrite_map in_dom_m_lookup_iff option.sel)
    using GC_remap_ran_m_remap rtranclp_GC_remap_ran_m_no_rewrite by fastforce
  done

lemma GC_remap_ran_m_exists_earlier:
  ‹GC_remap (old, m, new) (old', m', new')  ⟹ C ∈# dom_m new' ⟹ C ∉# dom_m new ⟹
         ∃D. m' D = Some C ∧ D ∈# dom_m old ∧
         fmlookup new' C = fmlookup old D›
  by (induction rule: GC_remap.induct[split_format(complete)]) auto


lemma rtranclp_GC_remap_ran_m_exists_earlier:
  ‹GC_remap** (old, m, new) (old', m', new')  ⟹ C ∈# dom_m new' ⟹ C ∉# dom_m new ⟹
         ∃D. m' D = Some C ∧ D ∈# dom_m old ∧
         fmlookup new' C = fmlookup old D›
  apply (induction rule: rtranclp_induct[of r ‹(_, _, _)› ‹(_, _, _)›, split_format(complete), of for r])
  apply (auto dest: GC_remap_ran_m_exists_earlier)
  apply (case_tac ‹C ∈# dom_m b›)
  apply (auto elim!: GC_remapE split: if_splits)
  apply blast
  using rtranclp_GC_remap_ran_m_no_new_map rtranclp_GC_remap_ran_m_no_rewrite
  by (metis fst_conv)

(*TODO Move + romove dup*)
lemma all_all_init_atms_all_init_lits:
  ‹set_mset (ℒall (all_init_atms N NE)) = set_mset (all_init_lits N NE)›
  unfolding all_all_init_atms ..

lemma rewatch_heur_st_correct_watching:
  assumes
    pre: ‹cdcl_GC_clauses_pre_wl y› and
    S_T: ‹(S, T) ∈ isasat_GC_clauses_rel y›
  shows ‹rewatch_heur_st S ≤ ⇓ (twl_st_heur_restart''' (length (get_clauses_wl_heur S)))
    (rewatch_spec T)›
proof -
  obtain M N D NE UE NS US Q W where
    T: ‹T = (M, N, D, NE, UE, NS, US, Q, W)›
    by (cases T) auto

  obtain M' N' D' j W' vm clvls cach lbd outl stats fast_ema slow_ema ccount
       vdom avdom lcount opts where
    S: ‹S = (M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, (fast_ema, slow_ema, ccount),
       vdom, avdom, lcount, opts)›
    by (cases S) auto

  have
    valid: ‹valid_arena N' N (set vdom)› and
    dist: ‹distinct vdom› and
    dom_m_vdom: ‹set_mset (dom_m N) ⊆ set vdom› and
    W: ‹(W', W) ∈ ⟨Id⟩map_fun_rel (D0 (all_init_atms_st T))› and
    empty: ‹⋀L. L ∈# all_init_lits_st y ⟹ W L = []› and
    NUE:‹get_unit_init_clss_wl y = NE ›
      ‹get_unit_learned_clss_wl y = UE›
      ‹get_trail_wl y = M›
      ‹get_subsumed_init_clauses_wl y = NS›
      ‹get_subsumed_learned_clauses_wl y = US›
    using assms by (auto simp: twl_st_heur_restart_def S T)
  obtain m where
    m: ‹GC_remap** (get_clauses_wl y, Map.empty, fmempty)
             (fmempty, m, N)›
    using assms by (auto simp: twl_st_heur_restart_def S T)
  obtain x xa xb where
    y_x: ‹(y, x) ∈ Id› ‹x = y› and
    lits_y: ‹literals_are_ℒin' y› and
    x_xa: ‹(x, xa) ∈ state_wl_l None› and
    ‹correct_watching'' x› and
    xa_xb: ‹(xa, xb) ∈ twl_st_l None› and
    ‹twl_list_invs xa› and
    struct_invs: ‹twl_struct_invs xb› and
    ‹get_conflict_l xa = None› and
    ‹clauses_to_update_l xa = {#}› and
    ‹count_decided (get_trail_l xa) = 0› and
    ‹∀L∈set (get_trail_l xa). mark_of L = 0›
    using pre
    unfolding cdcl_GC_clauses_pre_wl_def
      cdcl_GC_clauses_pre_def
    by blast
  have [iff]:
    ‹distinct_mset (mset (watched_l C) + mset (unwatched_l C)) ⟷ distinct C› for C
    unfolding mset_append[symmetric]
    by auto

  have ‹twl_st_inv xb›
    using xa_xb struct_invs
    by (auto simp: twl_struct_invs_def
      cdclW_restart_mset.cdclW_all_struct_inv_def)
  then have A:
    ‹⋀C. C ∈# dom_m (get_clauses_wl x) ⟹ distinct (get_clauses_wl x ∝ C) ∧ 2 ≤ length (get_clauses_wl x ∝ C)›
    using xa_xb x_xa
    by (cases x; cases ‹irred (get_clauses_wl x) C›)
      (auto simp: twl_struct_invs_def twl_st_inv.simps
        twl_st_l_def state_wl_l_def ran_m_def conj_disj_distribR
        Collect_disj_eq Collect_conv_if
      dest!: multi_member_split
      split: if_splits)
  have struct_wf:
    ‹C ∈# dom_m N ⟹ distinct (N ∝ C) ∧ 2 ≤ length (N ∝ C)› for C
    using rtranclp_GC_remap_ran_m_exists_earlier[OF m, of ‹C›] A y_x
    by (auto simp: T dest: )

  have eq_UnD: ‹A = A' ∪ A'' ⟹ A' ⊆ A› for A A' A''
      by blast

  have eq3: ‹all_init_lits (get_clauses_wl y) (NE+NS) = all_init_lits N (NE+NS)›
    using rtranclp_GC_remap_init_clss_l_old_new[OF m]
    by (auto simp: all_init_lits_def)
  moreover have ‹all_lits_st y = all_lits_st T›
    using rtranclp_GC_remap_init_clss_l_old_new[OF m] rtranclp_GC_remap_learned_clss_l_old_new[OF m]
    apply (auto simp: all_init_lits_def T NUE all_lits_def)
    by (metis NUE(1) NUE(2) all_clss_l_ran_m all_lits_def get_unit_clauses_wl_alt_def)
  ultimately have lits: ‹literals_are_in_ℒin_mm (all_init_atms N (NE+NS)) (mset `# ran_mf N)›
    using literals_are_ℒin'_literals_are_ℒin_iff(3)[OF x_xa xa_xb struct_invs] lits_y
      rtranclp_GC_remap_init_clss_l_old_new[OF m]
      rtranclp_GC_remap_learned_clss_l_old_new[OF m]
    by (auto simp: literals_are_in_ℒin_mm_def all_all_init_atms_all_init_lits
      y_x literals_are_ℒin'_def literals_are_ℒin_def all_lits_def[of N] T
      get_unit_clauses_wl_alt_def all_lits_def atm_of_eq_atm_of
      is_ℒall_def NUE all_init_atms_def all_init_lits_def all_atms_def conj_disj_distribR
      in_all_lits_of_mm_ain_atms_of_iff atms_of_ms_def atm_of_all_lits_of_mm
      ex_disj_distrib Collect_disj_eq atms_of_def all_atm_of_all_lits_of_mm
      dest!: multi_member_split[of _ ‹ran_m _›]
      split: if_splits
      simp del: all_init_atms_def[symmetric] all_atms_def[symmetric])

  have eq: ‹set_mset (ℒall (all_init_atms N (NE+NS))) = set_mset (all_init_lits_st y)›
    using rtranclp_GC_remap_init_clss_l_old_new[OF m]
    by (auto simp: T all_init_lits_def NUE
      all_all_init_atms_all_init_lits)
  then have vd: ‹vdom_m (all_init_atms N (NE+NS)) W N ⊆ set_mset (dom_m N)›
    using empty dom_m_vdom
    by (auto simp: vdom_m_def)
  have ‹{#i ∈# clause_to_update L (M, N, get_conflict_wl y, NE, UE, NS, US, {#}, {#}).
         i ∈# dom_m N#} =
       {#i ∈# clause_to_update L (M, N, get_conflict_wl y, NE, UE, NS, US, {#}, {#}).
         True#}› for L
       by (rule filter_mset_cong2)  (auto simp: clause_to_update_def)
  then have corr2: ‹correct_watching'''
        ({#mset (fst x). x ∈# init_clss_l (get_clauses_wl y)#} + NE + NS)
        (M, N, get_conflict_wl y, NE, UE, NS, US, Q, W'a) ⟹
       correct_watching' (M, N, get_conflict_wl y, NE, UE,NS, US, Q, W'a)› for W'a
     using rtranclp_GC_remap_init_clss_l_old_new[OF m]
     by (auto simp: correct_watching'''.simps correct_watching'.simps)
  have eq2: ‹all_init_lits (get_clauses_wl y) (NE+NS) = all_init_lits N (NE+NS)›
    using rtranclp_GC_remap_init_clss_l_old_new[OF m]
    by (auto simp: T all_init_lits_def NUE
      all_all_init_atms_all_init_lits)
  have ‹i ∈# dom_m N ⟹ set (N ∝ i) ⊆ set_mset (all_init_lits N (NE+NS))› for i
    using lits by (auto dest!: multi_member_split split_list
      simp: literals_are_in_ℒin_mm_def ran_m_def
        all_lits_of_mm_add_mset all_lits_of_m_add_mset
        all_all_init_atms_all_init_lits)
  then have blit2: ‹correct_watching'''
        ({#mset x. x ∈# init_clss_lf (get_clauses_wl y)#} + NE + NS)
        (M, N, get_conflict_wl y, NE, UE, NS, US, Q, W'a) ⟹
        blits_in_ℒin' (M, N, get_conflict_wl y, NE, UE, NS, US, Q, W'a)› for W'a
      using rtranclp_GC_remap_init_clss_l_old_new[OF m]
      unfolding  correct_watching'''.simps  blits_in_ℒin'_def eq2
        all_all_init_atms_all_init_lits all_init_lits_alt_def[symmetric]
      by (fastforce simp: correct_watching'''.simps  blits_in_ℒin'_def
        simp: eq all_all_init_atms eq2
        dest!: multi_member_split[of _ ‹all_init_lits N (NE+NS)›]
        dest: mset_eq_setD)
  have ‹correct_watching'''
        ({#mset x. x ∈# init_clss_lf (get_clauses_wl y)#} + (NE + NS))
        (M, N, get_conflict_wl y, NE, UE, NS, US, Q, W'a) ⟹
        vdom_m (all_init_atms N (NE+NS)) W'a N ⊆ set_mset (dom_m N)› for W'a
      unfolding  correct_watching'''.simps  blits_in_ℒin'_def
        all_all_init_atms_all_init_lits all_init_lits_def[symmetric]
        all_init_lits_alt_def[symmetric]
      using eq eq3
      by (force simp: correct_watching'''.simps vdom_m_def NUE
        all_all_init_atms)
  then have st: ‹(x, W'a) ∈ ⟨Id⟩map_fun_rel (D0 (all_init_atms N (NE+NS))) ⟹
     correct_watching'''
        ({#mset x. x ∈# init_clss_lf (get_clauses_wl y)#} + NE + NS)
        (M, N, get_conflict_wl y, NE, UE, NS, US, Q, W'a) ⟹
      ((M', N', D', j, x, vm, clvls, cach, lbd, outl, stats, (fast_ema,
         slow_ema, ccount), vdom, avdom, lcount, opts),
        M, N, get_conflict_wl y, NE, UE, NS, {#}, Q, W'a)
       ∈ twl_st_heur_restart› for W'a m x
       using S_T dom_m_vdom
       by (auto simp: S T twl_st_heur_restart_def y_x NUE ac_simps)
  have truc: ‹xa ∈# all_lits_of_mm ({#mset (fst x). x ∈# learned_clss_l N#} + (UE + US)) ⟹
       xa ∈# all_lits_of_mm ({#mset (fst x). x ∈# init_clss_l N#} + (NE + NS))› for xa
    using lits_y eq3 rtranclp_GC_remap_learned_clss_l[OF m]
    unfolding literals_are_ℒin'_def all_init_lits_def NUE
      all_lits_of_mm_union all_init_lits_def  all_all_init_atms_all_init_lits
    by auto

  show ?thesis
    supply [[goals_limit=1]]
    using assms
    unfolding rewatch_heur_st_def T S
    apply clarify
    apply (rule ASSERT_leI)
    subgoal by (auto dest!: valid_arena_vdom_subset simp: twl_st_heur_restart_def)
    apply (rule bind_refine_res)
    prefer 2
    apply (rule order.trans)
    apply (rule rewatch_heur_rewatch[OF valid _ dist dom_m_vdom W[unfolded T, simplified] lits])
    apply (solves simp)
    apply (rule vd)
    apply (rule order_trans[OF ref_two_step'])
    apply (rule rewatch_correctness[where M=M and N=N and NE=NE and UE=UE and C=D and Q=Q and
        NS=NS and US=US])
    apply (rule empty[unfolded all_init_lits_def]; assumption)
    apply (rule struct_wf; assumption)
    subgoal using lits eq2 by (auto simp: literals_are_in_ℒin_mm_def all_init_atms_def all_init_lits_def
         all_atm_of_all_lits_of_mm
       simp del: all_init_atms_def[symmetric])
    apply (subst conc_fun_RES)
    apply (rule order.refl)
    by (fastforce simp: rewatch_spec_def RETURN_RES_refine_iff NUE
        literals_are_in_ℒin_mm_def literals_are_ℒin'_def add.assoc
      intro: corr2 blit2 st dest: truc)
qed

lemma GC_remap_dom_m_subset:
  ‹GC_remap (old, m, new) (old', m', new') ⟹ dom_m old' ⊆# dom_m old›
  by (induction rule: GC_remap.induct[split_format(complete)]) (auto dest!: multi_member_split)

lemma rtranclp_GC_remap_dom_m_subset:
  ‹rtranclp GC_remap (old, m, new) (old', m', new') ⟹ dom_m old' ⊆# dom_m old›
  apply (induction rule: rtranclp_induct[of r ‹(_, _, _)› ‹(_, _, _)›, split_format(complete), of for r])
  subgoal by auto
  subgoal for old1 m1 new1 old2 m2 new2
    using GC_remap_dom_m_subset[of old1 m1 new1 old2 m2 new2] by auto
  done

lemma GC_remap_mapping_unchanged:
  ‹GC_remap (old, m, new) (old', m', new') ⟹ C ∈ dom m ⟹ m' C = m C›
  by (induction rule: GC_remap.induct[split_format(complete)]) auto

lemma rtranclp_GC_remap_mapping_unchanged:
  ‹GC_remap** (old, m, new) (old', m', new') ⟹ C ∈ dom m ⟹ m' C = m C›
  apply (induction rule: rtranclp_induct[of r ‹(_, _, _)› ‹(_, _, _)›, split_format(complete), of for r])
  subgoal by auto
  subgoal for old1 m1 new1 old2 m2 new2
    using GC_remap_mapping_unchanged[of old1 m1 new1 old2 m2 new2, of C]
    by (auto dest: GC_remap_mapping_unchanged simp: dom_def intro!: image_mset_cong2)
  done


lemma GC_remap_mapping_dom_extended:
  ‹GC_remap (old, m, new) (old', m', new') ⟹ dom m' = dom m ∪ set_mset (dom_m old - dom_m old')›
  by (induction rule: GC_remap.induct[split_format(complete)]) (auto dest!: multi_member_split)

lemma rtranclp_GC_remap_mapping_dom_extended:
  ‹GC_remap** (old, m, new) (old', m', new') ⟹ dom m' = dom m ∪ set_mset (dom_m old - dom_m old')›
  apply (induction rule: rtranclp_induct[of r ‹(_, _, _)› ‹(_, _, _)›, split_format(complete), of for r])
  subgoal by auto
  subgoal for old1 m1 new1 old2 m2 new2
    using GC_remap_mapping_dom_extended[of old1 m1 new1 old2 m2 new2]
    GC_remap_dom_m_subset[of old1 m1 new1 old2 m2 new2]
    rtranclp_GC_remap_dom_m_subset[of old m new old1 m1 new1]
    by (auto dest: GC_remap_mapping_dom_extended simp: dom_def mset_subset_eq_exists_conv)
  done

lemma GC_remap_dom_m:
  ‹GC_remap (old, m, new) (old', m', new') ⟹ dom_m new' = dom_m new + the `# m' `# (dom_m old - dom_m old')›
  by (induction rule: GC_remap.induct[split_format(complete)]) (auto dest!: multi_member_split)

lemma rtranclp_GC_remap_dom_m:
  ‹rtranclp GC_remap (old, m, new) (old', m', new') ⟹ dom_m new' = dom_m new + the `# m' `# (dom_m old - dom_m old')›
  apply (induction rule: rtranclp_induct[of r ‹(_, _, _)› ‹(_, _, _)›, split_format(complete), of for r])
  subgoal by auto
  subgoal for old1 m1 new1 old2 m2 new2
    using GC_remap_dom_m[of old1 m1 new1 old2 m2 new2] GC_remap_dom_m_subset[of old1 m1 new1 old2 m2 new2]
    rtranclp_GC_remap_dom_m_subset[of old m new old1 m1 new1]
    GC_remap_mapping_unchanged[of old1 m1 new1 old2 m2 new2]
    rtranclp_GC_remap_mapping_dom_extended[of old m new old1 m1 new1]
    by (auto dest: simp: mset_subset_eq_exists_conv intro!: image_mset_cong2)
  done

lemma isasat_GC_clauses_rel_packed_le:
  assumes
    xy: ‹(x, y) ∈ twl_st_heur_restart''' r› and
    ST: ‹(S, T) ∈ isasat_GC_clauses_rel y›
  shows ‹length (get_clauses_wl_heur S) ≤ length (get_clauses_wl_heur x)› and
     ‹∀C ∈ set (get_vdom S). C < length (get_clauses_wl_heur x)›
proof -
  obtain m where
    ‹(S, T) ∈ twl_st_heur_restart› and
    ‹∀L∈#all_init_lits_st y. get_watched_wl T L = []› and
    ‹get_trail_wl T = get_trail_wl y› and
    ‹get_conflict_wl T = get_conflict_wl y› and
    ‹get_unit_init_clss_wl T = get_unit_init_clss_wl y› and
    ‹get_unit_learned_clss_wl T = get_unit_learned_clss_wl y› and
    remap: ‹GC_remap** (get_clauses_wl y, Map.empty, fmempty)
      (fmempty, m, get_clauses_wl T)› and
    packed: ‹arena_is_packed (get_clauses_wl_heur S) (get_clauses_wl T)›
     using ST by auto
  have ‹valid_arena (get_clauses_wl_heur x) (get_clauses_wl y) (set (get_vdom x))›
    using xy unfolding twl_st_heur_restart_def by (cases x; cases y) auto
  from valid_arena_ge_length_clauses[OF this]
  have ‹(∑C∈#dom_m (get_clauses_wl  y). length (get_clauses_wl y ∝ C) +
              header_size (get_clauses_wl y ∝ C)) ≤ length (get_clauses_wl_heur x)›
   (is ‹?A ≤ _›) .
  moreover have ‹?A = (∑C∈#dom_m (get_clauses_wl T). length (get_clauses_wl T ∝ C) +
              header_size (get_clauses_wl T ∝ C))›
    using rtranclp_GC_remap_ran_m_remap[OF remap]
    by (auto simp: rtranclp_GC_remap_dom_m[OF remap] intro!: sum_mset_cong)
  ultimately show le: ‹length (get_clauses_wl_heur S) ≤ length (get_clauses_wl_heur x)›
    using packed unfolding arena_is_packed_def by simp

  have ‹valid_arena (get_clauses_wl_heur S) (get_clauses_wl T) (set (get_vdom S))›
    using ST unfolding twl_st_heur_restart_def by (cases S; cases T) auto
  then show ‹∀C ∈ set (get_vdom S). C < length (get_clauses_wl_heur x)›
    using le
    by (auto dest: valid_arena_in_vdom_le_arena)
qed

lemma isasat_GC_clauses_wl_D:
  ‹(isasat_GC_clauses_wl_D, cdcl_GC_clauses_wl)
    ∈ twl_st_heur_restart''' r →f ⟨twl_st_heur_restart'''' r⟩nres_rel›
  unfolding isasat_GC_clauses_wl_D_def cdcl_GC_clauses_wl_D_alt_def
  apply (intro frefI nres_relI)
  apply (refine_vcg isasat_GC_clauses_prog_wl_cdcl_remap_st[where r=r]
    rewatch_heur_st_correct_watching)
  subgoal unfolding isasat_GC_clauses_pre_wl_D_def by blast
  subgoal by fast
  subgoal by (rule isasat_GC_clauses_rel_packed_le)
  subgoal by (rule isasat_GC_clauses_rel_packed_le(2))
  apply assumption+
  subgoal by (auto)
  subgoal by (auto)
  done



definition cdcl_twl_full_restart_wl_D_GC_heur_prog where
‹cdcl_twl_full_restart_wl_D_GC_heur_prog S0 = do {
    S ← do {
      if count_decided_st_heur S0 > 0
      then do {
        S ← find_decomp_wl_st_int 0 S0;
        empty_Q S
      } else RETURN S0
    };
    ASSERT(length (get_clauses_wl_heur S) = length (get_clauses_wl_heur S0));
    T ← remove_one_annot_true_clause_imp_wl_D_heur S;
    ASSERT(length (get_clauses_wl_heur T) = length (get_clauses_wl_heur S0));
    U ← mark_to_delete_clauses_wl_D_heur T;
    ASSERT(length (get_clauses_wl_heur U) = length (get_clauses_wl_heur S0));
    V ← isasat_GC_clauses_wl_D U;
    RETURN V
  }›

lemma
    cdcl_twl_full_restart_wl_GC_prog_pre_heur:
      ‹cdcl_twl_full_restart_wl_GC_prog_pre T ⟹
        (S, T) ∈ twl_st_heur''' r ⟷ (S, T) ∈ twl_st_heur_restart_ana r› (is ‹_ ⟹ _ ?A›) and
     cdcl_twl_full_restart_wl_D_GC_prog_post_heur:
       ‹cdcl_twl_full_restart_wl_GC_prog_post S0 T ⟹
        (S, T) ∈ twl_st_heur ⟷ (S, T) ∈ twl_st_heur_restart›  (is ‹_ ⟹ _ ?B›)
proof -
  note cong = trail_pol_cong heuristic_rel_cong
      option_lookup_clause_rel_cong D0_cong isa_vmtf_cong phase_saving_cong
      cach_refinement_empty_cong vdom_m_cong isasat_input_nempty_cong
      isasat_input_bounded_cong

  show ‹cdcl_twl_full_restart_wl_GC_prog_pre T ⟹ ?A›
    supply [[goals_limit=1]]
    unfolding cdcl_twl_full_restart_wl_GC_prog_pre_def cdcl_twl_full_restart_l_GC_prog_pre_def
    apply normalize_goal+
    apply (rule iffI)
    subgoal for U V
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T U V]
        cong[of ‹all_atms_st T› ‹all_init_atms_st T›]
	vdom_m_cong[of ‹all_atms_st T› ‹all_init_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
      apply -
      apply (simp_all del: isasat_input_nempty_def isasat_input_bounded_def)
      apply (cases S; cases T)
      by (simp add: twl_st_heur_def twl_st_heur_restart_ana_def
        twl_st_heur_restart_def del: isasat_input_nempty_def)
    subgoal for U V
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T U V]
        cong[of ‹all_init_atms_st T› ‹all_atms_st T›]
	vdom_m_cong[of ‹all_init_atms_st T› ‹all_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
      apply -
      by (cases S; cases T)
         (simp add: twl_st_heur_def twl_st_heur_restart_ana_def
        twl_st_heur_restart_def del: isasat_input_nempty_def)
    done
  show ‹cdcl_twl_full_restart_wl_GC_prog_post S0 T ⟹ ?B›
    supply [[goals_limit=1]]
    unfolding cdcl_twl_full_restart_wl_GC_prog_post_def
       cdcl_twl_full_restart_wl_GC_prog_post_def
       cdcl_twl_full_restart_l_GC_prog_pre_def
    apply normalize_goal+
    subgoal for S0' T' S0''
    apply (rule iffI)
    subgoal
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T T']
        cong[of ‹all_atms_st T› ‹all_init_atms_st T›]
	vdom_m_cong[of ‹all_atms_st T› ‹all_init_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
        cdcl_twl_restart_l_invs[of S0' S0'' T']
      apply -
      apply (clarsimp simp del: isasat_input_nempty_def isasat_input_bounded_def)
      apply (cases S; cases T; cases T')
      apply (simp add: twl_st_heur_def twl_st_heur_restart_def del: isasat_input_nempty_def)
      using isa_vmtf_cong option_lookup_clause_rel_cong trail_pol_cong heuristic_rel_cong
      by presburger
    subgoal
      using literals_are_ℒin'_literals_are_ℒin_iff(3)[of T T']
        cong[of ‹all_init_atms_st T› ‹all_atms_st T›]
	vdom_m_cong[of ‹all_init_atms_st T› ‹all_atms_st T› ‹get_watched_wl T› ‹get_clauses_wl T›]
        cdcl_twl_restart_l_invs[of S0' S0'' T']
      apply -
      apply (cases S; cases T)
      by (clarsimp simp add: twl_st_heur_def twl_st_heur_restart_def
        simp del: isasat_input_nempty_def)
    done
    done

qed

lemma cdcl_twl_full_restart_wl_D_GC_heur_prog:
  ‹(cdcl_twl_full_restart_wl_D_GC_heur_prog, cdcl_twl_full_restart_wl_GC_prog) ∈
    twl_st_heur''' r →f ⟨twl_st_heur'''' r⟩nres_rel›
  unfolding cdcl_twl_full_restart_wl_D_GC_heur_prog_def
    cdcl_twl_full_restart_wl_GC_prog_def
  apply (intro frefI nres_relI)
  apply (refine_rcg cdcl_twl_local_restart_wl_spec0
    remove_one_annot_true_clause_imp_wl_D_heur_remove_one_annot_true_clause_imp_wl_D[where r=r, THEN fref_to_Down]
    mark_to_delete_clauses_wl_D_heur_mark_to_delete_clauses_wl2_D[where r=r, THEN fref_to_Down]
    isasat_GC_clauses_wl_D[where r=r, THEN fref_to_Down])
  apply (subst (asm) cdcl_twl_full_restart_wl_GC_prog_pre_heur, assumption)
  apply assumption
  subgoal
    unfolding cdcl_twl_full_restart_wl_GC_prog_pre_def
      cdcl_twl_full_restart_l_GC_prog_pre_def
    by normalize_goal+ auto
  subgoal by (auto simp: twl_st_heur_restart_ana_def)
  apply assumption
  subgoal by (auto simp: twl_st_heur_restart_ana_def)
  subgoal by (auto simp: twl_st_heur_restart_ana_def)
  subgoal by (auto simp: twl_st_heur_restart_ana_def)
  subgoal for x y
    by (blast dest: cdcl_twl_full_restart_wl_D_GC_prog_post_heur)
  done


definition end_of_restart_phase :: ‹restart_heuristics ⇒ 64 word› where
  ‹end_of_restart_phase = (λ(_, _, (restart_phase,_ ,_ , end_of_phase, _), _).
    end_of_phase)›

definition end_of_restart_phase_st :: ‹twl_st_wl_heur ⇒ 64 word› where
  ‹end_of_restart_phase_st = (λ(M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur,
       vdom, avdom, lcount, opts, old_arena).
      end_of_restart_phase heur)›


definition end_of_rephasing_phase_st :: ‹twl_st_wl_heur ⇒ 64 word› where
  ‹end_of_rephasing_phase_st = (λ(M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur,
       vdom, avdom, lcount, opts, old_arena).
      end_of_rephasing_phase_heur heur)›


text ‹Using \<^term>‹a + 1› ensures that we do not get stuck with 0.›
fun incr_restart_phase_end :: ‹restart_heuristics ⇒ restart_heuristics› where
  ‹incr_restart_phase_end (fast_ema, slow_ema, (ccount, ema_lvl, restart_phase, end_of_phase, length_phase), wasted) =
    (fast_ema, slow_ema, (ccount, ema_lvl, restart_phase, end_of_phase + length_phase, (length_phase * 3) >> 1), wasted)›

definition update_restart_phases :: ‹twl_st_wl_heur ⇒ twl_st_wl_heur nres› where
  ‹update_restart_phases = (λ(M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur,
       vdom, avdom, lcount, opts, old_arena). do {
     heur ← RETURN (incr_restart_phase heur);
     heur ← RETURN (incr_restart_phase_end heur);
     RETURN (M', N', D', j, W', vm, clvls, cach, lbd, outl, stats, heur,
         vdom, avdom, lcount, opts, old_arena)
  })›

definition update_all_phases :: ‹twl_st_wl_heur ⇒ nat ⇒ (twl_st_wl_heur × nat) nres› where
  ‹update_all_phases = (λS n. do {
     let lcount = get_learned_count S;
     end_of_restart_phase ← RETURN (end_of_restart_phase_st S);
     S ← (if end_of_restart_phase > of_nat lcount then RETURN S else update_restart_phases S);
     S ← (if end_of_rephasing_phase_st S > of_nat lcount then RETURN S else rephase_heur_st S);
     RETURN (S, n)
  })›


definition restart_prog_wl_D_heur
  :: "twl_st_wl_heur ⇒ nat ⇒ bool ⇒ (twl_st_wl_heur × nat) nres"
where
  ‹restart_prog_wl_D_heur S n brk = do {
    b ← restart_required_heur S n;
    if ¬brk ∧ b = FLAG_GC_restart
    then do {
       T ← cdcl_twl_full_restart_wl_D_GC_heur_prog S;
       RETURN (T, n+1)
    }
    else if ¬brk ∧ b = FLAG_restart
    then do {
       T ← cdcl_twl_restart_wl_heur S;
       RETURN (T, n+1)
    }
    else update_all_phases S n
  }›

lemma restart_required_heur_restart_required_wl:
  ‹(uncurry restart_required_heur, uncurry restart_required_wl) ∈
    twl_st_heur ×f nat_rel →f ⟨restart_flag_rel⟩nres_rel›
    unfolding restart_required_heur_def restart_required_wl_def uncurry_def Let_def
      restart_flag_rel_def FLAG_GC_restart_def FLAG_restart_def FLAG_no_restart_def
      GC_required_heur_def
    by (intro frefI nres_relI)
      (auto simp: twl_st_heur_def get_learned_clss_wl_def RETURN_RES_refine_iff)

lemma restart_required_heur_restart_required_wl0:
  ‹(uncurry restart_required_heur, uncurry restart_required_wl) ∈
    twl_st_heur''' r ×f nat_rel →f ⟨restart_flag_rel⟩nres_rel›
    unfolding restart_required_heur_def restart_required_wl_def uncurry_def Let_def
      restart_flag_rel_def  FLAG_GC_restart_def FLAG_restart_def FLAG_no_restart_def
      GC_required_heur_def
    by (intro frefI nres_relI)
     (auto simp: twl_st_heur_def get_learned_clss_wl_def RETURN_RES_refine_iff)

(*heuristic_rel (all_atms cd (cf + cg + ch + ci))
     (incr_restart_phase_end
       (incr_restart_phase*)
lemma heuristic_rel_incr_restartI[intro!]:
  ‹heuristic_rel 𝒜 heur ⟹ heuristic_rel 𝒜 (incr_restart_phase_end heur)›
  by (auto simp: heuristic_rel_def)

lemma update_all_phases_Pair:
  ‹(uncurry update_all_phases, uncurry (RETURN oo Pair)) ∈
    twl_st_heur'''' r ×f nat_rel →f ⟨twl_st_heur'''' r ×f nat_rel⟩nres_rel›
proof -
  have [refine0]: ‹(S, S') ∈ twl_st_heur'''' r ⟹ update_restart_phases S ≤ SPEC(λS. (S, S') ∈ twl_st_heur'''' r)›
    for S :: twl_st_wl_heur and S' :: ‹nat twl_st_wl›
    unfolding update_all_phases_def update_restart_phases_def
    by (auto simp: twl_st_heur'_def twl_st_heur_def
        intro!: rephase_heur_st_spec[THEN order_trans]
        simp del: incr_restart_phase_end.simps incr_restart_phase.simps)
  have [refine0]: ‹(S, S') ∈ twl_st_heur'''' r ⟹ rephase_heur_st S ≤ SPEC(λS. (S, S') ∈ twl_st_heur'''' r)›
    for S :: twl_st_wl_heur and S' :: ‹nat twl_st_wl›
    unfolding update_all_phases_def rephase_heur_st_def
    apply (cases S')
    apply (refine_vcg rephase_heur_spec[THEN order_trans, of ‹all_atms_st S'›])
    apply (clarsimp_all simp: twl_st_heur'_def twl_st_heur_def)
    done
  have Pair_alt_def: ‹RETURN ∘∘ Pair = (λS n. do {S ← RETURN S; S ← RETURN S; RETURN (S, n)})›
    by (auto intro!: ext)

  show ?thesis
    supply[[goals_limit=1]]
    unfolding update_all_phases_def Pair_alt_def
    apply (subst (1) bind_to_let_conv)
    apply (subst (1) Let_def)
    apply (subst (1) Let_def)
    apply (intro frefI nres_relI)
    apply (case_tac x rule:prod.exhaust)
    apply (simp only: uncurry_def prod.case)
    apply refine_vcg
    subgoal by simp
    subgoal by simp
    subgoal by simp
    done
qed

lemma restart_prog_wl_D_heur_restart_prog_wl_D:
  ‹(uncurry2 restart_prog_wl_D_heur, uncurry2 restart_prog_wl) ∈
    twl_st_heur''' r ×f nat_rel ×f bool_rel  →f ⟨twl_st_heur'''' r ×f nat_rel⟩nres_rel›
proof -
  have [refine0]: ‹GC_required_heur S n ≤ SPEC (λ_. True)› for S n
    by (auto simp: GC_required_heur_def)
  show ?thesis
   supply RETURN_as_SPEC_refine[refine2 del]
    unfolding restart_prog_wl_D_heur_def restart_prog_wl_def uncurry_def
    apply (intro frefI nres_relI)
    apply (refine_rcg
        restart_required_heur_restart_required_wl0[where r=r, THEN fref_to_Down_curry]
        cdcl_twl_restart_wl_heur_cdcl_twl_restart_wl_D_prog[where r=r, THEN fref_to_Down]
        cdcl_twl_full_restart_wl_D_GC_heur_prog[where r=r, THEN fref_to_Down]
        update_all_phases_Pair[where r=r, THEN fref_to_Down_curry, unfolded comp_def])
    subgoal by auto
    subgoal by (auto simp: restart_flag_rel_def FLAG_GC_restart_def FLAG_restart_def
      FLAG_no_restart_def)
    subgoal by auto
    subgoal by auto
    subgoal by (auto simp: restart_flag_rel_def FLAG_GC_restart_def FLAG_restart_def
      FLAG_no_restart_def)
    subgoal by auto
    subgoal by auto
    subgoal
      by auto
    done
 qed


lemma restart_prog_wl_D_heur_restart_prog_wl_D2:
  ‹(uncurry2 restart_prog_wl_D_heur, uncurry2 restart_prog_wl) ∈
  twl_st_heur ×f nat_rel ×f bool_rel →f ⟨twl_st_heur ×f nat_rel⟩nres_rel›
  apply (intro frefI nres_relI)
  apply (rule_tac r2 = ‹length(get_clauses_wl_heur (fst (fst x)))› and x'1 = ‹y› in
    order_trans[OF restart_prog_wl_D_heur_restart_prog_wl_D[THEN fref_to_Down]])
  apply fast
  apply (auto intro!: conc_fun_R_mono)
  done

definition isasat_trail_nth_st :: ‹twl_st_wl_heur ⇒ nat ⇒ nat literal nres› where
‹isasat_trail_nth_st S i = isa_trail_nth (get_trail_wl_heur S) i›

lemma isasat_trail_nth_st_alt_def:
  ‹isasat_trail_nth_st = (λ(M, _) i.  isa_trail_nth M i)›
  by (auto simp: isasat_trail_nth_st_def intro!: ext)

definition get_the_propagation_reason_pol_st :: ‹twl_st_wl_heur ⇒ nat literal ⇒ nat option nres› where
‹get_the_propagation_reason_pol_st S i = get_the_propagation_reason_pol (get_trail_wl_heur S) i›

lemma get_the_propagation_reason_pol_st_alt_def:
  ‹get_the_propagation_reason_pol_st = (λ(M, _) i.  get_the_propagation_reason_pol M i)›
  by (auto simp: get_the_propagation_reason_pol_st_def intro!: ext)


definition rewatch_heur_st_pre :: ‹twl_st_wl_heur ⇒ bool› where
‹rewatch_heur_st_pre S ⟷ (∀i < length (get_vdom S). get_vdom S ! i ≤ sint64_max)›

lemma isasat_GC_clauses_wl_D_rewatch_pre:
  assumes
    ‹length (get_clauses_wl_heur x) ≤ sint64_max› and
    ‹length (get_clauses_wl_heur xc) ≤ length (get_clauses_wl_heur x)› and
    ‹∀i ∈ set (get_vdom xc). i ≤ length (get_clauses_wl_heur x)›
  shows ‹rewatch_heur_st_pre xc›
  using assms
  unfolding rewatch_heur_st_pre_def all_set_conv_all_nth
  by auto

lemma li_uint32_maxdiv2_le_unit32_max: ‹a ≤ uint32_max div 2 + 1 ⟹ a ≤ uint32_max›
  by (auto simp: uint32_max_def)

end